First-Class Type Classes

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

Structures in Coq January 27th 2014

Introduction to dependent types in 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 :=

Preuves Interactives et Applications

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

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

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

Embedding logics in Dedukti

User-Defined Algebraic Data Types

λ calculus is inconsistent

Modular implicits for OCaml how to assert success. Gallium Seminar,

Lecture #23: Conversion and Type Inference

Vectors are records, too!

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

Heq: a Coq library for Heterogeneous Equality

Certificates for incremental type checking

Inductive Definitions, continued

Inheritance and Overloading in Agda

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

Theorem Proving Principles, Techniques, Applications Recursion

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

Basic Foundations of Isabelle/HOL

Type checking by theorem proving in IDRIS

Programming with dependent types: passing fad or useful tool?

Programming Languages and Compilers (CS 421)

Mathematics for Computer Scientists 2 (G52MC2)

Calculus of Inductive Constructions

Built-in Module BOOL. Lecture Note 01a

Unifying Nominal and Structural Ad-Hoc Polymorphism

From Types to Sets in Isabelle/HOL

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

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

Higher-Order Logic. Specification and Verification with Higher-Order Logic

Inductive data types

A constructive denotational semantics for Kahn networks in Coq

Idris: Implementing a Dependently Typed Programming Language

An update on XML types

Logical Verification Course Notes. Femke van Raamsdonk Vrije Universiteit Amsterdam

Programming with (co-)inductive types in Coq

Proving OCaml s type system? Jacques Garrigue Nagoya University

Lean 2 Quick Reference

Functional Programming with Isabelle/HOL

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

l e t print_name r = p r i n t _ e n d l i n e ( Name : ^ r. name)

Overloading, Type Classes, and Algebraic Datatypes

Coq in a Hurry. Yves Bertot

Infinite Objects and Proofs 1

Exercise 1 ( = 22 points)

Programming with Math and Logic

Importing HOL-Light into Coq

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

Modules Matter Most. Robert Harper Carnegie Mellon University. MacQueen Fest Chicago May 2012

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.

Course year Typeclasses and their instances

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

Type Processing by Constraint Reasoning

Tracing Ambiguity in GADT Type Inference

On Agda JAIST/AIST WS CVS/AIST Yoshiki Kinoshita, Yoriyuki Yamagata. Agenda

Type Classes and Overloading in Higher-Order Logic

A Coq tutorial for confirmed Proof system users

Overview. Declarative Languages D7012E. Overloading. Overloading Polymorphism Subtyping

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

Harvard School of Engineering and Applied Sciences Computer Science 152

Formalization of Array Combinators and their Fusion Rules in Coq

Simplifying and Improving Qualified Types

Soundness and principality of type inference for structural polymorphism

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

CSE 3302 Programming Languages Lecture 8: Functional Programming

type classes & locales

Type Classes for Mathematics in Type Theory

Equations: a tool for dependent pattern-matching

Advanced Features: Type Classes and Relations

Autosubst Manual. May 20, 2016

Combining Coq and Gappa for Certifying FP Programs

Generalised Recursion in ML and Mixins

Derek Dreyer. WG2.8 Meeting, Iceland July 16-20, 2007

The Next 700 Syntactic Models of Type Theory

CIS 500: Software Foundations

10 Years of Partiality and General Recursion in Type Theory

Goal. CS152: Programming Languages. Lecture 15 Parametric Polymorphism. What the Library Likes. What The Client Likes. Start simpler.

Contents. A Source Code 32 A.1 ASCII versions of mathematical symbols A.2 Definitions of used library functions... 32

CSE-321 Programming Languages 2014 Final

Formal Verification of a Floating-Point Elementary Function

Simon Peyton Jones Microsoft Research August 2013

Three Applications of Strictness in the Efficient Implementaton of Higher-Order Terms

Intro to Haskell Notes: Part 5

Type-indexed functions in Generic Haskell

Haskell Scripts. Yan Huang

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

Types-2. Polymorphism

Lecture 13: Subtyping

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

Typed Lambda Calculus

Ideas over terms generalization in Coq

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

Expr_annotated.v. Expr_annotated.v. Printed by Zach Tatlock

Lecture 19: Functions, Types and Data Structures in Haskell

Exercise 1 ( = 24 points)

Transcription:

First-Class Type Classes Matthieu Sozeau Joint work with Nicolas Oury LRI, Univ. Paris-Sud - Démons Team & INRIA Saclay - ProVal Project Gallium Seminar November 3rd 2008 INRIA Rocquencourt

Solutions for overloading Intersection types: overloading by declaring multiple signatures for a single constant (e.g. CDuce). Bounded quantification and class-based overloading. Overloading circumscribed by a subtyping relation (e.g. structural subtyping à la OCaml).

Solutions for overloading Intersection types: overloading by declaring multiple signatures for a single constant (e.g. CDuce). Bounded quantification and class-based overloading. Overloading circumscribed by a subtyping relation (e.g. structural subtyping à la OCaml). Our objective: Modularity: separate definitions of the specializations The setting is Coq: no intentional type analysis, no latitude on the kernel language!

Making ad-hoc polymorphism less ad hoc class Eq A where (==) :: A A Bool instance Eq Bool where x == y = if x then y else not y

Making ad-hoc polymorphism less ad hoc 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

Parameterized instances instance (Eq A) Eq [A] where [] == [] = True (x : xs) == (y : ys) = x == y && xs == ys == = False

A structuring concept class Num A where (+) :: A A A... class (Num A) Fractional A where (/) :: A A A... class (Fractional A) Floating A where exp :: A A... The MLer point of view A system of modules and functors with sugar for implicit instantiation and functorization.

Motivations Overloading: in programs, specifications and proofs.

Motivations Overloading: in programs, specifications and proofs. A safer Haskell Proofs are part of classes, added guarantees. Better extraction. Class Eq A := eqb : A A bool ; eq eqb : x y, reflects (eq x y) (eqb x y).

Motivations Overloading: in programs, specifications and proofs. A safer Haskell Proofs are part of classes, added guarantees. Better extraction. Class Eq A := eqb : A A bool ; eq eqb : x y, reflects (eq x y) (eqb x y). Extensions: dependent types give new power to type classes. Class Reflexive A (R : relation A) := reflexive : x, R x x.

Outline 1 Type Classes in Coq A cheap implementation Example: Numbers and monads 2 Superclasses and substructures The power of Pi Example: Categories 3 Extensions Dependent classes Logic Programming 4 Summary, Related, Current and Future Work

Ingredients Dependent records: a singleton inductive type containing each component and some projections. 8 / 34

Ingredients Dependent records: a singleton inductive type containing each component and some projections. Implicit arguments: inferring the value of arguments (e.g. types). Definition id {A : Type} (a : A) : A := a. Check (@id : Π A, A A). Check (@id nat : nat nat). 8 / 34

Ingredients Dependent records: a singleton inductive type containing each component and some projections. Implicit arguments: inferring the value of arguments (e.g. types). Definition id {A : Type} (a : A) : A := a. Check (@id : Π A, A A). Check (@id nat : nat nat). Check (@id : nat nat). 8 / 34

Ingredients Dependent records: a singleton inductive type containing each component and some projections. Implicit arguments: inferring the value of arguments (e.g. types). Definition id {A : Type} (a : A) : A := a. Check (@id : Π A, A A). Check (@id nat : nat nat). Check (@id : nat nat). Check (id : nat nat). 8 / 34

Ingredients Dependent records: a singleton inductive type containing each component and some projections. Implicit arguments: inferring the value of arguments (e.g. types). Definition id {A : Type} (a : A) : A := a. Check (@id : Π A, A A). Check (@id nat : nat nat). Check (@id : nat nat). Check (id : nat nat). Check (id 3). 8 / 34

Implementation Parameterized dependent records Class Id (α 1 : τ 1 ) (α n : τ n ) := f 1 : φ 1 ; ; f m : φ m. 9 / 34

Implementation Parameterized dependent records Record Id (α 1 : τ 1 ) (α n : τ n ) := {f 1 : φ 1 ; ; f m : φ m }. 9 / 34

Implementation Parameterized dependent records Record Id (α 1 : τ 1 ) (α n : τ n ) := {f 1 : φ 1 ; ; f m : φ m }. Instances are just definitions of type Id t n. 9 / 34

Implementation Parameterized dependent records Record Id (α 1 : τ 1 ) (α n : τ n ) := {f 1 : φ 1 ; ; f m : φ m }. Instances are just definitions of type Id t n. Custom implicit arguments of projections f 1 : α n : τ n, Id α n φ 1 9 / 34

Implementation Parameterized dependent records Record Id (α 1 : τ 1 ) (α n : τ n ) := {f 1 : φ 1 ; ; f m : φ m }. Instances are just definitions of type Id t n. Custom implicit arguments of projections f 1 : { α n : τ n }, {Id α n } φ 1 9 / 34

Implementation Parameterized dependent records Record Id (α 1 : τ 1 ) (α n : τ n ) := {f 1 : φ 1 ; ; f m : φ m }. Instances are just definitions of type Id t n. Custom implicit arguments of projections f 1 : { α n : τ n }, {Id α n } φ 1 Proof-search tactic with instances as lemmas A : Type, eqa : Eq A? : Eq (list A) 9 / 34

Elaboration with classes, an example (λx y : bool. eqb x y) 10 / 34

Elaboration with classes, an example (λx y : bool. eqb x y) { implicit arguments } (λx y : bool. @eqb (? A : Type) (? eq : Eq? A ) x y) 10 / 34

Elaboration with classes, an example (λx y : bool. eqb x y) { implicit arguments } (λx y : bool. @eqb (? A : Type) (? eq : Eq? A ) x y) { unification } (λx y : bool. @eqb bool (? eq : Eq bool) x y) 10 / 34

Elaboration with classes, an example (λx y : bool. eqb x y) { implicit arguments } (λ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) 10 / 34

Outline 1 Type Classes in Coq A cheap implementation Example: Numbers and monads 2 Superclasses and substructures The power of Pi Example: Categories 3 Extensions Dependent classes Logic Programming 4 Summary, Related, Current and Future Work

Numeric overloading Class Num α := zero : α ; one : α ; plus : α α α. 12 / 34

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. 12 / 34

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. 12 / 34

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 (λ x : nat, x + (1 + 0 + x)). Check (λ x : Z, x + (1 + 0 + x)). 12 / 34

Monad Class Monad (η : Type Type) := unit : {α}, α η α ; bind : {α β}, η α (α η β) η β ; bind unit left : α β (x : α) (f : α η β), bind (unit x) f = f x ; bind unit right : α (x : η α), bind x unit = x ; bind assoc : α β δ (x : η α) (f : α η β) (g : β η δ), bind x (fun a : α bind (f a) g) = bind (bind x f ) g. 13 / 34

Monad Class Monad (η : Type Type) := unit : {α}, α η α ; bind : {α β}, η α (α η β) η β ; bind unit left : α β (x : α) (f : α η β), bind (unit x) f = f x ; bind unit right : α (x : η α), bind x unit = x ; bind assoc : α β δ (x : η α) (f : α η β) (g : β η δ), bind x (fun a : α bind (f a) g) = bind (bind x f ) g. Infix >= := bind (at level 55). Notation x T ; E := (bind T (fun x : E)) (at level 30, right associativity). Notation return t := (unit t) (at level 20). 13 / 34

Definitions Program Instance identity monad : Monad id := unit α a := a ; bind α β m f := f m. 14 / 34

Definitions Program Instance identity monad : Monad id := unit α a := a ; bind α β m f := f m. Section Monad Defs. Context [ mon : Monad η ]. 14 / 34

Definitions Program Instance identity monad : Monad id := unit α a := a ; bind α β m f := f m. Section Monad Defs. Context [ mon : Monad η ]. Definition ap {α β} (f : α β) (x: η α) : η β := a x ; return (f a). Definition join {α} (x : η (η α)) : η α := x >= id. 14 / 34

Proofs Lemma do return eta : α (u : η α), x u ; return x = u. Proof. intros α u. rewrite (eta expansion unit). η : Type Type mon : Monad η α : Type u : η α ============================ u >= unit = u 15 / 34

Proofs Lemma do return eta : α (u : η α), x u ; return x = u. Proof. intros α u. rewrite (eta expansion unit). η : Type Type mon : Monad η α : Type u : η α ============================ u >= unit = u apply bind unit right. Qed. End Monad Defs. 15 / 34

Outline 1 Type Classes in Coq A cheap implementation Example: Numbers and monads 2 Superclasses and substructures The power of Pi Example: Categories 3 Extensions Dependent classes Logic Programming 4 Summary, Related, Current and Future Work

Fields or Parameters? When one doesn t have manifest types and with constraints... Class Functor := A : Type; B : Type; src : Category A ; dst : Category B ;... Class Functor obj obj := src : Category obj ; dst : Category obj ;... Class Functor obj (src : Category obj) obj (dst : Category obj ) :=...??? 17 / 34

Sharing by equalities Definition adjunction (F : Functor) (G : Functor), src F = dst G dst F = src G... Obfuscates the goals and the computations, awkward to use. 18 / 34

Sharing by parameters Class {(C : Category obj, D : Category obj )} Functor :=... Class Functor {(C : Category obj, D : Category obj )} :=... 19 / 34

Sharing by parameters Class {(C : Category obj, D : Category obj )} Functor :=... Class Functor {(C : Category obj, D : Category obj )} :=... Record Functor {obj} (C : Category obj) {obj }(D : Category obj ) :=... 19 / 34

Sharing by parameters Class {(C : Category obj, D : Category obj )} Functor :=... Class Functor {(C : Category obj, D : Category obj )} :=... Record Functor {obj} (C : Category obj) {obj }(D : Category obj ) :=... Definition adjunction [ C : Category obj, D : Category obj ] (F : Functor C D) (G : Functor D C) :=... Uses the dependent product and named, first-class instances. 19 / 34

Implicit Generalization An old convention: the free variables of a statement are implicitly universally quantified. E.g., when defining a set of equations: x + y = y + x x + 0 = 0 x + S y = S (x + y) 20 / 34

Implicit Generalization An old convention: the free variables of a statement are implicitly universally quantified. E.g., when defining a set of equations: x + y = y + x x + 0 = 0 x + S y = S (x + y) We introduce new syntax to automatically generalize the free variables of a given term or binder: Γ (t) : Type Γ Π FV(t)\Γ, t Γ (t) : T : Type Γ λ FV(t)\Γ, t (x i : τ i ) {(y : T )} (x i : τ i ) {(FV(T ) \ x i )} (y : T ) (x i : τ i ) ((y : T )) (x i : τ i ) (FV(T ) \ x i ) (y : T ) 20 / 34

Substructures A superclass becomes a parameter, a substructure is a method which is also an instance. Class Monoid A := monop : A A A ;... Class Group A := grp mon :> Monoid A ;... 21 / 34

Substructures A superclass becomes a parameter, a substructure is a method which is also an instance. Class Monoid A := monop : A A A ;... Class Group A := grp mon : Monoid A ;... Instance grp mon [ Group A ] : Monoid A. 21 / 34

Substructures A superclass becomes a parameter, a substructure is a method which is also an instance. 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. 21 / 34

Outline 1 Type Classes in Coq A cheap implementation Example: Numbers and monads 2 Superclasses and substructures The power of Pi Example: Categories 3 Extensions Dependent classes Logic Programming 4 Summary, Related, Current and Future Work

Category Class Category (obj : Type) (hom : obj obj Type) := morphisms :> a b, Setoid (hom a b) ; id : a, hom a a; compose : {a b c}, hom a b hom b c hom a c; id unit left : ((f : hom a b)), compose f (id b) == f ; id unit right : ((f : hom a b)), compose (id a) f == f ; assoc : a b c d (f : hom a b) (g : hom b c) (h : hom c d), compose f (compose g h) == compose (compose f g) h. Notation x y := (compose y x) (left associativity, at level 40). 23 / 34

Abstract instances Definition opposite (X : Type) := X. Program Instance opposite category {( Category obj hom )} : Category (opposite obj) (flip hom). 24 / 34

Abstract instances Definition opposite (X : Type) := X. Program Instance opposite category {( Category obj hom )} : Category (opposite obj) (flip hom). Class {(C : Category obj hom)} Terminal (one : obj) := bang : x, hom x one ; unique : x (f g : hom x one), f == g. 24 / 34

An abstract proof Definition isomorphic [ Category obj hom ] a b := { f : hom a b & { g : hom b a f g == id b g f == id a } }. Lemma terminal isomorphic [ C : Category obj hom ] : ( Terminal C x Terminal C y isomorphic x y ). Proof. intros. red. do 2 (bang ). split ; apply unique. Qed. 25 / 34

Outline 1 Type Classes in Coq A cheap implementation Example: Numbers and monads 2 Superclasses and substructures The power of Pi Example: Categories 3 Extensions Dependent classes Logic Programming 4 Summary, Related, Current and Future Work

Dependent classes (demo script) Class Reflexive {A} (R : relation A) := refl : Π x, R x x. Print Reflexive. Instance eq refl A : Reflexive (@eq A). Proof. red. apply refl equal. Qed. Instance iff refl : Reflexive iff. Proof. red. tauto. Qed. Goal Π P, P P. Proof. apply refl. Qed. Goal Π A (x : A), x = x. Proof. intros A ; apply refl. Qed. Ltac reflexivity := apply refl. Lemma foo [ Reflexive nat R ] : R 0 0. Proof. intros. reflexivity. Qed. 27 / 34

Boolean formulas Inductive formula := cst : bool formula not : formula formula and : formula formula formula or : formula formula formula impl : formula formula formula. 28 / 34

Boolean formulas 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. 28 / 34

Reification Class Reify (prop : Prop) := reification : formula ; reify correct : interp reification prop. 29 / 34

Reification Class Reify (prop : Prop) := reification : formula ; reify correct : interp reification prop. Check (@reification : Π prop : Prop, Reify prop formula). Implicit Arguments reification [[Reify]]. 29 / 34

Reification Class Reify (prop : Prop) := reification : formula ; reify correct : interp reification prop. Check (@reification : Π 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). 29 / 34

Reification Class Reify (prop : Prop) := reification : formula ; reify correct : interp reification prop. Check (@reification : Π 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)))). 29 / 34

Outline 1 Type Classes in Coq A cheap implementation Example: Numbers and monads 2 Superclasses and substructures The power of Pi Example: Categories 3 Extensions Dependent classes Logic Programming 4 Summary, Related, Current and Future Work

Summary A lightweight and general implementation of type classes, available in Coq v8.2 A type-theoretic explanation and extension of type-classes concepts On top of that: Realistic test-case: a new setoid-rewriting tactic built on top of classes. A system to automatically infer instances by Matthias Puech. 31 / 34

How does it compare to Canonical Structures? Declaration of a Class and instances instead of using implicit coercions + declaration of some canonical structures. Class Coercion (from to : Type) := coerce : from to. 32 / 34

How does it compare to Canonical Structures? Declaration of a Class and instances instead of using implicit coercions + declaration of some canonical structures. Indexing on parameters only but less sensible to the shape of unification problems (simpler to explain!). 32 / 34

How does it compare to Canonical Structures? Declaration of a Class and instances instead of using implicit coercions + declaration of some canonical structures. Indexing on parameters only but less sensible to the shape of unification problems (simpler to explain!). Based on an extensible resolution system instead of recursive unification of head constants. 32 / 34

Ongoing and future work Debugging! Refined, parameterized proof-search (ambiguity checking, mode declarations, discrimination nets... ) Integration with the proof shell: move to open terms Improve extraction and embedding of Haskell programs 33 / 34

The End http://coq.inria.fr/v8.2beta/