Inductive data types

Similar documents
Programming with (co-)inductive types in Coq

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

Inductive data types (I)

Infinite Objects and Proofs 1

Inductive Definitions, continued

CIS 500: Software Foundations

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

Certified Programming with Dependent Types

CS 456 (Fall 2018) Scribe Notes: 2

Induction in Coq. Nate Foster Spring 2018

CIS 500: Software Foundations

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.

4 Programming with Types

3 Pairs and Lists. 3.1 Formal vs. Informal Proofs

Theorem Proving Principles, Techniques, Applications Recursion

CIS 500: Software Foundations

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

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

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

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.

Introduction to Co-Induction in Coq

Mathematics for Computer Scientists 2 (G52MC2)

Programming with dependent types: passing fad or useful tool?

Coq in a Hurry. Yves Bertot

A Simple Supercompiler Formally Verified in Coq

Isabelle s meta-logic. p.1

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

Coq in a Hurry. Yves Bertot. October 2008

LASER 2011 Summerschool Elba Island, Italy Basics of COQ

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

Functional Programming and Modeling

A Coq tutorial for confirmed Proof system users

Coq in a Hurry. Yves Bertot. To cite this version: HAL Id: inria v6

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

Exercise 1 (2+2+2 points)

Introduction to the Coq proof assistant First part

Structures in Coq January 27th 2014

Equations: a tool for dependent pattern-matching

Proving Properties on Programs From the Coq Tutorial at ITP 2015

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

Logical Verification Course Notes. Femke van Raamsdonk Vrije Universiteit Amsterdam

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

Coq in a Hurry. Yves Bertot. HAL Id: inria v5

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

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

Topic 7: Algebraic Data Types

Now that we have keys defined for the maps, we can start talking about maps. The first of these are Total Maps. Total Maps are defined as follows:

Calculus of Inductive Constructions

First-Class Type Classes

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

Coq Summer School. Yves Bertot

Advanced Features: Type Classes and Relations

type classes & locales

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

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

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

Haskell 98 in short! CPSC 449 Principles of Programming Languages

Heq: a Coq library for Heterogeneous Equality

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

Programming with Universes, Generically

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

Why OCaml? Introduction to Objective Caml. Features of OCaml. Applications written in OCaml. Stefan Wehr

Data types for mcrl2

Introduction to Objective Caml

Name: Entry: Gp: 1. CSL 102: Introduction to Computer Science. Minor 2 Mon 21 Mar 2005 VI 301(Gps 1-6)& VI 401(Gps 7-8) 9:30-10:30 Max Marks 40

The Coq Proof Assistant A Tutorial

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

Exercise 1 ( = 22 points)

Coq projects for type theory 2018

L06_exercise.v. L06_exercise.v

Inductive Data Types

The Coq Proof Assistant A Tutorial

Vector Clocks in Coq

Pattern Matching and Abstract Data Types

1 Problem Description

Exercise 1 ( = 18 points)

A constructive denotational semantics for Kahn networks in Coq

Fall 2013 Midterm Exam 10/22/13. This is a closed-book, closed-notes exam. Problem Points Score. Various definitions are provided in the exam.

Automated Reasoning. Natural Deduction in First-Order Logic

10 Years of Partiality and General Recursion in Type Theory

An intuitionistic proof of a discrete form of the Jordan Curve Theorem formalized in Coq with combinatorial hypermaps

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

Bidirectional Certified Programming

Combining Programming with Theorem Proving

Autosubst Manual. May 20, 2016

A CRASH COURSE IN SEMANTICS

Preuves Interactives et Applications

4.5 Pure Linear Functional Programming

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

Induction Schemes. Math Foundations of Computer Science

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

Exercise 1 ( = 24 points)

Denotational Semantics. Domain Theory

A Tutorial on [Co-]Inductive Types in Coq

Overview. A Compact Introduction to Isabelle/HOL. Tobias Nipkow. System Architecture. Overview of Isabelle/HOL

Standard ML. Data types. ML Datatypes.1

Lecture 19: Functions, Types and Data Structures in Haskell

Embedding logics in Dedukti

Transcription:

Inductive data types Assia Mahboubi 9 juin 2010

In this class, we shall present how Coq s type system allows us to define data types using inductive declarations.

Generalities Inductive declarations An arbitrary type as assumed by : Variable T : Type. gives no a priori information on the nature, the number, or the properties of its inhabitants.

Generalities Inductive declarations An inductive type declaration explains how the inhabitants of the type are built, by giving names to each construction rule :

Generalities Inductive declarations An inductive type declaration explains how the inhabitants of the type are built, by giving names to each construction rule : Print bool.

Generalities Inductive declarations An inductive type declaration explains how the inhabitants of the type are built, by giving names to each construction rule : Print bool. Inductive bool : Set := true : bool false : bool.

Generalities Inductive declarations An inductive type declaration explains how the inhabitants of the type are built, by giving names to each construction rule : Print bool. Inductive bool : Set := true : bool false : bool. Print nat.

Generalities Inductive declarations An inductive type declaration explains how the inhabitants of the type are built, by giving names to each construction rule : Print bool. Inductive bool : Set := true : bool false : bool. Print nat. Inductive nat : Set := O : nat S : nat -> nat.

Generalities Inductive declarations An inductive type declaration explains how the inhabitants of the type are built, by giving names to each construction rule : Print bool. Inductive bool : Set := true : bool false : bool. Print nat. Inductive nat : Set := O : nat S : nat -> nat. Each such rule is called a constructor.

Generalities Inductive declarations in Coq Inductive types in Coq can be seen as the generalization of similar type constructions in more common programming languages. They are in fact an extremely rich way of defining data-types, operators, connectives, specifications,... They are at the core of powerful programming and reasoning techniques.

Enumerated types Enumerated types Enumerated types are types which list and name exhaustively their inhabitants. Inductive bool : Set := true : bool false : bool.

Enumerated types Enumerated types Enumerated types are types which list and name exhaustively their inhabitants. Inductive bool : Set := true : bool false : bool. Inductive prevert_enum : Set := one_stone : prevert_enum two_houses : prevert_enum some_flowers : prevert_enum dozen_of_oysters : prevert_enum an_other_racoon : prevert_enum. Labels refer to distinct elements.

Enumerated types Enumerated types : program by case analysis Inspect the enumerated type inhabitants and assign values : Definition my_negb (b : bool) := match b with true => false false => true.

Enumerated types Enumerated types : program by case analysis Inspect the enumerated type inhabitants and assign values : Definition my_negb (b : bool) := match b with true => false false => true. Definition is_prevert_animal (x : prevert_enum) : bool := match x with dozen_of_oyster => true an_other_racoon => true _ => false end.

Enumerated types Enumerated types : program by case analysis Inspect the enumerated type inhabitants and assign values : Definition my_negb (b : bool) := match b with true => false false => true. Definition is_prevert_animal (x : prevert_enum) : bool := match x with dozen_of_oyster => true an_other_racoon => true _ => false end. Eval compute in (is_prevert_animal one_stone).

Enumerated types Enumerated types : program by case analysis Inspect the enumerated type inhabitants and assign values : Definition my_negb (b : bool) := match b with true => false false => true. Definition is_prevert_animal (x : prevert_enum) : bool := match x with dozen_of_oyster => true an_other_racoon => true _ => false end. Eval compute in (is_prevert_animal one_stone). = false : bool

Enumerated types Enumerated types : reason by case analysis Inspect the enumerated type inhabitants and build proofs : Lemma bool_case : forall b : bool, b = true b = false. Proof. intro b. case b. left; reflexivity. right; reflexivity. Qed.

Enumerated types Enumerated types : reason by case analysis Inspect the enumerated type inhabitants and build proofs : Lemma is_prevert_animalp : forall x : prevert_enum, is_prevert_animal x = true -> x = dozen_of_oysters x = an_other_racoon. Proof. (* Case analysis + computation *) intro x; case x; simpl; intro e. (* In the three first cases: e: false = true *) discriminate e. discriminate e. discriminate e. (* Now: e: true = true *) left; reflexivity. right; reflexivity. Qed.

Enumerated types Enumerated types : reason by case analysis Two important tactics, not specific to enumerated types : simpl : makes computation progress (pattern matching applied to a term starting with a constructor) discriminate : allows to use the fact that constructors are distincts : discriminate H : closes a goal featuring a hypothesis H like (H : true = false) ; discriminate : closes a goal like (O <> S n).

Inductive nat : Set := O : nat S : nat -> nat.

Inductive nat : Set := O : nat S : nat -> nat. Inductive list (A : Type) := nil : list A cons : A -> list A -> list A.

Inductive nat : Set := O : nat S : nat -> nat. Inductive list (A : Type) := nil : list A cons : A -> list A -> list A. Base case constructors do not feature self-reference to the type. Recursive case constructors do.

Let us craft new inductive types : Inductive natbintree : Set :=

Let us craft new inductive types : Inductive natbintree : Set := Leaf : nat -> natbintree

Let us craft new inductive types : Inductive natbintree : Set := Leaf : nat -> natbintree Node : nat ->

Let us craft new inductive types : Inductive natbintree : Set := Leaf : nat -> natbintree Node : nat -> natbintree -> natbintree

Let us craft new inductive types : Inductive natbintree : Set := Leaf : nat -> natbintree Node : nat -> natbintree -> natbintree -> natbintree

Let us craft new inductive types : Inductive natbintree : Set := Leaf : nat -> natbintree Node : nat -> natbintree -> natbintree -> natbintree Inductive term : Set := Zero : term One : term Plus : term -> term -> term Mult : term -> term -> term. An inhabitant of a recursive type is built from a finite number of constructor applications.

: program by case analysis We have already seen some examples of such pattern matching : Definition isnottwo x := match x with S (S O) => false _ => true end.

: program by case analysis We have already seen some examples of such pattern matching : Definition isnottwo x := match x with S (S O) => false _ => true end. Definition is_single_nbt (t : natbintree) := match t with Leaf _ => true _ => false end.

: proofs by case analysis Lemma is_single_nbtp : forall t, is_single_nbt t = true -> exists n : nat, t = Leaf n. Proof.

: proofs by case analysis Lemma is_single_nbtp : forall t, is_single_nbt t = true -> exists n : nat, t = Leaf n. Proof. (* We use the possibility to destruct the tree while introducing *) intros [ nleaf nnode t1 t2] h.

: proofs by case analysis Lemma is_single_nbtp : forall t, is_single_nbt t = true -> exists n : nat, t = Leaf n. Proof. (* We use the possibility to destruct the tree while introducing *) intros [ nleaf nnode t1 t2] h. (* First case: we use the available label *) exists nleaf. reflexivity.

: proofs by case analysis Lemma is_single_nbtp : forall t, is_single_nbt t = true -> exists n : nat, t = Leaf n. Proof. (* We use the possibility to destruct the tree while introducing *) intros [ nleaf nnode t1 t2] h. (* First case: we use the available label *) exists nleaf. reflexivity. (* Second case: the test evaluates to false *) simpl in h. discriminate.

: proofs by case analysis Lemma is_single_nbtp : forall t, is_single_nbt t = true -> exists n : nat, t = Leaf n. Proof. (* We use the possibility to destruct the tree while introducing *) intros [ nleaf nnode t1 t2] h. (* First case: we use the available label *) exists nleaf. reflexivity. (* Second case: the test evaluates to false *) simpl in h. discriminate. Qed.

Constructors are injective : Lemma inj_leaf : forall x y, Leaf x = Leaf y -> x = y. Proof. intros x y hlxly. injection hlxly. trivial. Qed.

: structural induction Let us go back to the definition of natural numbers : Inductive nat : Set := O : nat S : nat -> nat.

: structural induction Let us go back to the definition of natural numbers : Inductive nat : Set := O : nat S : nat -> nat. The Inductive keyword means that at definition time, this system geneates an induction principle : nat_ind : forall P : nat -> Prop, P 0 ->

: structural induction Let us go back to the definition of natural numbers : Inductive nat : Set := O : nat S : nat -> nat. The Inductive keyword means that at definition time, this system geneates an induction principle : nat_ind : forall P : nat -> Prop, P 0 -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n

: structural induction To prove that for P : term -> Prop, the theorem forall t : term, P t holds, it is sufficient to :

: structural induction To prove that for P : term -> Prop, the theorem forall t : term, P t holds, it is sufficient to : Prove that the property holds for the base cases : (P Zero) (P One)

: structural induction To prove that for P : term -> Prop, the theorem forall t : term, P t holds, it is sufficient to : Prove that the property holds for the base cases : (P Zero) (P One) Prove that the property is transmitted inductively : forall t1 t2 : term, P t1 -> P t2 -> P (Plus t1 t2) forall t1 t2 : term, P t1 -> P t2 -> P (Mult t1 t2)

: structural induction To prove that for P : term -> Prop, the theorem forall t : term, P t holds, it is sufficient to : Prove that the property holds for the base cases : (P Zero) (P One) Prove that the property is transmitted inductively : forall t1 t2 : term, P t1 -> P t2 -> P (Plus t1 t2) forall t1 t2 : term, P t1 -> P t2 -> P (Mult t1 t2) The type term is the smallest type containing Zero and One, and closed under Plus and Mult.

: structural induction The induction principles generated at definition time by the system allow to : Program by recursion (Fixpoint) Prove by induction (induction)

: program by structural induction We can compute some information on the size of a term : Fixpoint height (t : natbintree) : nat := match t with Leaf _ => 0 Node _ t1 t2 => Max.max (height t1) (height t2) + 1 end.

: program by structural induction We can compute some information on the size of a term : Fixpoint height (t : natbintree) : nat := match t with Leaf _ => 0 Node _ t1 t2 => Max.max (height t1) (height t2) + 1 end. Fixpoint size (t : natbintree) : nat := match t with

: program by structural induction We can compute some information on the size of a term : Fixpoint height (t : natbintree) : nat := match t with Leaf _ => 0 Node _ t1 t2 => Max.max (height t1) (height t2) + 1 end. Fixpoint size (t : natbintree) : nat := match t with Leaf _ => 1

: program by structural induction We can compute some information on the size of a term : Fixpoint height (t : natbintree) : nat := match t with Leaf _ => 0 Node _ t1 t2 => Max.max (height t1) (height t2) + 1 end. Fixpoint size (t : natbintree) : nat := match t with Leaf _ => 1 Node _ t1 t2 => (size t1) + (size t2) + 1 end.

: program by structural induction We can access some information contained in a term : Require Import List. Fixpoint label_at_occ (dflt : nat) (t : natbintree)(u : list bool) := match u, t with nil, _ => (match t with Leaf n => n Node n => n end) b :: tl, t => match t with Leaf _ => dflt Node _ t1 t2 => if b then label_at_occ t2 tl dflt else label_at_occ t1 tl dflt end end.

: proofs by structural induction We have already seen induction at work on nats and lists. Here its goes on binary trees : Lemma le_height_size : forall t : natbintree, height t <= size t. Proof. induction t; simpl. auto. apply plus_le_compat_r. apply max_case. apply (le_trans _ IHt1). apply le_plus_l. apply (le_trans _ IHt2). apply le_plus_r. Qed.

They are also inductive types! Option types A polymorphic (like list) non recursive type : Print option. Inductive option (A : Type) : Type := Some : A -> option A None : option A

They are also inductive types! Option types A polymorphic (like list) non recursive type : Print option. Inductive option (A : Type) : Type := Some : A -> option A None : option A Use it to lift a type to version with default value : Fixpoint olast (A : Type)(l : list A) : option A := match l with nil => None a :: nil => Some a a :: l => olast A l end.

They are also inductive types! Pairs & co A polymorphic (like list) pair construction : Print pair. Inductive prod (A B : Type) : Type := pair : A -> B -> A * B

They are also inductive types! Pairs & co A polymorphic (like list) pair construction : Print pair. Inductive prod (A B : Type) : Type := pair : A -> B -> A * B The notation A * B denotes (prod A B). The notation (x, y) denotes (pair x y) (implicit argument). Check (2, 4). : nat * nat Check (true, 2 :: nil). : bool * (list nat) Fetching the components : Eval compute in (fst (0, true)). = 0 : nat Eval compute in (snd (0, true)). = true : bool

They are also inductive types! Pairs & co Pairs can be nested : Check (0, 1, true). : nat * nat * bool Eval compute in (fst (0, 1, true)). = (0, 1) : nat * nat This can also be adapted to polymorphic n-tuples : Inductive triple (T1 T2 T3 : Type) := Triple T1 -> T2 -> T3 -> triple T1 T2 T3.

They are also inductive types! Record types A record type bundles pieces of data you wish to gather in a single type. Record admin_person := MkAdmin { id_number : nat; date_of_birth : nat * nat * nat; place_of_birth : nat; sex : bool} They are also inductive types with a single constructor!

They are also inductive types! Record types You can access to the fields : Variable t : admin_person. Check (id_number t). : nat Check id_number. fun a : admin_person => let (id_number, _, _, _) := a in id_number : admin_person -> nat In proofs, you can break an element of record type with tactics case/destruct. Warning : this is pure functional programming...