<<

Lecture 3: Typed Lambda and Curry-Howard

H. Geuvers

Radboud University Nijmegen, NL

21st Estonian Winter School in Winter 2016

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 1 / 65 Outline

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 2 / 65 Typed λ calculus as a basis for

λ- : type M:A program : proof : formula program : (full) specification

Aim: • as an integrated system for proving and programming. • Type Theory as a basis for proof assistants and interactive proving.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 3 / 65 Simple type theory

Simplest system: λ→ or simple type theory, STT. Just arrow types

Typ := TVar | (Typ → Typ)

• Examples: (α → β) → α,(α → β) → ((β → γ) → (α → γ)) • Brackets associate to the right and outside brackets are omitted: (α → β) → (β → γ) → α → γ • Types are denoted by A, B,....

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 4 / 65 Simple type theory `ala Church

Formulation with contexts to declare the free variables:

x1 : A1, x2 : A2,..., xn : An

is a , usually denoted byΓ. Derivation rules of λ→ (`ala Church):

x:A ∈ Γ Γ ` M : A → B Γ ` N : A Γ, x:A ` P : B Γ ` x : A Γ ` MN : B Γ ` λx:A.P : A → B

Γ `λ→ M : A if there is a derivation using these rules with conclusion Γ ` M : A

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 5 / 65 Examples

` λx : A.λy : B.x : A → B → A ` λx : A → B.λy : B → .λz : A.y (x z):(A→B)→(B→C)→A→C ` λx : A.λy :(B → A) → A.y(λz : B.x): A → ((B → A) → A) → A

Not for every type there is a closed term of that type:

(A → A) → A is not inhabited

That is: there is no term M such that

` M :(A → A) → A.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 6 / 65 Typed Terms versus Type Assignment

• With typed terms also called typing `ala Church, we have terms with type information in the λ-abstraction

λx : A.x : A → A

• Terms have unique types, • The type is directly computed from the type info in the variables. • With typed assignment also called typing `ala Curry, we assign types to untyped λ-terms

λx.x : A → A

• Terms do not have unique types, • A principal type can be computed using unification.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 7 / 65 Church vs. Curry typing

• The Curry formulation is especially interesting for programming: you want to write as little type information as possible; let the compiler infer the types for you. • The Church formulation is especially interesting for proof checking: terms are created interactively; type structure is so intricate that type inference is undecidable (if you start from an untyped term). [ This lecture]

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 8 / 65 Formulas-as-Types (Curry, Howard)

Recall: there are two readings of a judgement M : A 1 term as /program, type as specification: M is a of type A 2 type as a , term as its proof: M is a proof of the proposition A

• There is a one-to-one correspondence: typable terms in λ→ ' derivations in minimal proposition logic

• x1 : B1, x2 : B2,..., xn : Bn ` M : A can be read as M is a proof of A from the assumptions B1, B2,..., Bn.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 9 / 65 Example

[A → B → C]3 [A]1 [A → B]2 [A]1 B → C B C 1 ' A → C 2 (A → B) → A → C 3 (A → B → C) → (A → B) → A → C

λx:A → B → C.λy:A → B.λz:A.x z (y z) :(A → B → C) → (A → B) → A → C

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 10 / 65 Example

[x : A → B → C]3 [z : A]1 [y : A → B]2 [z : A]1 x z : B → C y z : B x z (y z): C 1 λz:A.x z (y z): A → C 2 λy:A → B.λz:A.x z (y z):( A → B) → A → C 3 λx:A → B → C.λy:A → B.λz:A.x z (y z):( A→B→C)→(A→B)→A→C Exercise: Give the derivation that corresponds to

λx:C → E.λy:(C → E) → E.y(λz.y x): (C → E) → ((C → E) → E) → E

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 11 / 65 Typed We have seen Combinatory Logic with the axioms for I, K and S. We now know their typed definition in λ→: I := λx : A.x : A → A K := λx : A.λy : B.x : A → B → A S := λx:A → B → C.λy:A → B.λz:A.x z (y z) :(A → B → C) → (A → B) → A → C

• The three axiom schemes A → A, A → B → A and (A → B → C) → (A → B) → A → C together with the derivation rule Modus Ponens is exactly Hilbert style minimal proposition logic. • The typed CL terms are exactly the derivations in this logic. • Modus Ponens corresponds with Application in CL Exercise: Show that the scheme A → A is derivable. Cast in CL terminology: I can be defined in terms of S and K. To be precise: I = SKK. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 12 / 65 = Cut-elimination

• β-reduction:(λx:A.M)P →β M[x := P] Cut-elimination in minimal logic= β-reduction in λ→.

[A]1 D D1 2 A B D2 −→ 1 D A → B A 1 B B [x : A]1 D1 D2 D P : A M : B 2 −→ 1 β D λx:A.M : A → B P : A 1 M[x := P]: B (λx:A.M)P : B

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 13 / 65 Example

Proof of A → A → B, (A → B) → A ` B with a cut.

[A]1 A → A → B [A]1 A → A → B [A]1 A → B [A]1 A → B B B (A → B) → A A → B A → B A B It contains a cut: a →-i directly followed by an →-e.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 14 / 65 Example proof with term information

[x : A]1 p : A → A → B [y : A]1 p : A → A → B [x : A]1 p x : A → B [y : A]1 p y : A → B p x x : B p y y : B q :(A → B) → A λx:A.p x x : A → B λy:A.p y y : A → B q(λx:A.p x x): A (λy:A.p y y)(q(λx:A.p x x)): B

Term contains a β-redex: (λx:A.p x x)(q(λx:A.p x x))

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 15 / 65 with other connectives

Adding product types × to λ→. (Proposition logic with conjunction ∧.)

Γ ` M : A × B Γ ` M : A × B Γ ` P : A Γ ` Q : B

Γ ` π1M : A Γ ` π2M : B Γ ` hP, Qi : A × B With reduction rules

π1hP, Qi → P

π2hP, Qi → Q

Similar rules can be given for sum-types A + B, corresponding to disjunction A ∨ B.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 16 / 65 Extension to logic

• First order language: domain D, with variables x, y, z : D and possibly functions over D, e.g. f : D → D, g : D → D → D. • Rules for ∀x:D.φ and ∃x:D.φ. • NB There are two “kinds” of variables: the first order variables (ranging over the domain D) and the “proof variables” (used as [local] assumptions of formulas). • Formulas and domain are both types. What is the type of a predicate or relation? • A predicate P is a from D to the collection of types, ∗ • P : D → ∗ for P a predicate and R : D → D → ∗ for R a binary relation on D. • We will have to make this more precise . . .

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 17 / 65 Idea of extending to ∀

Term rules for the ∀-quantifier in predicate logic.

Γ ` M : ∀x:D.A Γ ` M : A if t : D x not free in Γ Γ ` M t : A[x := t] Γ ` λx:D.M : ∀x:D.A With the usual β-reduction rule

(λx:D.M)t → M[x := t]

. This conforms with cut-elimination (or “detour elimination”) on logical derivations.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 18 / 65 Example

Deriving irreflexivity from anti-symmetry

AntiSym R := ∀x, y:D.(Rxy) → (Ryx) → ⊥ Irrefl R := ∀x:D.(Rxx) → ⊥

Derivation in predicate logic: ∀x, y:D.R x y → R y x → ⊥ ∀y:D.R x y → R y x → ⊥ R x x → R x x → ⊥ [R x x]1 R x x → ⊥ [R x x]1 ⊥ 1 R x x → ⊥ ∀x:D.R x x → ⊥

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 19 / 65 Example derivation in type theory, with terms

H : ∀x, y:D.R x y → R y x → ⊥ H x : ∀y:D.R x y → R y x → ⊥ H x x : R x x → R x x → ⊥ [H0 : R x x]1 H x x H0 : R x x → ⊥ [H0 : R x x]1 H x x H0 H0 : ⊥ 1 λH0:(R x x).H x x H0 H0 : R x x → ⊥ λx:A.λH0:(R x x).H x x H0 H0 : ∀x:D.R x x → ⊥

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 20 / 65 Theory

• We have seen informally “dependent types at work” in the predicate logic example. • Now: the rules

With dependent types: • everything depends on everything • we can’t first define the types, then the terms • two universes: ∗ and  • ∗ is the universe of types • We can’t have ∗ : ∗, so we have another universe: ∗ : . NB The Coq system uses “Set” and “Prop” for what I call ∗ and “Type” for what I call .

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 21 / 65 First order Dependent Type theory, λP Derive judgements of the form

Γ ` M : B

• Γ is a context

x1 : B1, x2 : B2,..., xn : Bn

• M and B are terms taken from the set of pseudoterms

T ::= Var | ∗ |  | (T T) | (λx:T.T) | Πx:T.T

Auxiliary judgement Γ ` denoting that Γ is a correct context. H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 22 / 65 Derivation rules of λP

s ranges over {∗, }. Γ ` A : s Γ ` (base) ∅ ` (ctxt) if x not in Γ (ax) Γ, x:A ` Γ `∗ :  Γ ` Γ ` A : ∗ Γ, x:A ` B : s (proj) if x:A ∈ Γ (Π) Γ ` x : A Γ ` Πx:A.B : s

Γ, x:A ` M : B Γ ` Πx:A.B : s Γ ` M :Πx:A.B Γ ` N : A (λ) (app) Γ ` λx:A.M :Πx:A.B Γ ` MN : B[x := N]

Γ ` M : B Γ ` A : s (conv) A =βη B Γ ` M : A Notation: write A → B for Πx:A.B if x ∈/ FV(B).

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 23 / 65 The use of the Π-type

• The Π rule allows to form two forms of function types. Γ, x:A ` B : s Γ ` A : ∗ (Π) Γ ` Πx:A.B : s

Πx:A.B '{f | ∀a : A(f a : B[x := a])} Write A → B if x ∈/ FV(B) • With s = ∗, we can formD → D andΠ x:D.x = x, etc. • With s = , we can formD → D → ∗ andD → ∗.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 24 / 65 Representation of PRED (minimal predicate logic) into λP

Represent both the domains of the logic and the formulas as types.

A: ∗, P :A → ∗, R :A → A → ∗,

Now implication is represented as → and ∀ is represented as Π:

∀x:A.P x 7→ Πx:A.P x

Intro and elim rules are just λ-abstraction and application

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 25 / 65 Example

A: ∗, R :A → A → ∗ ` λz:A.λh:(Πx, y:A.R x y).h z z :Πz:A.(Πx, y:A.R x y) → R z z

This term is a proof of ∀z:A.(∀x, y:A.R(x, y)) → R(z, z) Exercise: Find terms of the following types (NB → binds strongest)

(Πx:A.P x → Q x) → (Πx:A.P x) → Πx:A.Q x

and

(Πx:A.P x → Πz.R z z) → (Πx:A.P x) → Πz:A.R z z).

Also write down the contexts in which these terms are typed.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 26 / 65 Direct embedding of logic in type theory

For λ→ and λP we have seen

Direct representations of logic in type theory.

• Connectives each have a counterpart in the type theory: implication ∼→ -type universal quantification ∼∀ -type • Logical rules have their direct counterpart in type theory λ-abstraction ∼→ -introduction application ∼→ - elimination λ-abstraction ∼∀ -introduction application ∼∀ -elimination • Context declares signature, local varibales and assumptions.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 27 / 65 LF embedding of logic in type theory

Second way of interpreting logic in type theory De Bruijn:

Logical framework encoding of logic in type theory.

• Type theory used as a meta system for encoding ones own logic.

• Choose an appropriate context ΓL, in which the logic L (including its proof rules) is declared. • Context used as a signature for the logic. • Use the as the ‘meta’ calculus for dealing with and .

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 28 / 65 Direct and LF embedding

proof formula direct embedding λx:A.x A → A LF embedding imp intr AA λx:TA.x T (A ⇒ A) • Direct representation: One type system : One logic, Logical rules ∼ type theoretic rules • LF encoding One type system : Many , Logical rules ∼ context declarations

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 29 / 65 Examples of the Deep embedding

The encoding of logics in a logical framework is shown by three examples: 1 Minimal proposition logic 2 Minimal predicate logic (just {⇒, ∀}) 3 Untyped λ-calculus

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 30 / 65 Minimal propositional logic

Fix the signature (context) of minimal propositional logic.

prop : ∗ imp: prop → prop → prop

Notation: A ⇒ B for imp AB The type prop is the type of ‘names’ of . NB : A term of type propcan not be inhabited (proved), as it is not a type. We ‘lift’ a name p : prop to the type of its proofs by introducing the following map:

T: prop → ∗.

Intended meaning ofT p is ‘the type of proofs of p’. We interpret ‘p is valid’ by ‘Tp is inhabited’.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 31 / 65 Encoding of derivations

To derive Tp we also encode the logical derivation rules

imp intr:Π p, q : prop.(Tp → Tq) → T(p ⇒ q), imp el:Π p, q : prop.T(p ⇒ q) → Tp → Tq.

New phenomenon: Π-type:

Πx:A.B(x) ' the type of functions f such that f a : B(a) for all a:A

imp intr takes two (names of) propositions p and q and a term f :T p → T q and returns a term of type T(p ⇒ q) Indeed A ⇒ A, becomes valid:

imp intrAA(λx:T A.x) : T(A ⇒ A)

Exercise: Construct a term of type T(A ⇒ (B ⇒ A))

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 32 / 65 Signature of PROP in LF To encode proposition logic in LF we need a context (signature) ΣPROP: prop : ∗ ⇒ : prop → prop → prop T: prop → ∗ imp intr : (A, B : prop)(T A → T B) → T(A ⇒ B) imp el : (A, B : prop)T(A ⇒ B) → T A → T B.

Desired properties of the encoding: • Adequacy(soundness) of the encoding:

`PROP A =⇒ ΣPROP, a1:prop,..., an:prop ` p :T A for some p.

{a, ..., an} is the set of proposition variables in A. • Faithfulness (or completeness) is the converse. It also holds, but more involved to prove.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 33 / 65 Minimal predicate logic over one domain A

Signature:

prop : ∗, A: ∗, T: prop → ∗ f : A → A, R:A → A → prop, ⇒ : prop → prop → prop, imp intr : Πp, q : prop.(Tp → Tq) → T(p ⇒ q), imp el : Πp, q : prop.T(p ⇒ q) → Tp → Tq.

Now encode ∀: ∀ takes a P :A → prop and returns a proposition, so we add: ∀ : (A → prop) → prop

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 34 / 65 Minimal predicate logic over one domain A

Signature: ΣPRED prop : ∗, A: ∗, . . imp intr : Πp, q : prop.(Tp → Tq) → T(p ⇒ q), imp el : Πp, q : prop.T(p ⇒ q) → Tp → Tq.

Now encode ∀: ∀ takes a P :A → prop and returns a proposition, so: ∀ : (A → prop) → prop Universal quantification is translated as follows.

∀x:A.(Px) 7→∀ (λx:A.(Px))

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 35 / 65 Intro and Elim rules for ∀

∀ : (A → prop) → prop, ∀ intr : ΠP:A → prop.(Πx:A.T(Px)) → T(∀P), ∀ elim : ΠP:A → prop.T(∀P) → Πx:A.T(Px).

The proof of ∀z:A(∀x, y:A.Rxy) ⇒ Rzz is now mirrored by the proof-term

∀ intr[ ]( λz:A.imp intr[ ][ ](λh:T(∀x, y:A.Rxy). ∀ elim[ ](∀ elim[ ]hz)z))

We have replaced the instantiations of the Π-type by [ ]. This term is of type

T(∀(λz:A.imp(∀(λx:A.(∀(λy:A.Rxy))))(Rzz)))

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 36 / 65 Intro and Elim rules for ∀

∀ : (A → prop) → prop, ∀ intr : ΠP:A → prop.(Πx:A.T(Px)) → T(∀P), ∀ elim : ΠP:A → prop.T(∀P) → Πx:A.T(Px).

The proof of ∀z:A(∀x, y:A.Rxy) ⇒ Rzz is now mirrored by the proof-term

∀ intr[ ]( λz:A.imp intr[ ][ ](λh:T(∀x, y:A.Rxy). ∀ elim[ ](∀ elim[ ]hz)z))

Exercise: Construct a proof-term that mirrors the (obvious) proof of ∀x(P x ⇒ Q x) ⇒ ∀x.P x ⇒ ∀x.Q x

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 37 / 65 Untyped λ-calculus

SignatureΣ lambda :D: ∗; app : D → (D → D); abs : (D → D) → D.

• A x in λ-calculus becomes x : D in the type system. • The translation [−]:Λ → Term(D) is defined as follows.

[x] = x; [PQ] = app [P][Q]; [λx.P] = abs (λx:D.[P]).

Examples: [λx.xx] := abs(λx:D.app x x) [(λx.xx)(λy.y)] := app(abs(λx:D.app x x))(abs(λy:D.y)).

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 38 / 65 Introducing β-equality

eq:D → D → ∗.

Notation P = Q for eq PQ. Rules for proving equalities.

refl:Π x:D.x = x, sym:Π x, y:D.x = y → y = x, trans:Π x, y, z:D.x = y → y = z → x = z, mon:Π x, x0, z, z0:D.x = x0 → z = z0 → (app z x) = (app z0 x0), xi:Π f , g:D → D.(Πx:D.(fx) = (gx)) → (abs f ) = (abs g), beta:Π f :D → D.Πx:D.(app(abs f )x) = (fx).

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 39 / 65 Properties of λP

• Uniqueness of types If Γ ` M : σ and Γ ` M : τ, then σ=βητ. • Subject Reduction If Γ ` M : σ and M →βη N, then Γ ` N : σ. • Strong Normalization If Γ ` M : σ, then all βη-reductions from M terminate. Proof of SN is by defining a reduction preserving map from λP to λ→.

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 40 / 65 Questions

Γ ` M : σ? TCP Γ ` M :? TSP Γ `?: σ TIP For λP: • TIP is undecidable • TCP/TSP: simultaneously with Context checking

H. Geuvers - Radboud Univ. EWSCS 2016 Typed λ-calculus 41 / 65 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 ∼ . . . + exceptions

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