Synthesis of distributed mobile programs using monadic types in Coq

Similar documents
Coq, a formal proof development environment combining logic and programming. Hugo Herbelin

Introduction to dependent types in Coq

From natural numbers to the lambda calculus

Functional Programming with Isabelle/HOL

Beluga: A Framework for Programming and Reasoning with Deductive Systems (System Description)

Mathematics for Computer Scientists 2 (G52MC2)

Types Summer School Gothenburg Sweden August Dogma oftype Theory. Everything has a type

Meta-programming with Names and Necessity p.1

Importing HOL-Light into Coq

c constructor P, Q terms used as propositions G, H hypotheses scope identifier for a notation scope M, module identifiers t, u arbitrary terms

Coq quick reference. Category Example Description. Inductive type with instances defined by constructors, including y of type Y. Inductive X : Univ :=

Mathematics for Computer Scientists 2 (G52MC2)

λ calculus is inconsistent

A NEW PROOF-ASSISTANT THAT REVISITS HOMOTOPY TYPE THEORY THE THEORETICAL FOUNDATIONS OF COQ USING NICOLAS TABAREAU

Refinement Types as Proof Irrelevance. William Lovas with Frank Pfenning

Preuves Interactives et Applications

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

Functional Programming in Coq. Nate Foster Spring 2018

CIS 194: Homework 8. Due Wednesday, 8 April. Propositional Logic. Implication

Logical Verification Course Notes. Femke van Raamsdonk Vrije Universiteit Amsterdam

Formalization of Array Combinators and their Fusion Rules in Coq

Introduction to Co-Induction in Coq

Certified Programming with Dependent Types

Contractive Signatures with Recursive Types, Type Parameters, and Abstract Types

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

The Truth about Types. Bartosz Milewski

Programming Languages Lecture 14: Sum, Product, Recursive Types

Less naive type theory

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

Numerical Computations and Formal Methods

CS 6110 S11 Lecture 25 Typed λ-calculus 6 April 2011

Lecture slides & distribution files:

1 Introduction. 3 Syntax

Calculus of Inductive Constructions

CS152: Programming Languages. Lecture 11 STLC Extensions and Related Topics. Dan Grossman Spring 2011

GADTs meet Subtyping

Programming with dependent types: passing fad or useful tool?

Modular dependent induction in Coq, Mendler-style. Paolo Torrini

Autosubst Manual. May 20, 2016

Introduction to Coq Proof Assistant

Combining Programming with Theorem Proving

Type- & Example-Driven Program Synthesis. Steve Zdancewic WG 2.8, August 2014

Mutable References. Chapter 1

Embedding logics in Dedukti

Chapter 1. Introduction

Coq. LASER 2011 Summerschool Elba Island, Italy. Christine Paulin-Mohring

Dependent Types and Irrelevance

Programming Languages Lecture 15: Recursive Types & Subtyping

Adam Chlipala University of California, Berkeley ICFP 2006

Lesson 4 Typed Arithmetic Typed Lambda Calculus

From Types to Sets in Isabelle/HOL

CIS 500: Software Foundations

Inductive data types

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

A Simple Supercompiler Formally Verified in Coq

Generic Constructors and Eliminators from Descriptions

A Pronominal Account of Binding and Computation

A Reflection-based Proof Tactic for Lattices in COQ

CS 4110 Programming Languages & Logics. Lecture 28 Recursive Types

Equations: a tool for dependent pattern-matching

CSE505, Fall 2012, Midterm Examination October 30, 2012

Program Verification Through Characteristic Formulae

Lambda Calculus and Type Inference

The design of a programming language for provably correct programs: success and failure

The pitfalls of protocol design

Induction in Coq. Nate Foster Spring 2018

CSE-321 Programming Languages 2010 Final

Proofs-Programs correspondance and Security

Part VI. Imperative Functional Programming

A CRASH COURSE IN SEMANTICS

Assistant for Language Theory. SASyLF: An Educational Proof. Corporation. Microsoft. Key Shin. Workshop on Mechanizing Metatheory

Inductive data types (I)

Runtime Behavior of Conversion Interpretation of Subtyping

Natural deduction environment for Matita

type classes & locales

Typed Lambda Calculus for Syntacticians

Typed Compilation Against Non-Manifest Base Classes

A Certified Implementation of a Distributed Security Logic

No.4, Section 2, North Jianshe Road, , Sichuan, Chengdu, P.R. China.

Isabelle/HOL:Selected Features and Recent Improvements

Inductive datatypes in HOL. lessons learned in Formal-Logic Engineering

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

Integration of SMT Solvers with ITPs There and Back Again

Com S 541. Programming Languages I

Provably Correct Software

Verification in Coq. Prof. Clarkson Fall Today s music: Check Yo Self by Ice Cube

Recitation 8: Dynamic and Unityped Languages : Foundations of Programming Languages

Vector Clocks in Coq

Second-Order Type Systems

Arbitrary-rank polymorphism in (GHC) Haskell

Mechanizing Metatheory with LF and Twelf

Mechanized metatheory revisited

Programming Languages

Inductive Beluga: Programming Proofs

Logical Mobility and Locality Types

Computational Higher-Dimensional Type Theory

CS 456 (Fall 2018) Scribe Notes: 2

1.3. Conditional expressions To express case distinctions like

Research Statement. Daniel R. Licata

Nominal Techniques in Coq

Transcription:

Synthesis of distributed mobile programs using monadic types in Coq Marino Miculan Marco Paviotti Dept. of Mathematics and Computer Science University of Udine ITP 2012 August 13th, 2012 1 / 22

The problem The extraction of certified functional and effect-free programs is a well-know practice in the field of Type Theory, however: There are many other computational effects (and corresponding Type Theories, possibly) These scenarios would greatly benefit from a mechanisms for extraction Languages implementing these aspects usually do not support the Curry-Howard isomorphism Implementing a specific proof-assistant would be a daunting task anyway. 2 / 22

Our contribution We propose: a general methodology for circumventing this problem using the existing technology (Coq) + encapsulate non-functional aspects in monadic types + implement a post-extraction compiler for realizing monadic constructors in the target language example: distributed programs with effects in Erlang. 3 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations Use the Extraction facility to get the functional code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations Use the Extraction facility to get the functional code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations Use the Extraction facility to get the functional code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations Use the Extraction facility to get the functional code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations Use the Extraction facility to get the functional code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations Use the Extraction facility to get the functional code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations Use the Extraction facility to get the functional code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations E Use the Extraction facility to get the functional code Target Code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations E Use the Extraction facility to get the functional code Target Code 4 / 22

A general methodology Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations E Use the Extraction facility to get the functional code Target Code 4 / 22

Extraction of distributed code We define the distributed monad in the Calculus of Inductive Constructions forall w: World and A: IO w A :Set 5 / 22

Extraction of distributed code We define the distributed monad in the Calculus of Inductive Constructions A computation localized on the specified host forall w: World and A: IO w A :Set 5 / 22

Extraction of distributed code We define the distributed monad in the Calculus of Inductive Constructions forall w: World and A: IO w A :Set By Curry-Howard Isomorphism, the (constructive) proofs of these specifications are turned into decorated Haskell code IO w A Extraction H 5 / 22

Extraction of distributed code We define the distributed monad in the Calculus of Inductive Constructions forall w: World and A: IO w A :Set By Curry-Howard Isomorphism, the (constructive) proofs of these specifications are turned into decorated Haskell code IO w A Extraction These decorations are exploited by the Haskell-Erlang Compiler H E : H E 5 / 22

Extraction of distributed code We define the distributed monad in the Calculus of Inductive Constructions forall w: World and A: IO w A :Set By Curry-Howard Isomorphism, the (constructive) proofs of these specifications are turned into decorated Haskell code IO w A Extraction These decorations are exploited by the Haskell-Erlang Compiler H E : H E Haskell annotated code Erlang distributed code 5 / 22

Monads in Coq We define a family of monads indexed by worlds from Set to Set. Given a world w a monad is a functor defined as IO w A = S ((R w A) + Error) 6 / 22

Monads in Coq We define a family of monads indexed by worlds from Set to Set. Given a world w a monad is a functor defined as IO w A = S ((R w A) + Error) A Localized computation 6 / 22

Monads in Coq We define a family of monads indexed by worlds from Set to Set. Given a world w a monad is a functor defined as IO w A = S ((R w A) + Error) Function Space from the global store to the results plus the error state 6 / 22

Monads in Coq We define a family of monads indexed by worlds from Set to Set. Given a world w a monad is a functor defined as IO w A = S ((R w A) + Error) Monadic Operators IOget w A : IO remote A IO w A λ κ σ. κ(σ) (Operator s implementation) 6 / 22

Monads in Coq We define a family of monads indexed by worlds from Set to Set. Given a world w a monad is a functor defined as IO w A = S ((R w A) + Error) Monadic Operators IOget w A : IO remote A IO w A λ κ σ. κ(σ) (Operator s implementation) Other monadic operators IOreturn w A : A IO w A IObind w A B : IO w A (A IO w B) IO w B IOlookup w A : Ref w (N IO w A) IO w A IOupdate w A : Ref w N IO w A IO w A IOnew w A : N (Ref w IO w A) IO w A 6 / 22

Extraction Lemma (Remote Procedure Call) w w, (N IO w bool) (N IO w bool) Proof. simpl; introv f. intro n. apply* IOget. Qed. Haskell rpc w w f n = ioget w w (f n) 7 / 22

Extraction Lemma (Remote Procedure Call) w w, (N IO w bool) (N IO w bool) Proof. simpl; introv f. intro n. apply* Given IOget. two worlds w w Qed. Haskell rpc w w f n = ioget w w (f n) 7 / 22

Extraction Lemma (Remote Procedure Call) w w, (N IO w bool) (N IO w bool) Proof. simpl; introv f. intro n. apply* IOget. Qed. Given a function f Haskell rpc w w f n = ioget w w (f n) 7 / 22

Extraction Lemma (Remote Procedure Call) w w, (N IO w bool) (N IO w bool) Proof. simpl; introv f. intro n. apply* IOget. Qed. Given a value, say n Haskell rpc w w f n = ioget w w (f n) 7 / 22

Extraction Lemma (Remote Procedure Call) w w, (N IO w bool) (N IO w bool) Proof. simpl; introv f. intro n. apply* IOget. Qed. Haskell rpc w w f n = ioget w w (f n) Apply IOget to (f n) (World parameters are inferred) 7 / 22

Extraction Lemma (Remote Procedure Call) w w, (N IO w bool) (N IO w bool) Proof. simpl; introv f. intro n. apply* IOget. Qed. Haskell rpc w w f n = ioget w w (f n) Notice: The annotated Haskell code is not runnable yet! 7 / 22

The HEC Compiler: the mobility fragment M ioget A 1 A 2 (F A 3 ) ρ = spawn(element(2, E A 1 ρ ), ρ(η), dispatcher,[fun () -> E F ρ E A 3 ρ end, E A 2 ρ, {self(), node()}]), receive{result, Z} -> Z Erlang Primitives remind: Pid = spawn (Host, Module, Function, Parameters) (Code Mobilty) receive {Pattern} -> Expression (Receive) Pid! Expression (Send) 8 / 22

The HEC Compiler: the location fragment New M ionew A 1 A 2 A 3 ρ = (E A 3 ρ )(spawn(element(2, E A 1 ρ ), ρ("module name"), location, [E A 2 ρ ])) Update M ioupdate w A 1 A 2 A 3 ρ = E A 1 ρ!{update, E A 2 ρ }, E A 3 ρ Lookup M iolookup w A 1 A 2 ρ = E A 1 ρ!{get, {self(), node()}}, receive{result, Z} (E A 2 ρ )(Z) Erlang Primitives remind: Pid = spawn (Host, Module, Function, Parameters) (Code Mobilty) receive {Pattern} -> Expression (Receive) Pid! Expression (Send) 9 / 22

The rpc example Haskell rpc w w f n = ioget w w (f n) Erlang spawn(element(2, w), ρ(η), dispatcher, [fun () -> f(n) end, w, {self(), node()}]), receive{result, Z} -> Z 10 / 22

The rpc example Haskell rpc w w f n = ioget w w (f n) Erlang spawn(element(2, w), ρ(η), dispatcher, [fun () -> f(n) end, w, {self(), node()}]), receive{result, Z} -> Z M : H E Haskell annotations are exploited by the HEC compiler 10 / 22

The rpc example Haskell rpc w w f n = ioget w w (f n) W Erlang (3) Result, f(n) (2) Execute! λ. f(n) spawn(element(2, w), ρ(η), dispatcher, [fun () -> f(n) end, w, {self(), node()}]), receive{result, Z} -> Z Main (1) Dispatch! λ. f(n) W 10 / 22

Adding a Type Theory Step 2: Define a Monad in Coq covering the computational aspects Type Theory with effects < > CiC Sets + Monadic Types Extraction Step 1: Encode a given Type theory in Coq Step 3: Prove the soundness theorem Step 4: Define the translation function E Functional Code (e.g Haskell) with computational annotations E Target Code Use the Extraction facility to get the functional code 11 / 22

λ XD : A Type Theory for distributed computations We give a Type theory similar to the Intuitionistic Hybrid Modal Logic of Licata and Harper 1 Syntax (Hybrid Modal Logic IS5) Types A ::= Unit Bool Nat A B A B Ref (Locations) w.a (Necessarily A) w.a (Possibly A) A@w (Nominal) A (Lax Modality) 1 A monadic formalization of ML5. In Proc. LFMTP, EPTCS 34, 2010. 12 / 22

λ XD : Type System A is inhabited with M at world w Γ M : A[w] (Judgment) Environment Term Type World Fragment Monadic Fragment 13 / 22

λ XD : Type System A is inhabited with M at world w Γ M : A[w] (Judgment) Environment Term Type World Fragment Γ M : A[w] Γ some w M : w.a[w] (some) (A small subset) Intuitionistic a world w as Witness a Proof of A(w) Γ M : u.a[w] Γ, x : A{z/u}[w] N : C[w ] z fresh in Γ Γ letd (z, x) = M in N : C[w ] (letd) World Fragment Monadic Fragment 13 / 22

λ XD : Type System A is inhabited with M at world w Γ M : A[w] (Judgment) Environment Term Type Monadic Fragment (A small subset) Mobile A Γ M : A[w ] Γ get w M : A[w] (Get) World Fragment Monadic Fragment 13 / 22

Soundness We restrict the mobile types which are movable: b {Unit, Nat, Bool} Mobile A Mobile B (Basic) (Prod) Mobile b Mobile A B Mobile A@w (At) Mobile A Mobile w.a (Forall) Mobile A Mobile w.a (Exists) Lemma (Mobility) A Type, if Mobile A, then w, w W, A <w> = A <w >. [Coq proof] 14 / 22

λ XD : Translation into CIC Sets <w> : Type λxd Type Unit <w> = unit Nat <w> = nat A B <w> = A <w> * B <w> w.a <w> = forall w, (A w ) <w> Bool <w> = bool Ref <w> = ref w A@w <w> = A <w > Theorem (Soundness) A B <w> = A <w> -> B <w> w.a <w> = { w : world & (A w ) <w> } A <w> = IO w (A <w>) Let Γ Ctxt, t Term, A Type, w World, let {w 1,..., w n } be all free worlds in Γ, A, t, w. Γ XD t : A[w] w 1 : W,..., w n : W : Γ A <w>. [Coq proof] 15 / 22

Example: Remote Write (W1) Dispatcher (2) Execute Lemma (Remote Write) w w, N@w XD (Ref @w ) <w > Proof. simpl; introv value. apply IOget with (remote := w2). apply IOnew. exact value. intros address. apply IOret. exact address. Defined. (1) Spawn Main (5) Pid W2 Executer (3) Value (4) Pid Location(X) The host W2 contains two process, the spawned location and the executer 16 / 22

Example: Remote Read (W1) Dispatcher (2) Execute Lemma (Remote Read) w w, Ref @w XD (N) <w > Proof. simpl; introv address. apply IOget with (remote := w2). apply* IOlookup. intros. apply* IOret. Defined. (1) Spawn Main (5) Value W2 Executer (3) get Location(X) (4) value The location is already present in W2 and replies to set and get requests 17 / 22

Update Lemma A suitable lemma can be stated and proved in order to guarantee the system behaves correctly Lemma (Update operation is correct) w w, N, s : Store, getvalue ((IObind ((@remotewrite w w value )) (λ l. (@remoteread w w l))) s) = Some value. [Coq proof] 18 / 22

Update lemma: proof deals with monadic constructors Proof is by unfolding the Definitions of the monadic operators 19 / 22

Conclusions Summary We presented a generic methodology for code extraction which works in principle with every computational aspect + Define a front-end type theory on top (optional) + Define a IO monad over Set + Implement the back-end compiler We shown how to extract distributed code by defining a Distributed Monad IO w A We defined the λ XD on top 20 / 22

Conclusions Future work Other computational aspects / target languages is the compiler correct? Compiler Certification could be achieved by using the implementation of monadic operators as the specification of how the target code must behave Alternatively, an Axiomatization for the operators is needed + which takes into account the mobility + not available yet (AFAIK) + but could be achived extending similar work (see e.g., Power and Plotkin Axioms ) 21 / 22

Thanks for your attention. Questions? 22 / 22

Appendix: λ XD Typesystem - World Fragment Γ M : A[w ] w fresh in Γ Γ Λw.M : w.a[w (Box) ] Γ M : w.a[w] Γ unbox w M : A[w] (Unbox) Γ x : A[w] Γ Γ x : A[w] (Var) Γ M : A[w] Γ some w M : w.a[w] (Some) Γ M : u.a[w] Γ, x : A{z/u}[w] N : C[w ] z fresh in Γ Γ letd (z, x) = M in N : C[w ] (LetD) Γ M : A[w] Γ hold M : A@w[w ] (Hold) Γ M : A@z[w] Γ, x : A[z] N : C[w] Γ leta x = M in N : C[w] (LetA) Return 23 / 22

Appendix: λ XD Typesystem - Monadic fragment Γ M : A[w] Γ return M : A[w] (Ret) Γ M : A[w] Γ N : A B[w] ( Γ bind M N : B[w] Mobile A Γ M : A[w ] Γ get w M : A[w] (Get) Γ M : Nat[w] Γ N : Ref A[w] Γ new M N : A[w] (New) Γ M : Ref [w] Γ N : Nat A Γ lookup M N : A[w] (Lookup) Γ T 1 : Ref [w] Γ T 2 : Nat[w] Γ T 3 : A[w] Γ update T 1 T 2 T 3 : A[w] (Update) Return 24 / 22

Appendix: On the axiomatization Rules for the lookup and update are borrowed from Plotkin and Power 1 : 1.l loc (u loc,v (x))v = x 2.l loc (l loc (t vv ) v ) v = l loc (t vv ) v 3.u loc,v (u loc,v (x)) = u loc,v (x) 4.u loc,v (l loc (t v ) v ) = u loc,v (t v ) 5.l loc (l loc (t vv ) v ) v = l loc (l loc(t vv ) v ) v where loc = loc 6.u loc,v (u loc,v (x)) = u loc,v (u loc,v (x))where loc = loc Return 1 G. D. Plotkin and J. Power. Notions of computation determine monads. In Proc. FoSSaCS, LNCS 2303, 2002. 25 / 22