Inductive data types (I)

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

Inductive data types

Induction in Coq. Nate Foster Spring 2018

Inductive Definitions, continued

CIS 500: Software Foundations

Programming with (co-)inductive types in Coq

CIS 500: Software Foundations

Theorem Proving Principles, Techniques, Applications Recursion

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

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.

CIS 500: Software Foundations

Coq Summer School. Yves Bertot

The University of Nottingham

Infinite Objects and Proofs 1

Introduction to Co-Induction in Coq

Programming with dependent types: passing fad or useful tool?

Structures in Coq January 27th 2014

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

A Simple Supercompiler Formally Verified in Coq

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

3 Pairs and Lists. 3.1 Formal vs. Informal Proofs

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

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

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

Introduction to dependent types in Coq

4 Programming with Types

MLW. Henk Barendregt and Freek Wiedijk assisted by Andrew Polonsky. March 26, Radboud University Nijmegen

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

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

Programming Languages and Techniques (CIS120)

Programming Languages and Techniques (CIS120)

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

Coq projects for type theory 2018

Programming Language. Control Structures: Selection (switch) Eng. Anis Nazer First Semester

CS 456 (Fall 2018) Scribe Notes: 2

Coq in a Hurry. Yves Bertot

CS115 - Module 10 - General Trees

Advanced Features: Type Classes and Relations

Inductive prod (A B : Set) : Set := pair : A -> B -> prod A B. Inductive sum (A B : Set) : Set := inl : A -> sum A B inr : B -> sum A B.

Calculus of Inductive Constructions

Lesson 4 Typed Arithmetic Typed Lambda Calculus

Foundations of Computer Science Spring Mathematical Preliminaries

Functional Programming and Modeling

Ideas over terms generalization in Coq

Chapter 2 The Language PCF

Haskell 98 in short! CPSC 449 Principles of Programming Languages

Built-in Module BOOL. Lecture Note 01a

Heq: a Coq library for Heterogeneous Equality

LASER 2011 Summerschool Elba Island, Italy Basics of COQ

Programming Languages Lecture 14: Sum, Product, Recursive Types

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

type classes & locales

Exercise 1 (2+2+2 points)

Equations: a tool for dependent pattern-matching

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

First-Class Type Classes

Types and Programming Languages. Lecture 5. Extensions of simple types

Propositional Calculus: Boolean Functions and Expressions. CS 270: Mathematical Foundations of Computer Science Jeremy Johnson

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

Mathematics for Computer Scientists 2 (G52MC2)

Inductive Data Types

A Coq tutorial for confirmed Proof system users

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

Logical Verification Course Notes. Femke van Raamsdonk Vrije Universiteit Amsterdam

CHAPTER 8. Copyright Cengage Learning. All rights reserved.

Provably Correct Software

Metaprogramming assignment 3

Lists. Michael P. Fourman. February 2, 2010

Handout 2 August 25, 2008

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

CSCE 314 Programming Languages

Certified Programming with Dependent Types

Coq in a Hurry. Yves Bertot. October 2008

Denotational Semantics. Domain Theory

Functional programming with Common Lisp

Lecture 19: Functions, Types and Data Structures in Haskell

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

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

Adam Chlipala University of California, Berkeley ICFP 2006

λ calculus is inconsistent

Functional Programming in Coq. Nate Foster Spring 2018

Bidirectional Certified Programming

From natural numbers to the lambda calculus

Lecture. Tutorial. Term Rewriting Systems. Tuesdays and Thursdays 16:40 18:10 (Rafael Peñaloza) Wednesdays 14:50 16:20 (Marcel Lippmann)

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

Recap. ANSI C Reserved Words C++ Multimedia Programming Lecture 2. Erwin M. Bakker Joachim Rijsdam

Exercise 1 ( = 18 points)

Topic 5: Enumerated Types and Switch Statements

CS 4110 Programming Languages & Logics. Lecture 28 Recursive Types

A Canonical 1 Locally Named Representation of Binding. α -equivalence is identity. Randy Pollack. Masahiko Sato. LFCS, University of Edinburgh

Exercise 1 ( = 22 points)

Proving Properties of Recursive Functions and Data Structures. CS 270 Math Foundations of CS Jeremy Johnson

These notes are intended exclusively for the personal usage of the students of CS352 at Cal Poly Pomona. Any other usage is prohibited without

Data types for mcrl2

COMPUTER SCIENCE TRIPOS

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

Concepts of programming languages

L04_in_class.v. L04_in_class.v. Printed by Zach Tatlock

Mathematics for Computer Scientists 2 (G52MC2)

Transcription:

5th Asian-Pacific Summer School on Formal Methods August 5-10, 2013, Tsinghua University, Beijing, China Inductive data types (I) jean-jacques.levy@inria.fr 2013-8-6 http://sts.thss.tsinghua.edu.cn/coqschool2013 Notes adapted from Assia Mahboubi (coq school 2010, Paris) and Benjamin Pierce (software foundations course, UPenn)

Recap A! B 8x:A, B when x 62 FVar(B)

Recap definition of functions fun x => M notation for anonymous functions functional kernel of Coq is a typed λ-calculus all calculations are finite every Coq term has a unique normal form Enumerated (finite) types

Plan 今天 小菜 一碟 recap recursive types recursive definitions example 1: natural numbers example 2: day lists example 3: binary trees

Recap Coq commands / keywords: - Definition for functions definitions - Check to show types - Compute to show values - Eval compute in to show values - Inductive to define a new data type - match... with for case analysis on constructors - Type set of all types - simpl to compute normal form - reflexivity to conclude with trivial equality - discriminate to conclude with distinct constructors Example neq_on_days : monday <> tuesday. Proof. discriminate. Qed.

typed λ-calculus with recursion (and a bit of arithmetic)

PCF language (1/3) [Plotkin 1975] Terms M, N, P ::= x, y, z,... (variables) λx.m (M as function of x) M ( N ) (M applied to N) n (natural integer constant) M N (arithmetic op) ifz P then M else N (conditionnal) Y (recursion)

PCF language (1/3) [Plotkin 1975] Terms M, N, P ::= x, y, z,... (variables) λx.m (M as function of x) M ( N ) (M applied to N) n (natural integer constant) M N (arithmetic op) ifz P then M else N (conditionnal) Y (recursion) Calculations ( reductions ) ( x.m)(n) M{x := N} m n m n ifz 0 then M else N ifz n+1 then M else N M N Yf f (Yf )

PCF language (2/3) Fact(3) Fact = Y ( f. x. ifz x then 1 else x f (x 1)) Thus following term: ( Fact. Fact(3)) (Y ( f. x. ifz x then 1 else x? f (x 1)))

PCF language (3/3) Some computations terminate, but not all. (normalization, but not strong normalization) Let F = f. x. ifz x then 1 else x? f (x 1). Then ( Fact. Fact(3)) (YF ) ( Fact. Fact(3)) (F (YF )) 6 ( Fact. Fact(3)) (F (F (YF ))) ( Fact. Fact(3)) (F n (YF )) Quite common in usual programming languages In Coq, we do have strong normalization.

PCF language (3/3) In Coq, we do have strong normalization.

Computability Any most general model of computation has non terminating programs. [Kleene, 1950] Coq cannot express all computable functions but the power of Coq typing allows many of them

Recursive data types

Recursive types (1/7) Recursive types (1/6) Inductive nat : Set := O : nat S : nat -> nat. Inductive daylist : Type := nil : daylist cons : day -> daylist -> daylist. Base case constructors do not feature self-reference to the type. Recursive case constructors do. Definition weekend_days := cons saturday (cons sunday nil)).

Recursive types (2/7) 0, 1 = S(0), 2 = S(S(0), 3 = S(S(S(0),... (unary representation) cons tuesday (cons wednesday (cons friday (cons sunday nil))) tuesday wednesday friday sunday

Recursive types (3/7) Recursive types (2/6)... Coq language can handle notations for infix operators. Notation "x :: l" := (cons x l) (at level 60, right associativity). Notation "[ ]" := nil. Notation "[ x,.., y ]" := (cons x.. (cons y nil)..). Notation "x + y" := (plus x y) (at level 50, left associativity). Therefore weekend days can be also written: Definition weekend_days := saturday :: sunday :: nil. or Definition weekend_days := [saturday, sunday]. saturday sunday

Recursive types (4/7) Recursive types (3/6)... with recursive definitions of functions. Fixpoint length (l:daylist) {struct l} : nat := match l with nil => O d :: l => S (length l ) end. Fixpoint repeat (d:day) (count:nat) {struct count} : daylist := match count with O => nil S count => d :: (repeat d count ) end. The decreasing argument is precised as hint for termination. to insure strong normalization

Recursive types (5/7) Recursive types (4/6)... with recursive definitions of functions. Fixpoint app (l1 l2 : daylist) {struct l1} : daylist := match l1 with nil => l2 d :: t => d :: (app t l2) end. Notation "x ++ y" := (app x y) (right associativity, at level 60). Example test_app1: [monday,tuesday,wednesday] ++ [thursday,friday] = [monday,tuesday,wednesday,thursday,friday]. Proof. reflexivity. Qed. Example test_app2: nil ++ [monday,wednesday] = [monday,wednesday]. Proof. reflexivity. Qed. Example test_app3: [monday,wednesday] ++ nil = [monday,wednesday]. Proof. reflexivity. Qed. d1 d2 d3 d4 d5 ++

Recursive types (5/7) Recursive types (4/6)... with recursive definitions of functions. Fixpoint app (l1 l2 : daylist) {struct l1} : daylist := match l1 with nil => l2 d :: t => d :: (app t l2) end. Notation "x ++ y" := (app x y) (right associativity, at level 60). Example test_app1: [monday,tuesday,wednesday] ++ [thursday,friday] = [monday,tuesday,wednesday,thursday,friday]. Proof. reflexivity. Qed. Example test_app2: nil ++ [monday,wednesday] = [monday,wednesday]. Proof. reflexivity. Qed. Example test_app3: [monday,wednesday] ++ nil = [monday,wednesday]. Proof. reflexivity. Qed.

Recursive types (5/7) Recursive types (4/6)... with recursive definitions of functions. Fixpoint app (l1 l2 : daylist) {struct l1} : daylist := match l1 with nil => l2 d :: t => d :: (app t l2) end. Notation "x ++ y" := (app x y) (right associativity, at level 60). Example test_app1: [monday,tuesday,wednesday] ++ [thursday,friday] = [monday,tuesday,wednesday,thursday,friday]. Proof. reflexivity. Qed. Example test_app2: nil ++ [monday,wednesday] = [monday,wednesday]. Proof. reflexivity. Qed. Example test_app3: [monday,wednesday] ++ nil = [monday,wednesday]. Proof. reflexivity. Qed. d1 d2 d3 d4 d5

Recursive types (6/7) Recursive types (5/6)... with recursive definitions of functions. Definition bag := daylist. Definition eq_day (d:day)(d :day) : bool := match d, d with monday, monday tuesday, tuesday wednesday, wednesday => true thursday, thursday friday, friday => true saturday, saturday => true sunday, sunday => true _, _ => false end. Fixpoint count (d:day) (s:bag) {struct s} : nat := match s with nil => 0 h :: t => if eq_day d h then 1 + count d t else count d t end.

Recursive types (7/7) Recursive types (6/6) Exercice 4 Show following propositions: Example test_count1: count sunday [monday, sunday, friday, sunday] = 2. Example test_count2: count sunday [monday, tuesday, friday, friday] = 0. Exercice 5 Define union of two bags of days. Exercice 6 Define add of one day to a bag of days. Exercice 7 Define remove one day from a bag of days. Exercice 8 Define remove all occurences of a day from a bag of days. Exercice 9 Define member to test if a day is member of a bag of days. Exercice 10 Define subset to test if a bag of days is a subset of another bag of days.

Remark on constructors Remark on constructors I Constructors are injective: Lemma inj_succ : forall n m, S n = S m -> n = m. Proof. intros n m H. injection H. easy. Qed. I Constructors are all distinct.

Recap Coq commands / keywords: - Definition for functions definitions - Check to show types - Compute to show values - Eval compute in to show values - Inductive to define a new data type - match... with for case analysis on constructors - Type set of all types - simpl to compute normal form - reflexivity to conclude with trivial equality - discriminate to conclude with distinct constructors - Fixpoint for recursive functions definitions - struct to hint for termination

Other recursive datatypes (1/2) Recursive types and structural induction (8/9) Another recursive type: binary trees. Inductive natbintree : Type := Leaf : nat -> natbintree Node : nat -> natbintree -> natbintree -> natbintree. Abstract Syntax Trees for terms. Inductive term : Set := Zero : term One : term Plus : term -> term -> term Mult : term -> term -> term.

Other recursive datatypes (2/2) Recursive types and structural induction (9/9) Counting leaves and nodes in binary trees. Fixpoint count_leaves (t : natbintree) {struct t} : nat := match t with leaf n => 1 node n t1 t2 => (count_leaves t1) + (count_leaves t2) end. Fixpoint count_nodes (t : natbintree) {struct t} : nat := match t with leaf n => 0 node n t1 t2 => 1 + (count_nodes t1) + (count_nodes t2) end. Exercice 13 Show Lemma leaves_and_nodes : forall t : natbintree, count_leaves t = 1 + count_nodes t.