Inductive Types an Proof Assistants
Total Page:16
File Type:pdf, Size:1020Kb
Inductive Types Proof Assistants Homotopy Type Theory Radboud University Lecture 4: Inductive types an Proof Assistants H. Geuvers Radboud University Nijmegen, NL 21st Estonian Winter School in Computer Science Winter 2016 H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 1 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Outline Inductive Types Proof Assistants Homotopy Type Theory H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 2 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Curry-Howard-de Bruijn logic ∼ type theory formula ∼ type proof ∼ term detour elimination ∼ β-reduction proposition logic ∼ simply typed λ-calculus predicate logic ∼ dependently typed λ-calculus λP intuitionistic logic ∼ . + inductive types higher order logic ∼ . + higher types and polymorphism classical logic ∼ . + exceptions H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 3 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Inductive types by examples from the Coq system • booleans • natural numbers • integers • pairs • linear lists • binary trees • logical operations H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 5 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Inductive types • types • recursive functions • definition using pattern matching • ι-reduction = evaluation of recursive functions • recursion principle • proof by cases proof by induction • induction principle H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 6 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Booleans: type Coq definition: Inductive bool : Set := | true : bool | false : bool. Coq produces: bool rec. bool rect. bool ind. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 7 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Booleans: elimination Check bool ind. bool ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b bool is the type with exactly two elements: true and false Check bool rec. bool rec : forall P : bool -> Set, P true -> P false -> forall b : bool, P b bool rec is for defining functions on bool Defining functions over bool = Constructing proofs over bool !! H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 8 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Booleans: recursive functions using pattern matching In practice we don't use use bool rec, but we program functions using pattern matching Definition of negation: Definition neg (b : bool) : bool := match b with | true => false | false => true end. ι-reduction neg true !! false H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 9 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Booleans: proof by cases For proofs we don't use use bool ind, but we do proofs by induction and by cases. forall b : bool, neg (neg b) = b To prove this in Coq we use tactics: • elim b. • destruct b. • simpl. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 10 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Natural numbers: type Coq definition: Inductive nat : Set := |O : nat |S: nat -> nat. Coq produces: nat rec. nat rect. nat ind. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 11 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Natural numbers: recursive function Fixpoint plus (n m : nat) fstruct ng : nat := match n with | O => m | S p => S (plus p m) end. • We write Fixpoint instead of Definition because we define plus recursively • In the definition body, we can only call plus on a structurally smaller element ι-reduction plus (S O) (S O) !! S (plus O (S O)) !! S (S O) H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 12 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Natural numbers: induction principle nat ind: forall P : nat -> Prop, P 0 -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n nat ind : 8P: P(0) ) (8n 2 N: P(n) ) P(n + 1)) ) 8n 2 N: P(n) H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 13 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Natural numbers: proof by induction forall n : nat, plus n O = n tactics • induction n. • destruct n. • rewrite IHn. • simpl. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 14 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Truth: type The inductively defined proposition with exactly one proof. Coq definition: Inductive True: Prop := I : True. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 15 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Falsity: type The inductively defined proposition with no proof. Coq definition: Inductive False : Prop := . H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 16 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Falsity: induction principle False ind: forall P : Prop, False -> P . ? E? P H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 17 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Lists: type Coq definition: Inductive natlist : Set := | nil : natlist | cons : nat -> natlist -> natlist the list1 ; 2; 3; 4 is encoded by cons 1 (cons 2 (cons 3 (cons 4 nil))) H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 18 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Lists: recursive function Definition of append: Fixpoint append (l k : natlist) fstruct lg : natlist := match l with | nil => k | cons n l' => cons n (append l' k) end. ι-reduction append (cons 0 nil) nil !! cons 0 (append nil nil) !! cons 0 nil H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 19 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Lists: induction principle natlist ind: forall P : natlist -> Prop, P nil -> (forall (n : nat) (l : natlist), P l -> P (cons n l)) -> forall l : natlist, P l H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 20 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Coq only accepts well-defined inductive types You cannot write Inductive Lambda : Set := | var : nat -> Lambda | app : Lambda -> Lambda -> Lambda | abs:(Lambda -> Lambda) -> Lambda. The type X you are defining should occur strictly positive in all σi , if constr : σ1 ! ::: ! σn ! X is a constructor declaration. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 21 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Coq only accepts well-defined recursive functions Recursive call must be to a structurally smaller argument. You cannot write Fixpointf (n : nat) fstruct ng : nat := match n with | O => O | S p =>f (f p) end. Fixpointg (n : nat) fstruct ng : nat := match n with | O => 1 | S p =>g (p mod 2) end. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 22 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University What are Proof Assistants { History John McCarthy (1927 { 2011) 1961, Computer Programs for Checking Mathematical Proofs Proof-checking by computer may be as important as proof generation. It is part of the definition of formal system that proofs be machine checkable. ... For example, instead of trying out computer programs on test cases until they are debugged, one should prove that they have the desired properties. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 24 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University What are Proof Assistants { History Around 1970 five new systems / projects / ideas • Automath De Bruijn (Eindhoven) now: Coq • Nqthm Boyer, Moore (Austin, Texas) now: ACL2, PVS • LCF Milner (Stanford; Edinburgh) now: HOL, Isabelle • Mizar Trybulec (Bia lystok, Poland) • Evidence Algorithm Glushkov (Kiev, Oekrain) H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 25 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University HOL Light LCF tradition (Milner): LCF ! HOL ! HOL Light Stanford, US ! Cambridge, UK ! Portland, US Based on: higher order logic John Harrison proves correctness of floating point hardware at Intel formalises mathematics in his spare time very simple and elegant system easy to extend (add your own tactics) not user friendly H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 26 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Isabelle 'successor' of HOL Based on: higher order logic cooperation between two universities: Cambridge, UK focus: computer security M¨unchen,Duitsland focus: mathematics and programming languages balanced system nice proof language powerful automation H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 27 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Coq Based on: type theory INRIA en Microsoft Institut National de Recherche en Informatique et en Automatique system with the most impressive formalisation so far system used most at Nijmegen integrated programming language ≈ Haskell mathematically expressive the built in logic is intu¨ıtionistic H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 28 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Mizar Andrzej Trybulec Bia lystok, Polen also: Nagano, Japan Based on: set theory most mathematical of all proof assistants largest library of formalised mathematics 2,1 miljon lines of code user friendly sometimes