From Math to Machine A formal derivation of an executable Krivine Machine Wouter Swierstra Brouwer Seminar

Similar documents
A formal derivation of an executable Krivine machine

Accurate Step Counting

Fall 2013 Midterm Exam 10/22/13. This is a closed-book, closed-notes exam. Problem Points Score. Various definitions are provided in the exam.

Implementing functional languages with abstract machines

Concepts of programming languages

BRICS. A Concrete Framework for Environment Machines. BRICS RS-06-3 Biernacka & Danvy: A Concrete Framework for Environment Machines

From Reduction-based to Reduction-free Normalization

arxiv:cs/ v1 [cs.lo] 8 Aug 2005

Programming Languages Lecture 14: Sum, Product, Recursive Types

Lecture 13: Subtyping

MPRI course 2-4 Functional programming languages Exercises

Accurate Step Counting

BRICS. A Syntactic Correspondence between Context-Sensitive Calculi and Abstract Machines

Last class. CS Principles of Programming Languages. Introduction. Outline

A Typed Calculus Supporting Shallow Embeddings of Abstract Machines

CSE-321 Programming Languages 2010 Midterm

CIS 500 Software Foundations Fall September 25

Lecture Notes on Data Representation

λ calculus Function application Untyped λ-calculus - Basic Idea Terms, Variables, Syntax β reduction Advanced Formal Methods

M. Snyder, George Mason University LAMBDA CALCULUS. (untyped)

λ calculus is inconsistent

Lesson 4 Typed Arithmetic Typed Lambda Calculus

Type Systems Winter Semester 2006

Programming Languages

Lambda calculus. Wouter Swierstra and Alejandro Serrano. Advanced functional programming - Lecture 6

Towards Reasoning about State Transformer Monads in Agda. Master of Science Thesis in Computer Science: Algorithm, Language and Logic.

An experiment with variable binding, denotational semantics, and logical relations in Coq. Adam Chlipala University of California, Berkeley

CS 6110 S11 Lecture 12 Naming and Scope 21 February 2011

The Lambda Calculus. 27 September. Fall Software Foundations CIS 500. The lambda-calculus. Announcements

Parsing Scheme (+ (* 2 3) 1) * 1

A Derivational Approach to the Operational Semantics of Functional Languages

COMP 1130 Lambda Calculus. based on slides by Jeff Foster, U Maryland

COMP 4161 NICTA Advanced Course. Advanced Topics in Software Verification. Toby Murray, June Andronick, Gerwin Klein

` e : T. Gradual Typing. ` e X. Ronald Garcia University of British Columbia

CSE-321 Programming Languages 2012 Midterm

Hutton, Graham and Bahr, Patrick (2016) Cutting out continuations. In: WadlerFest, April 2016, Edinburgh, Scotland.

Programming Languages Lecture 15: Recursive Types & Subtyping

Programming Languages Fall 2014

Exercise 1 (2+2+2 points)

Trees. Solution: type TreeF a t = BinF t a t LeafF 1 point for the right kind; 1 point per constructor.

A native compiler for Coq: current status, perspectives and discussion 1

dynamically typed dynamically scoped

Programming Language Features. CMSC 330: Organization of Programming Languages. Turing Completeness. Turing Machine.

CMSC 330: Organization of Programming Languages

Polymorphism and System-F (OV)

Toward explicit rewrite rules in the λπ-calculus modulo

CSE505, Fall 2012, Midterm Examination October 30, 2012

Lambda Calculus: Implementation Techniques and a Proof. COS 441 Slides 15

A Simple Supercompiler Formally Verified in Coq

Exercise 1 ( = 22 points)

4.5 Pure Linear Functional Programming

λ-calculus Lecture 1 Venanzio Capretta MGS Nottingham

CIS 500 Software Foundations Midterm I

Chapter 13: Reference. Why reference Typing Evaluation Store Typings Safety Notes

Part III Chapter 15: Subtyping

Operational Semantics 1 / 13

Graphical Untyped Lambda Calculus Interactive Interpreter

Journal of Computer and System Sciences

A Small Interpreted Language

Exercise 1 ( = 24 points)

Mikiβ : A General GUI Library for Visualizing Proof Trees

J. Barkley Rosser, 81, a professor emeritus of mathematics and computer science at the University of Wisconsin who had served in government, died

Theorem Proving Principles, Techniques, Applications Recursion

The University of Nottingham SCHOOL OF COMPUTER SCIENCE A LEVEL 4 MODULE, SPRING SEMESTER MATHEMATICAL FOUNDATIONS OF PROGRAMMING ANSWERS

Part III. Chapter 15: Subtyping

CMSC 336: Type Systems for Programming Languages Lecture 5: Simply Typed Lambda Calculus Acar & Ahmed January 24, 2008

More Untyped Lambda Calculus & Simply Typed Lambda Calculus

Lecture 5: The Untyped λ-calculus

Generalized Refocusing: From Hybrid Strategies to Abstract Machines

Introduction to Functional Programming in Haskell 1 / 56

Exercise 1 ( = 18 points)

Denotational Semantics. Domain Theory

Lambda Calculus and Type Inference

Functional Languages. Hwansoo Han

BRICS. A Functional Correspondence between Monadic Evaluators and Abstract Machines for Languages with Computational Effects

Lambda Calculus. Concepts in Programming Languages Recitation 6:

Functions as data. Massimo Merro. 9 November Massimo Merro The Lambda language 1 / 21

MLW. Henk Barendregt and Freek Wiedijk assisted by Andrew Polonsky. March 26, Radboud University Nijmegen

CSE-321 Programming Languages 2011 Final

CS131 Typed Lambda Calculus Worksheet Due Thursday, April 19th

CIS 500 Software Foundations Fall October 2

Graph reduction. Simple graph reduction with visualisation

Metaprogramming assignment 3

TYPE INFERENCE. François Pottier. The Programming Languages Mentoring ICFP August 30, 2015

CITS3211 FUNCTIONAL PROGRAMMING. 14. Graph reduction

Verifying Program Invariants with Refinement Types

On Meaning Preservation of a Calculus of Records

Type checking. Jianguo Lu. November 27, slides adapted from Sean Treichler and Alex Aiken s. Jianguo Lu November 27, / 39

Type Checking and Type Inference

Abstract register machines Lecture 23 Tuesday, April 19, 2016

Formal Semantics. Aspects to formalize. Lambda calculus. Approach

Verified Characteristic Formulae for CakeML. Armaël Guéneau, Magnus O. Myreen, Ramana Kumar, Michael Norrish April 27, 2017

Recursion and Induction

11/6/17. Outline. FP Foundations, Scheme. Imperative Languages. Functional Programming. Mathematical Foundations. Mathematical Foundations

Universes. Universes for Data. Peter Morris. University of Nottingham. November 12, 2009

Types for References, Exceptions and Continuations. Review of Subtyping. Γ e:τ τ <:σ Γ e:σ. Annoucements. How s the midterm going?

Introduction to Lambda Calculus. Lecture 5 CS 565 1/24/08

Lecture 15 CIS 341: COMPILERS

Typed Lambda Calculus and Exception Handling

9/23/2014. Why study? Lambda calculus. Church Rosser theorem Completeness of Lambda Calculus: Turing Complete

Transcription:

From Math to Machine A formal derivation of an executable Krivine Machine Wouter Swierstra Brouwer Seminar

β reduction (λx. t0) t1 t0 {t1/x}

Substitution Realizing β-reduction through substitution is a terrible idea! Instead, modern compilers evaluate lambda terms using an abstract machine, such as Haskell s STG or OCaml s CAM. Such abstract machines are usually described as tail-recursive functions/finite state machines.

Abstract machines Abstract machines have many applications: trace and debug evaluation; study operational behaviour of new language constructs (continuations, threads, etc.); a framework for program analysis.

Who comes up with these things?

Olivier Danvy and his many students and collaborators

Most of our implementations of the abstract machines raise compiler warnings about nonexhaustive matches. These are inherent to programming abstract machines in an ML-like language Mads Sig Ager, Dariusz Biernacki, Olivier Danvy, Jan Midtgaard

Outline 1. Define a small step evaluator for the simply typed lambda calculus in Agda; 2. Show that this evaluator terminates; 3. From this evaluator, derive a Krivine machine in two short steps; 4. Prove that these preserve the semantics and termination behaviour of the original.

Small step evaluation

Types data Ty : Set where O : Ty _=>_ : Ty -> Ty -> Ty Context : Set Context = List Ty

Terms data Term : Context -> Ty -> Set where Lam : Term (Cons u Γ) v -> Term Γ (u => v) App : Term Γ (u => v) -> Term Γ u -> Term Γ v Var : Ref Γ u -> Term Γ u

Closed terms data Closed : Ty -> Set where Closure : Term Γ u -> Env Γ -> Closed u Clapp : Closed (u => v) -> Closed u -> Closed v data Env : Context -> Set where Nil : Env Nil _ _ : Closed u -> Env Γ -> Env (Cons u Γ)

Values isval : Closed u -> Set isval (Closure (Lam body) env) = Unit isval _ = Empty data Value (u : Ty) : Set where Val : (c : Closed u) -> isval c -> Value u

Reduction rules

Head reduction in three steps Decompose the term into a redex and evaluation context; Contract the redex; Plug the result back into the context.

Redex data Redex : Ty -> Set where Lookup : Ref Γ u -> Env Γ -> Redex u App : Term Γ (u => v) -> Term Γ u -> Env Γ -> Redex v Beta : Term (Cons u Γ) v -> Env Γ -> Closed u -> Redex v

Contraction contract : Redex u -> Closed u contract (Lookup i env) = env! i contract (App f x env) = Clapp (Closure f env) (Closure x env) contract (Beta body env arg) = Closure body (arg env)

Decomposition as a view Idea: every closed term is: a value; or a redex in some evaluation context. Define a view on closed terms.

Views: example Natural numbers are typically defined using the Peano axioms. But sometimes you want to use the fact that every number is even or odd, e.g. when converting to a binary representation; or proving 2 is irrational. But why is that a valid proof principle?

Views: example How can we derive even-odd induction from Peano induction? Define a data type EvenOdd : Nat -> Set Define a covering function evenodd : (n : Nat) -> EvenOdd n

The view data type data EvenOdd : Nat -> Set where IsEven : (k : Nat) -> EvenOdd (double k) IsOdd : (k : Nat) -> EvenOdd (Succ (double k))

Covering function evenodd : (n : Nat) -> EvenOdd n evenodd Zero = IsEven Zero evenodd (Succ Zero) = IsOdd Zero evenodd (Succ (Succ k)) with evenodd k... IsEven k' = IsEven (Succ k')... IsOdd k' = IsOdd (Succ k')

Example example: Nat -> Bin example n with evenodd n example.(double k) IsEven k =... example.(succ (double k)) IsOdd k =...

Decomposition as a view Idea: every closed term is: a value; or a redex in some evaluation context. Define a view on closed terms.

Evaluation contexts data EvalContext : Context -> Set where MT : EvalContext Nil ARG : Closed u -> EvalContext -> EvalContext (Cons u ) _==>_ : List Ty -> Ty -> Ty Nil ==> u = u (Cons v ) ==> u = v => ( ==> u)

Plug plug : EvalContext -> Closed ( ==> u) -> Closed u plug MT f = f plug (ARG x ctx) f = plug ctx (Clapp f x)

Decomposition data Decomposition : Closed u -> Set where Val : (t : Closed u) -> isval t -> Decomposition t Decompose : (r : Redex ( ==> v)) -> (ctx : EvalContext ) -> Decomposition (plug ctx (fromredex r))

Decompose decacc : (ctx : EvalContext ) (c : Closed ( ==> u)) -> Decomposition (plug ctx c) decacc MT (Closure (Lam body) env) = Val (Closure (Lam body) env) unit decacc (ARG x ctx) (Closure (Lam body) env) = Decompose (Beta body env x) ctx decacc ctx (Closure (App f x) env) = Decompose (App f x env) ctx decacc ctx (Closure (Var i) env) = Decompose (Lookup i env) ctx decacc ctx (Clapp f x) = decacc (ARG x ctx) f decompose : (c : Closed u) -> Decomposition c decompose c = decacc MT c

Head-reduction headreduce : Closed u -> Closed u headreduce c with decompose c... Val val p = val... Decompose redex ctx = plug ctx (contract redex)

Iterated head reduction evaluate : Closed u -> Value u evaluate c = iterate (decompose c) where iterate : Decomposition c -> Value u iterate (Val val p) = Val val p iterate (Decompose r ctx) = iterate (decompose (plug ctx (contract r)))

Iterated head reduction evaluate : Closed u -> Value u evaluate c = iterate (decompose c) where iterate : Decomposition c -> Value u iterate (Val val p) = Val val p iterate (Decompose r ctx) = iterate (decompose (plug ctx (contract r))) Does not pass the termination check!

The Bove-Capretta method

Bove-Capretta data Trace : Decomposition c -> Set where Done : (val : Closed u) -> (p : isval val) -> Trace (Val val p) Step : Trace (decompose (plug ctx (contract r))) -> Trace (Decompose r ctx)

Iterated head reduction, again iterate : {u : Ty} {c : Closed u} -> (d : Decomposition c) -> Trace d -> Value u iterate.(val val p) Done = Val val p iterate.(decompose r ctx) (Step step) = let d' = decompose (plug ctx (contract r)) in iterate d' step

Nearly done We still need to find a trace for every term... (c : Closed u) -> Trace (decompose c)

Nearly done We still need to find a trace for every term... (c : Closed u) -> Trace (decompose c) Fail

Nearly done We still need to find a trace for every term... (c : Closed u) -> Trace (decompose c) Fail Yet we know that the simply typed lambda calculus is strongly normalizing...

Logical relation Reducible : (u : Ty) -> (t : Closed u) -> Set Reducible O t = Trace (decompose t) Reducible (u => v) t = Pair (Trace (decompose t)) ((x : Closed u) -> Reducible u x -> Reducible (Clapp t x))

Required lemmas lemma1 : (t : Closed u) -> Reducible (headreduce t) -> Reducible t lemma2 : (t : Term G u) (env : Env G) -> ReducibleEnv env -> Reducible (Closure t env)

Result! theorem : (c : Closed u) -> Reducible c theorem (Closure t env) = lemma2 t env (envtheorem env) theorem (Clapp f x) = snd (theorem f) x (theorem x) termination : (c : Closed u) -> Trace (decompose c)...an easy corollary

Finally, evaluation evaluate : Closed u -> Value u evaluate t = iterate (decompose t) (termination t)

The story so far... Data types for terms, closed terms, values, redexes, evaluation contexts. Defined a three step head-reduction function: decompose, contract, plug. Proven that iterated head reduction yields a normal form...... and used this to define a normalization function.

What s next? Use Danvy & Nielsen s refocusing transformation to define a pre-abstract machine; Inline the iterate function (and one or two minor changes), yields the Krivine machine. Prove that each transformation preserves the termination behaviour and semantics.

A term

A redex

Contract

Plug

Repeat

The drawback To contract a single redex, we need to: traverse the term to find a redex; contract the redex; traverse the context to plug back the contractum.

Refocusing The refocusing transformation (Danvy & Nielsen) avoids these traversals. Instead, given a decomposition, it navigates to the next redex immediately. Intuitively, refocusing behaves just the same as decompose. plug

Refocus summary refocus : (ctx : EvalContext Δ) -> (c : Closed (Δ ==> u)) -> Decomposition (plug ctx c) refocuscorrect : (ctx : EvalContext ) -> (c : Closed ( ==> u)) -> refocus ctx c == decompose (plug ctx c)

What else? It is easy to prove that iteratively refocusing and contracting redexes produces the same result as the small step evaluator. And that if the Trace data type is inhabited, then so is the corresponding data type for the refocussing evaluator.

The Krivine machine Now inline the iterate function; and disallow closed applications; and compress corridor transitions.

The Krivine machine refocus : (ctx : EvalContext Δ) -> (t : Term Γ (Δ ==> u)) -> (env : Env Γ) -> Value u refocus ctx (Var i) env = let Closure t env = lookup i env q in refocus ctx t env refocus ctx (App f x) env = refocus (ARG (Closure x env) ctx) f env refocus (ARG x ctx) (Lam body) env = refocus ctx body (x env) refocus MT (Lam body) env = Val (Closure (Lam body) env)

Once again... We need to prove that this function terminates...... by adapting the proof we saw for the refocusing evaluator.... and show that it produces the same value as our previous evaluation functions.

Proofs? decomposeplug : (ctx : EvalContext us) -> (c : Closed (us ==> u)) -> decompose (plug ctx c) == decacc ctx c decomposeplug MT c = Refl decomposeplug (ARG x ctx) t rewrite decomposeplug ctx (Clapp t x) decomposeclapp ctx t x = Refl

Proofs?? refocuscorrect : (ctx : EvalContext D) (c : Closed (D ==> u)) -> refocus ctx c == decompose (plug ctx c) refocuscorrect MT (Closure (Lam body) env) = Refl refocuscorrect (ARG x ctx) (Closure (Lam body) env) rewrite decomposeplug ctx (Clapp (Closure (Lam body) env ) x) decomposeclapp ctx (Closure (Lam body) env) x = Refl refocuscorrect MT (Closure (Var i) env) = Refl refocuscorrect (ARG x ctx) (Closure (Var i) env) rewrite decomposeplug ctx (Clapp (Closure (Var i) env) x) decomposeclapp ctx (Closure (Var i) env) x = Refl refocuscorrect MT (Closure (App f x) env) = Refl refocuscorrect (ARG arg ctx) (Closure (App f x) env) rewrite decomposeplug ctx (Clapp (Closure (App f x) env) arg) decomposeclapp ctx (Closure (App f x) env) arg = Refl refocuscorrect MT (Clapp f x) = refocuscorrect (ARG x MT) f refocuscorrect (ARG arg ctx) (Clapp f x) = refocuscorrect (ARG x (ARG arg ctx)) f

Proofs??? termination : (t : Term Nil u) -> Trace t Nil MT termination t = lemma MT t Nil (unit, unit) (Section5.termination (Closure t Nil)) where lemma : (ctx : EvalContext D) (t : Term G (D ==> u)) (env : Env G) -> (invariant ctx env) -> Section5.Trace (Section5.refocus ctx (Closure t env)) -> Trace t env ctx lemma MT (Lam body) env p (Section5.Done.(Closure (Lam body) env).(record { })) = Done body lemma MT (App f x) env p (Section5.Step y0) = App f x (lemma (ARG (Closure x env) MT) f env ((fst p), p) y0) lemma MT (Var i) env (p1, p2) (Section5.Step y') rewrite lookupclosure env i p1 = Lookup i p1 (lemma MT (getterm (lookup i env p1)) (getenv (lookup i env p1)) (lookuplemma env i p1, p2) y') lemma (ARG (Closure arg env') ctx) (Lam body) env (p1, (p2, p3)) (Section5.Step step) = Beta ctx arg env' body (lemma ctx body ((Closure arg env') env) ((p2, p1), p3) step) lemma (ARG (Clapp y y') ctx) (Lam y0) env (y1, ()) (Section5.Step y3) lemma (ARG arg ctx) (App (Lam y) x) env p (Section5.Step step) = App ((Lam y)) x (lemma (ARG (Closure x env) (ARG arg ctx)) (Lam y) env ((fst p), p) step) lemma (ARG arg ctx) (App (App f y) x) env p (Section5.Step step) = App (App f y) x (lemma (ARG (Closure x env) (ARG arg ctx)) (App f y) env ((fst p), p) step) lemma (ARG arg ctx) (App (Var i) x) env p (Section5.Step step) rewrite lookupclosure env i (fst p) = let c = lookup i env (fst p) in App (Var i) x (lemma (ARG (Closure x env) (ARG arg ctx)) (Var i) env ((fst p), p) step) lemma (ARG arg ctx) (Var i) env p (Section5.Step step) rewrite lookupclosure env i (fst p) = let c = lookup i env (fst p) in Lookup i (fst p) (lemma (ARG arg ctx) (getterm c) (getenv c) ((lookuplemma env i (fst p)), (snd p)) step)

Lines of code vs. lines of proof In total, 585 lines of Agda (288/297) 50 lines Prelude (50/0) 100 lines data types & basics (100/0) 100 lines lemmas (0/100) 100 lines reducibility proof (33/67) 100 lines pre-abstract machine (25/75) 135 lines Krivine machine (80/55)

Conclusions You can reason about functions with nontrivial recursive behaviour in type theory. You can repeat the same trick on many other derivations of abstract machines. All our evaluators are executable; we do not need any additional postulates.