Coq support for HoTT Matthieu Sozeau Inria Paris & PPS, Universit´eParis 7 Diderot

HoTT/UF Workshop June 29th 2015 Warsaw, Poland Embracing proof-relevant

Much work on DTT in Coq focused on propositional equality, assumed proof-irrelevant.

Two examples which rely deeply on equality: I A generalized tactic I A toolbox for handling definitions by dependent pattern-matching and well-founded recursion

Today: Adapting these tools to the new setting

Matthieu Sozeau - Coq support for HoTT 2 Coq support for HoTT

1 Proof-relevant rewriting strategies Generalized rewriting Rewriting with Type-valued relations

2 Equations Intro Dependent pattern-matching compilation Recursion Reasoning support

Matthieu Sozeau - Coq support for HoTT 3 Generalized Rewriting

Why generalized rewriting when we have Id-elimination?

I Id is not the only interesting ...

I Even with a univalent equality, Id-elim is not enough: capturing rewrites under binders.

Matthieu Sozeau - Coq support for HoTT 4 Goal ∀ APQ ,( ∀ x : A, Px ↔ Qx ) → (∃ x, ¬ Px ) → (∃ x, ¬ Qx ). Proof. intros APQH HnP . setoid rewrite ← H . exact HnP. Qed.

Higher-order morphisms

Inductive ex {A : Type} (P : A → Prop): Prop := ex intro: ∀ x : A, Px → ex P. Instance ex iff P A : Proper( pointwise relation A iff ++> iff) (@ex A).

Matthieu Sozeau - Coq support for HoTT 5 Proof. intros APQH HnP . setoid rewrite ← H . exact HnP. Qed.

Higher-order morphisms

Inductive ex {A : Type} (P : A → Prop): Prop := ex intro: ∀ x : A, Px → ex P. Instance ex iff P A : Proper( pointwise relation A iff ++> iff) (@ex A).

Goal ∀ APQ ,( ∀ x : A, Px ↔ Qx ) → (∃ x, ¬ Px ) → (∃ x, ¬ Qx ).

Matthieu Sozeau - Coq support for HoTT 5 Higher-order morphisms

Inductive ex {A : Type} (P : A → Prop): Prop := ex intro: ∀ x : A, Px → ex P. Instance ex iff P A : Proper( pointwise relation A iff ++> iff) (@ex A).

Goal ∀ APQ ,( ∀ x : A, Px ↔ Qx ) → (∃ x, ¬ Px ) → (∃ x, ¬ Qx ). Proof. intros APQH HnP . setoid rewrite ← H . exact HnP. Qed.

Matthieu Sozeau - Coq support for HoTT 5 One can build a set of combinators to rewrite in depth: HOL conversions [Paulson 83], ELAN strategies.

I Congruence. ap : Π AB (f : A → B)(xy : A), x = y → fx = fy  Applies at the toplevel only  Smaller proof term: mentions the changed terms only  Generalizes to n-ary, parallel rewriting  Still restricted to equality

Rewriting in Type Theory

Moving from substitution to congruence.

I Built-in substitution: Leibniz equality. Π A (P : A → Type)(xy : A), Px → x = y → Py .  Applies to any context  Iterated rewrites result in large proof terms: repeats the context that depends on x  Restricted to equality, one rewrite at a time

Matthieu Sozeau - Coq support for HoTT 6 One can build a set of combinators to rewrite in depth: HOL conversions [Paulson 83], ELAN strategies.

Rewriting in Type Theory

Moving from substitution to congruence.

I Built-in substitution: Leibniz equality. Π A (P : A → Type)(xy : A), Px → x = y → Py .  Applies to any context  Iterated rewrites result in large proof terms: repeats the context that depends on x  Restricted to equality, one rewrite at a time

I Congruence. ap : Π AB (f : A → B)(xy : A), x = y → fx = fy  Applies at the toplevel only  Smaller proof term: mentions the changed terms only  Generalizes to n-ary, parallel rewriting  Still restricted to equality

Matthieu Sozeau - Coq support for HoTT 6 Rewriting in Type Theory

Moving from substitution to congruence.

I Built-in substitution: Leibniz equality. Π A (P : A → Type)(xy : A), Px → x = y → Py .  Applies to any context  Iterated rewrites result in large proof terms: repeats the context that depends on x  Restricted to equality, one rewrite at a time

I Congruence. ap : Π AB (f : A → B)(xy : A), x = y → fx = fy  Applies at the toplevel only  Smaller proof term: mentions the changed terms only  Generalizes to n-ary, parallel rewriting  Still restricted to equality One can build a set of combinators to rewrite in depth: HOL conversions [Paulson 83], ELAN strategies.

Matthieu Sozeau - Coq support for HoTT 6 Requires proof search:

I Heuristic in NuPRL based on subrelations (impl ⊂ iff)

I Complete procedure in Coq.

Generalized Rewriting in Type Theory

D. Basin [NuPRL, 94], C. Sacerdoti Coen [Coq, 04], Sozeau [Coq, 09]

I Generalized to any relation Proper( iff ++> iff) not , Π PQ , P ↔ Q → ¬ P ↔ ¬ Q I Multiple signatures for a given constant Proper( impl −→ impl) not

Matthieu Sozeau - Coq support for HoTT 7 Generalized Rewriting in Type Theory

D. Basin [NuPRL, 94], C. Sacerdoti Coen [Coq, 04], Sozeau [Coq, 09]

I Generalized to any relation Proper( iff ++> iff) not , Π PQ , P ↔ Q → ¬ P ↔ ¬ Q I Multiple signatures for a given constant Proper( impl −→ impl) not Requires proof search:

I Heuristic in NuPRL based on subrelations (impl ⊂ iff)

I Complete procedure in Coq.

Matthieu Sozeau - Coq support for HoTT 7 Our Algorithm

Two phases: 1 Constraint generation in ML. Recursive descent on the term to find the rewrites to perform, building a proof skeleton. 2 Constraint solving using resolution. Depth-first search to solve the constraints with the declared hints / typeclass instances.

Matthieu Sozeau - Coq support for HoTT 8 Unify

0 0 unifyρ(Γ, ψ, t) ⇑ ψ , ρ : R t u R 0 Γ | ψ ` t ρ0 u a ψ Atom ∗ unifyρ(Γ, ψ, t) ⇓ τ , type(Γ, ψ, t) 0 ψ , {?S :Γ ` relation τ, ?m :Γ ` Proper τ ?S t} Γ | ψ ` t ?S t a ψ ∪ ψ0 ?m

The two main rules

R 0 0 Γ | ψ ` τ p τ a ψ

Matthieu Sozeau - Coq support for HoTT 9 Atom ∗ unifyρ(Γ, ψ, t) ⇓ τ , type(Γ, ψ, t) 0 ψ , {?S :Γ ` relation τ, ?m :Γ ` Proper τ ?S t} Γ | ψ ` t ?S t a ψ ∪ ψ0 ?m

The two main rules

R 0 0 Γ | ψ ` τ p τ a ψ

Unify

0 0 unifyρ(Γ, ψ, t) ⇑ ψ , ρ : R t u R 0 Γ | ψ ` t ρ0 u a ψ

Matthieu Sozeau - Coq support for HoTT 9 The two main rules

R 0 0 Γ | ψ ` τ p τ a ψ

Unify

0 0 unifyρ(Γ, ψ, t) ⇑ ψ , ρ : R t u R 0 Γ | ψ ` t ρ0 u a ψ Atom ∗ unifyρ(Γ, ψ, t) ⇓ τ , type(Γ, ψ, t) 0 ψ , {?S :Γ ` relation τ, ?m :Γ ` Proper τ ?S t} Γ | ψ ` t ?S t a ψ ∪ ψ0 ?m

Matthieu Sozeau - Coq support for HoTT 9 I Algebraic presentation, supporting higher-order functions and polymorphism:

Π ABCR 0 R1 R2 , Proper(( R1 ++> R2 ) ++> (R0 ++> R1 ) ++> (R0 ++> R2 )) (@compose ABC )

I Generic morphism declarations.

I Subrelations, quotienting the signatures.

I Rewriting on operators/functions, parallel rewrites. . .

Proper instances

I Extensible signatures (shallow embedding) all : ∀ A : Type,(A → Prop) → Prop Π A, Proper( pointwise relation A iff ++> iff) (@all A)

Matthieu Sozeau - Coq support for HoTT 10 I Generic morphism declarations.

I Subrelations, quotienting the signatures.

I Rewriting on operators/functions, parallel rewrites. . .

Proper instances

I Extensible signatures (shallow embedding) all : ∀ A : Type,(A → Prop) → Prop Π A, Proper( pointwise relation A iff ++> iff) (@all A)

I Algebraic presentation, supporting higher-order functions and polymorphism:

Π ABCR 0 R1 R2 , Proper(( R1 ++> R2 ) ++> (R0 ++> R1 ) ++> (R0 ++> R2 )) (@compose ABC )

Matthieu Sozeau - Coq support for HoTT 10 I Subrelations, quotienting the signatures.

I Rewriting on operators/functions, parallel rewrites. . .

Proper instances

I Extensible signatures (shallow embedding) all : ∀ A : Type,(A → Prop) → Prop Π A, Proper( pointwise relation A iff ++> iff) (@all A)

I Algebraic presentation, supporting higher-order functions and polymorphism:

Π ABCR 0 R1 R2 , Proper(( R1 ++> R2 ) ++> (R0 ++> R1 ) ++> (R0 ++> R2 )) (@compose ABC )

I Generic morphism declarations.

Matthieu Sozeau - Coq support for HoTT 10 I Rewriting on operators/functions, parallel rewrites. . .

Proper instances

I Extensible signatures (shallow embedding) all : ∀ A : Type,(A → Prop) → Prop Π A, Proper( pointwise relation A iff ++> iff) (@all A)

I Algebraic presentation, supporting higher-order functions and polymorphism:

Π ABCR 0 R1 R2 , Proper(( R1 ++> R2 ) ++> (R0 ++> R1 ) ++> (R0 ++> R2 )) (@compose ABC )

I Generic morphism declarations.

I Subrelations, quotienting the signatures.

Matthieu Sozeau - Coq support for HoTT 10 Proper instances

I Extensible signatures (shallow embedding) all : ∀ A : Type,(A → Prop) → Prop Π A, Proper( pointwise relation A iff ++> iff) (@all A)

I Algebraic presentation, supporting higher-order functions and polymorphism:

Π ABCR 0 R1 R2 , Proper(( R1 ++> R2 ) ++> (R0 ++> R1 ) ++> (R0 ++> R2 )) (@compose ABC )

I Generic morphism declarations.

I Subrelations, quotienting the signatures.

I Rewriting on operators/functions, parallel rewrites. . .

Matthieu Sozeau - Coq support for HoTT 10 Coq support for HoTT

1 Proof-relevant rewriting strategies Generalized rewriting Rewriting with Type-valued relations

2 Equations Intro Dependent pattern-matching compilation Recursion Reasoning support Equations Elimination principle Eliminating calls

Matthieu Sozeau - Coq support for HoTT 11 Impredicativity helped

All fine with relations in Prop, how about Type-valued relations?

Proper :ΠA : Typei, (A → A → Typej) → A → Typej.

Need to show, under A : Typei:

Proper ((A → A → Typej) → A → Typej) (iso rel A −→ eq A −→ iso) (Proper A)

Requires: Typemax(i,j+1) ≤ Typei i.e. j < i. But then iso A : Typei 6< Typej ⇒ inconsistency.

Matthieu Sozeau - Coq support for HoTT 12 Universe Polymorphism

With universe polymorphism (Sozeau & Tabareau [ITP’14]):

Properi j :ΠA : Typei, (A → A → Typej) → A → Typej

We can show, under A : Typei:

Properi0 j0 ((A → A → Typej) → A → Typej) (iso rel A −→ eq A −→ iso) (Properi j A) The constraint max(i, j + 1) ≤ i0 is satisfiable.

Actually, crelation(A : Typei) := A → A → Typej is already problematic: no relation equivalence or subrelation definition possible.

Matthieu Sozeau - Coq support for HoTT 13 Computational relations

Generalized rewriting will now handle:

I The general function space morphism between types.

I Type-level identity, and equivalences of types

I Computationally relevant relations like CoRN’s appartness relation on reals.

I Hom-types of categories which are not Prop-based setoids, e.g. groupoids. . .

Matthieu Sozeau - Coq support for HoTT 14 Automation

User-definable strategies: rewrite strat.

I A tactic/strategy language: bottom-up, innermost, with composition, disjunction, rewrite hint databases. . .

I Much faster than autorewrite.

I More control on the shape of proof terms.

Matthieu Sozeau - Coq support for HoTT 15 An example

Suppose the theory of monoids on T. A goal: x y : T |- x • (( • y) • ).

I autorewrite with monoids will do two rewrites with both unit laws, the proof term will be roughtly twice the goal size.

I rewrite_strat (topdown (repeat (hints monoids))) will first rewrite  • y to y and directly after, y •  to y, resulting in a proof term of size roughly that of the initial goal, and will be twice as fast as well.

Matthieu Sozeau - Coq support for HoTT 16 Coq support for HoTT

1 Proof-relevant rewriting strategies Generalized rewriting Rewriting with Type-valued relations

2 Equations Intro Dependent pattern-matching compilation Recursion Reasoning support Equations Elimination principle Eliminating calls

Matthieu Sozeau - Coq support for HoTT 17 Dependent Pattern-Matching

A word of caution:

I Strong elimination: match x return if x then unit else nat with | true => tt | false => 0 end ⇒ intensional character of inductive values.

I Dependent Pattern-Matching: match (v : vector bool (S 0)) return bool with | Vcons a ?(0) v => a end ⇒ extends to the propositional equational theory of inductive types.

Matthieu Sozeau - Coq support for HoTT 18 Dependent Pattern-Matching and Equality

1992 Pattern Matching With Dependent Types – Coquand 1999 Dependently Typed Functional Programs and Their Proofs – McBride 2004 The View From The Left – McBride and McKinna 2006 Eliminating Dependent Pattern-Matching – Goguen, McBride and McKinna. ≈ 2010 GADTs in the programming languages community. 2014 Pattern Matching without K – Cockx, Devriese and Piessens

Matthieu Sozeau - Coq support for HoTT 19 Overview of Equations

http://github.com/mattam82/Coq-Equations (opam package coq:equations)

I Agda/Epigram-style definitions (including with)

I Purely logical handling of recursion.

I Propositional equations for definitional equalities and rewriting.

I Function graph and elimination principle derivation (w/ support for applying it). Entirely elaborated to the vanilla kernel!

Matthieu Sozeau - Coq support for HoTT 20 DEMO

Matthieu Sozeau - Coq support for HoTT 21 2 Translation from the splitting tree to Coq terms with holes 3 Proofs of the obligations using a mix of ML and Ltac code 4 Derivation of auxilliary structures from the completed splitting tree

Heavily inspired by Goguen, McBride and McKinna (2006) and Norell (2007).

Compilation setup

Elaboration into CIC + K (as an axiom or a user-provided proof)

1 Generation of a splitting tree from the clauses

Matthieu Sozeau - Coq support for HoTT 22 3 Proofs of the obligations using a mix of ML and Ltac code 4 Derivation of auxilliary structures from the completed splitting tree

Heavily inspired by Goguen, McBride and McKinna (2006) and Norell (2007).

Compilation setup

Elaboration into CIC + K (as an axiom or a user-provided proof)

1 Generation of a splitting tree from the clauses 2 Translation from the splitting tree to Coq terms with holes

Matthieu Sozeau - Coq support for HoTT 22 4 Derivation of auxilliary structures from the completed splitting tree

Compilation setup

Elaboration into CIC + K (as an axiom or a user-provided proof)

1 Generation of a splitting tree from the clauses 2 Translation from the splitting tree to Coq terms with holes 3 Proofs of the obligations using a mix of ML and Ltac code

Heavily inspired by Goguen, McBride and McKinna (2006) and Norell (2007).

Matthieu Sozeau - Coq support for HoTT 22 Compilation setup

Elaboration into CIC + K (as an axiom or a user-provided proof)

1 Generation of a splitting tree from the clauses 2 Translation from the splitting tree to Coq terms with holes 3 Proofs of the obligations using a mix of ML and Ltac code 4 Derivation of auxilliary structures from the completed splitting tree

Heavily inspired by Goguen, McBride and McKinna (2006) and Norell (2007).

Matthieu Sozeau - Coq support for HoTT 22 Searching for a splitting tree

pattern p ::= x | C −→p | ?(t) context map c ::= ∆ ` −→p :Γ splitting spl ::= Split(c, x, (spl?)n) | Compute(c, rhs) node rhs ::= Program(t) | Refine(c, t, spl)

Goal −−−→ Starting with f∆: τ := −→p . . ., find a covering of the context map ∆ ` ∆ : ∆ by −→p .

Matthieu Sozeau - Coq support for HoTT 23 Proof search example

Overlapping clauses with first-match semantics.

Equations equal (nm : nat): { n = m } + { n 6= m } := equal OO := left eq refl; equal (S n)(S m) with equal nm := { equal (S n)(S ?(n)) (left eq refl) := left eq refl; equal (S n)(S m)(right p) := right } ; equal xy := right .

Split(n m : nat ` n m : n m : nat, n,[ Split(m : nat ` O m : n m : nat, m,[ Compute(` OO : n m : nat, Program(left eq refl)), Compute(m : nat ` O (S m): n m : nat, Program(right ))]), Split(n m : nat ` (S n) m : n m : nat, m,[ Compute(n : nat ` (S n) O : n m : nat, . . . ), Compute(n m : nat ` (S n)(S m): n m : nat, Refine(equal nm , idsubst(n m : nat, x : {n = m} + {n 6= m}), `). . . )])])

Matthieu Sozeau - Coq support for HoTT 24 I Split(c, x, s): witnessed by dependent elimination. dependent destruction, using JMeq or user-given K/hSet proofs.

I Program(t): witnessed by t (w/ some substitution). I Refine(t, c, s): witnessed by: 1 inserting a let-definition in the context, 2 strengthening it, 3 abstracting it and clearing its body, 4 applying the compiled term for the subprogram with one additional variable.

Translation from the splitting to Coq

For each node ∆ ` ps :Γ Π ∆, fcomp ps.

Matthieu Sozeau - Coq support for HoTT 25 I Program(t): witnessed by t (w/ some substitution). I Refine(t, c, s): witnessed by: 1 inserting a let-definition in the context, 2 strengthening it, 3 abstracting it and clearing its body, 4 applying the compiled term for the subprogram with one additional variable.

Translation from the splitting to Coq

For each node ∆ ` ps :Γ Π ∆, fcomp ps. I Split(c, x, s): witnessed by dependent elimination. dependent destruction, using JMeq or user-given K/hSet proofs.

Matthieu Sozeau - Coq support for HoTT 25 I Refine(t, c, s): witnessed by: 1 inserting a let-definition in the context, 2 strengthening it, 3 abstracting it and clearing its body, 4 applying the compiled term for the subprogram with one additional variable.

Translation from the splitting to Coq

For each node ∆ ` ps :Γ Π ∆, fcomp ps. I Split(c, x, s): witnessed by dependent elimination. dependent destruction, using JMeq or user-given K/hSet proofs.

I Program(t): witnessed by t (w/ some substitution).

Matthieu Sozeau - Coq support for HoTT 25 Translation from the splitting to Coq

For each node ∆ ` ps :Γ Π ∆, fcomp ps. I Split(c, x, s): witnessed by dependent elimination. dependent destruction, using JMeq or user-given K/hSet proofs.

I Program(t): witnessed by t (w/ some substitution). I Refine(t, c, s): witnessed by: 1 inserting a let-definition in the context, 2 strengthening it, 3 abstracting it and clearing its body, 4 applying the compiled term for the subprogram with one additional variable.

Matthieu Sozeau - Coq support for HoTT 25 1 Proof-relevant rewriting strategies Generalized rewriting Rewriting with Type-valued relations

2 Equations Intro Dependent pattern-matching compilation Recursion Reasoning support Equations Elimination principle Eliminating calls

Matthieu Sozeau - Coq support for HoTT 26 In comparison with sized types (e.g. Agda’s size annotations):  More general.  Avoid extending the type system and the metatheory. / Relies on the reduction of a well-foundedness proof, necessary for SN. In turn, relies on logical information on the computational behavior to be available.

Recursion

I Syntactic guardness checks are fragile (and buggy)

I Incompatible with abstraction/modularity

I In Coq’s case, restricted to structural recursion on a single argument

Idea Use the logic and well-founded recursion instead!

Matthieu Sozeau - Coq support for HoTT 27 Recursion

I Syntactic guardness checks are fragile (and buggy)

I Incompatible with abstraction/modularity

I In Coq’s case, restricted to structural recursion on a single argument

Idea Use the logic and well-founded recursion instead!

In comparison with sized types (e.g. Agda’s size annotations):  More general.  Avoid extending the type system and the metatheory. / Relies on the reduction of a well-foundedness proof, necessary for SN. In turn, relies on logical information on the computational behavior to be available.

Matthieu Sozeau - Coq support for HoTT 27 I General definition of direct subterm: Isubfull : Π ∆l ∆r, I ∆l → I ∆r → Prop I Wrap the inductive type in a sigma and define an homogeneous relation on the sigma type: Isub : relation (Σ ∆, I ∆) I Extracts efficiently to a general fixpoint (assuming accessibility is defined in Prop).

Moving to Type is easy (thoughI subfull is not always a proposition).

Subterm relations and well-founded recursion

Well-founded recursion on the subterm relation for inductive familiesI : Π ∆, Type.

Matthieu Sozeau - Coq support for HoTT 28 I Wrap the inductive type in a sigma and define an homogeneous relation on the sigma type: Isub : relation (Σ ∆, I ∆) I Extracts efficiently to a general fixpoint (assuming accessibility is defined in Prop).

Subterm relations and well-founded recursion

Well-founded recursion on the subterm relation for inductive familiesI : Π ∆, Type.

I General definition of direct subterm: Isubfull : Π ∆l ∆r, I ∆l → I ∆r → Prop

Moving to Type is easy (thoughI subfull is not always a proposition).

Matthieu Sozeau - Coq support for HoTT 28 I Extracts efficiently to a general fixpoint (assuming accessibility is defined in Prop).

Subterm relations and well-founded recursion

Well-founded recursion on the subterm relation for inductive familiesI : Π ∆, Type.

I General definition of direct subterm: Isubfull : Π ∆l ∆r, I ∆l → I ∆r → Prop I Wrap the inductive type in a sigma and define an homogeneous relation on the sigma type: Isub : relation (Σ ∆, I ∆)

Moving to Type is easy (thoughI subfull is not always a proposition).

Matthieu Sozeau - Coq support for HoTT 28 Subterm relations and well-founded recursion

Well-founded recursion on the subterm relation for inductive familiesI : Π ∆, Type.

I General definition of direct subterm: Isubfull : Π ∆l ∆r, I ∆l → I ∆r → Prop I Wrap the inductive type in a sigma and define an homogeneous relation on the sigma type: Isub : relation (Σ ∆, I ∆) I Extracts efficiently to a general fixpoint (assuming accessibility is defined in Prop).

Moving to Type is easy (thoughI subfull is not always a proposition).

Matthieu Sozeau - Coq support for HoTT 28 Inductive vector direct subterm (A : Type) : ∀ nn 0 : nat, vector An → vector An 0 → Prop := vector direct subterm 1 1 : ∀ (h : A)(n : nat)(v : vector An ), vector direct subtermAn (Sn ) v (Vcons hv )

Check vector subterm : ∀ A : Type, relation {n : nat & vector An }. Equations unzip {AB }{n} (v : vector (A × B) n) : vector An × vector Bn := unzip ABnv by recv := unzip AB ?(O) nil := (nil, nil) ; unzip AB ?(Sn )(cons (pairxy ) nv) with unzip v := { | (pair xs ys) := (cons x xs, cons y ys) }.

Example: vectors

Derive Signature for vector. Derive Subterm for vector.

Matthieu Sozeau - Coq support for HoTT 29 Equations unzip {AB }{n} (v : vector (A × B) n) : vector An × vector Bn := unzip ABnv by recv := unzip AB ?(O) nil := (nil, nil) ; unzip AB ?(Sn )(cons (pairxy ) nv) with unzip v := { | (pair xs ys) := (cons x xs, cons y ys) }.

Example: vectors

Derive Signature for vector. Derive Subterm for vector.

Inductive vector direct subterm (A : Type) : ∀ nn 0 : nat, vector An → vector An 0 → Prop := vector direct subterm 1 1 : ∀ (h : A)(n : nat)(v : vector An ), vector direct subtermAn (Sn ) v (Vcons hv )

Check vector subterm : ∀ A : Type, relation {n : nat & vector An }.

Matthieu Sozeau - Coq support for HoTT 29 Example: vectors

Derive Signature for vector. Derive Subterm for vector.

Inductive vector direct subterm (A : Type) : ∀ nn 0 : nat, vector An → vector An 0 → Prop := vector direct subterm 1 1 : ∀ (h : A)(n : nat)(v : vector An ), vector direct subtermAn (Sn ) v (Vcons hv )

Check vector subterm : ∀ A : Type, relation {n : nat & vector An }. Equations unzip {AB }{n} (v : vector (A × B) n) : vector An × vector Bn := unzip ABnv by recv := unzip AB ?(O) nil := (nil, nil) ; unzip AB ?(Sn )(cons (pairxy ) nv) with unzip v := { | (pair xs ys) := (cons x xs, cons y ys) }.

Matthieu Sozeau - Coq support for HoTT 29 Vectors with decidable equality

Using dependent elimination on decidable indices. Equations unzip dec {AB } ‘{EqDec A} ‘{EqDec B} {n} (v : vector (A × B) n): vector An × vector Bn := unzip dec AB nv by recv := unzip dec AB ?(O) nil := (nil, nil) ; unzip dec AB ?(Sn )(cons (pairxy ) nv) with unzip dec v := { | pair xs ys := (cons x xs, cons y ys) }. Print Assumptions unzip dec. Closed under the global context

Matthieu Sozeau - Coq support for HoTT 30 1 Proof-relevant rewriting strategies Generalized rewriting Rewriting with Type-valued relations

2 Equations Intro Dependent pattern-matching compilation Recursion Reasoning support Equations Elimination principle Eliminating calls

Matthieu Sozeau - Coq support for HoTT 31 Equations

I Equations derived from the splitting tree hold definitionally in CCI (assuming no use of K).

I Equations for with nodes are just proxies to helper functions.

I All put together in a rewrite database, f can be opacified.

I For well-founded definitions, fixpoint unfolding lemma. This requires funext and showing that accessibility is an hProp.

Matthieu Sozeau - Coq support for HoTT 32 Generating a function graph

Equations filter {A} (l : list A)(p : A → bool): list A := filter A nil p := nil; filter A (cons al ) p with pa := { | true := a :: filter lp ; | false := filter lp }.

Matthieu Sozeau - Coq support for HoTT 33 Generated mutual induction principle

Check(filter ind mut : ∀ (P : ∀ (A : Type)(l : list A)(p : A → bool), filter comp lp → Prop) (P0 : ∀ (A : Type)(a : A)(l : list A)(p : A → bool), bool → filter comp (a :: l) p → Prop),

(∀ Ap , PA [] p []) →

(∀ Aalp , filter ind 1 Aalp (pa )(filter obligation 2 (@filter) Aalp (pa )) → P0Aalp (pa )(filter obligation 2 (@filter) Aalp (pa )) → PA (a :: l) p (filter obligation 2 (@filter) Aalp (pa ))) →

(∀ Aalp , filter ind Alp (filter lp ) → PAlp (filter lp ) → P0Aalp true( a :: filter lp )) → (∀ Aalp , filter ind Alp (filter lp ) → PAlp (filter lp ) → P0Aalp false( filter lp )) →

∀ Alp (f3 : filter comp lp ), filter ind Alp f3 → PAlp f3 ).

Matthieu Sozeau - Coq support for HoTT 34 Elimination principle

We prove filter respects the generated graph and derive:

Check (filter elim : ∀ P : ∀ (A : Type)(l : list A)(p : A → bool), filter comp lp → Prop, let P0 := fun (A : Type)(a : A)(l : list A)(p : A → bool) (refine : bool)(H : filter comp (a :: l) p) ⇒ pa = refine → PA (a :: l) pH in (∀ (A : Type)(p : A → bool), PA [] p []) → (∀ (A : Type)(a : A)(l : list A)(p : A → bool), PAlp (filter lp ) → P0Aalp true( a :: filter lp )) → (∀ (A : Type)(a : A)(l : list A)(p : A → bool), PAlp (filter lp ) → P0Aalp false( filter lp )) → ∀ (A : Type)(l : list A)(p : A → bool), PAlp (filter lp )).

Matthieu Sozeau - Coq support for HoTT 35 Using the “abstraction by equalities” technique again, we can abstract:

0 0 (λ (ll : list A)(r : appcomp ll ), l 0 = [] → r = app l [] → app l [] = l) l [] (app l [])

Directly apply the elimination principle and simplify the equations.

Eliminating calls

The elimination principle can only be applied usefully to calls with solely variable arguments.

Π A (l : list A), app l [] = l

Matthieu Sozeau - Coq support for HoTT 36 Eliminating calls

The elimination principle can only be applied usefully to calls with solely variable arguments.

Π A (l : list A), app l [] = l

Using the “abstraction by equalities” technique again, we can abstract:

0 0 (λ (ll : list A)(r : appcomp ll ), l 0 = [] → r = app l [] → app l [] = l) l [] (app l [])

Directly apply the elimination principle and simplify the equations.

Matthieu Sozeau - Coq support for HoTT 36 Summary

A function definition package handling:

I Full, nested dependent pattern-matching

I Structural and well-founded recursion on dependent types

I Generation of useful support lemmas for reasoning a posteriori

I No axioms if you provide the right proofs.

Benchmarked on a bit-fiddling library and a proof of (relative) consistency for Predicative System F (de Bruijn style, Mangin & Sozeau [LFMTP’15]).

Matthieu Sozeau - Coq support for HoTT 37 Discussion

I We moved to dependent equality (i.e. equality in sigma types) instead of JMeq. This is necessary to use hProp/hSet hypotheses (JMeq requires K on type equalities).

I Efficiency of computation? Consequences of moving to type-valued equality.

Matthieu Sozeau - Coq support for HoTT 38 Perspectives

I Non-constructor indices and unsolved constraints, e.g.: 0 = x + y, with a subsequent splitting on x.

I Support for views/arbitrary eliminators (e.g. McBride’s by).

I Structural and well-founded mutual recursion.

I Better support for the encode-decode method?

I HITs: low-level and high-level syntax for pattern-matching and solving higher equality obligations (Barras & Mangin).

Matthieu Sozeau - Coq support for HoTT 39 The End

Thanks!

Matthieu Sozeau - Coq support for HoTT 40 Using the clauses e we then build a subcovering s of the identity context map

t c = idsubst(∆ , xt : τ∆, ∆t[t/xt ])

and return Refine(t, c, s). Compilation produces t −→ `.n : Π ∆ (xt : τ∆) ∆t[t/xt ], (fcomp p )[t/xt ], we build

t −→ (λ∆, `.n ∆ t ∆t) : Π ∆, fcomp p

With nodes in detail

Consider a current problem ∆ ` −→p :Γ and a user clause −→ f up with tpre := { e } matching it. We typecheck tpre into t : τ and use strenghtening and abstraction to find a new context

t t ∆ , xt : τ, ∆t[t/xt ] such that ∆ , ∆t ∼ ∆

Matthieu Sozeau - Coq support for HoTT 41 Compilation produces t −→ `.n : Π ∆ (xt : τ∆) ∆t[t/xt ], (fcomp p )[t/xt ], we build

t −→ (λ∆, `.n ∆ t ∆t) : Π ∆, fcomp p

With nodes in detail

Consider a current problem ∆ ` −→p :Γ and a user clause −→ f up with tpre := { e } matching it. We typecheck tpre into t : τ and use strenghtening and abstraction to find a new context

t t ∆ , xt : τ, ∆t[t/xt ] such that ∆ , ∆t ∼ ∆

Using the clauses e we then build a subcovering s of the identity context map

t c = idsubst(∆ , xt : τ∆, ∆t[t/xt ])

and return Refine(t, c, s).

Matthieu Sozeau - Coq support for HoTT 41 With nodes in detail

Consider a current problem ∆ ` −→p :Γ and a user clause −→ f up with tpre := { e } matching it. We typecheck tpre into t : τ and use strenghtening and abstraction to find a new context

t t ∆ , xt : τ, ∆t[t/xt ] such that ∆ , ∆t ∼ ∆

Using the clauses e we then build a subcovering s of the identity context map

t c = idsubst(∆ , xt : τ∆, ∆t[t/xt ])

and return Refine(t, c, s). Compilation produces t −→ `.n : Π ∆ (xt : τ∆) ∆t[t/xt ], (fcomp p )[t/xt ], we build

t −→ (λ∆, `.n ∆ t ∆t) : Π ∆, fcomp p

Matthieu Sozeau - Coq support for HoTT 41 Elimination principle: inductive graph

−→ −→ For f.` : Π ∆, fcomp t we generate f.`ind : Π ∆, fcomp t → Prop and prove Π ∆, f.`ind ∆ (f.` ∆).

AbsRec(f, t) abstracts all the calls to fcomp proj from the term t, returning a new derivation Γ0 ` t0 where Γ0 contains bindings of −→ the form x : Π ∆, fcomp t for all the recursive calls.

Define Hyps(Γ) by a map to produce the corresponding inductive −→ hyps of the form Hx : Π ∆, find t (x ∆).

Matthieu Sozeau - Coq support for HoTT 42 Inductive graph constructors

Direct translation from the splitting tree:

I Split(c, x, s), Rec(v, s) : collect the constructors for the subsplitting(s) s, if any. −→ I Compute(∆ ` p :Γ, rhs) : By case on rhs: 0 I Program(t) : Compute Ψ ` t = AbsRec(f, t) and return the statement −→ 0 Π ∆ Ψ Hyps(Ψ), f.`ind p t 0 −→x −→ x I Refine(t, ∆ ` v , x, v x : ∆ , x : τ, ∆x, `.n)s : Compute Ψ ` t0 = AbsRec(f, t) and return: −→ Π ∆ Ψ Hyps(Ψ) (res :f comp p ) x 0 −→ f.`.nind ∆ t ∆x res → f.`ind p res

We continue with the generation of the f.`.nind graph.

Matthieu Sozeau - Coq support for HoTT 43