c constructor P, Q terms used as propositions G, H hypotheses scope identifier for a notation scope M, module identifiers t, u arbitrary terms

Similar documents
Coq quick reference. Category Example Description. Inductive type with instances defined by constructors, including y of type Y. Inductive X : Univ :=

Inductive Definitions, continued

Introduction to dependent types in Coq

CS 456 (Fall 2018) Scribe Notes: 2

Certified Programming with Dependent Types

Infinite Objects and Proofs 1

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

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.

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

Lean 2 Quick Reference

The Coq Proof Assistant A Tutorial

LASER 2011 Summerschool Elba Island, Italy Basics of COQ

A Coq tutorial for confirmed Proof system users

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

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

Logical Verification Course Notes. Femke van Raamsdonk Vrije Universiteit Amsterdam

The Coq Proof Assistant A Tutorial

Autosubst Manual. May 20, 2016

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

Induction in Coq. Nate Foster Spring 2018

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

The Coq Proof Assistant A Tutorial

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:

CIS 500: Software Foundations

Programming with (co-)inductive types in Coq

Equations: a tool for dependent pattern-matching

Inductive data types

Coq in a Hurry. Yves Bertot

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

Mathematics for Computer Scientists 2 (G52MC2)

Coq in a Hurry. Yves Bertot. October 2008

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.

Automated Reasoning. Natural Deduction in First-Order Logic

A Simple Supercompiler Formally Verified in Coq

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.

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

4 Programming with Types

Introduction to Co-Induction in Coq

Coq projects for type theory 2018

First-Class Type Classes

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

CMSC 330: Organization of Programming Languages

Basic Foundations of Isabelle/HOL

Calculus of Inductive Constructions

CIS 500: Software Foundations

Idris: Implementing a Dependently Typed Programming Language

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

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

CIS 500: Software Foundations

Theorem Proving Principles, Techniques, Applications Recursion

Heq: a Coq library for Heterogeneous Equality

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

L06_exercise.v. L06_exercise.v

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

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

Introduction to the Coq proof assistant First part

A Tutorial on [Co-]Inductive Types in Coq

3 Pairs and Lists. 3.1 Formal vs. Informal Proofs

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

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

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

Type checking by theorem proving in IDRIS

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

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

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

Proving Theorems with Athena

Lee Pike. June 3, 2005

3.4 Deduction and Evaluation: Tools Conditional-Equational Logic

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

CMSC 330: Organization of Programming Languages. Formal Semantics of a Prog. Lang. Specifying Syntax, Semantics

The Isar Proof Language in 2016

Structures in Coq January 27th 2014

Mtac2: Typed Tactics for Backward Reasoning in Coq

Theorem proving. PVS theorem prover. Hoare style verification PVS. More on embeddings. What if. Abhik Roychoudhury CS 6214

MPRI course 2-4 Functional programming languages Exercises

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

An Ssreflect Tutorial

An introduction to small scale reflection in Coq

Beluga: A Framework for Programming and Reasoning with Deductive Systems (System Description)

Operational Semantics

Functional Programming in Coq. Nate Foster Spring 2018

Meta-programming with Names and Necessity p.1

Chapter 3. The While programming language

Lesson 4 Typed Arithmetic Typed Lambda Calculus

CIS 500 Software Foundations Midterm I

Mathematics for Computer Scientists 2 (G52MC2)

Preuves Interactives et Applications

The pitfalls of protocol design

Advanced Features: Type Classes and Relations

To figure this out we need a more precise understanding of how ML works

CMSC 330: Organization of Programming Languages. Functional Programming with Lists

CS 4110 Programming Languages & Logics. Lecture 28 Recursive Types

CS4215 Programming Language Implementation. Martin Henz

Programming Languages Third Edition

Built-in Module BOOL. Lecture Note 01a

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

Abella: A System for Reasoning about Relational Specifications

Inductive data types (I)

(Refer Slide Time: 4:00)

Adding GADTs to OCaml the direct approach

Transcription:

Coq quick reference Meta variables Usage Meta variables Usage c constructor P, Q terms used as propositions db identifier for a hint database s string G, H hypotheses scope identifier for a notation scope M, module identifiers t, u arbitrary terms m module values tactic Ltac terms with tactic values n numbers v, w, x, y identifiers p patterns V, W, X, Y terms used as types Category Example Description Sorts Prop Universe of logical propositions. Set Type Universe of program types or specifications. Type of Set and Prop. Types forall x : X, Y : Univ Dependent product type (Y can refer to x). Declarations and definitions X -> Y : Univ CoFixpoint x (y : Y) : X := t. CoInductive X : Univ := v : V.... Declare Module M : V. Definition x : X := t. Definition x (v : V) : X := t. Arrow type (non-dependent product), same as forall _: X, Y. As for Fixpoint but defines co-recursive functions. As for Inductive but defines co-inductive type. Declare a module of a particular module type. Constant definition. Non-recursive function definition. Surround arg with braces instead of parens to make it implicit. Use @x to get version of function x with all arguments made explicit. End M. Finish definition of module or module type M. Fixpoint x (v : V) : X := t. As for Definition but t can make recursive calls to x. Fixpoint x (v : V) : X := t with y (w : W) : Y := u. Import M. Mutually recursive function definitions. Import the names of module M so they can be used without qualification. Version 0.4, December 2012 for Coq version 8.4! 1

Category Example Description Inductive X : Univ := v : V.... Inductive X : Univ := v : V... with Y : Univ := w : W.... Let x : X := t. Inductive type with instances defined by constructors, including y of type Y. Mutually recursive inductive types X and Y. Local definition. Module M (x : X). Start definition of parameterised module M. Module Type V. Start definition of module type V. Module M := m. Parameter x : X. Record x (v : V) : Univ := y { w : W;... }. Define M to be the module m Global constant declaration (also Axiom). Define a dependent record type, i.e., an inductive type x with a single constructor y which has arguments w. The w are also projection functions for extracting argument values from a value of this type. Require M. Load the file containing module M. Require Import M. As for Require and import names of M. Require Export M. Variable x : X. Variables x : X y : Y. As for Require Import and indicate that M should be imported if the current module is imported. Local variable declaration (also Hypothesis). Local variable declarations (also Hypotheses). Terms fun x : X => t Non-recursive function. fix x (v : V) : X := t t t1... tn _ let x := t in u match t return with p1 => t1... end Recursive function. Function call. Anonymous variable in pattern or ask for argument to be inferred in function call. Local binding. Dependent pattern matching term selection where return expresses the type of the match expression. return can be omitted if it is the common type of all of the branches. return T Return type is T. Version 0.4, December 2012 for Coq version 8.4! 2

Category Example Description as x in U u1 u2... as x in U u1 u2... if t then u else v admit Prepend to return type so that T can refer to the matched term t by name x. Prepend to return type so that T can refer to the type U u1 u2... of the matched term t. Combination of above. If t is of type Bool, same as match t with true => u false => v end. Works for all types with exactly two constructors. Term that fits in any hole. Evaluation Eval tactic in t. Normalisation of t using conversion tactic chosen from red, hnf, compute, simpl, cbv, lazy, unfold, fold, pattern. Theorems Theorem x : X. Theorem definition (also Example, Lemma). Proofs Proof. Start proof of theorem (optional). Admitted. Defined. Guarded. Qed. Admission (stands for proof). As for Qed but mark definition as transparent so it can be unfolded. Report if it is possible to finish a proof yielding a properly guarded co-inductive proof term. Build, check and save opaque proof term. Proof steps Show. Print current goals. Show n. Show Proof. tactic. n:tactic. Print nth goal. Print the current proof term. Apply tactic to first goal. Apply tactic to nth goal. tactic in H. Apply tactic to hypothesis H. tactic in *. Apply tactic to first goal and all hypotheses. Undo [n]. Undo n single steps (default: 1). Restart. Abort. Go back to beginning of proof attempt. Abort proof attempt. Printing Check t. Check if well-formed, print type. Print x. Print Assumptions x. Print value and type. Print the axioms, parameters and variables that the theorem or definition x depends on. Version 0.4, December 2012 for Coq version 8.4! 3

Category Example Description Searching SearchAbout t. List results whose conclusion is (t...) where t is usually just an identifier. SearchPattern (p). SearchRewrite (p). List theorems whose conclusion or final hypthesis matches p (e.g, _ + _ + <= _). List theorems whose conclusion is an equality where one side matches p.?id Meta-variable in non-linear pattern. _ Wildcard pattern. Hints Hint hintdef : db... Declare the specified hints in the given databases (default: core). Hint Local hintdef : db... Constructors id Extern n [pattern] => tactic Declare a local hint that is not to be exported. Add all of the constructors of type id. Add tactic as a hint for when pattern matches (with cost n). Immediate t Add apply t; trivial (cost 1). Opaque x Add an opacity hint about x. Resolve t Add apply t (cost number of subgoals generated). Rewrite t1 t2... Add equalities ti to rewrite database. Also Rewrite <- t1 t2... to add rewrite from right to left. Transparent x Add a transparency hint about x. Unfold x Add unfold x (cost 4). Notation Notation... := t (at level n, left/ right associativity) : scope Notation... := t. Reserved Notation... (at level n, left/right associativity) : scope Inductive... where... := t : scope Define new notation in given scope, syntax from string pattern, defined to be term t with priority n (from 0 to 100) and given associativity. Define new notation in current scope with default priority and associativity. Reserve notation without giving its definition so it can be used to define its own meaning. Provide a declaration of a reserved notation while using it in an inductive definition Scopes Close Scope scope. Remove all occurrences of scope from the notation interpretation stack. Same modifiers as Open Scope. Locate *. Locate interpretations of notation *. Version 0.4, December 2012 for Coq version 8.4! 4

Category Example Description Environments and Sections Open Scope scope. Print Scope scope. t %scope Reset Initial. Add scope to the stack of notation interpretation scopes. As shown, the change does not survive the end of a section where it occurs. Add Global modifier to survive sections or Local modifier to avoid exporting from module. Print notations from scope. Term with notation interpreted in scope. Reset to initial environment and empty context. Reset x. Reset everything since and including x. Section x. Begin nested section called x. End x. End nested section x and add variables and hypotheses from the section as extra arguments to definitions from the section. Extraction Extraction x. Extract a single constant or module. Recursive Extraction x y.... Extraction file x y.... Extraction Library x. Recursive Extraction Library x. Extraction Language language. Recursively extract ids and their dependencies. Recursively extraction to a file. Extract file x.v to x.ml. Recursive extraction of library x and dependencies. Set language for extraction to Ocaml, Haskell or Scheme. Miscellaneous Implicit Arguments x [v w...]. Make arguments v w... of x implicit. x... (v := t)... Call x giving an explicit value for the implicit argument v. Options Set/Unset option. Toggle an option. Printing All Whether to show full details of terms. Printing Notations Whether to print using notations or not. Implicit Arguments Auto infer arguments at function definition. Version 0.4, December 2012 for Coq version 8.4! 5

Library Type Notation Constructors and (P Q: Prop) : Prop P /\ Q conj : P -> Q -> P /\ Q bool : Set true : bool false : bool eq (X : Type) (x : X) : X -> Prop x = y eq_refl : x = x ex (X : Set) (P : X -> Prop) exists x : X, P ex_intro : forall x : X, P x -> exists x, P x False : Prop nat : Set Numerics in nat_scope O : nat S : nat -> nat or (P Q: Prop) P \/ Q or_introl : P -> P \/ Q or_intror : Q -> P \/ Q sig (X : Type) (P : X -> Prop) : Type { x : X P } exist : forall x : X, P x -> sig P sumbool (P : Prop) (Q : Prop) : Set { P } + { Q } left : P -> { P } + { Q } right : Q -> { P } + { Q } sumor (X : Type) (P : Prop) : Type X + { P } inleft : X -> X + { P } inright : P -> X + { P } unit : Set True : Prop tt : Set I : True Definition Type all : forall X : Type, (X -> Prop) -> Prop iff : Prop -> Prop -> Prop not : Prop -> Prop proj1 : forall P Q : Prop, P /\ Q -> P proj1_sig : forall (X : Type) (P : X -> Prop), {x P x} -> X proj2 : forall P Q : Prop, P /\ Q -> Q proj2_sig : forall (X : Type) (P : X -> Prop) (e : {x P x}), P (proj1_sig e) Description Universal quantification If and only if Logical negation Left projection Left projection from dependent sum Right projection Right projection from dependent sum Version 0.4, December 2012 for Coq version 8.4! 6

Tactics Tactics work on the conclusion of the current goal by default. Use tactic in H to apply to hypothesis H or tactic in * to apply to conclusion and all hypotheses. Tactics including apply, assumption, auto, constructor, destruct, exact, exists, induction, left, split and right also have existential versions (e.g., eapply, eassumption, etc) which do not fail if an instantiation of a variable cannot be found. Rather, they introduce new existential variables in those positions. Tactic Hypotheses Conclusion Description apply H H : u -> t2 t1 Apply deduction backwards, where t1 unifies with t2. New conclusion is u. Can also use theorems or constructors apply H with x := t... apply G in H H : t1 G : t2 -> u As for apply H but with specified bindings in H. Can omit x := if unambiguous Apply deduction forwards in H, where t1 unifies with t2 and replace with u. assert t as H u New subgoal t called H which is to be proven and then used to prove u. assert (H : t) Same as assert t as H. assumption H : t t Look for conclusion in hypotheses; if found, same as apply H. auto [n] auto using x, y,... auto with db... autorewrite with db... case t cbv flags Solve if possible with combination of intros, apply and reflexivity. Uses current hypotheses plus core hint database. Only go to a depth of n (default: 5). auto with x, y,... added to hint database. auto using given hint databases as well as default core. databases can be * to use all databases Rewrite according to rewrite databases (built with Hint Rewrite). Simpler than destruct and like elim but uses case-analysis elimination principle not a recursive one. Call by value normalisation where flags are from beta, delta, iota or zeta. If no flag is given all are assumed. change t u Change conclusion to t if t and u are convertible. clear H H :... Remove H from context. cofix x compute Primitive tactic to start proof by coinduction using hypothesis x. Synonym for cbv beta delta iota zeta. Version 0.4, December 2012 for Coq version 8.4! 7

Tactic Hypotheses Conclusion Description congruence t = u Decision procedure for ground equalities and uninterpreted functions. constructor Look for constructor c that can be applied to solve current goal; if found, same as apply c. contradiction H : False Look for hypothesis that is equivalent to False; if found, solve goal. cut t u New goals: t -> u and t. decide equality forall x y : X, {x = y} + {x <> y} Solve a decidable equality. destruct t H :... t...... t... Destruct t by cases. discriminate t <> u Show that t and u are not equal if they are made using different constructors. elim t Basic induction that eliminates by choosing appropriate destructor and applying it. exact t u Prove u by giving exact proof term t. exists t ex X P Witness x : X and P : X -> Prop, conclusion becomes P x. fail Always fail to apply f_equal f t1 t2 = f u1 u2 Replace with goals t1 = u1, t2 = u2. field firstorder fix x n fold t Extension of ring tactic to rational expressions. First order logic decision procedure. Primitive tactic to start proof by induction calling induction hypothesis x and acting on premise n of current goal. Replace occurrences of the reduction of t (using red). generalize t... t... Generalize a goal sub-term t, replacing it with a new variable id and adding forall x to the goal. generalize dependent t hnf idtac [s...] Same as generalize, except also generalizes all hypotheses that depend on x, clearing them. Replace current conclusion with its head normal form. Identity, print s if present. induction t Induction on t. Version 0.4, December 2012 for Coq version 8.4! 8

Tactic Hypotheses Conclusion Description induction t using x Induction on t with non-standard induction principle x. injection H H : t = u Use injectivity of constructors in H to infer equalities between sub-terms of t and u. instantiate n := t...?x... Refine the nth existential variable in the conclusion (numbering from right) with the term t. intro H t -> u New hypothesis H : t, conclusion u. intros [H1 H2...] t -> u ->... -> v intro H1, intro H2,... intros [x y] forall x : X, y :Y, t New hypotheses x : X, y : Y conclusion t. intuition intuition auto with *. intuition tactic Intuitionistic propositional calculus decision procedure and apply tactic to the generated subgoals. inversion H H : c x = c y Invert injective constructor c to get x = y. inversion H H : c x = d y Invert different constructors to get contradiction. inversion H H : t Invert inductively defined t, creating one subgoal for each way in which t could be constructed (with associated assumptions). lazy flags Call by need normalisation where flags are from beta, delta, iota or zeta. If no flag is given all are assumed left E.g., P \/ Q Apply first constructor of a type with two constructors, e.g., apply or_introl. omega pattern t pose x := t Decision procedure for first-order logic with Presburger arithmetic (numeric constants, +, -, S, pred, * by constant, =, <>, <=, /\, \/, ~ and ->). Perform beta expansion on t which should be a free sub-term of the current conclusion. Add local definition x := t to hypotheses. red forall (x : X), y x Where y is a transparent constant, replace y by its definition then beta, iota and zeta reduce. refine t u Give exact proof term like exact but with some holes marked with _. reflexivity t = u Simplify then remove equality. Version 0.4, December 2012 for Coq version 8.4! 9

Tactic Hypotheses Conclusion Description remember t as x... t... Add hypothesis x = t to all subgoals and replace t in goal with x. rename H into G H:... Rename hypothesis H to G. replace t with u... u... Replace all occurrences of t by u and add new subgoal t = u. rewrite H H : t = u... t... t... Same as rewrite -> H rewrite -> H H : t = u... t... u... Rewrite t to u. H can also be a previously proven theorem. rewrite <- H H : t = u... u... u... Rewrite u to t. right E.g., P \/ Q Apply second constructor of a type with two constructors, e.g., apply or_intror. ring Normalise polynomials over any ring or semi-ring structure. set x := t H :... t...... t... Replace t with x in the current goal and add local definition x := t to the context. simpl Simplify by applying beta and iota reduction and expanding transparent constants. specialize (H t u...) H : P ->... -> Q Specialise H by applying it to terms t u... split intros; apply conj. subst x x = t or t = x... x... Replace x with t, clear assumption. subst Perform substitution of all equality assumptions. symmetry t = u Switch sides to get u = t. tauto trivial unfold x, y,... intuition fail. Automatically apply basic tactics using only hypotheses with a single premise. Delta reduce by x, y,... and simplify. Version 0.4, December 2012 for Coq version 8.4! 10

Tacticals tactic1; tactic2 tactic1; [tactic2 tactic3...] tactic1 tactic2 abstract tactic do n tactic first [tactic1 tactic2...] info tactic progress tactic repeat tactic solve tactic solve [tactic1 tactic2... ] timeout n tactic try tactic Apply tactic1 then apply tactic2 to all goals that result. Apply tactic1 then apply tactic2 to the first resulting goal, tactic3 to second goal and so on. Apply tactic1 and if it fails apply tactic2 to the original goal. Similar to solve but also saves auxiliary lemma. Apply tactic to the goal n times. Apply the first tactic in the sequence that doesn t fail. Behave as tactic but display operations for complex tactics such as auto. Fail if tactic succeeds without making progress. Repeatedly apply tactic until no goals left or it fails on all subgoals. If tactic solves the goal, behave just like tactic, otherwise do nothing. As for solve but try to solve with the listed tactics in order. Stop tactic if it hasn t finished after n seconds. Same as tactic idtac. (... ) Grouping. Version 0.4, December 2012 for Coq version 8.4! 11

Ltac Ltac x v w... := t. defines a new tactic x with arguments v w... defined by the Ltac term t. Terms constr:( t ) Construct the Gallina term t. fresh s Evaluate to an identifier unbound in the goal with the new name being s with a number appended to make it unique if necessary. x Value of the bound variable x.?x Value of variable?x previously bound in a pattern. let [rec] x v w... := t in u match goal with p => tactic... end. match reverse goal with... end. match t with p => tactic... end. Bind x with arguments v w... to t in u. If rec is present t can contain recursive references to x. Match the current proof state against the patterns in the order given and if a match is found, apply the corresponding tactic. If the tactic fails, backtrack to find other matches for the current pattern, if any. If there are no (more) matches to a pattern, go to the next clause. Hypotheses are matched in newest to oldest order. As for match goal with but match hypotheses from oldest to newest. Non-linear pattern matching of an Ltac term. type of t The type of Gallina term t. Patterns - p Match against top of the current conclusion. - context [ p ] Match against any sub-term of the current conclusion. p1 - p2 t Match p1 against the hypotheses and p2 against the conclusion. Term-valued pattern.?x Variable pattern. H1 : p1, H2 : p2,... Match patterns from right to left against hypotheses. Version 0.4, December 2012 for Coq version 8.4! 12