Structures in Coq January 27th 2014

Similar documents
First-Class Type Classes

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

Advanced Features: Type Classes and Relations

Programming with dependent types: passing fad or useful tool?

Programming with (co-)inductive types in Coq

Inductive data types

Heq: a Coq library for Heterogeneous Equality

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

Induction in Coq. Nate Foster Spring 2018

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

Inductive data types (I)

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

A Coq tutorial for confirmed Proof system users

Introduction to dependent types in Coq

Equations: a tool for dependent pattern-matching

Lecture #23: Conversion and Type Inference

Infinite Objects and Proofs 1

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

Conversion vs. Subtyping. Lecture #23: Conversion and Type Inference. Integer Conversions. Conversions: Implicit vs. Explicit. Object x = "Hello";

Chapter 3 Linear Structures: Lists

Inheritance and Overloading in Agda

Theorem Proving Principles, Techniques, Applications Recursion

CIS 500: Software Foundations

Numerical Computations and Formal Methods

Inductive Definitions, continued

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

Exercise 1 ( = 24 points)

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

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

Coq Summer School. Yves Bertot

Exercise 1 ( = 18 points)

HOL DEFINING HIGHER ORDER LOGIC LAST TIME ON HOL CONTENT. Slide 3. Slide 1. Slide 4. Slide 2 WHAT IS HIGHER ORDER LOGIC? 2 LAST TIME ON HOL 1

Advanced Functional Programming

10 Years of Partiality and General Recursion in Type Theory

The Worker/Wrapper Transformation

The Worker/Wrapper Transformation

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

Chapter 3 Linear Structures: Lists

Programming with Math and Logic

CIS 500: Software Foundations

The Next 700 Safe Tactic Languages

Canonical Structures for the working Coq user

Basic Foundations of Isabelle/HOL

Overview. Declarative Languages D7012E. Overloading. Overloading Polymorphism Subtyping

Exercise 1 ( = 22 points)

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

CIS 500: Software Foundations

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

Embedding logics in Dedukti

LASER 2011 Summerschool Elba Island, Italy Basics of COQ

Coq projects for type theory 2018

Inductive Data Types

User-Defined Algebraic Data Types

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.

Refinements for free! 1

Lecture 19: Functions, Types and Data Structures in Haskell

An introduction introduction to functional functional programming programming using usin Haskell

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

Type checking by theorem proving in IDRIS

Dependent Polymorphism. Makoto Hamana

Built-in Module BOOL. Lecture Note 01a

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

Certified Programming with Dependent Types

1. true / false By a compiler we mean a program that translates to code that will run natively on some machine.

A general introduction to Functional Programming using Haskell

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

9/21/17. Outline. Expression Evaluation and Control Flow. Arithmetic Expressions. Operators. Operators. Notation & Placement

Refinements for free!

Haske k ll An introduction to Functional functional programming using Haskell Purely Lazy Example: QuickSort in Java Example: QuickSort in Haskell

Lecture #13: Type Inference and Unification. Typing In the Language ML. Type Inference. Doing Type Inference

Type families and data kinds

THE AGDA STANDARD LIBRARY

Functional Programming with Isabelle/HOL

MPRI course 2-4 Functional programming languages Exercises

CS 320: Concepts of Programming Languages

Lecture 4: Higher Order Functions

Coq in a Hurry. Yves Bertot

Haskell & functional programming, some slightly more advanced stuff. Matteo Pradella

Course year Typeclasses and their instances

Idris: Implementing a Dependently Typed Programming Language

CSCI.6962/4962 Software Verification Fundamental Proof Methods in Computer Science (Arkoudas and Musser) Sections p.

Structural polymorphism in Generic Haskell

Formal Verification of Monad Transformers

Calculus of Inductive Constructions

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

λ calculus is inconsistent

Vectors are records, too!

From natural numbers to the lambda calculus

Problem Set CVO 103, Spring 2018

Isabelle s meta-logic. p.1

Ornaments in ML. Thomas Williams, Didier Rémy. April 18, Inria - Gallium

Pragmatic Quotient Types in Coq

CSci 450: Org. of Programming Languages Overloading and Type Classes

Ideas over terms generalization in Coq

Pattern Matching and Abstract Data Types

Harvard School of Engineering and Applied Sciences CS 152: Programming Languages

4 Programming with Types

Programming Languages Fall 2013

DECIDING KLEENE ALGEBRAS IN COQ THOMAS BRAIBANT AND DAMIEN POUS

Transcription:

MPRI 2-7-2 Proof assistants ~sozeau/teaching/mpri-2-7-2-270114.pdf! Structures in Coq January 27th 2014 Matthieu Sozeau! "r² project team! Inria Paris & PPS! matthieu.sozeau@inria.fr! www.pps /~sozeau

Announcement Talk by Hendrik Tews, founder of FireEye, Inc. tomorrow at 2pm.! PhD defense of J. Cretin on the 30th, at 2.30pm, room 227C of P7. Erasable coercions: a unified approach to type systems

Structures in Coq Outline 1. Record Types! 2. Mathematical Structures and Coercions! 3. Type Classes and Canonical Structures! 4. Monadic Programming with Type Classes! 5. Interfaces and Implementations

Last time: Inductive Types Least Fixpoints of Polynomial Functors.! Introduction and elimination! Guardedness and strict positivity condition! Dependent elimination and restricted eliminations! Examples: ordinals and W-type.! Questions?

Record Types 1. Record Types! 2. Mathematical Structures and Coercions! 3. Type Classes and Canonical Structures! 4. Monadic Programming using Type Classes! 5. Interfaces and Implementations

Record Types Record R (A : Type) : s := { f1 : τ1;..; fn : τn }. Inductive R (A : Type) : s := Build_R (f1 : τ1).. (fn : τn) : R A.! No indices, no recursion.

Projection is pattern-matching Inductive R (A : Type) : s := Build_R (f1 : τ1).. (fn : τn) : R A.! Definition f_i (A : Type) (r : R A) : τi[f_i r/f_i] := match r with Build_R t1.. tn => ti end.

Definitional Equality By construction:! fi (Build_rec t1.. tn) βδηι ti But no η law:! Build_R (p1 t).. (pn t) t. Only derivable for Leibniz equality.! Exercises (10min):! Define the sigma/dependent sum/pair type as a record.! Prove the η rule (surjective pairing) for it.! Define Q as a record of two integers.! Define the equivalence of two rationals as a proposition.! Definition addition of two rationals and show that it respects the equivalence (on pairwise equivalent arguments, it gives equivalent results).

Mathematical Structures and Coercions 1. Record Types! 2. Mathematical Structures and Coercions! 3. Type Classes and Canonical Structures! 4. Monadic Programming using Type Classes! 5. Interfaces and Implementations

Mathematical Structures Defining composite structures as first-class objects:! Record monoid := { carrier : Type; unit : carrier; op : carrier -> carrier carrier; left_unit : x, op x unit = x; }. Generic programming for any monoid:! Definition monsquare (m : monoid) (x : carrier m) : carrier m := op m x x. Type and value dependency is essential to make this typecheck.!

Mathematical Structures and Coercions Exercise (10min) Define the interface of a monoid given partially above, and the (0, +) and (1, *) monoids on N.! Define a generic exponentiation operation aˆn where n : N and a is in a monoid.! Show that the exponentiation of the unit is always equal to the unit.

Coercions Record monoid := { carrier : Type; unit : carrier; }. Generic programming for any monoid:! Definition monsquare (m : monoid) (x : m) : m := op m x x. Coercion carrier : monoid >-> Sortclass. The term m is coerced to a sort (its carrier) in the types.

Coercions Coercions are not displayed (by default, see Set Printing Coercions)! The actual core term is the same as before! Coherence of coercions (only one path up to between two types) and decidability ensured by a syntactic check.! Uses: types embedded in records, injections (i.e. from N to Z), modeling a subtyping relation! Definition Z_of_nat (n : nat) : Z := Zpos n. Coercion Z_of_nat : nat >-> Z.

Mathematical Structures and Coercions Exercice (10min) Define the coercion from N to B, sending 0 to false and everything else to the true constructor.! Define the converse coercion from B to N, giving value 1 to true.! Show that B N B is the identity (pointwise)! Which one should be a coercion?! Redo the monoid exercise using a coercion for carrier

Type Classes 1. Record Types! 2. Mathematical Structures and Coercions! 3. Type Classes and Canonical Structures! 4. Monadic Programming using Type Classes! 5. Interfaces and Implementations

Type Classes Introduced in Haskell (Wadler & Blott, POPL 89) and in Isabelle (Nipkow & Snelting, FPCA 91).!! class Eq a where (==) :: a a Bool instance Eq Bool where x == y = if x then y else not y in :: Eq a a [a] Bool in x [] = False in x (y : ys) = x == y in x ys

Type Classes Parametrized instances! instance (Eq a) => Eq [a] where [] == [] = True (x : xs) == (y : ys) = x == y && xs == ys _ == _! Super-classes! class Num a where (+) :: a a a = False class (Num a) Fractional a where (=) :: a a a

Type Classes in Coq I Parametrized dependent records Record Id ( 1 : 1 ) ( n : n ) := {f 1 : 1 ; ; f m : m}. Instances are just definitions of conclusion Id! t n. I Custom implicit arguments of projections f 1 : 8{! n : n }, {Id! n }! 1

Elaboration example x y : bool. eqb x y { Implicit arguments } x y : bool. @eqb x y { Typing } x y : bool. @eqb (? A : Type) (? eq : Eq? A ) x y { Unification } x y : bool. @eqb bool (? eq : Eq bool) x y { Proof search for Eq bool returns Eq bool } x y : bool. @eqb bool Eq bool x y

Type Classes Proof-search à la Prolog! An elaboration (no kernel change) (Sozeau and Oury, 2008).

Inheritance Class Monoid A := { monop : A! A! A ;... } Class Group A := { grp mon :> Monoid A ;... } Substructures become subinstances: Class Monoid A := { monop : A! A! A ;... } Class Group A := { grp mon : Monoid A ;... } Instance grp mon {Group A} : Monoid A. Definition foo {Group A} (x : A) :A := monop x x. Similar to the existing Structures based on coercive subtyping.

Example: numeric overloading Class Num := { zero : ; one : ; plus :!! }. Instance nat num : Num nat := { zero := 0%nat ; one := 1%nat ; plus := Peano.plus }. Instance Z num : Num Z := { zero := 0%Z ; one := 1%Z ; plus := Zplus }. Notation 0 := zero. Notation 1 := one. Infix + := plus. Check ( Check ( x : nat, x + (1 + 0 + x)). x : Z, x + (1 + 0 + x)). (* Defaulting *) Check ( x, x + 1).

Example: value-dependent classes Class Reflexive {A} (R : relation A) := refl : 8 x, R x x. Instance eq refl A : Reflexive (@eq A) :=@refl equal A. Instance i refl : Reflexive i. Proof. red. tauto. Qed. Goal 8 P, P $ P. Proof. apply refl. Qed. Goal 8 A (x : A), x = x. Proof. intros A ; apply refl. Qed. Ltac refl := apply refl. Lemma foo {Reflexive nat R} : R 0 0. Proof. intros. refl. Qed.

Example: reification Inductive formula := cst : bool! formula not : formula! formula and : formula! formula! formula or : formula! formula! formula impl : formula! formula! formula. Fixpoint interp f := match f with cst b ) if b then True else False not b ) interp b and a b ) interp a ^ interp b or a b ) interp a _ interp b impl a b ) interp a! interp b end.

Example: reification Class Reify (prop : Prop) := { reification : formula ; reify correct : interp reification $ prop }. Check (@reification : 8 prop : Prop, Reify prop! formula). Implicit Arguments reification [[Reify]]. Program Instance true reif : Reify True := { reification := cst true }. Program Instance not reif (Rb : Reify b) : Reify ( b) := { reification := not (reification b) }. Example example prop := reification ((True ^ False)! False). Check (refl equal : example prop = impl (and (cst true) (not (cst false))) (not (not (cst false)))).

Exercise (30min) Define a reifier for expressions in the monoid class, a reflexive tactic for simplifying monoidal expressions (unit laws), and an overloaded lemma to apply it.

Canonical Structures Another way to use records to represent structures, capable of doing much the same as type classes, but based on unification instead of general proof search. At the basis of the ssreflect plugin.! See Mahboubi and Tassi [ITP 13] for a gentle introduction and comparison, and Ziliani et al [ICFP 11] for the previous examples using both representations.

Monadic Programming with Type Classes 1. Record Types! 2. Mathematical Structures and Coercions! 3. Type Classes and Canonical Structures! 4. Monadic Programming with Type Classes! 5. Interfaces and Implementations

Monadic programming: ML vs Haskell type NatState α = nat α nat val counter : unit nat let counter = let x = ref 0 in fun () => let x =!x in x := x + 1; x return :: α NatState α bind :: NatState α (α NatState β) NatState β get :: NatState nat put :: nat NatState () counter :: NatState nat counter = do x <- get; put (x + 1); return x counter :: () nat counter () =???

Monad In our setting, a monad will be defined as:! Record Monad (M : Type Type) := { return : {A}, A M A; bind : {A B}, M A (A M B) M B; bind_assoc {A B f g} (m : M A) : bind (bind m f) g = bind m (fun x : A bind (f x) g); bind_return_left : ; bind_return_right : }

Monad Instances Instances of the monad interface:! id (identity monad) option (partiality monad) list ( nondeterminism monad, unit is [x], bind is concat map) continuations, computation

Exercises (±20min)! Define Monad as a type class with overloaded bind and return/unit operations! Define monadic multiplication mult of type M (M A) M A for any monad. State and prove what is its relation to the unit.! Define the identity instance.! Define the option instance (return is injection, bind propagates errors represented as None), there is a corresponding error / nothing action to define. To prove the laws, you will need to use the axiom in Coq.Logic.FunctionalExtensionality.! Define the generic mapm operator for any monad (it processes the effects from left to right). Generalize it to a foldm.! Optional: Using this, write a partial function turning an integer to a natural, and a function that takes a list of integers and returns the sum of their positive numbers, if they re all positive only.! Optional: Using the state monad, verify a tree labeling algorithm:!! http://www.pps.univ-paris-diderot.fr/~sozeau/teaching/classes/monad.v

Interfaces and Implementations: Fibonacci 1. Record Types! 2. Mathematical Structures and Coercions! 3. Type Classes and Canonical Structures! 4. Monadic Programming with Type Classes! 5. Interfaces and Implementations

Interfaces vs Implementations Software engineering principle applied to proof engineering:!! Work with interfaces instead of implementations.! An illustrative example:!! Verifying a fast implementation of Fibonacci.

A correct exponentiation algorithm The following definition is very naïve, but obviously correct: Fixpoint power (a : Z) (n : nat) := match n with 0%nat ) 1 S p ) a power a p end. Eval vm compute in power 2 40. = 1099511627776 : Z

An efficient exponentiation algorithm This one is more e cient but relies on a more elaborate property: Function binary power mult (acc x : Z) (n : nat) {measure (fun i)i) n} : Z := match n with 0%nat ) acc ) if Even.even odd dec n then binary power mult acc (x x) (div2 n) else binary power mult (acc x) (x x) (div2 n) end. Definition binary power (x:z) (n:nat) := binary power mult 1 x n. Eval vm compute in binary power 2 40. = 1099511627776 : Z Goal binary power 2 234 = power 2 234. Proof. reflexivity. Qed.

How to proceed? Is binary_power correct (w.r.t. power)?! Is it worth proving this correctness only for powers of integers?! And prove it again for powers of real numbers, matrices?! NO!! Program with interfaces, here a monoid.!

Support for overloading Quantification on parameters: Definition two {A dot one} {M :@Monoid A dot one} := dot one one. Using implicit generalization: Generalizable Variables A dot one. Definition three {Monoid A dot one} := dot two one. Global names for parameters: Definition monop {Monoid A dot one} := dot. Definition monunit {Monoid A dot one} := one. Generic notations: Infix :=monop. Notation 1 := monunit.

A generic exponentiation algorithm Generic power and binary power. Section Power. Context {Monoid A dot one}. Fixpoint power (a : A) (n : nat) := match n with 0%nat ) 1 S p ) a (power a p) end. Lemma power of unit : 8 n : nat, power 1 n = 1. Proof.... Qed.

A generic, efficient exponentiation algorithm Function binary power mult (acc x : A) (n : nat) {measure (fun i)i) n} : A := match n with 0%nat ) acc ) if Even.even odd dec n then binary power mult acc (x x) (div2 n) else binary power mult (acc x) (x x) (div2 n) end. Definition binary power (x : A) (n : nat) := binary power mult 1 x n. Lemma binary spec x n : power x n = binary power x n. Proof.... Qed. End Power.

Exercise (1h, for next week) Get ~sozeau/teaching/mpri-2-7-2-270114-fibonacci.v! Define the monoid instance for multiplication on Z! Define a commutative semiring type class, comprising 2 monoids, one for multiplication and one for addition and the additional laws (commutativity and one distributive law).! Define an instance for Z (you can leave some obligations admitted).! Define square 2x2 α-valued matrices, for α supporting a commutative semiring structure.! Define the multiplicative and additive monoids on these matrices and the semiring structure.! Using the characterization of the Fibonacci sequence below, define Fibonacci in terms of exponentiation of matrices.! Show that this fast version is equivalent to the naive definition of Fibonacci.! Experiment with vm_compute and compute on these different implementations (using power, binary power)