Inductive Types Proof Assistants Homotopy 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 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) {struct n} : 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 : ∀P. P(0) ⇒ (∀n ∈ N. P(n) ⇒ P(n + 1)) ⇒ ∀n ∈ 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) {struct l} : 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) {struct n} : nat := match n with | O => O | S p =>f (f p) end. Fixpointg (n : nat) {struct n} : 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 (Bialystok, 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 Bialystok, 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 hard to follow

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 29 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Proof Assistants for software verification

Holy Grail ‘Things like even software verification, this has been the Holy Grail of computer science for many decades but now in some very key areas, for example, driver verification we’re building tools that can do actual proof about the software and how it works in order to guarantee the reliability.’ Bill Gates, 18 april 2002

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 30 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Why would we believe a proof assistant?

. . . a proof assistant is just another program . . . To attain the utmost level of reliability: • Description of the rules and the logic of the system. • A small “kernel”. All proofs can be reduced to a small number of basic proof steps. high level steps are defined in terms of the small ones.

LCF approach [Milner]: Have an abstract data type of theorems thm, where the only constants of this data type are the axioms and the only functions to this data type are the inference rules of the logic.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 31 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Why would we believe a proof assistant?

. . . a proof assistant is just another program . . . Other possibilities to increase the reliability of the proof assistant • Check the proof checker. Verify the correctness of the proof assistant in a proof assistant (e.g. the system itself). Example Coq in Coq: Construct a model of Coq in Coq itself and show that all tactics are sound with respect to this model NB. G¨odel’sincompleteness . . . , so we need to assume something. • The De Bruijn criterion A proof assistant satisfies the D.B. criterion if it generates proof objects that can be checked independently of the system that created it using a simple program that a skeptical user can write him/herself.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 32 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Why would we believe a proof assistant?

Separating the simple proof checker from the powerful proof engine

Proof Assistant (Interactive Theorem Prover)

Tactics

Goals User Proof assistant

Proof Assistant with a small kernel that satisfies the De Bruijn criterion Proof Assistant Tactics

Proof object OK Proof checker User Goals Proof Engine

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 33 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University How would we believe a proof assistant?

Q: Does the formula on the screen correspond to what we have proven? Given that I trust the proof assistant, how much proof code (definitions) do I need to read (and understand) to believe that the final theorem is the one I wanted to see proven? That’s an issue . . . The situation seems different between mathematics and computer science.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 34 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University How to believe that we have actually proven a theorem? The 4 colour theorem Kenneth Appel en Wolfgang Haken, 1976 Neil Robertson e.a., 1996 Coq: Georges Gonthier, Microsoft-INRIA 2004

Can every map be coloured with only 4 different colours? • Gonthier has two pages of Coq definitions and notations that are all that’s needed to fully and precisely understand his statement of the 4 colour theorem. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 35 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University How to believe that we have actually proven a theorem?

The verified C compiler Compcert Coq: Leroy et al. INRIA 2006

Just stating what the correctness of a C-compiler means takes many pages already . . .

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 36 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Mathematical users of Proof Assistants

Some cool formalization of mathematics projects: • Mathematical Components: Feit-Thompson theorem • Flyspeck: Kepler Conjecture • Homotopy Type Theory

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 37 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University ‘Mathematical Components’ project of INRIA-Microsoft Coq version of the 4 color theorem

• formal version of the proof • program correctness Coq as a functional programming language Georges Gonthier new technologies: 2005 • mathematical language Ssreflect Feit-Thompson theorem • hypermaps

Assia Mahboubi

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 38 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Example of a formal Coq proof Lemma no_minSimple_odd_group (gT : minSimpleOddGroupType) : False. Proof. have [/forall_inP | [S [T [_ W W1 W2 defW pairST]]]] := FTtypeP_pair_cases gT. exact/negP/not_all_FTtype1. have xdefW: W2 \x W1 = W by rewrite dprodC. have pairTS := typeP_pair_sym xdefW pairST. pose p := #|W2|; pose q := #|W1|. have p’q: q != p. have [[[ctiW _ _] _ _ _ _] /mulG_sub[sW1W sW2W]] := (pairST, dprodW defW). have [cycW _ _] := ctiW; apply: contraTneq (cycW) => eq_pq. rewrite (cyclic_dprod defW) ?(cyclicS _ cycW) // -/q eq_pq. by rewrite /coprime gcdnn -trivg_card1; have [] := cycTI_nontrivial ctiW. without loss{p’q} ltqp: S T W1 W2 defW xdefW pairST pairTS @p @q / q < p. move=> IH_ST; rewrite neq_ltn in p’q. by case/orP: p’q; [apply: (IH_ST S T) | apply: (IH_ST T S)]. have [[_ maxS maxT] _ _ _ _] := pairST. have [[U StypeP] [V TtypeP]] := (typeP_pairW pairST, typeP_pairW pairTS). have Stype2: FTtype S == 2 := FTtypeP_max_typeII maxS StypeP ltqp. have Ttype2: FTtype T == 2 := FTtypeP_min_typeII maxS maxT StypeP TtypeP ltqp. have /mmax_exists[L maxNU_L]: ’N(U) \proper setT. have [[_ ntU _ _] cUU _ _ _] := compl_of_typeII maxS StypeP Stype2. by rewrite mFT_norm_proper // mFT_sol_proper abelian_sol. have /mmax_exists[M maxNV_M]: ’N(V) \proper setT. have [[_ ntV _ _] cVV _ _ _] := compl_of_typeII maxT TtypeP Ttype2. by rewrite mFT_norm_proper // mFT_sol_proper abelian_sol. have [[maxL sNU_L] [maxM sNV_M]] := (setIdP maxNU_L, setIdP maxNV_M). have [frobL sUH _] := FTtypeII_support_facts maxS StypeP Stype2 pairST maxNU_L. etceterahave [frobM _ _] := FTtypeII_support_facts maxT TtypeP Ttype2 pairTS maxNV_M. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 39 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Kepler conjecture

Johannes Kepler 1611

cubic face-centred packing

strena seu de nive sexangula on the hexagonal snow flake

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 40 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University the Flyspeck project • 1996: proof of the Kepler conjecture book of 334 pages giga bytes data and days of computer calculations • reviewers of the Mathematische Annalen: Tom Hales can’t find any mistakes, but too complicated to fully check

2003: Flyspeck project= complete formal version of the proof HOL Light proof assistant (+ Isabelle proof assistant) • 2014: formal proof of the Kepler conjecture impossible that there is a mistake somewhere

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 41 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University example of a formal HOL proof

w~ 6= ~0 ∧ ~u · w~ = ~v · w~ ⇒ ∠(~u × w~ ,~u × w~ ) = ∠(~u,~v)

let VECTOR_ANGLE_DOUBLECROSS = prove (‘!u v w. John Harrison ~(w = vec 0) /\ u dot w = &0 /\ v dot w = &0 ==> vector_angle (u cross w) (v cross w) = vector_angle u v‘, REPEAT GEN_TAC THEN ASM_CASES_TAC ‘u:real^3 = vec 0‘ THENL [ASM_REWRITE_TAC[vector_angle; CROSS_0]; ALL_TAC] THEN ASM_CASES_TAC ‘v:real^3 = vec 0‘ THENL [ASM_REWRITE_TAC[vector_angle; CROSS_0]; ALL_TAC] THEN STRIP_TAC THEN SUBGOAL_THEN ‘~(u cross w = vec 0) /\ ~(v cross w = vec 0)‘ ASSUME_TAC THENL [REPEAT(POP_ASSUM MP_TAC) THEN REWRITE_TAC[GSYM DOT_EQ_0] THEN VEC3_TAC; ALL_TAC] THEN ASM_SIMP_TAC[VECTOR_ANGLE_EQ] THEN SIMP_TAC[vector_norm; GSYM SQRT_MUL; DOT_POS_LE] THEN ASM_REWRITE_TAC[DOT_CROSS; REAL_MUL_LZERO; REAL_SUB_RZERO] THEN REWRITE_TAC[REAL_ARITH ‘(x * y) * (z * y):real = (y * y) * x * z‘] THEN SIMP_TAC[SQRT_MUL; DOT_POS_LE; REAL_LE_SQUARE; REAL_LE_MUL] THEN SIMP_TAC[SQRT_POW_2; DOT_POS_LE; GSYM REAL_POW_2] THEN REAL_ARITH_TAC);;

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 42 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Dedekind cuts and Cauchy sequences

Richard Dedekind Augustin-Louis Cauchy

a left and a equivalence class of right set converging sequences

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 43 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Homotopy type theory (HoTT) Fields medal 2002 • homotopy theory algebraic varieties • formulation of motivistic cohomology

Vladimir Voevodsky mathematics independent of specific definitions 2006 homotopy type theory • homotopy is the ‘proper’ notion of equality • homotopy = continuous transformation

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 44 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Homotopy Theory

Part of Algebraic Topology dealing with homotopy groups: associating groups to topological spaces to classify them.

• an equality is a path from one object to another (continuous transformation) • higher equality = transformation between paths = a path between paths.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 46 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Identity type

Identity is an inductive type Inductive identity (A : Type) : A -> A -> Type := refl : forall a : A, a = a. The smallest binary relation on A containing {(x, x) | x : A}. Giving refl : forall (A : Type)(a : A), a = a. and ident ind : forall(A : Type)(P : forall a b : A, a=b -> Prop), (forall a : A, P a a refl) -> forall (x y : A) (i : x = y), P x y i + a reduction rule

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 47 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Identity type in more mathematical presentation

The identity elimination is usually called the J-rule. We have

refl : ∀A : Type, ∀a : A, a = a

P : ∀ab : A, a = b → Prop r : ∀a : A, P a a refl J J r : ∀xy : A, ∀i : x = y, P x y i with computation rule

J a a (refl a) → r.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 48 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Properties of the Identity type The J-rule gives: • Identity is an equivalence relation: sym: a = b → b = a (symmetry) and trans: a = b → b = c → a = c (transitivity). • Substitutivity (Leibniz property) t : Q(a) r : a = b t0 : Q(b) But: t0 is not just t. (In fact t0 ≡ J a b r t.) Note: Type theory is intensional: terms contain a lot of explicit information to guarantee that type checking is decidable. The following “simpler” Leibniz rule would render type checking undecidable t : Q(a) r : a = b t : Q(b)

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 49 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Properties of the Identity type The J-rule does not give: • Function extensionality f , g : A → B r : ∀a : A, f a = g a FunExt t : f = g for some term t. • Proof Irrelevance (all proofs are equal). A : Prop a : A b : A PI t : a = b for some term t. • Uniqueness of Identity Proofs (UIP).

a, b : A q0, q1 : a = b UIP t : q0 = q1 for some term t. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 50 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Proof Irrelevance In type theory, terms may depend on proofs, for example the reciprocal in R takes a proof argument: 1 : ∀x : .x 6= 0 → − R R Now, if x : R and q : x 6= 0 and r :6= 0, we want 1 1 = x, p x, r which we get from PI (Proof Irrelevance), because then p = q. Often we can program/proof around this. In this case we have as 1 equation for − (for all x, p : x 6= 0): 1 1 x · = 1 = · x. x, p x, p And so: 1 1 1 1 1 1 = · (x · ) = ( · x) · = . x, p x, p x, q x, p x, q x, q

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 51 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Uniqueness of Identity Proofs (UIP) Isn’t UIP derivable??

a, b : A q0, q1 : a = b UIP t : q0 = q1 for some term t. The intuition of the type a = b is that the only term of this type is refl (and then a and b should be the same). UIP is equivalent to the K-rule: a : A q : a = a K t : q = refl a a for some term t. This rule may look even more natural . . . . There is a countermodel to K (and UIP): M. Hofmann and Th. Streicher, The groupoid interpretation of type theory, 1998. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 52 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Types are topological spaces, equality proofs are paths Voevodsky

: A type A is a topological space and if a, b : A with p : a = b, then

p is a continuous path from a to b in A.

If p, q : a = b and h : p = q, then

h is a continuous transformation from p to q in A

also called a homotopy.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 53 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Equality proofs are paths, path-equalities are higher paths

Which of the following hold under the topological interpretation? Note: A property P : ∀ab : A, a = b → Prop should be closed under continuous transformations of points and paths. P : ∀ab : A, a = b → Prop r : ∀a : A, P a a refl J J r : ∀xy : A, ∀i : x = y, P x y i

a, b : A q0, q1 : a = b UIP (for some term t). t : q0 = q1

a : A q : a = a K (for some term t). t : q = refl a a

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 54 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Equality proofs are paths, derivable rules? Exercise: Which of the following hold under the topological interpretation?

P : ∀ab : A, a = b → Prop r : ∀a, b : A, ∃q : a = b, P a b q (some term t) t : ∀xy : A, ∀i : x = y, P x y i

P : ∀ab : A, a = b → Prop r : ∀a : A, ∃b : A, ∃q : a = b, P a b q (some t) t : ∀xy : A, ∀i : x = y, P x y i

P : ∀ab : A, a = b → Prop r : ∃a, b : A, ∃q : a = b, P a b q (some term t) t : ∀xy : A, ∀i : x = y, P x y i

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 55 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Homotopy Type Theory

Voevodsky’s Homotopy Type Theory (HoTT): • We need to add: Univalence Axiom: for all types A and B:

(A = B) ' (A ' B)

where A ' B denotes that A and B are isomorphic: there are f : A → B and g : B → A such that ∀x : A, g(f x) = x etc. • HoTT is the internal language for homotopy theory. All proofs in homotopy theory should be formalised in type theory. (Agda and Coq give support for that.) Challenge: Is there a computational interpretation of the axiom UA?

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 56 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Why Univalence is useful

• From UA one can prove function extensionality: f , g : A → B r : ∀a : A, f a = g a FunExt t : f = g

• UA allows to treat isomorphic data types as equal data types. If you have an algorithm (proof) about binary N, you have an algoritm (proof) about unary N.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 57 / 58 Inductive Types Proof Assistants Homotopy Type Theory Radboud University Higher Inductive Types (HITs)

An alernative to UA are higher inductive types: inductive types + path constructors. Inductive circle : Type := | base : circle | loop : base == base. with elimination and computation rules. Inductive torus : Type := | base : torus | meridian : base == base | equator : base == base | surf : meridian ◦ equator = equator ◦ meridian Challenge: What are the proper general rules for higher inductive types?

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 58 / 58