Type checking by theorem proving in IDRIS

Similar documents
Idris: Implementing a Dependently Typed Programming Language

Idris. Programming with Dependent Types. Edwin Brady University of St Andrews, Scotland,

Embedded Domain Specific Language Implementation using Dependent Types

Idris, a language with dependent types Extended Abstract

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

Embedded Domain Specific Languages in Idris

Type Driven Development in Idris

EDIT-TIME TACTICS IN IDRIS

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

First-Class Type Classes

Lecture 8: Summary of Haskell course + Type Level Programming

Haskell Scripts. Yan Huang

Overloading, Type Classes, and Algebraic Datatypes

CS 320: Concepts of Programming Languages

CS 11 Haskell track: lecture 1

Type-driven Development of Communicating Systems in Idris

Generic Constructors and Eliminators from Descriptions

The Idris Tutorial. Release The Idris Community

Documentation for the Idris Language. Version 1.0

Advanced Type System Features Tom Schrijvers. Leuven Haskell User Group

Side Effects (3B) Young Won Lim 11/20/17

Side Effects (3A) Young Won Lim 1/13/18

Type Processing by Constraint Reasoning

Side Effects (3B) Young Won Lim 11/23/17

Side Effects (3B) Young Won Lim 11/27/17

Ideas over terms generalization in Coq

Advanced Functional Programming

Provably Correct Software

Reasoning with the HERMIT

Programming in IDRIS: A Tutorial

Documentation for the Idris Language. Version 0.99

Ivor, a Proof Engine

The Idris Tutorial. Release The Idris Community

Functional Programming and Modeling

CS 320 Midterm Exam. Fall 2018

Encoding functions in FOL. Where we re at. Ideally, would like an IF stmt. Another example. Solution 1: write your own IF. No built-in IF in Simplify

Coq projects for type theory 2018

Dependently Typed Functional Programming with Idris

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

Lightweight Invariants with Full Dependent Types

Alonzo a Compiler for Agda

Equations: a tool for dependent pattern-matching

Background Operators (1E) Young Won Lim 7/7/18

Simon Peyton Jones Microsoft Research August 2013

Resource-dependent Effects for Communication Protocols

Types and Type Inference

Lecture #23: Conversion and Type Inference

Monad Background (3A) Young Won Lim 11/18/17

Overview. Declarative Languages D7012E. Overloading. Overloading Polymorphism Subtyping

Introduction to Haskell

Embedded Domain Specific Languages in Idris Lecture 3: State, Side Effects and Resources

Introduction to Programming, Aug-Dec 2006

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

Haskell Overview II (2A) Young Won Lim 8/9/16

Dependent Polymorphism. Makoto Hamana

Types and Type Inference

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

Monad Background (3A) Young Won Lim 11/8/17

Advanced features of Functional Programming (Haskell)

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

Lecture 19: Functions, Types and Data Structures in Haskell

CSc 372 Comparative Programming Languages

Data Types The ML Type System

IA014: Advanced Functional Programming

Polymorphism Overview (1A) Young Won Lim 2/20/18

CSc 372. Comparative Programming Languages. 8 : Haskell Function Examples. Department of Computer Science University of Arizona

Type families and data kinds

RSL Reference Manual

Lambda Calculus and Type Inference

Structures in Coq January 27th 2014

G Programming Languages - Fall 2012

Background Type Classes (1B) Young Won Lim 6/28/18

Isabelle s meta-logic. p.1

CSCE 314 Programming Languages

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

Programming with Universes, Generically

Type Systems, Type Inference, and Polymorphism

Shell CSCE 314 TAMU. Higher Order Functions

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

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

Fundamentals of Prolog

3 Pairs and Lists. 3.1 Formal vs. Informal Proofs

3. Functional Programming. Oscar Nierstrasz

Principles of Programming Languages

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

Introduction to dependent types in Coq

CS 457/557: Functional Languages

Lecture C-10: Parser Combinators - Introduction

CS 360: Programming Languages Lecture 10: Introduction to Haskell

Haskell: From Basic to Advanced. Part 2 Type Classes, Laziness, IO, Modules

An introduction introduction to functional functional programming programming using usin Haskell

Principles of Programming Languages

Correct-by-Construction Concurrency: using Dependent Types to Verify Implementations of Effectful Resource Usage Protocols

The Typed Racket Guide

CS 320: Concepts of Programming Languages

Functional Logic Programming. Kristjan Vedel

Simple Unification-based Type Inference for GADTs

An Introduction to Programming and Proving in Agda (incomplete draft)

Autosubst Manual. May 20, 2016

Transcription:

Type checking by theorem proving in IDRIS p. 1 Type checking by theorem proving in IDRIS Scottish Theorem Proving, 10th February 2012 ecb10@st-andrews.ac.uk University of St Andrews Edwin Brady

Type checking by theorem proving in IDRIS p. 2 Introduction This talk is about IDRIS, a programming language with dependent types. cabal update; cabal install idris http://idris-lang.org/ http://idris-lang.org/documentation/ In particular: How I tried to write a type checker, and accidentally wrote a theorem prover instead

Type checking by theorem proving in IDRIS p. 3 The IDRIS Programming Language IDRIS is a general purpose pure functional programming language, with support for theorem proving. Features include: Full dependent types, dependent pattern matching Dependent records Type classes (Haskell style) Numeric overloading, Monads, do-notation, idiom brackets,... Tactic based theorem proving High level constructs: where, case, with, monad comprehensions, syntax overloading Totality checking, cumulative universes Interfaces for systems programming (e.g. C libraries)

Type checking by theorem proving in IDRIS p. 4 Dependent Types in IDRIS IDRIS syntax is influenced by Haskell. Some data types: data Nat = O S Nat infixr 5 :: -- Define an infix operator data Vect : Set -> Nat -> Set where -- List with size Nil : Vect a 0 (::) : a -> Vect a k -> Vect a (1 + k)

Type checking by theorem proving in IDRIS p. 4 Dependent Types in IDRIS IDRIS syntax is influenced by Haskell. Some data types: data Nat = O S Nat infixr 5 :: -- Define an infix operator data Vect : Set -> Nat -> Set where -- List with size Nil : Vect a 0 (::) : a -> Vect a k -> Vect a (1 + k) For example: S (S (S O)) : Nat Nil : Vect a 0 2 :: 3 :: 5 :: Nil : Vect Int 3 -- overloaded literal

Type checking by theorem proving in IDRIS p. 4 Dependent Types in IDRIS IDRIS syntax is influenced by Haskell. Some data types: data Nat = O S Nat infixr 5 :: -- Define an infix operator data Vect : Set -> Nat -> Set where -- List with size Nil : Vect a 0 (::) : a -> Vect a k -> Vect a (1 + k) The type of a function over vectors describes invariants of the input/output lengths, e.g.. append : Vect a n -> Vect a m -> Vect a (n + m) append Nil ys = ys append (x :: xs) ys = x :: append xs ys

Type checking by theorem proving in IDRIS p. 5 Dependent Types in IDRIS We can use Haskell style type classes to constrain polymorphic functions, e.g., pairwise addition a vectors of numbers: total vadd : Num a => Vect a n -> Vect a n -> Vect a n vadd Nil Nil = Nil vadd (x :: xs) (y :: ys) = x + y :: vadd xs ys (Aside: The total keyword means that it is an error if the totality checker cannot determine that the function is total)

Type checking by theorem proving in IDRIS p. 6 The Core Language, TT High level IDRIS programs elaborate to a core language, TT: TT allows only data declarations and top level pattern matching definitions Limited syntax: Variables, application, binders (λ,, let, patterns), constants All terms fully explicit Advantage: type checker is small ( 500 lines) so less chance of errors Challenge: how to build TT programs from IDRIS programs?

Type checking by theorem proving in IDRIS p. 7 Elaboration example Consider again pairwise addition: vadd : Num a => Vect a n -> Vect a n -> Vect a n vadd Nil Nil = Nil vadd (x :: xs) (y :: ys) = x + y :: vadd xs ys

Type checking by theorem proving in IDRIS p. 7 Elaboration example Consider again pairwise addition: vadd : Num a => Vect a n -> Vect a n -> Vect a n vadd Nil Nil = Nil vadd (x :: xs) (y :: ys) = x + y :: vadd xs ys Step 1: Add implicit arguments vadd : (a : _) -> (n : _) -> Num a -> Vect a n -> Vect a n -> Vect a n vadd c (Nil _) (Nil _) = (Nil _) vadd c ((::) x xs) ((::) y ys) = (::) ((+) _ x y) (vadd _ xs ys)

Type checking by theorem proving in IDRIS p. 7 Elaboration example Consider again pairwise addition: vadd : Num a => Vect a n -> Vect a n -> Vect a n vadd Nil Nil = Nil vadd (x :: xs) (y :: ys) = x + y :: vadd xs ys Step 2: Solve implicit arguments vadd : (a : Set) -> (n : Nat) -> Num a -> Vect a n -> Vect a n -> Vect a n vadd a O c (Nil a) (Nil a) = (Nil a) vadd a (S k) c ((::) a k x xs) ((::) a k y ys) = (::) a k ((+) c x y) (vadd a k c xs ys)

Type checking by theorem proving in IDRIS p. 7 Elaboration example Consider again pairwise addition: vadd : Num a => Vect a n -> Vect a n -> Vect a n vadd Nil Nil = Nil vadd (x :: xs) (y :: ys) = x + y :: vadd xs ys Step 3: Make pattern bindings explicit vadd : (a : Set) -> (n : Nat) -> Num a -> Vect a n -> Vect a n -> Vect a n pat a : Set, c : Num a. vadd a O c (Nil a) (Nil a) = (Nil a) pat v : Set, k : Nat, c : Num a. pat x : a, xs : Vect a k, y : a, ys : Vect a k. vadd a (S k) c ((::) a k x xs) ((::) a k y ys) = (::) a k ((+) c x y) (vadd a k c xs ys)

Type checking by theorem proving in IDRIS p. 8 Implementing Elaboration IDRIS programs may contain several high level constructs not present in TT: Implicit arguments, type classes where clauses, with and case structures, pattern matching let,... Types often left locally implicit We want the high level language to be as expressive as possible, while remaining translatable to TT.

Type checking by theorem proving in IDRIS p. 9 An observation Consider Coq/Lego style theorem proving (with tactics) and Agda style (by pattern matching). Pattern matching is a convenient abstraction for humans to write programs Tactics are a convenient abstraction for building programs by refinement i.e. explaining programming to a machine Idea: High level program structure directs tactics to build TT programs by refinement The Elaborator is implemented as an Embedded Domain Specific Language in Haskell

Type checking by theorem proving in IDRIS p. 10 Implementing Elaboration Proof State The proof state is encapsulated in a monad, Elab, and contains: Current proof term (including holes) Holes are incomplete parts of the proof term (i.e. sub-goals) Sub-goal in focus Global context (definitions) We distinguish terms which have been typechecked from those which have not: RawTerm has not been type checked (and may contain placeholders, _) Term has been type checked (Type is a synonym)

Type checking by theorem proving in IDRIS p. 11 Implementing Elaboration Operations Some primitive operations: Type checking check :: Normalisation normalise :: RawTerm -> Elab (Term, Type) Term -> Elab Term Unification unify :: Term -> Term -> Elab () Querying proof state Get the local environment get_env :: Elab [(Name, Type)] Get the current proof term get_proofterm :: Elab Term

Type checking by theorem proving in IDRIS p. 12 Implementing Elaboration Tactics A tactic is a function which updates a proof state, for example by: Updating the proof term Solving a sub-goal Changing focus For example: focus :: Name -> Elab () claim :: Name -> RawTerm -> Elab () forall :: Name -> RawTerm -> Elab () exact :: RawTerm -> Elab () apply :: RawTerm -> [RawTerm] -> Elab ()

Type checking by theorem proving in IDRIS p. 13 Implementing Elaboration Tactics Tactics can be combined to make more complex tactics By sequencing, with do-notation By combinators: try :: Elab a -> Elab a -> Elab a If first tactic fails, use the second tryall :: [Elab a] -> Elab a Try all tactics, exactly one must succeed Used to disambiguate overloaded names Effectively, we can use the Elab monad to write proof scripts (c.f. Coq s Ltac language)

Type checking by theorem proving in IDRIS p. 14 Elaborating Applications Given an IDRIS application of a function f to arguments args: Type check f Yields types for each argument, ty_i For each arg_i : ty_i, invent a name n_i and run the tactic claim n_i ty_i Apply f to ns For each non-placeholder arg, focus on the corresponding n and elaborate arg. (Complication: elaborating an argument may affect the type of another argument!)

Type checking by theorem proving in IDRIS p. 15 Elaborating Applications For example, recall append append : (a : Set) -> (n : Nat) -> (m : Nat) -> Vect a n -> Vect a m -> Vect a (n + m) To build an application append _ Nil (1 :: 2 :: Nil) do claim a Set ; claim n Nat ; claim m Nat claim xs (Vect a n) ; claim ys (Vect a m) apply (append a n m xs ys) focus xs; elab Nil focus ys; elab (1 :: 2 :: Nil) Elaborating each sub-term (and running apply) also runs the unify operation, which fills in the _

Type checking by theorem proving in IDRIS p. 16 Elaborating Bindings Given a binder and its scope, say (x : S) -> T Check that the current goal type is a Set Create a hole for S claim n_s Set Create a binder with forall x n_s Elaborate S and T

Type checking by theorem proving in IDRIS p. 17 Elaborating Bindings For example, to build (n : Nat) -> Vect Int n do claim n_s Set forall n n_s focus n_s; elab Nat elab (Vect Int n)

Type checking by theorem proving in IDRIS p. 18 Running the Elaborator runelab :: Name -> Type -> Elab a -> Idris a build :: Pattern -> PTerm -> Elab Term Elaboration is type-directed The Idris monad encapsulates system state The Pattern argument indicates whether this is the left hand side of a definition Free variables on the lhs in IDRIS become pattern bindings in TT patbind :: Name -> Elab () convert current goal to pattern binding PTerm is the representation of the high-level syntax

Type checking by theorem proving in IDRIS p. 19 Elaborating Declarations Given a function declaration of the following form: f : S1 ->... -> Sn -> T f x1... xn = e Elaborate the type, and add f to the context Elaborate the lhs Any out of scope names are assumed to be pattern variables Elaborate the rhs in the scope of the pattern variables from the lhs Check that the lhs and rhs have the same type

Type checking by theorem proving in IDRIS p. 20 Elaborating where Given a function declaration with auxiliary local definitions: f : S1 ->... -> Sn -> T f x1... xn = e where f_aux =... Elaborate the lhs of f Lift the auxiliary definitions to top level functions by adding the pattern variables from the lhs Elaborate the auxiliary definitions Elaborate the rhs of f as normal

Type checking by theorem proving in IDRIS p. 21 Adding Type Classes IDRIS has type classes, for example: class Show a where show : a -> String instance Show Nat where show O = "O" show (S k) = "s" ++ show k Type classes translate directly into data types; instances translate directly into functions.

Type checking by theorem proving in IDRIS p. 22 Adding Type Classes This type class translates to: data Show : (a : Set) -> Set where ShowInstance : (show : a -> String) -> Show a show : Show a => a -> String show {{ShowInstance show }} x = show x instanceshownat : Show Nat instanceshownat = ShowInstance show where show : Nat -> String show O = "O" show (S k) = "s" ++ show k

Type checking by theorem proving in IDRIS p. 23 Adding Type Classes Type class constraints are a special kind of implicit argument (c.f. Agda s instance arguments) Ordinary implicit arguments solved by unification Constraint arguments solved by a tactic resolvetc :: Elab () Looks for a local solution first Then looks for globally defined instances May give rise to further constraints

Type checking by theorem proving in IDRIS p. 24 Conclusions Given basic components (type checking, unification, evaluation) we have built an elaborator for a dependently typed language. Building a DSL for elaboration first: flexible, expressive Access to environment, proof-term, ability to create sub-goals,... High level features have proved easy to add Type classes, records, overloading, tactic scripts,... Efficient enough (meaning: outperforms previous version!) Prelude: 3000 lines in 6.5 secs Profiling suggests some easy efficiency gains to be had Don t forget to cabal install idris :-)