Section 7: Recursive Definitions and Environments

 

 

 

The Y Combinator

We want to devise a λ-expression (a program in LC) that computes the least-fixed point of any function mapping some set of data values S to itself. In simple cases, we can let S be the natural numbers or the integers. One of the most surprising properties of the lambda-calculus (LC) is that there are a multitude (countably infinite) number of fundamentally different functions (with regard to equivalence under β-conversion [a finite sequence of applications of the β rule combined with changing variable names whenever convenient]) in the lambda-calculus that compute the least-fixed point of a function. Consider the λ-expression

  Y = λf. [λx. f (x x)] [λx. f (x x)]

(where square brackets are used as alternate notation for parentheses for improved readability). It is easy to show that Y behaves like a fixed-point operator, i.e., that

  f (Y f) = Y f.

Y is typically called the least fixed-point combinator. In the λ-calculus, λ expressions with no free variables are called combinators. If we perform β-reduction on the term

  f (Y f),

we get:

  f (Y f)  ->  f ([λx. f (x x)] [λx. f (x x)])

Similarly, if we perform two β-reduction steps on the term

  Y f,

we get:

  Y f  ->  (λx. f (x x)) (λx. f (x x)))  ->  f ([λx. f (x x)] [λx. f (x x)])

In call-by-name functional languages, the λ-expression Y defines the least-fixed-point operator that we defined in the last section of our notes.

Now consider the λ-expression

  B = λf. λx. f (x f x).

If we perform two β-reduction steps on the term

  B f B,

we get:

  B f B  ->  (λx. f (x f x)) B  ->  f (B f B)

Hence,

  Y' = λf. B f B

is another fixed-point operator. It is easy to confirm that:

  Y' f  =  f (Y' f)

Both the Y and Y' combinators take a function f and find its fixed point in call-by-name languages (where β-reduction is always valid). Suppose we want to find the fixed point of the function FACT defined by:

  λfact. λn. if n = 0 then 1 else n*(fact n-1)

Now we can apply the Y combinator to this function to get its fixed-point (assuming call-by-name):

  Y FACT -> (λx. FACT (x x)) (λx. FACT (x x)) ->  FACT ((λx. FACT (x x)) (λx. FACT (x x)))
          = λn. if n = 0 then 1 else n*((Y FACT) n-1)

where the final term uses the "collapsed" notation (Y FACT) instead of its expansion using β-reduction.

If we use Y' instead of Y, we can avoid the use of collapsed notation:

  Y' FACT -> B FACT B -> (λx. FACT (x FACT x)) B -> FACT (B FACT B) ->
          -> λn. if n = 0 then 1 else n*((B FACT B) n-1)

Note that we must are presuming call-by-name semantics here. Under call-by-value β-reduction

  Y f -> (λx. f (x x)) (λx. f (x x)))  ->  f ([λx. f (x x)] [λx. f (x x)])
      -> f (f ([λx. f (x x)] [λx. f (x x)])) -> ...

and

  Y' f -> B f B  ->  f(B f B)  ->  f(f(B f B)) ...

which both diverge. In call-by-value computations involving Y, the evaluation of

  (λx. f (x x)) (λx. f (x x)))

must be suspended by wrapping it in a &lambda-abstraction:

  λz. [(λx. f (x x)) (λx. f (x x)) z]

Alternatively, in call-by-value computations involving Y', the evaluation of

  B f

must be suspended by wrapping it in a λ-abstraction:

  λx. f (λn.(x f x) n)

Coping with Call-by-value

We can work around the pathologies of call-by-value semantics by defining appropriate variant of the Y [or Y'] combinator that constructs the least-fixed point of an argument function of the specified type using call-by-value semantics. The trick is to wrap the term f (x x) [or (x f x)] in an appropriate η-conversion (which is explained below). By appropriate we mean an η-conversion of the same type as the domain (and range) of f. In the pure λ-calculus, every value is a unary function, but in functional programming languages, the domain of f may be functions of some other arity or even a non-functional lazy domain like lazy lists.

In the pure λ-calculus, there are no atomic constants like natural numbers or booleans. The only constants are unary λ-expressions (functions). Hence, it is natural to assert for any term M that:

  M = λx. M x

or equivalently (in Scheme notation)

  M = (lambda (x) (M x))

The rewrite rule

  M  ->  λx. Mx

is called η-conversion. Technically, this rule can be applied in either direction (in contrast to β-reduction), but it is typically used in the direction shown above.

In the pure λ-calculus, there are no atomic constants like natural numbers or booleans. The only constants are unary λ-expressions (functions). Hence, it is natural to assert for any term M that:

  M = λx. Mx

or equivalently (in Scheme notation)

  M = (lambda (x) (M x))

In models of the pure λ-calculus, this rule is typically required to hold. But it is not implied by the equational axiom of (safe) β-conversion. It clearly does not hold for atomic (non-functional) constants, e.g.

  7 != λx. (7 x)

If the term M denotes a binary function (as is possible in Scheme or multi-ary generalizations of the λ-calculus), then

 M = (lambda (x y) (M x y))

which is not the same equation as standard η-conversion. Note that the implied type of M is different in "binary" η-conversion than it is in conventional ηconversion.

In realistic functional languages, conventional η-conversion does not hold. Rather, a restricted form of η-conversion holds for each type T. In other words, the appropriate η conversion for M depends on the type of M.

We can use conventional η-conversion to transform the call-by-name Y operator to a call-by-value least fixed-point operator for unary functions:

  λf. [λx. λy. f (x x) y] [λx. λy. f (x x) y]

(Recall that application associates to the left in the &lambda-calculus.) Similarly, we can transform the call-by-name Y' operator to a call-by-value least fixed-point operator for unary functions:

  λf. B' f B'

where

  B' = λf. λx. f (λn.(x f x) n).

For binary functions (functionals f mapping D x D -> D into itself), Y has the form:

λf. [λx. λy,z. f (x x) y] [λx. λy,z. f (x x) y]

or, equivalently, in Scheme notation

(lambda (f) ((lambda (x) (lambda (y z) (f (x x) y))) (lambda (x) (lambda (y z) (f (x x) y)))))

Recursive Definitions and Environments

Slogan: Definitions build environments.

Let us add a recursive binding mechanism (akin to let) to our language where we restrict right-hand sides to lambda expressions.

(define-struct rec-let (lhs   ; variable
                        rhs   ; required to be a lambda-expression
                        body))

where lhs is the new local variable, rhs is the lambda-expression defining the value of the new variable, and body is an expression that can use the new local variable. The new variable lhs is visible in both rhs and body. The code for it in the interpreter might look like

((rec-let? M) ... (MEval (rec-let-body M)
                         (extend env
                                 (rec-let-lhs M)
                                 (make-closure (rec-let-rhs M) E))))

To turn the rhs expression of M into a recursive closure, we desire that E be exactly like the environment that we are in the process of constructing. In other words, we would like a special kind of environment, env*, with the following property:

env* = (extend env (rec-let-lhs M) (make-closure (rec-let-rhs M) env*))

Using the following call-by-name procedure

(define FENV
  (lambda (env*)
    (extend env (rec-let-lhs M) (make-closure (rec-let-rhs M) env*))))

the equation for env* can be rewritten as

env* = (FENV env*)

which shows that we want the environment to be the fixed-point of the function FENV (from environments to environments).

Environments via Functions

If environments are implemented as functions (symbol -> Value), the definition of fix-env, the function that finds the fixed-point of FENV (and any other function mapping environments to enviroments), is simple:

(define fix-env
  (lambda (F)
    (local
       [(define renv (F (lambda (id) (lookup id renv))))]
      renv)))

If Scheme were a call-by-name functional language, the code would be even simpler:

(define fix-env
  (lambda (F)
    (local
       [(define renv (F renv))]
      renv)))

In a call-by-name language where local is a recursive binding form (as in Scheme),

(local
   [(define renv (F renv))]
  renv)))

binds renv to a suspension (closure of no arguments). When renv is evaluated in the body of the local, the suspension evaluates (F renv) in the environment constructed by local, which includes the new binding of renv. Hence,

  (fix-env FENV)
= (local 
     [(define renv (FENV renv))]
    renv)
= (local 
     [(define renv 
        (extend env (rec-let-lhs M) (make-closure (rec-let-rhs M) renv)))]
    renv)

If

M = (make-rec-let x proc ast),

then

  (local 
     [(define renv 
        (extend env (rec-let-lhs M) (make-closure (rec-let-rhs M) renv)))]
    renv)
= (local 
     [(define renv 
        (extend env x (make-closure proc renv)))]
    renv)
= (define renv' (extend env (make-closure proc renv')))
  renv'

which is exactly the recursively defined environment that we need. The (call-by-value) Scheme solution requires the expression

(lambda (id) (lookup id renv))

instead of renv because call-by-value evaluation requires a lambda form to suspend evaluation. On inputs of type environment (symbol -> Value), the function

(lambda (env) (lambda (id) (lookup id env))

is the identity function except on improper values of env (undefined, divergence, or errors). On improper values for env,

(lambda (id) (lookup id env))

is still a legal environment (which will blow up only if it is applied to an Id).

In the absence of this suspension trick, the evaluation of the right hand side of the definition

     (define renv 
       (extend env (rec-let-lhs M) (make-closure (rec-let-rhs M) renv)))

in Scheme, blows up because renv has no value (is undefined) until the evaluation of the definition is complete. The suspension trick wraps the argument renv inside a λ-expresssion without changing its meaning for proper values of env.

We have used the recursive binding mechanism local in Scheme (and call-by-name Scheme) to implement rec-let. Note that we have placed a strong requirement on the rhs of rec-let in LC. Therefore, this interpreter does not completely explain how local works in Scheme.

We could write an interpreter in Scheme that uses the Y operator instead of explicit recursion to define renv. Of course, the standard Y operator presumes call-by-name but we could use a call-by-value variant of Y (which wraps f (x x) in an appropriate η-conversion) to make our interpreter work for conventional Scheme call-by-value semantics. But this reduction only provides mathematical insight into the meaning of rec-let. It does not embody a translation that we can directly use in mapping rec-let to machine code. We will investigate this issue later in the course.

Environments via Structures

If we represent environments as structures, how do we write fix-env? We need to construct an environment that contains references to copies of itself. In a "lazy" functional language like Haskell or Miranda, we could easily build the required structures using "lazy" constructors that defer the evaluation of their arguments. For example if we added lazy-cons to LC, the code would

(rec-let x (lazy-cons 0 x) ...)

define x as a list consisting of 0 followed by x.

Instead of storing a value in each field of the constructed structure, a lazy constructor stores a suspension specifying how to evaluate the field when it is needed. This is exactly the same mechanism used to support call-by-name argument passing to program defined functions. Note that this approach formulates self referential environments as infinite data structures rather data structures with cycles.

In functional Scheme, we cannot use ordinary data structures (non-functions) to solve the recursion equation

env* = (FENV env*)

because these structures cannot be infinite. However, in imperative Scheme, we can create the the necessary self-references using cycles created by brute-force. We can overwrite the environment field inside the new closure to point to the extended environment. If we adhere to the list representation of environments introduced in Lecture 4, we must abandon the explicit definition the fix-env operation because finding the references to env* that we have to patch in (FENV env*) is either impossible (if F uses constructors that we have not anticipated) or inefficient (because we have to traverse all of the data structure (F env*)).

Fortunately, we don't need a a general function fix-env for our interpreter; all we need is a function (rec-extend M env) that builds the fixed-point for FENV:

(FENV (rec-extend M env)) = (rec-extend M env)

We can define rec-extend as follows

(define rec-extend
  (lambda (M env)
    (local 
       [(define var (rec-let-lhs M))
        (define rhs (rec-let-rhs M))
        (define clsr (make-closure rhs null))
	(define renv (extend env var clsr))]
      ; Is renv the appropriate environment yet?  No.
      (set-closure-env! clsr renv))))

Given the environment env and the rec-let expression M of a function, (rec-extend env M) extends the environment with the new function binding in M such that the environment embedded in the function binding is the extended environment. The rec-let clause of our interpreter can now be written

((rec-let? M) ... (MEval (rec-let-body M) (rec-extend env M)))

Exercise: Devise a minor change to the representation of environments as structures that makes it easy to write a general fix-env function.
Hint: Scheme boxes.

More General Recursive Binding

In LC we restricted the right hand sides of recursive bindings to lambda-expressions. What happens if we allow arbitrary expressions as right hand sides. To evaluate such a right hand side rhs, we must call (MEval rhs env*) where env* is our self-referential environment. If this evaluation actually tries to lookup the value of var, the computation should abort because the value of var is still being computed. (DrScheme follows this convention, but unfortunately, it is left unspecified in the Scheme standard.) For LC, we can rewrite rec-extend to support this generalization as follows:

(define rec-extend
  (lambda (M env)
    (local 
       [(define var (rec-let-lhs M))
        (define rhs (rec-let-rhs M))
	(define renv (extend env var (void)))]
      ; (void) is a dummy value for var
      ; Is renv the appropriate environment yet?  No.
      (set-first-pair-val! renv (MEval rhs renv)))))

(define set-first-pair-val!
  (lambda (env val)
     (local
        [(define first-pair (first env))]
       (set-pair-val! first-pair val))))

To make MEval abort on the inspection of (void) we also need to add the clause

   [(void? M) (error 'MEval 
                 "variable referenced before definition in rec-let")]

to the definition of MEval.

Exercise: Generalize the code from the previous exercise to allow an arbitrary expression as the right hand side of a rec-let.
Hint: It is much simpler than the code above.

Memory Management, Part 0

At this point in the course, we have explained several language constructs, such as procedures and applications, in terms of their Scheme counterparts. A notable exception is environments, which we have represented explicitly. Indeed, since we explicitly manage variable-value associations in our interpreters, we can analyze some of the lower-level behavior of the language such as its memory management.

Consider the evaluation of

(let (x 1)
  (let (y 2)
    (let (z 3)
      ...)))

We begin with the empty environment; as we encounter each let, we add a new level to our environment; as we leave the let body, we remove that level. At the end of the entire expression, our environment will be empty again. Hence, our environment will have behaved like a stack. Do environments always behave in this manner?

Consider this expression:

((lambda (x) (lambda (y) (y x))) 10)

This evaluates to an ``apply to 10'' function. This expression

((lambda (z) (lambda (w) (+ z w))) 20)

evaluates to a procedure that adds 20 to its argument. If we apply the former to the latter, what happens to the environment at each stage?

We begin with the empty environment; the first application is performed and the environment is extended with the binding [x = 10]. The closure is now created, and the environment is emptied. Then we evaluate the second expression, which adds the binding [z = 20]. The second closure (call it C2) is produced, and the environment is emptied again.

At this point, we are ready to perform the desired application. When we do, the environment has the bindings [x = 10] and [y = C2]. Now C2 is applied, which has the effect of replacing the current environment with its own, which contains the bindings [z = 20]. The application adds the binding [w = 10], at which point the addition is performed, 30 is returned, and the environment is emptied again.

The moral of this example is that environments for LC programs, no matter how we choose to represent them, branch out like trees, and a simple stack discipline is insufficient for maintaining them. In many programming languages (such as Fortran, Pascal, and C/C++), it is often the programmer's job to manage memory and the collection of environments in existence at any point in the program must be representable by a stack. Such languages prohibit language constructions that can create a collection of environments that cannot represented by a stack. In our interpreters for LC, we have left the task of managing the storage for LC environments to Scheme's run-time system. Otherwise, our interpreters would have been much more complex programs.

Exercise: Impose restrictions on LC so that it remains a procedural language yet does not require a tree-based management of environments
Hint: Follow the example of Pascal.

Implementing Recursion in OCaml (Optional)

The Y combinator is not necessarily the best way to actually implement recursion. Nearly all functional languages build in a fixed-point operator and use a conventional call stack to implement it. Of course, implementations of the Y combinator rely on the same call stack, but there are no cyclic data structures in the stack.

The new expression type is:

type debruijn =
    Cn of int
  | Vr of int
  | Lm of debruijn
  | Ap of debruijn * debruijn
  | Fi of debruijn

With substitution semantics for variable bindings, the semantics of this new language are:

exception Error of string;;

let rec plus e m =
  match e with
      Cn i -> Cn i
    | Vr n -> Vr (if n<m then n else n+1)
    | Ap (e1,e2) -> Ap (plus e1 m, plus e2 m)
    | Lm (e) -> Lm (plus e (m+1))
    | Fi (e) -> Fi (plus e (m+1))

let rec subst i v e =
  match e with
      Cn i -> Cn i
    | Vr n -> (if n==i then v
               else if n>i then Vr (n-1)
               else Vr n)
    | Ap (e1,e2) -> Ap (subst i v e1, subst i v e2)
    | Lm (e) -> Lm (subst (i+1) (plus v 0) e)
    | Fi (e) -> Fi (subst (i+1) (plus v 0) e)

(* Big-step semantics *)

let rec ev5 e =
  match e with
      Cn i -> Cn i
    | Vr x ->
        raise (Error "Unbound variable")
    | Lm e' -> e
    | Ap (e1,e2) ->
        (match (ev5 e1) with
             Lm e  ->
               let v = ev5 e2 in
                 ev5 (subst 0 v e)
           | _ -> raise (Error "Applying non-function"))
    | Fi e' -> ev5 (subst 0 (Fi e') e')

The expression Fi e evaluates the contained expression e with e bound to a variable. The substitution replaces occurrences of the newly bound variable with the expression Fi e, delaying evaluation until it is needed.

With environment semantics, we define the fixed point operator as follows:

type debValue =
    V_Cn of int
  | V_Lm of (debValue list) * debruijn
  | V_Fi of (debValue list) * debruijn

let rec ev6 env e =
  match e with
      Cn i -> V_Cn i
    | Vr 0 ->
        (match env with
             v::vs ->
               (match v with
                    V_Fi (env',e') -> ev6 (v::env') e'
                  | _ -> v)
           | _ -> raise (Error "Unbound variable"))
    | Vr n ->
        (match env with
             v::vs -> ev6 vs (Vr (n-1))
           | _ -> raise (Error "Unbound variable"))
    | Lm e' -> V_Lm (env, e')
    | Ap (e1,e2) ->
        (match (ev6 env e1) with
             (V_Lm (env',e))  ->
               let v = ev6 env e2 in
                 ev6 (v::env') e
           | _ -> raise (Error "Applying non-function"))
    | Fi e' -> ev6 ((V_Fi (env, e'))::env) e'

The fixed point expression Fi e evaluates to the evaluation of e with Fi e bound in the environment. Again, the expression e, and not the result of evaluating this expression, is stored in the environment; evaluation is delayed until it is needed.

In the rest of these notes, we consider in Scheme some additional ways to implement recursive definitions.

Back to course website