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.

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

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

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

Induction in Coq. Nate Foster Spring 2018

CIS 500: Software Foundations

Theorem Proving Principles, Techniques, Applications Recursion

L06_exercise.v. L06_exercise.v

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

A Simple Supercompiler Formally Verified in Coq

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

CIS 500: Software Foundations

CIS 500: Software Foundations

Functional Programming - 2. Higher Order Functions

Introduction to dependent types in Coq

Inductive data types

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

4 Programming with Types

Inductive Definitions, continued

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.

Infinite Objects and Proofs 1

Induction and Semantics in Dafny

Logic and Computation Lecture 20 CSU 290 Spring 2009 (Pucella) Thursday, Mar 12, 2009

Printed by Zachary Tatlock Nov 04, 16 8:59 Page 1/9. (** Call By Value << v ::= \ x. e e1 > e1 e1 e2 > e1 e2 e2 > e2 v e2 > v e2

Inductive data types (I)

Certified Programming with Dependent Types

GADTs. Wouter Swierstra. Advanced functional programming - Lecture 7. Faculty of Science Information and Computing Sciences

(ii) Define a function ulh that takes a list xs, and pairs each element with all other elements in xs.

3 Pairs and Lists. 3.1 Formal vs. Informal Proofs

Programming Languages Fall 2014

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

2 Programming and Proving

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

Week 5 Tutorial Structural Induction

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:

GADTs. Wouter Swierstra and Alejandro Serrano. Advanced functional programming - Lecture 7. [Faculty of Science Information and Computing Sciences]

Introduction to Functional Programming in Haskell 1 / 56

Functional Programming and Modeling

Printed by Zach Tatlock Dec 08, 15 22:18 Page 1/11. Lemma subst_only_free: forall e1 e2 x e3, Subst e1 e2 x e3 > ~ free e1 x > e1 = e3. Proof.

GADTs. Alejandro Serrano. AFP Summer School. [Faculty of Science Information and Computing Sciences]

IA014: Advanced Functional Programming

Programming and Proving in Isabelle/HOL

CS 456 (Fall 2018) Scribe Notes: 2

CIS 500 Software Foundations Midterm I

Work-in-progress: "Verifying" the Glasgow Haskell Compiler Core language

Exercise 1 ( = 22 points)

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

Isabelle s meta-logic. p.1

Talen en Compilers. Jurriaan Hage , period 2. November 13, Department of Information and Computing Sciences Utrecht University

CS 320: Concepts of Programming Languages

Cornell University 12 Oct Solutions. (a) [9 pts] For each of the 3 functions below, pick an appropriate type for it from the list below.

10 Years of Partiality and General Recursion in Type Theory

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

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

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

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

Induction Schemes. Math Foundations of Computer Science

Provably Correct Software

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

Logic - CM0845 Introduction to Haskell

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

Introduction to Co-Induction in Coq

CSCC24 Functional Programming Scheme Part 2

An introduction introduction to functional functional programming programming using usin Haskell

First-Class Type Classes

Accurate Step Counting

Lecture 4: Higher Order Functions

CS115 - Module 10 - General Trees

Work-in-progress: Verifying the Glasgow Haskell Compiler Core language

Review. CS152: Programming Languages. Lecture 11 STLC Extensions and Related Topics. Let bindings (CBV) Adding Stuff. Booleans and Conditionals

Module Title: Informatics 1 Functional Programming (first sitting) Exam Diet (Dec/April/Aug): December 2014 Brief notes on answers:

Warm-up and Memoization

Lecture 6: Sequential Sorting

Programming with (co-)inductive types in Coq

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

Programming Languages Lecture 14: Sum, Product, Recursive Types

Preuves Interactives et Applications

Lesson 4 Typed Arithmetic Typed Lambda Calculus

CS 320 Homework One Due 2/5 11:59pm

Exercise 1 (2+2+2 points)

CS152: Programming Languages. Lecture 2 Syntax. Dan Grossman Spring 2011

Booleans (aka Truth Values) Programming Languages and Compilers (CS 421) Booleans and Short-Circuit Evaluation. Tuples as Values.

Combining Programming with Theorem Proving

15 212: Principles of Programming. Some Notes on Continuations

A CRASH COURSE IN SEMANTICS

Homework 3 COSE212, Fall 2018

Advanced Functional Programming

Advanced Features: Type Classes and Relations

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

Induction for Data Types

Coq Summer School. Yves Bertot

Australian researchers develop typeface they say can boost memory, that could help students cramming for exams.

Lecture 5: Lazy Evaluation and Infinite Data Structures

Fall 2018 Lecture N

CS1102: Macros and Recursion

CSCE 314 Programming Languages

Accurate Step Counting

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

Shell CSCE 314 TAMU. Higher Order Functions

Pattern Matching and Abstract Data Types

Transcription:

Oct 07, 16 18:11 Page 1/10 * Lecture 02 Set Implicit Arguments. Inductive list (A: Type) : Type := nil : list A cons : A > list A > list A. Fixpoint length (A: Type) (l: list A) : nat := nil _ => O cons x xs => S (length xs) so far, Coq will not infer the type argument for nil: << Check (cons 1 nil). Error: The term "nil" has type "forall A : Type, list A" while it is expected to have type "list nat". >> Check (cons 1 (nil nat)). we can tell Coq to always try though Arguments nil {A}. Check (cons 1 nil). Eval cbv in (length (cons 1 nil)). Eval cbv in (length (cons 2 (cons 1 nil))). Eval cbv in (length nil). Fixpoint countdown (n: nat) := O => cons n nil S m => cons n (countdown m) Eval cbv in (countdown 0). Eval cbv in (countdown 3). Eval cbv in (countdown 10). Fixpoint map (A B: Type) (f: A > B) (l: list A) : list B := cons x xs => cons (f x) (map f xs) Eval cbv in (map (plus 1) (countdown 3)). Eval cbv in (map (fun _ => true) (countdown 3)). Definition is_zero (n: nat) : bool := O => true S m => false Eval cbv in (map is_zero (countdown 3)). Fixpoint is_even (n: nat) : bool := O => true S O => false S (S m) => is_even m Eval cbv in (map is_even (countdown 3)). Oct 07, 16 18:11 Page 2/10 Lemma map_length: forall (A B: Type) (f: A > B) (l: list A), length (map f l) = length l. discuss how does induction works Definition compose (A B C: Type) (f: B > C) (g: A > B) : A > C := fun (x : A) => f (g x). Lemma map_map_compose: forall (A B C: Type) (g: A > B) (f: B > C) (l: list A), map f (map g l) = map (compose f g) l. + (* map f (map g (cons a l)) > map f (cons (g a) (map g l)) > cons (f (g a)) (map f (map g l)) rewrite IHl. need to "unfold" compose to simpl unfold compose. Fixpoint foldr (A B: Type) (f: A > B > B) (l: list A) (b: B) : B := nil => b cons x xs => f x (foldr f xs b) foldr f (cons 1 (cons 2 (cons 3 nil))) x > f 1 (f 2 (f 3 x)) Check plus. Print plus. Eval cbv in (foldr plus (countdown 10) 0). Fixpoint fact (n: nat) : nat := O => 1 S m => mult n (fact m) Eval cbv in (fact 0). Eval cbv in (fact 1). Eval cbv in (fact 2). Eval cbv in (fact 3). Eval cbv in (fact 4). Definition fact (n: nat) : nat := O => 1 S m => foldr mult (map (plus 1) (countdown m)) 1 Monday October 10, 2016 lec02/ 1/5

Oct 07, 16 18:11 Page 3/10 Eval cbv in (fact 0). Eval cbv in (fact 1). Eval cbv in (fact 2). Eval cbv in (fact 3). Eval cbv in (fact 4). Lemma fact_fact : fact n = fact n. challenge problem Admitted. we can also define map using fold Definition map (A B: Type) (f: A > B) (l: list A) : list B := foldr (fun x acc => cons (f x) acc) l nil. Lemma map_map : forall (A B: Type) (f: A > B) (l: list A), map f l = map f l. + rewrite IHl. again, need to unfold so simpl can work unfold map. note: very sensitive to order of rewrite and unroll! another flavor of fold Fixpoint foldl (A B: Type) (f: A > B > B) (l: list A) (b: B) : B := nil => b cons x xs => foldl f xs (f x b) add one list to the end of another Fixpoint app (A: Type) (l1: list A) (l2: list A) : list A := match l1 with nil => l2 cons x xs => cons x (app xs l2) Eval cbv in (app (cons 1 (cons 2 nil)) (cons 3 nil)). Theorem app_nil: app l nil = l. Theorem app_assoc: forall A (l1 l2 l3: list A), app (app l1 l2) l3 = app l1 (app l2 l3). + rewrite IHl1. Oct 07, 16 18:11 Page 4/10 simple but inefficient way to reverse a list Fixpoint rev (A: Type) (l: list A) : list A := cons x xs => app (rev xs) (cons x nil) Eval cbv in (countdown 5). Eval cbv in (rev (countdown 5)). tail recursion is faster, but more complicated Fixpoint fast_rev_aux (A: Type) (l: list A) (acc: list A) : list A := nil => acc cons x xs => fast_rev_aux xs (cons x acc) Definition fast_rev (A: Type) (l: list A) : list A := fast_rev_aux l nil. let s make sure we got that right Theorem rev_ok: fast_rev l = rev l. + reduces rev, but does nothing to rev_fast unfold fast_rev. unfold fast_rev to fast_rev_aux now we can simplify the term TIP: if simpl doesn t work, try unfolding! + unfold fast_rev in *. this looks like it could be trouble... rewrite < IHl. STUCK! need to know about the rev_aux accumulator (acc) TIP: if your IH seems weak, try proving something more general Abort. Lemma fast_rev_aux_ok: fast_rev_aux l1 l2 = app (rev l1) l2. + STUCK AGAIN! need to know for *any* l2 TIP: if your IH seems weak, only intro up to the induction variable Abort. Lemma fast_rev_aux_ok: fast_rev_aux l1 l2 = app (rev l1) l2. intros A l1. + + rename l2 into foo. rewrite IHl1. rewrite app_assoc. now we can prove rev_ok as a special case of rev_aux_ok Lemma fast_rev_ok: Monday October 10, 2016 lec02/ 2/5

Oct 07, 16 18:11 Page 5/10 fast_rev l = rev l. unfold fast_rev. rewrite fast_rev_aux_ok. rewrite app_nil. << >> ~.,,,; ~.~.~ (.../ ~.~.~.~.~. } o~, ~.~.~.~.~.~. (/ \ ~.~.~.~.~.~.~. ; \ ~.~.~.~.~.~.~. ; {_.~.~.~.~.~.~.~ ;:. ~ ~.~.~.~.~. ;.: :._ ~.~.~.~.~ ;::.._ ~.~.~.~ ;::..,~.~.~. ;::::.. ;::;;:,: " / ~ ~" add an element to the end of a list Fixpoint snoc (A: Type) (l: list A) (x: A) : list A := nil => cons x nil cons y ys => cons y (snoc ys x) Theorem snoc_app_singleton: forall A (l: list A) (x: A), snoc l x = app l (cons x nil). Theorem app_snoc_l: forall A (l1: list A) (l2: list A) (x: A), app (snoc l1 x) l2 = app l1 (cons x l2). + rewrite IHl1. Theorem app_snoc_r: forall A (l1: list A) (l2: list A) (x: A), app l1 (snoc l2 x) = snoc (app l1 l2) x. Oct 07, 16 18:11 Page 6/10 + rewrite IHl1. simple but inefficient way to reverse a list Fixpoint rev_snoc (A: Type) (l: list A) : list A := cons x xs => snoc (rev_snoc xs) x Lemma fast_rev_aux_ok_snoc: fast_rev_aux l1 l2 = app (rev_snoc l1) l2. intros A l1. + + rewrite IHl1. rewrite app_snoc_l. Lemma fast_rev_ok_snoc: fast_rev l = rev_snoc l. unfold fast_rev. rewrite fast_rev_aux_ok_snoc. rewrite app_nil. Lemma length_app: length (app l1 l2) = plus (length l1) (length l2). + rewrite IHl1. Lemma plus_1_s: plus n 1 = S n. induction n. + rewrite IHn. Lemma rev_length: length (rev l) = length l. + rewrite length_app. rewrite plus_1_s. rewrite IHl. Lemma rev_app: rev (app l1 l2) = app (rev l2) (rev l1). Monday October 10, 2016 lec02/ 3/5

Oct 07, 16 18:11 Page 7/10 + rewrite app_nil. + rewrite IHl1. rewrite app_assoc. Lemma rev_involutive: rev (rev l) = l. + rewrite rev_app. rewrite IHl. * END LECTURE 02 SYNTAX We can define a programming language as an inductive datatype. include string library for variable names Require Import String. E ::= N V E + E E * E E? E Inductive expr : Type := Const : nat > expr Var : string > expr Add : expr > expr > expr Mul : expr > expr > expr Cmp : expr > expr > expr. S ::= Skip V < E S ;; S IF E THEN S ELSE S WHILE E {{ S }} Inductive stmt : Type := Skip : stmt Asgn : string > expr > stmt Seq : stmt > stmt > stmt Cond : expr > stmt > stmt > stmt While : expr > stmt > stmt. Programs are just elements of type stmt. Definition prog_skip : stmt := Skip. Definition prog_set_x : stmt := Asgn "x" (Const 1). Definition prog_incr_x_forever : stmt := While (Const 1) (Asgn "x" (Const 1)). Definition prog_xth_fib_in_y : stmt := Seq (Asgn "y" (Const 0)) ( Seq (Asgn "y0" (Const 1)) ( Seq (Asgn "y1" (Const 0)) ( Seq (Asgn "i" (Const 0)) ( While (Cmp (Var "i") (Var "x")) ( Seq (Asgn "y" (Add (Var "y0") (Var "y1"))) ( Seq (Asgn "y0" (Var "y1")) ( Oct 07, 16 18:11 Page 8/10 Seq (Asgn "y1" (Var "y")) ( (Asgn "i" (Add (Var "i") (Const 1))) )))) )))). but nobody wants to write programs like this, so Coq provides a "Notation" mechanism Notation " C X" := (Const X) (at level 80). Notation " V X" := (Var X) (at level 81). Notation "X <+> Y" := (Add X Y) (at level 83, left associativity). Notation "X <*> Y" := (Mul X Y) (at level 82, left associativity). Notation "X <?> Y" := (Cmp X Y) (at level 84). Notation "X < Y" := (Asgn X Y) (at level 86). Notation "X ;; Y" := (Seq X Y) (at level 87, left associativity). Notation " IF X THEN Y ELSE Z" := (Cond X Y Z) (at level 88). Notation " WHILE X {{ Y }}" := (While X Y) (at level 89). Definition prog_fib : stmt := "y" < C 0;; "y0" < C 1;; "y1" < C 0;; "i" < C 0;; WHILE (V"i" <?> V"x") {{ "y" < V"y0" <+> V"y1";; "y0" < V"y1";; "y1" < V"y";; "i" < V"i" <+> C 1 }}. Notation provides us with "concrete" syntax which "desugars" to the underlying "abstract syntax tree". Fixpoint nconsts (e: expr) : nat := Const _ => 1 Var _ => 0 Add e1 e2 => plus (nconsts e1) (nconsts e2) Mul e1 e2 => plus (nconsts e1) (nconsts e2) Cmp e1 e2 => plus (nconsts e1) (nconsts e2) Lemma has_3_consts: exists e, nconsts e = 3. exists (Add (Const 1) (Add (Const 2) (Const 3))). Fixpoint expr_with_n_consts (n: nat) : expr := O => Var "x" S m => Add (Const 0) (expr_with_n_consts m) Lemma has_n_consts: exists e, nconsts e = n. exists (expr_with_n_consts n). induction n. + rewrite IHn. Definition orb (b1 b2: bool) : bool := Monday October 10, 2016 lec02/ 4/5

Oct 07, 16 18:11 Page 9/10 match b1 with true => true false => b2 Fixpoint has_const (e: expr) : bool := Const _ => true Var _ => false Add e1 e2 => orb (has_const e1) (has_const e2) Mul e1 e2 => orb (has_const e1) (has_const e2) Cmp e1 e2 => orb (has_const e1) (has_const e2) Fixpoint has_var (e: expr) : bool := Const _ => false Var _ => true Add e1 e2 => orb (has_var e1) (has_var e2) Mul e1 e2 => orb (has_var e1) (has_var e2) Cmp e1 e2 => orb (has_var e1) (has_var e2) Lemma expr_bottoms_out: forall e, orb (has_const e) (has_var e) = true. induction e. + const + var + add e1 has a const e1 does not have a const * e2 has a const * e2 does not have a const we also want to simplify in the hypotheses and rewrite with the results + mul e1 has a const e1 does not have a const * e2 has a const * e2 does not have a const we also want to simplify in the hypotheses and rewrite with the results + cmp e1 has a const Oct 07, 16 18:11 Page 10/10 e1 does not have a const * e2 has a const * e2 does not have a const we also want to simplify in the hypotheses and rewrite with the results THE ABOVE PROOF IS VERY BAD. Make it better! Hint: think about how to rearrange the orbs. Some interesting types Inductive True : Prop := I : True. Inductive False : Prop :=. Lemma bogus: False > 1 = 2. inversion H. Lemma also_bogus: 1 = 2 > False. discriminate. Inductive yo : Prop := yolo : yo > yo. Lemma yoyo: yo > False. inversion H. well, that didn t work induction H. assumption. but that did! Monday October 10, 2016 lec02/ 5/5