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

Similar documents
SOFTWARE VERIFICATION AND COMPUTER PROOF (lesson 1) Enrico Tassi Inria Sophia-Antipolis

Adam Chlipala University of California, Berkeley ICFP 2006

Inductive Definitions, continued

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 :=

Programming with dependent types: passing fad or useful tool?

Coq Summer School, Session 9 : Dependent programs with logical parts. Pierre Letouzey

Introduction to the Calculus of Inductive Definitions

First-Class Type Classes

Induction in Coq. Nate Foster Spring 2018

Ideas over terms generalization in Coq

Numerical Computations and Formal Methods

Introduction to dependent types in Coq

CIS 500: Software Foundations

Intrinsically Typed Reflection of a Gallina Subset Supporting Dependent Types for Non-structural Recursion of Coq

Denotational Semantics. Domain Theory

Introduction to Co-Induction in Coq

Programming with (co-)inductive types in Coq

Why3 where programs meet provers

Improving Coq Propositional Reasoning Using a Lazy CNF Conversion

A Simple Supercompiler Formally Verified in Coq

Inductive data types

CIS 500 Software Foundations. Midterm I. (Standard and advanced versions together) October 1, 2013 Answer key

10 Years of Partiality and General Recursion in Type Theory

Type-directed Syntax Transformations for ATS-style Programs with Proofs

CMSC 330: Organization of Programming Languages. Objects and Abstract Data Types

Structures in Coq January 27th 2014

Outline. Introduction Concepts and terminology The case for static typing. Implementing a static type system Basic typing relations Adding context

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

Inductive data types (I)

Provably Correct Software

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

Theorem Proving Principles, Techniques, Applications Recursion

Clock-directed Modular Code-generation for Synchronous Data-flow Languages

A Coq tutorial for confirmed Proof system users

Coq projects for type theory 2018

CMSC 330: Organization of Programming Languages. Formal Semantics of a Prog. Lang. Specifying Syntax, Semantics

Embedding logics in Dedukti

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

Programming Languages Lecture 14: Sum, Product, Recursive Types

Combining Coq and Gappa for Certifying FP Programs

Combining Static and Dynamic Contract Checking for Curry

Coq with Classes. Matthieu Sozeau. Journées PPS 2011 September 5th 2011 Trouville, France. Project Team πr 2 INRIA Paris

Coq Summer School. Yves Bertot

Preuves Interactives et Applications

MoreIntro.v. MoreIntro.v. Printed by Zach Tatlock. Oct 07, 16 18:11 Page 1/10. Oct 07, 16 18:11 Page 2/10. Monday October 10, 2016 lec02/moreintro.

MoreIntro_annotated.v. MoreIntro_annotated.v. Printed by Zach Tatlock. Oct 04, 16 21:55 Page 1/10

Why. an intermediate language for deductive program verification

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

Deductive Verification in Frama-C and SPARK2014: Past, Present and Future

1 Introduction. 3 Syntax

Semi-Persistent Data Structures

CMSC 330: Organization of Programming Languages. OCaml Imperative Programming

Programming Languages and Techniques (CIS120)

CMSC 330: Organization of Programming Languages. OCaml Imperative Programming

The Substitution Model

CIS 500: Software Foundations

From natural numbers to the lambda calculus

CSCI-GA Scripting Languages

Coq Summer School, Session 2 : Basic programming with numbers and lists. Pierre Letouzey

CIS 500: Software Foundations

Deductive Program Verification with Why3

Simple proofs of simple programs in Why3

Lecture 14: Recursive Types

Programming Languages Lecture 15: Recursive Types & Subtyping

Lecture Notes on Induction and Recursion

Producing All Ideals of a Forest, Formally (Verification Pearl)

A constructive denotational semantics for Kahn networks in Coq

Work-in-progress: "Verifying" the Glasgow Haskell Compiler Core language

Analysis of dependent types in Coq through the deletion of the largest node of a binary search tree

August 5-10, 2013, Tsinghua University, Beijing, China. Polymorphic types

Reasoning About Imperative Programs. COS 441 Slides 10

Introduction to ML. Based on materials by Vitaly Shmatikov. General-purpose, non-c-like, non-oo language. Related languages: Haskell, Ocaml, F#,

Equations: a tool for dependent pattern-matching

CIS 500 Software Foundations. Final Exam. May 3, Answer key

Why3 A Multi-Prover Platform for Program Verification

COS 320. Compiling Techniques

Reminder of the last lecture. Aliasing Issues: Call by reference, Pointer programs. Introducing Aliasing Issues. Home Work from previous lecture

Infinite Objects and Proofs 1

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

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

Deductive Program Verification with Why3, Past and Future

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

CSCI.6962/4962 Software Verification Fundamental Proof Methods in Computer Science (Arkoudas and Musser) Chapter 11 p. 1/38

Testing. Wouter Swierstra and Alejandro Serrano. Advanced functional programming - Lecture 2. [Faculty of Science Information and Computing Sciences]

CS131 Typed Lambda Calculus Worksheet Due Thursday, April 19th

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

Typed Scheme: Scheme with Static Types

Catamorphism Generation and Fusion Using Coq

Continuation Passing Style. Continuation Passing Style

Why3 Where Programs Meet Provers

n n Try tutorial on front page to get started! n spring13/ n Stack Overflow!

Induction and Semantics in Dafny

Induction for Data Types

L3 Programming September 19, OCaml Cheatsheet

CS3110 Spring 2017 Lecture 9 Inductive proofs of specifications

Calculus of Inductive Constructions

Adding native specifications to JML

Pattern Matching and Abstract Data Types

Automated Reasoning. Natural Deduction in First-Order Logic

Transcription:

Coq LASER 2011 Summerschool Elba Island, Italy Christine Paulin-Mohring http://www.lri.fr/~paulin/laser Université Paris Sud & INRIA Saclay - Île-de-France September 2011 Lecture 4 : Advanced functional programming with COQ Example minimal.v, Monads.v Challenges section 5.2 ListChallenges.v

Outline Introduction Basics of COQ system Using simple inductive definitions Functional programming with COQ Recursive functions Partiality and dependent types Simple programs on lists Well-founded induction and recursion Imperative programming

Representing a partial function f : A B defined only on a subset dom of A. extend f in an arbitrary way (need a default value in B) use an option type: Coq? Print option. Inductive option (A : Type) : Type := Some : A -> option A None : option A need to consider the None case explicitely can be hidden with monadic notations Consider the function as a relation of type A B Prop Prove functionality F x y instead of f x no computation

Introducing logic in types Types and properties can be mixed freely. add an explicit precondition to the function: f : x : A, dom x B each call to f a requires a proof p of dom a internally: f a p. partially hide the proof in a subset type: f : {x : A dom x} B internally f (a, p) Coq? Print sig. Inductive sig (A : Type) (P : A -> Prop) : Type := exist : forall x : A, P x -> sig P use high-level tools like Program or type classes to separate programming from solving proof obligations.

Using subset types for specifications Proposition can appear also to restrict images: S : nat {n : nat 0 < n} restrictions can depend on the input: next : n : nat, {m : nat n < m} Dependent types

Other useful dependent types Coq? Print sumbool. Inductive sumbool (A B : Prop) : Set := left : A -> {A} + {B} right : B -> {A} + {B} Coq? Print sumor. Inductive sumor (A : Type) (B : Prop) : Type := inleft : A -> A + {B} inright : B -> A + {B} Coq? Check forall n m, {n <= m}+{m < n}. Coq? Check forall n, { m m < n }+{n=0}. Use objects in sumbool as booleans value: Coq? Check zerop. zerop : forall n : nat, {n = 0} + {0 < n} Coq? Definition choice (n x y:nat) := Coq? if zerop n then x else y. choice is defined

Annotated programs Advantages: Develop program and proof simultaneously. The specification is available each time the program is used. Possibly discovery of a program from a proof. Coq? Goal forall n m, {n <= m} + {m <= n}. Drawbacks: Inside Coq, the program contains proof-terms: printing, reduction can become awful. Some proof-irrelevance mechanism needed (primitive in PVS)

Outline Introduction Basics of COQ system Using simple inductive definitions Functional programming with COQ Recursive functions Partiality and dependent types Simple programs on lists Well-founded induction and recursion Imperative programming

Lists Coq? Require Import List. Coq? Print list. Inductive list (A : Type) : Type := nil : list A cons : A -> list A -> list A For nil: Argument A is implicit and maximally inserted First challenge on computing sum and max proving: sum l length l max l

Outline Introduction Basics of COQ system Using simple inductive definitions Functional programming with COQ Recursive functions Partiality and dependent types Simple programs on lists Well-founded induction and recursion Imperative programming

Well-founded induction and recursion Only structural recursion is accepted But it can be structural on complex inductive definitions Loop: let f n = if p n then n else f (n + 1) Fixpoint for well-founded relations: let f x = t(x, f ) Any call to f y in t is such that y < x for a well-founded relation (no infinite decreasing sequence).

Analysing the loop construction let f n = if p n then n else f (n + 1) Terminates only if there exists m n such that p m = true. See minimal.v

Typing well-founded fixpoint let f x = t(x, f ) To get f : A B we need: t : A (A B) B To ensure calls are done on smaller instances: t : x : A, ( y : A, y < x B) B generalisation to f : x : A, P(x) t : x : A, ( y : A, y < x P(y)) P(x)

Well-founded fixpoint in COQ Coq? Check Fix. Fix : forall (A : Type) (R : A -> A -> Prop), well_founded R -> forall P : A -> Type, (forall x : A, (forall y : A, R y x -> P y) -> P x) -> forall x : A, P x Property of Fix Coq? Check Fix_eq. Fix_eq : forall (A : Type) (R : A -> A -> Prop) (Rwf : well_founded R) (P : A -> Type) (F : forall x : A, (forall y : A, R y x -> P y) -> P x), (forall (x : A) (f g : forall y : A, R y x -> P y), (forall (y : A) (p : R y x), f y p = g y p) -> F x f = F x g) forall x : A, Fix Rwf P F x = F x (fun (y : A) (_ : R y x) => Fix Rwf P F y)

Accessibility How is it possible? The fixpoint is on the proof of well-foundness of the relation. Coq? Print Acc. Inductive Acc (x : A) : Prop := Acc_intro : (forall y : A, R y x -> Acc R y) -> Acc R x Coq? Print Acc_inv. Acc_inv = fun (x : A) (H : Acc R x) => match H with Acc_intro H0 => H0 end : forall (x : A), Acc R x -> forall y : A, R y x -> Acc R y When H : Acc R x, Acc_inv x H is structurally smaller than H. Coq? Print Fix_F. Fix_F (P : A -> Type) (F : forall x : A, (forall y : A, R y x -> P y) -> P x) = fix Fix_F (x : A) (a : Acc R x) {struct a} : P x := F x (fun (y : A) (h : R y x) => Fix_F y (Acc_inv a h))

Summary Fixpoint in COQ are syntactically restricted but expressive High-level tools help define recursive functions and reason on them: Program, Function... Types may contain logical parts. Functions in COQ are computable, may help doing proofs. Extraction of Ocaml or Haskell. Formally proved versions of tools developed this way: Compilers (C, lustre) Verifiers (prover traces, register allocation, LR-automata... ) Kernel of an SMT solver (alt-ergo)

Outline Introduction Basics of COQ system Using simple inductive definitions Functional programming with COQ Recursive functions Partiality and dependent types Simple programs on lists Well-founded induction and recursion Imperative programming

Monads to handle non functional behavior Used in Haskell to simulate imperative features. A monad is given by comp : Type Type, comp A represents computations of type A; return : A comp A (aka unit) return v represents the value v seen as the result of a computation; bind : comp A (A comp B) comp B, bind passes the result of the first computation to the second one. Syntax: do p <- e1; e2 means bind e 1 (fun p e 2 ). See Monads.v

Memory representation The previous monadic approach for states does not consider aliases in programs. Functional representation of memory and adresses: See LinkedLists.v

Representation of memory Coq? Definition adr := option Z. Coq? Definition null : adr := None. Coq? Record node : Type Coq? := mknode { value : nat ; next : adr}. Coq? Definition heap := Z -> option node. Coq? Definition val (h : heap) (a : adr) : option node Coq? := match a with None => None Some z => h z end. The state of the program will be the heap.

Alternative representation of memory More static separation (model à la Burstall-Bornat) Coq? Definition alloc := Z -> bool. Coq? Definition value_m := Z -> nat. Coq? Definition next_m := Z -> adr. Coq? Definition val (h:alloc) (vm:value_m) (nm:next_m) (a:adr) Coq? : option node Coq? := match a with None => None Coq? Some z => Coq? if h z then Some (mknode (vm z) (nm z)) Coq? else None Coq? end.

Frame properties Need to specify logically validity of adresses, separation Separation logic can be encoded See Ynot (Harvard, Morrisett) : a COQ library to reason on imperative programs with separation logic.