SlideShare a Scribd company logo
Journal of Functional Programming, 33(e1), January 2023
Is Sized Typing for Coq Practical?
Jonathan Chan
University of British Columbia,
University of Pennsylvania
Yufeng (Michael) Li
University of Waterloo,
University of Cambridge
William J. Bowman
University of British Columbia
1. Sized Types
2. Contributions
3. Implementation
2
Termination Checking: Guardedness
Fixpoint minus n m : nat :=
match n, m with
| S n′, S m′ => minus n′ m′
| _, _ => O
end.
3
Sized Types Contributions Implementation
✅
Termination Checking: Guardedness
4
Sized Types Contributions Implementation
Fixpoint div n m : nat :=
match n with
| S n′ => S (div minus n′ m m)
| O => O
end.
❌
Fixpoint minus n m : nat :=
match n, m with
| S n′, S m′ => minus n′ m′
| _, _ => O
end.
Termination Checking: Guardedness
5
Sized Types Contributions Implementation
Fixpoint div n m : nat :=
match n with
| S n′ => S (div minus n′ m m)
| O => O
end.
❌
Fixpoint minus n m : nat :=
match n, m with
| S n′, S m′ => minus n′ m′
| _, _ => O n
end.
❌
✅
Termination Checking: Sized Typing
6
Sized Types Contributions Implementation
Γ ⊢ n : nats+1
Γ ⊢ e1
: P O
Γ ⊢ n : nats
Γ, m : nats
⊢ e2
: P (S m)
───────────── ─────────────── ─────────────────────────
Γ ⊢ O : nats+1
Γ ⊢ S n : nats+1
Γ ⊢ match n with .
| O => e1
| S m => e2
end : P n
s ⩴ v | s+1 | ∞
Termination Checking: Sized Typing
7
Sized Types Contributions Implementation
Γ ⊢ n : nats+1
Γ ⊢ e1
: P O
Γ ⊢ n : nats
Γ, m : nats
⊢ e2
: P (S m)
───────────── ─────────────── ─────────────────────────
Γ ⊢ O : nats+1
Γ ⊢ S n : nats+1
Γ ⊢ match n with .
| O => e1
| S m => e2
end : P n
Termination Checking: Sized Typing
8
Sized Types Contributions Implementation
Fixpoint minus : natv
-> nat -> natv
.
Fixpoint div (n : natv+1
) (m : nat) : natv+1
:=
match n with
| S n′ => S (div (minus n′ m) m)
| O => O
end.
natv
natv
Past Work
• dependent types
• size inference algorithm
• nested (co)inductives
• universes
9
Sized Types Contributions Implementation
Hughes & Pareto & Sabry 1996
𝒞𝒞ℛ (Giménez 1998)
Amadio & Coupet-Grimal 1998
λ^ (Barthe et al. 2002; Frade 2004)
Λμ+
(Abel 2004)
λfixμν
(Abel 2003)
F^ (Barthe et al. 2005)
F^ω (Abel 2006)
F^× (Barthe et al. 2008)
MiniAgda (Abel 2010, 2012)
CIC^ (Barthe et al. 2006)
CIC^- (Grégoire & Sacchini 2010;
Sacchini 2011)
CCω
̂ (Sacchini 2013)
CIC^l (Sacchini 2014)
CIC^⊑ (Sacchini 2015)
Fcop
ω (Abel & Pientka 2016)
Abel 2017
─ ─ ─
Contributions
• dependent types
• size inference algorithm
• nested (co)inductives
• universes
10
Sized Types Contributions Implementation
CIC^ (Barthe et al. 2006)
CIC^- (Grégoire & Sacchini 2010;
Sacchini 2011)
CCω
̂ (Sacchini 2013)
CIC^*
(2019 – 2023)
─ ─ ─
Contributions
• dependent types
• annotation inference alg.
• nested (co)inductives
• universes
11
Sized Types Contributions Implementation
CIC^ (Barthe et al. 2006)
CIC^- (Grégoire & Sacchini 2010;
Sacchini 2011)
CCω
̂ (Sacchini 2013)
CIC^*
(2019 – 2023)
─ ─ ─
Contributions
• dependent types
• annotation inference alg.
• nested (co)inductives
• universes
12
Sized Types Contributions Implementation
CIC^ (Barthe et al. 2006)
CIC^- (Grégoire & Sacchini 2010;
Sacchini 2011)
CCω
̂ (Sacchini 2013)
CIC^*
(2019 – 2023)
─ ─ ─
Contributions
• dependent types
• annotation inference alg.
• nested (co)inductives
• cumulative universes
• global + local definitions
13
Sized Types Contributions Implementation
CIC^ (Barthe et al. 2006)
CIC^- (Grégoire & Sacchini 2010;
Sacchini 2011)
CCω
̂ (Sacchini 2013)
CIC^*
(2019 – 2023)
─ ─ ─
Contributions
• implementation in Coq fork
• dependent types
• annotation inference alg.
• nested (co)inductives
• cumulative universes
• global + local definitions
14
Sized Types Contributions Implementation
CIC^ (Barthe et al. 2006)
CIC^- (Grégoire & Sacchini 2010;
Sacchini 2011)
CCω
̂ (Sacchini 2013)
CIC^*
(2019 – 2023)
─ ─ ─
Goal: backward-compatible sized typing for Coq for
expressive, modular, efficient termination checking
15
Sized Types Contributions Implementation
Goal: backward-compatible sized typing for Coq for
expressive, modular, efficient termination checking
16
Sized Types Contributions Implementation
NOPE
17
Sized Types Contributions Implementation
Sized Typing Pipeline
CIC CIC^* + 𝒞 CIC^*
solve &
substitute
type checking
size inference
Coq kernel
18
Sized Types Contributions Implementation
Sized Typing Pipeline
Coq kernel
𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ ⇝ e″ : τ″
19
Sized Types Contributions Implementation
Sized Typing Pipeline
Coq kernel
𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ ⇝ e″ : τ″
20
Sized Types Contributions Implementation
Sized Typing Pipeline
Definition mynat [u] : Set ≔ natu
.
Fixpoint div [v, w, …] (n : mynatv
) (m : mynatw
) : mynatv
≔ …
Coq kernel
Size Inference: Fixpoints
𝒞; Γ, f : τv
⊢ … ⇝ 𝒞′; e : τv+1
SAT(𝒞′ ∪ {v ≠ ∞})
───────────────────────────────────────────────────
𝒞; Γ ⊢ … ⇝ 𝒞′; fix f : τ*
≔ e : τv
21
Sized Types Contributions
𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′
Inference Substitution Performance
Size Inference: Fixpoints
𝒞; Γ, f : τv
⊢ … ⇝ 𝒞′; e : τv+1
SAT(𝒞′ ∪ {v ≠ ∞})
───────────────────────────────────────────────────
𝒞; Γ ⊢ … ⇝ 𝒞′; fix f : τ*
≔ e : τv
22
Sized Types Contributions
𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′
Inference Substitution Performance
• nat*
→ nat → nat
• nat → nat*
→ nat
Size Inference: Fixpoints
𝒞; Γ, f : τv
⊢ … ⇝ 𝒞′; e : τv+1
SAT(𝒞′ ∪ {v ≠ ∞})
───────────────────────────────────────────────────
𝒞; Γ ⊢ … ⇝ 𝒞′; fix f : τ*
≔ e : τv
23
Sized Types Contributions
𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′
Inference Substitution Performance
• nat*
→ nat → nat*
• nat*
→ nat → nat
• nat → nat*
→ nat*
• nat → nat*
→ nat
Size Inference: Cumulativity
𝒞; Γ ⊢ … ⇝ 𝒞′; e : τ′ τ′ ≼ τ ⇝ 𝒞″
─────────────────────────────────────
𝒞; Γ ⊢ … ⇝ 𝒞′ ∪ 𝒞″; e : τ
24
Sized Types Contributions
𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′
Inference Substitution Performance
Size Inference: Constraints
nats₁
≼ nats₂
⇝ {s₁ ⊑ s₂}
conats₁
≼ conats₂
⇝ {s₂ ⊑ s₁}
25
Sized Types Contributions
τ1
≼ τ2
⇝ 𝒞
𝒞 ⩴ {v1
+ n1
⊑ v2
+ n2
, …}
Inference Substitution Performance
Size Inference: Constraint Satisfiability
nats₁
≼ nats₂
⇝ {s₁ ⊑ s₂}
conats₁
≼ conats₂
⇝ {s₂ ⊑ s₁}
─────── ──────
s ⊑ s+1 s ⊑ ∞
26
Sized Types Contributions
τ1
≼ τ2
⇝ 𝒞
Inference Substitution Performance
🆕 Constraint Solving
ρ = SAT(𝒞)
───────────────────
𝒞; e : τ ⇝ ρe : ρτ
27
Sized Types Contributions
𝒞′; e′ : τ′ ⇝ e″ : τ″
ρ ⩴ [v1
↦ s1
, …]
Inference Substitution Performance
🆕 Constraint Solving
ρ = SAT(𝒞)
───────────────────
𝒞; e : τ ⇝ ρe : ρτ
28
Sized Types Contributions
negative cycle
detection
e.g. Bellman–Ford
O(|𝒱||𝒞|)
𝒞′; e′ : τ′ ⇝ e″ : τ″
Inference Substitution Performance
Artificial Example: Exponentially Many `nat`s
Time Definition nats1 :=
(nat, nat, nat, nat, nat, nat, nat, nat).
Time Definition nats2 :=
(nats1, nats1, nats1, nats1).
...
Time Definition nats6 :=
(nats5, nats5, nats5, nats5).
29
Sized Types Contributions Inference Substitution Performance
Artificial Example: Exponentially Many `nat`s
Definition nats1 [v1, ..., v21] :
Set *v1
Set *v2
Set *v3
Set *v4
Set *v5
Set *v6
Set *v7
Set ≔
(natv8
, natv9
, natv10
, natv11
, natv12
, natv13
, natv14
, natv15
)v16,v17,v18,v19,v20,v21
.
Definition nats2 [...] ≔ (nats1v1,…,v21
, nats1v22,…,v42
, nats1…
, nats1…
).
...
𝒞nats1
= {v16 ⊑ v1, v17 ⊑ v2, …}
30
Sized Types Contributions Inference Substitution Performance
Artificial Example: Exponentially Many `nat`s
Definition nats1 [v1, ..., v21] :
Set *v1
Set *v2
Set *v3
Set *v4
Set *v5
Set *v6
Set *v7
Set ≔
(natv8
, natv9
, natv10
, natv11
, natv12
, natv13
, natv14
, natv15
)v16,v17,v18,v19,v20,v21
.
Definition nats2 [...] ≔ (nats1v1,…,v21
, nats1v22,…,v42
, nats1…
, nats1…
).
...
𝒞nats1
= {v16 ⊑ v1, v17 ⊑ v2, …}
31
Sized Types Contributions Inference Substitution Performance
Artificial Example: Exponentially Many `nat`s
Time Definition nats1 :=
(nat, nat, nat, nat,
nat, nat, nat, nat).
Time Definition nats2 :=
(nats1, nats1, nats1, nats1).
...
Time Definition nats6 :=
(nats5, nats5, nats5, nats5).
32
Sized Types Contributions
Definition |𝒱| Time (s)
nats1 21 0.004
nats2 93 0.020
nats3 381 0.177
nats4 1533 2.299
nats5 6141 35.385
nats6 24573 > 120
Inference Substitution Performance
Real Example: MSets/MSetList.v
Mean over five trials
33
Sized Types Contributions
Unsized compilation (s) 15.122 ± 0.073
Sized compilation (s) 83.660 ± 0.286
Slowdown 5.5×
SAT ops only (s) 64.600 ± 0.437
SAT ops only (%) 77.2%
Inference Substitution Performance
Real Example: MSets/MSetList.v
34
Sized Types Contributions
log(count)
log distribution of |𝒱| × |𝒞| during SAT operations
|𝒱| × |𝒞|
Inference Substitution Performance
~4K vs. × ~250 cs.
size inference
35
backward-compatible sized typing
definitions
constraint checking/solving
poor performance
Sized Types Contributions Implementation
Future Work?: Manual Size Annotation + Size Inference
36
Fixpoint minus [v] : natv
→ nat∞
→ natv
.
Fixpoint div [v] (n : natv
) (m : nat∞
) : natv
≔
match n with
| S n′ ⇒ S (div [?w] (minus [?u] n′ m) m)
| O ⇒ O
End.
𝒞div
= {?u = ?w, ?w+1 ⊑ v}
Sized Types Contributions Implementation
Thank you!
● Full spec of size
inference algorithm
& soundness wrt typing
● Metatheoretical results
& attempt at
set-theoretic model
● More detailed
performance profiling
& statistics
+ Closed draft PR:
https://github.com/coq/
coq/pull/12426
37
Check out the paper for:
38
bonus content
Concrete sized naturals
O : natv+1
S O : natv+2
S (S O) : natv+3
O : natv+2
S O : natv+3
S (S O) : natv+4
39
Real Example: setoid_ring/Field_theory.v
40
Mean over two trials (August 2023)
Unsized compilation (s) 17.815 ± 0.545
Sized compilation (s) 106.87 ± 1.94
Slowdown 6.0×
SAT ops only (s) 84.70
SAT ops only (%) 79.3% 17755 vs. × 14057 cs. ≈ 250M
Artificial Example: Universe Polymorphism
Set Printing Universes.
Set Universe Polymorphism.
Time Definition T1 : Type := Type -> Type
-> Type -> Type -> Type -> Type. Print T1.
Time Definition T2 : Type :=
T1 -> T1 -> T1 -> T1 -> T1 -> T1. Print T2.
Time Definition T3 : Type :=
T2 -> T2 -> T2 -> T2 -> T2 -> T2. Print T3.
Time Definition T4 : Type :=
T3 -> T3 -> T3 -> T3 -> T3 -> T3. Print T4.
Time Definition T5 : Type :=
T4 -> T4 -> T4 -> T4 -> T4 -> T4. Print T5.
Time Definition T6 : Type :=
T5 -> T5 -> T5 -> T5 -> T5 -> T5. Print T6.
Time Definition T7 : Type :=
T6 -> T6 -> T6 -> T6 -> T6 -> T6. Print T7.
41
Definition #u Time (s)
T1 7 ~ 0
T2 43 0.002
T3 259 0.026
T4 1555 0.057
T5 9331 0.374
T6 55987 3.300
T7 335921 18.170

More Related Content

Similar to Is Sized Typing for Coq Practical?

Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...
Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...
Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...
The Statistical and Applied Mathematical Sciences Institute
 
Asymptotics 140510003721-phpapp02
Asymptotics 140510003721-phpapp02Asymptotics 140510003721-phpapp02
Asymptotics 140510003721-phpapp02
mansab MIRZA
 
Test s velocity_15_5_4
Test s velocity_15_5_4Test s velocity_15_5_4
Test s velocity_15_5_4
Kunihiko Saito
 
Artificial software diversity: automatic synthesis of program sosies
Artificial software diversity: automatic synthesis of program sosiesArtificial software diversity: automatic synthesis of program sosies
Artificial software diversity: automatic synthesis of program sosies
FoCAS Initiative
 
Convolutional Neural Network (CNN) presentation from theory to code in Theano
Convolutional Neural Network (CNN) presentation from theory to code in TheanoConvolutional Neural Network (CNN) presentation from theory to code in Theano
Convolutional Neural Network (CNN) presentation from theory to code in Theano
Seongwon Hwang
 
CVPR2016 Fitting Surface Models to Data 抜粋
CVPR2016 Fitting Surface Models to Data 抜粋CVPR2016 Fitting Surface Models to Data 抜粋
CVPR2016 Fitting Surface Models to Data 抜粋
sumisumith
 
Applied machine learning for search engine relevance 3
Applied machine learning for search engine relevance 3Applied machine learning for search engine relevance 3
Applied machine learning for search engine relevance 3
Charles Martin
 
Signals and Systems Homework Help.pptx
Signals and Systems Homework Help.pptxSignals and Systems Homework Help.pptx
Signals and Systems Homework Help.pptx
Matlab Assignment Experts
 
Node Unique Label Cover
Node Unique Label CoverNode Unique Label Cover
Node Unique Label Cover
msramanujan
 
Optimal Meshing
Optimal MeshingOptimal Meshing
Optimal Meshing
Don Sheehy
 
ZK Study Club: Sumcheck Arguments and Their Applications
ZK Study Club: Sumcheck Arguments and Their ApplicationsZK Study Club: Sumcheck Arguments and Their Applications
ZK Study Club: Sumcheck Arguments and Their Applications
Alex Pruden
 
Q-Metrics in Theory And Practice
Q-Metrics in Theory And PracticeQ-Metrics in Theory And Practice
Q-Metrics in Theory And Practice
guest3550292
 
Q-Metrics in Theory and Practice
Q-Metrics in Theory and PracticeQ-Metrics in Theory and Practice
Q-Metrics in Theory and Practice
Magdi Mohamed
 
Cs6402 daa-2 marks set 1
Cs6402 daa-2 marks set 1Cs6402 daa-2 marks set 1
Cs6402 daa-2 marks set 1
Vai Jayanthi
 
Practical and Worst-Case Efficient Apportionment
Practical and Worst-Case Efficient ApportionmentPractical and Worst-Case Efficient Apportionment
Practical and Worst-Case Efficient Apportionment
Raphael Reitzig
 
Actors for Behavioural Simulation
Actors for Behavioural SimulationActors for Behavioural Simulation
Actors for Behavioural Simulation
ClarkTony
 
Lec7 deeprlbootcamp-svg+scg
Lec7 deeprlbootcamp-svg+scgLec7 deeprlbootcamp-svg+scg
Lec7 deeprlbootcamp-svg+scg
Ronald Teo
 
SVD and the Netflix Dataset
SVD and the Netflix DatasetSVD and the Netflix Dataset
SVD and the Netflix Dataset
Ben Mabey
 
Efficient Hill Climber for Constrained Pseudo-Boolean Optimization Problems
Efficient Hill Climber for Constrained Pseudo-Boolean Optimization ProblemsEfficient Hill Climber for Constrained Pseudo-Boolean Optimization Problems
Efficient Hill Climber for Constrained Pseudo-Boolean Optimization Problems
jfrchicanog
 
Asymptotic Notations.pptx
Asymptotic Notations.pptxAsymptotic Notations.pptx
Asymptotic Notations.pptx
SunilWork1
 

Similar to Is Sized Typing for Coq Practical? (20)

Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...
Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...
Program on Quasi-Monte Carlo and High-Dimensional Sampling Methods for Applie...
 
Asymptotics 140510003721-phpapp02
Asymptotics 140510003721-phpapp02Asymptotics 140510003721-phpapp02
Asymptotics 140510003721-phpapp02
 
Test s velocity_15_5_4
Test s velocity_15_5_4Test s velocity_15_5_4
Test s velocity_15_5_4
 
Artificial software diversity: automatic synthesis of program sosies
Artificial software diversity: automatic synthesis of program sosiesArtificial software diversity: automatic synthesis of program sosies
Artificial software diversity: automatic synthesis of program sosies
 
Convolutional Neural Network (CNN) presentation from theory to code in Theano
Convolutional Neural Network (CNN) presentation from theory to code in TheanoConvolutional Neural Network (CNN) presentation from theory to code in Theano
Convolutional Neural Network (CNN) presentation from theory to code in Theano
 
CVPR2016 Fitting Surface Models to Data 抜粋
CVPR2016 Fitting Surface Models to Data 抜粋CVPR2016 Fitting Surface Models to Data 抜粋
CVPR2016 Fitting Surface Models to Data 抜粋
 
Applied machine learning for search engine relevance 3
Applied machine learning for search engine relevance 3Applied machine learning for search engine relevance 3
Applied machine learning for search engine relevance 3
 
Signals and Systems Homework Help.pptx
Signals and Systems Homework Help.pptxSignals and Systems Homework Help.pptx
Signals and Systems Homework Help.pptx
 
Node Unique Label Cover
Node Unique Label CoverNode Unique Label Cover
Node Unique Label Cover
 
Optimal Meshing
Optimal MeshingOptimal Meshing
Optimal Meshing
 
ZK Study Club: Sumcheck Arguments and Their Applications
ZK Study Club: Sumcheck Arguments and Their ApplicationsZK Study Club: Sumcheck Arguments and Their Applications
ZK Study Club: Sumcheck Arguments and Their Applications
 
Q-Metrics in Theory And Practice
Q-Metrics in Theory And PracticeQ-Metrics in Theory And Practice
Q-Metrics in Theory And Practice
 
Q-Metrics in Theory and Practice
Q-Metrics in Theory and PracticeQ-Metrics in Theory and Practice
Q-Metrics in Theory and Practice
 
Cs6402 daa-2 marks set 1
Cs6402 daa-2 marks set 1Cs6402 daa-2 marks set 1
Cs6402 daa-2 marks set 1
 
Practical and Worst-Case Efficient Apportionment
Practical and Worst-Case Efficient ApportionmentPractical and Worst-Case Efficient Apportionment
Practical and Worst-Case Efficient Apportionment
 
Actors for Behavioural Simulation
Actors for Behavioural SimulationActors for Behavioural Simulation
Actors for Behavioural Simulation
 
Lec7 deeprlbootcamp-svg+scg
Lec7 deeprlbootcamp-svg+scgLec7 deeprlbootcamp-svg+scg
Lec7 deeprlbootcamp-svg+scg
 
SVD and the Netflix Dataset
SVD and the Netflix DatasetSVD and the Netflix Dataset
SVD and the Netflix Dataset
 
Efficient Hill Climber for Constrained Pseudo-Boolean Optimization Problems
Efficient Hill Climber for Constrained Pseudo-Boolean Optimization ProblemsEfficient Hill Climber for Constrained Pseudo-Boolean Optimization Problems
Efficient Hill Climber for Constrained Pseudo-Boolean Optimization Problems
 
Asymptotic Notations.pptx
Asymptotic Notations.pptxAsymptotic Notations.pptx
Asymptotic Notations.pptx
 

Recently uploaded

Deep Software Variability and Frictionless Reproducibility
Deep Software Variability and Frictionless ReproducibilityDeep Software Variability and Frictionless Reproducibility
Deep Software Variability and Frictionless Reproducibility
University of Rennes, INSA Rennes, Inria/IRISA, CNRS
 
What is greenhouse gasses and how many gasses are there to affect the Earth.
What is greenhouse gasses and how many gasses are there to affect the Earth.What is greenhouse gasses and how many gasses are there to affect the Earth.
What is greenhouse gasses and how many gasses are there to affect the Earth.
moosaasad1975
 
Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...
Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...
Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...
Ana Luísa Pinho
 
Leaf Initiation, Growth and Differentiation.pdf
Leaf Initiation, Growth and Differentiation.pdfLeaf Initiation, Growth and Differentiation.pdf
Leaf Initiation, Growth and Differentiation.pdf
Renu Jangid
 
Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...
Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...
Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...
University of Maribor
 
Topic: SICKLE CELL DISEASE IN CHILDREN-3.pdf
Topic: SICKLE CELL DISEASE IN CHILDREN-3.pdfTopic: SICKLE CELL DISEASE IN CHILDREN-3.pdf
Topic: SICKLE CELL DISEASE IN CHILDREN-3.pdf
TinyAnderson
 
SAR of Medicinal Chemistry 1st by dk.pdf
SAR of Medicinal Chemistry 1st by dk.pdfSAR of Medicinal Chemistry 1st by dk.pdf
SAR of Medicinal Chemistry 1st by dk.pdf
KrushnaDarade1
 
Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...
Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...
Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...
Travis Hills MN
 
Eukaryotic Transcription Presentation.pptx
Eukaryotic Transcription Presentation.pptxEukaryotic Transcription Presentation.pptx
Eukaryotic Transcription Presentation.pptx
RitabrataSarkar3
 
Toxic effects of heavy metals : Lead and Arsenic
Toxic effects of heavy metals : Lead and ArsenicToxic effects of heavy metals : Lead and Arsenic
Toxic effects of heavy metals : Lead and Arsenic
sanjana502982
 
Nucleic Acid-its structural and functional complexity.
Nucleic Acid-its structural and functional complexity.Nucleic Acid-its structural and functional complexity.
Nucleic Acid-its structural and functional complexity.
Nistarini College, Purulia (W.B) India
 
BREEDING METHODS FOR DISEASE RESISTANCE.pptx
BREEDING METHODS FOR DISEASE RESISTANCE.pptxBREEDING METHODS FOR DISEASE RESISTANCE.pptx
BREEDING METHODS FOR DISEASE RESISTANCE.pptx
RASHMI M G
 
20240520 Planning a Circuit Simulator in JavaScript.pptx
20240520 Planning a Circuit Simulator in JavaScript.pptx20240520 Planning a Circuit Simulator in JavaScript.pptx
20240520 Planning a Circuit Simulator in JavaScript.pptx
Sharon Liu
 
DMARDs Pharmacolgy Pharm D 5th Semester.pdf
DMARDs Pharmacolgy Pharm D 5th Semester.pdfDMARDs Pharmacolgy Pharm D 5th Semester.pdf
DMARDs Pharmacolgy Pharm D 5th Semester.pdf
fafyfskhan251kmf
 
bordetella pertussis.................................ppt
bordetella pertussis.................................pptbordetella pertussis.................................ppt
bordetella pertussis.................................ppt
kejapriya1
 
Unveiling the Energy Potential of Marshmallow Deposits.pdf
Unveiling the Energy Potential of Marshmallow Deposits.pdfUnveiling the Energy Potential of Marshmallow Deposits.pdf
Unveiling the Energy Potential of Marshmallow Deposits.pdf
Erdal Coalmaker
 
如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样
如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样
如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样
yqqaatn0
 
The use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptx
The use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptxThe use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptx
The use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptx
MAGOTI ERNEST
 
原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样
原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样
原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样
yqqaatn0
 
ANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptx
ANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptxANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptx
ANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptx
RASHMI M G
 

Recently uploaded (20)

Deep Software Variability and Frictionless Reproducibility
Deep Software Variability and Frictionless ReproducibilityDeep Software Variability and Frictionless Reproducibility
Deep Software Variability and Frictionless Reproducibility
 
What is greenhouse gasses and how many gasses are there to affect the Earth.
What is greenhouse gasses and how many gasses are there to affect the Earth.What is greenhouse gasses and how many gasses are there to affect the Earth.
What is greenhouse gasses and how many gasses are there to affect the Earth.
 
Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...
Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...
Deep Behavioral Phenotyping in Systems Neuroscience for Functional Atlasing a...
 
Leaf Initiation, Growth and Differentiation.pdf
Leaf Initiation, Growth and Differentiation.pdfLeaf Initiation, Growth and Differentiation.pdf
Leaf Initiation, Growth and Differentiation.pdf
 
Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...
Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...
Comparing Evolved Extractive Text Summary Scores of Bidirectional Encoder Rep...
 
Topic: SICKLE CELL DISEASE IN CHILDREN-3.pdf
Topic: SICKLE CELL DISEASE IN CHILDREN-3.pdfTopic: SICKLE CELL DISEASE IN CHILDREN-3.pdf
Topic: SICKLE CELL DISEASE IN CHILDREN-3.pdf
 
SAR of Medicinal Chemistry 1st by dk.pdf
SAR of Medicinal Chemistry 1st by dk.pdfSAR of Medicinal Chemistry 1st by dk.pdf
SAR of Medicinal Chemistry 1st by dk.pdf
 
Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...
Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...
Travis Hills' Endeavors in Minnesota: Fostering Environmental and Economic Pr...
 
Eukaryotic Transcription Presentation.pptx
Eukaryotic Transcription Presentation.pptxEukaryotic Transcription Presentation.pptx
Eukaryotic Transcription Presentation.pptx
 
Toxic effects of heavy metals : Lead and Arsenic
Toxic effects of heavy metals : Lead and ArsenicToxic effects of heavy metals : Lead and Arsenic
Toxic effects of heavy metals : Lead and Arsenic
 
Nucleic Acid-its structural and functional complexity.
Nucleic Acid-its structural and functional complexity.Nucleic Acid-its structural and functional complexity.
Nucleic Acid-its structural and functional complexity.
 
BREEDING METHODS FOR DISEASE RESISTANCE.pptx
BREEDING METHODS FOR DISEASE RESISTANCE.pptxBREEDING METHODS FOR DISEASE RESISTANCE.pptx
BREEDING METHODS FOR DISEASE RESISTANCE.pptx
 
20240520 Planning a Circuit Simulator in JavaScript.pptx
20240520 Planning a Circuit Simulator in JavaScript.pptx20240520 Planning a Circuit Simulator in JavaScript.pptx
20240520 Planning a Circuit Simulator in JavaScript.pptx
 
DMARDs Pharmacolgy Pharm D 5th Semester.pdf
DMARDs Pharmacolgy Pharm D 5th Semester.pdfDMARDs Pharmacolgy Pharm D 5th Semester.pdf
DMARDs Pharmacolgy Pharm D 5th Semester.pdf
 
bordetella pertussis.................................ppt
bordetella pertussis.................................pptbordetella pertussis.................................ppt
bordetella pertussis.................................ppt
 
Unveiling the Energy Potential of Marshmallow Deposits.pdf
Unveiling the Energy Potential of Marshmallow Deposits.pdfUnveiling the Energy Potential of Marshmallow Deposits.pdf
Unveiling the Energy Potential of Marshmallow Deposits.pdf
 
如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样
如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样
如何办理(uvic毕业证书)维多利亚大学毕业证本科学位证书原版一模一样
 
The use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptx
The use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptxThe use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptx
The use of Nauplii and metanauplii artemia in aquaculture (brine shrimp).pptx
 
原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样
原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样
原版制作(carleton毕业证书)卡尔顿大学毕业证硕士文凭原版一模一样
 
ANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptx
ANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptxANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptx
ANAMOLOUS SECONDARY GROWTH IN DICOT ROOTS.pptx
 

Is Sized Typing for Coq Practical?

  • 1. Journal of Functional Programming, 33(e1), January 2023 Is Sized Typing for Coq Practical? Jonathan Chan University of British Columbia, University of Pennsylvania Yufeng (Michael) Li University of Waterloo, University of Cambridge William J. Bowman University of British Columbia
  • 2. 1. Sized Types 2. Contributions 3. Implementation 2
  • 3. Termination Checking: Guardedness Fixpoint minus n m : nat := match n, m with | S n′, S m′ => minus n′ m′ | _, _ => O end. 3 Sized Types Contributions Implementation ✅
  • 4. Termination Checking: Guardedness 4 Sized Types Contributions Implementation Fixpoint div n m : nat := match n with | S n′ => S (div minus n′ m m) | O => O end. ❌ Fixpoint minus n m : nat := match n, m with | S n′, S m′ => minus n′ m′ | _, _ => O end.
  • 5. Termination Checking: Guardedness 5 Sized Types Contributions Implementation Fixpoint div n m : nat := match n with | S n′ => S (div minus n′ m m) | O => O end. ❌ Fixpoint minus n m : nat := match n, m with | S n′, S m′ => minus n′ m′ | _, _ => O n end. ❌ ✅
  • 6. Termination Checking: Sized Typing 6 Sized Types Contributions Implementation Γ ⊢ n : nats+1 Γ ⊢ e1 : P O Γ ⊢ n : nats Γ, m : nats ⊢ e2 : P (S m) ───────────── ─────────────── ───────────────────────── Γ ⊢ O : nats+1 Γ ⊢ S n : nats+1 Γ ⊢ match n with . | O => e1 | S m => e2 end : P n s ⩴ v | s+1 | ∞
  • 7. Termination Checking: Sized Typing 7 Sized Types Contributions Implementation Γ ⊢ n : nats+1 Γ ⊢ e1 : P O Γ ⊢ n : nats Γ, m : nats ⊢ e2 : P (S m) ───────────── ─────────────── ───────────────────────── Γ ⊢ O : nats+1 Γ ⊢ S n : nats+1 Γ ⊢ match n with . | O => e1 | S m => e2 end : P n
  • 8. Termination Checking: Sized Typing 8 Sized Types Contributions Implementation Fixpoint minus : natv -> nat -> natv . Fixpoint div (n : natv+1 ) (m : nat) : natv+1 := match n with | S n′ => S (div (minus n′ m) m) | O => O end. natv natv
  • 9. Past Work • dependent types • size inference algorithm • nested (co)inductives • universes 9 Sized Types Contributions Implementation Hughes & Pareto & Sabry 1996 𝒞𝒞ℛ (Giménez 1998) Amadio & Coupet-Grimal 1998 λ^ (Barthe et al. 2002; Frade 2004) Λμ+ (Abel 2004) λfixμν (Abel 2003) F^ (Barthe et al. 2005) F^ω (Abel 2006) F^× (Barthe et al. 2008) MiniAgda (Abel 2010, 2012) CIC^ (Barthe et al. 2006) CIC^- (Grégoire & Sacchini 2010; Sacchini 2011) CCω ̂ (Sacchini 2013) CIC^l (Sacchini 2014) CIC^⊑ (Sacchini 2015) Fcop ω (Abel & Pientka 2016) Abel 2017 ─ ─ ─
  • 10. Contributions • dependent types • size inference algorithm • nested (co)inductives • universes 10 Sized Types Contributions Implementation CIC^ (Barthe et al. 2006) CIC^- (Grégoire & Sacchini 2010; Sacchini 2011) CCω ̂ (Sacchini 2013) CIC^* (2019 – 2023) ─ ─ ─
  • 11. Contributions • dependent types • annotation inference alg. • nested (co)inductives • universes 11 Sized Types Contributions Implementation CIC^ (Barthe et al. 2006) CIC^- (Grégoire & Sacchini 2010; Sacchini 2011) CCω ̂ (Sacchini 2013) CIC^* (2019 – 2023) ─ ─ ─
  • 12. Contributions • dependent types • annotation inference alg. • nested (co)inductives • universes 12 Sized Types Contributions Implementation CIC^ (Barthe et al. 2006) CIC^- (Grégoire & Sacchini 2010; Sacchini 2011) CCω ̂ (Sacchini 2013) CIC^* (2019 – 2023) ─ ─ ─
  • 13. Contributions • dependent types • annotation inference alg. • nested (co)inductives • cumulative universes • global + local definitions 13 Sized Types Contributions Implementation CIC^ (Barthe et al. 2006) CIC^- (Grégoire & Sacchini 2010; Sacchini 2011) CCω ̂ (Sacchini 2013) CIC^* (2019 – 2023) ─ ─ ─
  • 14. Contributions • implementation in Coq fork • dependent types • annotation inference alg. • nested (co)inductives • cumulative universes • global + local definitions 14 Sized Types Contributions Implementation CIC^ (Barthe et al. 2006) CIC^- (Grégoire & Sacchini 2010; Sacchini 2011) CCω ̂ (Sacchini 2013) CIC^* (2019 – 2023) ─ ─ ─
  • 15. Goal: backward-compatible sized typing for Coq for expressive, modular, efficient termination checking 15 Sized Types Contributions Implementation
  • 16. Goal: backward-compatible sized typing for Coq for expressive, modular, efficient termination checking 16 Sized Types Contributions Implementation NOPE
  • 17. 17 Sized Types Contributions Implementation Sized Typing Pipeline CIC CIC^* + 𝒞 CIC^* solve & substitute type checking size inference Coq kernel
  • 18. 18 Sized Types Contributions Implementation Sized Typing Pipeline Coq kernel 𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ ⇝ e″ : τ″
  • 19. 19 Sized Types Contributions Implementation Sized Typing Pipeline Coq kernel 𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ ⇝ e″ : τ″
  • 20. 20 Sized Types Contributions Implementation Sized Typing Pipeline Definition mynat [u] : Set ≔ natu . Fixpoint div [v, w, …] (n : mynatv ) (m : mynatw ) : mynatv ≔ … Coq kernel
  • 21. Size Inference: Fixpoints 𝒞; Γ, f : τv ⊢ … ⇝ 𝒞′; e : τv+1 SAT(𝒞′ ∪ {v ≠ ∞}) ─────────────────────────────────────────────────── 𝒞; Γ ⊢ … ⇝ 𝒞′; fix f : τ* ≔ e : τv 21 Sized Types Contributions 𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ Inference Substitution Performance
  • 22. Size Inference: Fixpoints 𝒞; Γ, f : τv ⊢ … ⇝ 𝒞′; e : τv+1 SAT(𝒞′ ∪ {v ≠ ∞}) ─────────────────────────────────────────────────── 𝒞; Γ ⊢ … ⇝ 𝒞′; fix f : τ* ≔ e : τv 22 Sized Types Contributions 𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ Inference Substitution Performance • nat* → nat → nat • nat → nat* → nat
  • 23. Size Inference: Fixpoints 𝒞; Γ, f : τv ⊢ … ⇝ 𝒞′; e : τv+1 SAT(𝒞′ ∪ {v ≠ ∞}) ─────────────────────────────────────────────────── 𝒞; Γ ⊢ … ⇝ 𝒞′; fix f : τ* ≔ e : τv 23 Sized Types Contributions 𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ Inference Substitution Performance • nat* → nat → nat* • nat* → nat → nat • nat → nat* → nat* • nat → nat* → nat
  • 24. Size Inference: Cumulativity 𝒞; Γ ⊢ … ⇝ 𝒞′; e : τ′ τ′ ≼ τ ⇝ 𝒞″ ───────────────────────────────────── 𝒞; Γ ⊢ … ⇝ 𝒞′ ∪ 𝒞″; e : τ 24 Sized Types Contributions 𝒞; Γ ⊢ e : τ ⇝ 𝒞′; e′ : τ′ Inference Substitution Performance
  • 25. Size Inference: Constraints nats₁ ≼ nats₂ ⇝ {s₁ ⊑ s₂} conats₁ ≼ conats₂ ⇝ {s₂ ⊑ s₁} 25 Sized Types Contributions τ1 ≼ τ2 ⇝ 𝒞 𝒞 ⩴ {v1 + n1 ⊑ v2 + n2 , …} Inference Substitution Performance
  • 26. Size Inference: Constraint Satisfiability nats₁ ≼ nats₂ ⇝ {s₁ ⊑ s₂} conats₁ ≼ conats₂ ⇝ {s₂ ⊑ s₁} ─────── ────── s ⊑ s+1 s ⊑ ∞ 26 Sized Types Contributions τ1 ≼ τ2 ⇝ 𝒞 Inference Substitution Performance
  • 27. 🆕 Constraint Solving ρ = SAT(𝒞) ─────────────────── 𝒞; e : τ ⇝ ρe : ρτ 27 Sized Types Contributions 𝒞′; e′ : τ′ ⇝ e″ : τ″ ρ ⩴ [v1 ↦ s1 , …] Inference Substitution Performance
  • 28. 🆕 Constraint Solving ρ = SAT(𝒞) ─────────────────── 𝒞; e : τ ⇝ ρe : ρτ 28 Sized Types Contributions negative cycle detection e.g. Bellman–Ford O(|𝒱||𝒞|) 𝒞′; e′ : τ′ ⇝ e″ : τ″ Inference Substitution Performance
  • 29. Artificial Example: Exponentially Many `nat`s Time Definition nats1 := (nat, nat, nat, nat, nat, nat, nat, nat). Time Definition nats2 := (nats1, nats1, nats1, nats1). ... Time Definition nats6 := (nats5, nats5, nats5, nats5). 29 Sized Types Contributions Inference Substitution Performance
  • 30. Artificial Example: Exponentially Many `nat`s Definition nats1 [v1, ..., v21] : Set *v1 Set *v2 Set *v3 Set *v4 Set *v5 Set *v6 Set *v7 Set ≔ (natv8 , natv9 , natv10 , natv11 , natv12 , natv13 , natv14 , natv15 )v16,v17,v18,v19,v20,v21 . Definition nats2 [...] ≔ (nats1v1,…,v21 , nats1v22,…,v42 , nats1… , nats1… ). ... 𝒞nats1 = {v16 ⊑ v1, v17 ⊑ v2, …} 30 Sized Types Contributions Inference Substitution Performance
  • 31. Artificial Example: Exponentially Many `nat`s Definition nats1 [v1, ..., v21] : Set *v1 Set *v2 Set *v3 Set *v4 Set *v5 Set *v6 Set *v7 Set ≔ (natv8 , natv9 , natv10 , natv11 , natv12 , natv13 , natv14 , natv15 )v16,v17,v18,v19,v20,v21 . Definition nats2 [...] ≔ (nats1v1,…,v21 , nats1v22,…,v42 , nats1… , nats1… ). ... 𝒞nats1 = {v16 ⊑ v1, v17 ⊑ v2, …} 31 Sized Types Contributions Inference Substitution Performance
  • 32. Artificial Example: Exponentially Many `nat`s Time Definition nats1 := (nat, nat, nat, nat, nat, nat, nat, nat). Time Definition nats2 := (nats1, nats1, nats1, nats1). ... Time Definition nats6 := (nats5, nats5, nats5, nats5). 32 Sized Types Contributions Definition |𝒱| Time (s) nats1 21 0.004 nats2 93 0.020 nats3 381 0.177 nats4 1533 2.299 nats5 6141 35.385 nats6 24573 > 120 Inference Substitution Performance
  • 33. Real Example: MSets/MSetList.v Mean over five trials 33 Sized Types Contributions Unsized compilation (s) 15.122 ± 0.073 Sized compilation (s) 83.660 ± 0.286 Slowdown 5.5× SAT ops only (s) 64.600 ± 0.437 SAT ops only (%) 77.2% Inference Substitution Performance
  • 34. Real Example: MSets/MSetList.v 34 Sized Types Contributions log(count) log distribution of |𝒱| × |𝒞| during SAT operations |𝒱| × |𝒞| Inference Substitution Performance ~4K vs. × ~250 cs.
  • 35. size inference 35 backward-compatible sized typing definitions constraint checking/solving poor performance Sized Types Contributions Implementation
  • 36. Future Work?: Manual Size Annotation + Size Inference 36 Fixpoint minus [v] : natv → nat∞ → natv . Fixpoint div [v] (n : natv ) (m : nat∞ ) : natv ≔ match n with | S n′ ⇒ S (div [?w] (minus [?u] n′ m) m) | O ⇒ O End. 𝒞div = {?u = ?w, ?w+1 ⊑ v} Sized Types Contributions Implementation
  • 37. Thank you! ● Full spec of size inference algorithm & soundness wrt typing ● Metatheoretical results & attempt at set-theoretic model ● More detailed performance profiling & statistics + Closed draft PR: https://github.com/coq/ coq/pull/12426 37 Check out the paper for:
  • 39. Concrete sized naturals O : natv+1 S O : natv+2 S (S O) : natv+3 O : natv+2 S O : natv+3 S (S O) : natv+4 39
  • 40. Real Example: setoid_ring/Field_theory.v 40 Mean over two trials (August 2023) Unsized compilation (s) 17.815 ± 0.545 Sized compilation (s) 106.87 ± 1.94 Slowdown 6.0× SAT ops only (s) 84.70 SAT ops only (%) 79.3% 17755 vs. × 14057 cs. ≈ 250M
  • 41. Artificial Example: Universe Polymorphism Set Printing Universes. Set Universe Polymorphism. Time Definition T1 : Type := Type -> Type -> Type -> Type -> Type -> Type. Print T1. Time Definition T2 : Type := T1 -> T1 -> T1 -> T1 -> T1 -> T1. Print T2. Time Definition T3 : Type := T2 -> T2 -> T2 -> T2 -> T2 -> T2. Print T3. Time Definition T4 : Type := T3 -> T3 -> T3 -> T3 -> T3 -> T3. Print T4. Time Definition T5 : Type := T4 -> T4 -> T4 -> T4 -> T4 -> T4. Print T5. Time Definition T6 : Type := T5 -> T5 -> T5 -> T5 -> T5 -> T5. Print T6. Time Definition T7 : Type := T6 -> T6 -> T6 -> T6 -> T6 -> T6. Print T7. 41 Definition #u Time (s) T1 7 ~ 0 T2 43 0.002 T3 259 0.026 T4 1555 0.057 T5 9331 0.374 T6 55987 3.300 T7 335921 18.170