14 Commits

19 changed files with 1082 additions and 7 deletions

View File

@@ -17,12 +17,32 @@
border-bottom-right-radius: 0; border-bottom-right-radius: 0;
padding-left: 0.5em; padding-left: 0.5em;
padding-right: 0.5rem; padding-right: 0.5rem;
@include below-container-width {
@include bordered-block;
text-align: center;
border-bottom: none;
border-bottom-left-radius: 0;
border-bottom-right-radius: 0;
}
} }
&:last-child { &:last-child {
@include bordered-block; @include bordered-block;
border-top-left-radius: 0; border-top-left-radius: 0;
border-bottom-left-radius: 0; border-bottom-left-radius: 0;
@include below-container-width {
@include bordered-block;
border-top-left-radius: 0;
border-top-right-radius: 0;
}
}
}
tr {
@include below-container-width {
margin-bottom: 0.5rem;
} }
} }

179
code/dawn/Dawn.v Normal file
View File

@@ -0,0 +1,179 @@
Require Import Coq.Lists.List.
From Ltac2 Require Import Ltac2.
Inductive intrinsic :=
| swap
| clone
| drop
| quote
| compose
| apply.
Inductive expr :=
| e_int (i : intrinsic)
| e_quote (e : expr)
| e_comp (e1 e2 : expr).
Definition e_compose (e : expr) (es : list expr) := fold_left e_comp es e.
Inductive IsValue : expr -> Prop :=
| Val_quote : forall {e : expr}, IsValue (e_quote e).
Definition value := { v : expr & IsValue v }.
Definition value_stack := list value.
Definition v_quote (e : expr) := existT IsValue (e_quote e) Val_quote.
Inductive Sem_int : value_stack -> intrinsic -> value_stack -> Prop :=
| Sem_swap : forall (v v' : value) (vs : value_stack), Sem_int (v' :: v :: vs) swap (v :: v' :: vs)
| Sem_clone : forall (v : value) (vs : value_stack), Sem_int (v :: vs) clone (v :: v :: vs)
| Sem_drop : forall (v : value) (vs : value_stack), Sem_int (v :: vs) drop vs
| Sem_quote : forall (v : value) (vs : value_stack), Sem_int (v :: vs) quote ((v_quote (projT1 v)) :: vs)
| Sem_compose : forall (e e' : expr) (vs : value_stack), Sem_int (v_quote e' :: v_quote e :: vs) compose (v_quote (e_comp e e') :: vs)
| Sem_apply : forall (e : expr) (vs vs': value_stack), Sem_expr vs e vs' -> Sem_int (v_quote e :: vs) apply vs'
with Sem_expr : value_stack -> expr -> value_stack -> Prop :=
| Sem_e_int : forall (i : intrinsic) (vs vs' : value_stack), Sem_int vs i vs' -> Sem_expr vs (e_int i) vs'
| Sem_e_quote : forall (e : expr) (vs : value_stack), Sem_expr vs (e_quote e) (v_quote e :: vs)
| Sem_e_comp : forall (e1 e2 : expr) (vs1 vs2 vs3 : value_stack),
Sem_expr vs1 e1 vs2 -> Sem_expr vs2 e2 vs3 -> Sem_expr vs1 (e_comp e1 e2) vs3.
Definition false : expr := e_quote (e_int drop).
Definition false_v : value := v_quote (e_int drop).
Definition true : expr := e_quote (e_comp (e_int swap) (e_int drop)).
Definition true_v : value := v_quote (e_comp (e_int swap) (e_int drop)).
Theorem false_correct : forall (v v' : value) (vs : value_stack), Sem_expr (v' :: v :: vs) (e_comp false (e_int apply)) (v :: vs).
Proof.
intros v v' vs.
eapply Sem_e_comp.
- apply Sem_e_quote.
- apply Sem_e_int. apply Sem_apply. apply Sem_e_int. apply Sem_drop.
Qed.
Theorem true_correct : forall (v v' : value) (vs : value_stack), Sem_expr (v' :: v :: vs) (e_comp true (e_int apply)) (v' :: vs).
Proof.
intros v v' vs.
eapply Sem_e_comp.
- apply Sem_e_quote.
- apply Sem_e_int. apply Sem_apply. eapply Sem_e_comp.
* apply Sem_e_int. apply Sem_swap.
* apply Sem_e_int. apply Sem_drop.
Qed.
Definition or : expr := e_comp (e_int clone) (e_int apply).
Theorem or_false_v : forall (v : value) (vs : value_stack), Sem_expr (false_v :: v :: vs) or (v :: vs).
Proof with apply Sem_e_int.
intros v vs.
eapply Sem_e_comp...
- apply Sem_clone.
- apply Sem_apply... apply Sem_drop.
Qed.
Theorem or_true : forall (v : value) (vs : value_stack), Sem_expr (true_v :: v :: vs) or (true_v :: vs).
Proof with apply Sem_e_int.
intros v vs.
eapply Sem_e_comp...
- apply Sem_clone...
- apply Sem_apply. eapply Sem_e_comp...
* apply Sem_swap.
* apply Sem_drop.
Qed.
Definition or_false_false := or_false_v false_v.
Definition or_false_true := or_false_v true_v.
Definition or_true_false := or_true false_v.
Definition or_true_true := or_true true_v.
Fixpoint quote_n (n : nat) :=
match n with
| O => e_int quote
| S n' => e_compose (quote_n n') (e_int swap :: e_int quote :: e_int swap :: e_int compose :: nil)
end.
Theorem quote_2_correct : forall (v1 v2 : value) (vs : value_stack),
Sem_expr (v2 :: v1 :: vs) (quote_n 1) (v_quote (e_comp (projT1 v1) (projT1 v2)) :: vs).
Proof with apply Sem_e_int.
intros v1 v2 vs. simpl.
repeat (eapply Sem_e_comp)...
- apply Sem_quote.
- apply Sem_swap.
- apply Sem_quote.
- apply Sem_swap.
- apply Sem_compose.
Qed.
Theorem quote_3_correct : forall (v1 v2 v3 : value) (vs : value_stack),
Sem_expr (v3 :: v2 :: v1 :: vs) (quote_n 2) (v_quote (e_comp (projT1 v1) (e_comp (projT1 v2) (projT1 v3))) :: vs).
Proof with apply Sem_e_int.
intros v1 v2 v3 vs. simpl.
repeat (eapply Sem_e_comp)...
- apply Sem_quote.
- apply Sem_swap.
- apply Sem_quote.
- apply Sem_swap.
- apply Sem_compose.
- apply Sem_swap.
- apply Sem_quote.
- apply Sem_swap.
- apply Sem_compose.
Qed.
Ltac2 rec solve_basic () := Control.enter (fun _ =>
match! goal with
| [|- Sem_int ?vs1 swap ?vs2] => apply Sem_swap
| [|- Sem_int ?vs1 clone ?vs2] => apply Sem_clone
| [|- Sem_int ?vs1 drop ?vs2] => apply Sem_drop
| [|- Sem_int ?vs1 quote ?vs2] => apply Sem_quote
| [|- Sem_int ?vs1 compose ?vs2] => apply Sem_compose
| [|- Sem_int ?vs1 apply ?vs2] => apply Sem_apply
| [|- Sem_expr ?vs1 (e_comp ?e1 ?e2) ?vs2] => eapply Sem_e_comp; solve_basic ()
| [|- Sem_expr ?vs1 (e_int ?e) ?vs2] => apply Sem_e_int; solve_basic ()
| [|- Sem_e_quote ?vs1 (e_quote ?e) ?vs2] => apply Sem_quote
| [_ : _ |- _] => ()
end).
Theorem quote_2_correct' : forall (v1 v2 : value) (vs : value_stack),
Sem_expr (v2 :: v1 :: vs) (quote_n 1) (v_quote (e_comp (projT1 v1) (projT1 v2)) :: vs).
Proof. intros. simpl. solve_basic (). Qed.
Theorem quote_3_correct' : forall (v1 v2 v3 : value) (vs : value_stack),
Sem_expr (v3 :: v2 :: v1 :: vs) (quote_n 2) (v_quote (e_comp (projT1 v1) (e_comp (projT1 v2) (projT1 v3))) :: vs).
Proof. intros. simpl. solve_basic (). Qed.
Definition rotate_n (n : nat) := e_compose (quote_n n) (e_int swap :: e_int quote :: e_int compose :: e_int apply :: nil).
Lemma eval_value : forall (v : value) (vs : value_stack),
Sem_expr vs (projT1 v) (v :: vs).
Proof.
intros v vs.
destruct v. destruct i.
simpl. apply Sem_e_quote.
Qed.
Theorem rotate_3_correct : forall (v1 v2 v3 : value) (vs : value_stack),
Sem_expr (v3 :: v2 :: v1 :: vs) (rotate_n 1) (v1 :: v3 :: v2 :: vs).
Proof.
intros. unfold rotate_n. simpl. solve_basic ().
repeat (eapply Sem_e_comp); apply eval_value.
Qed.
Theorem rotate_4_correct : forall (v1 v2 v3 v4 : value) (vs : value_stack),
Sem_expr (v4 :: v3 :: v2 :: v1 :: vs) (rotate_n 2) (v1 :: v4 :: v3 :: v2 :: vs).
Proof.
intros. unfold rotate_n. simpl. solve_basic ().
repeat (eapply Sem_e_comp); apply eval_value.
Qed.
Theorem e_comp_assoc : forall (e1 e2 e3 : expr) (vs vs' : value_stack),
Sem_expr vs (e_comp e1 (e_comp e2 e3)) vs' <-> Sem_expr vs (e_comp (e_comp e1 e2) e3) vs'.
Proof.
intros e1 e2 e3 vs vs'.
split; intros Heval.
- inversion Heval; subst. inversion H4; subst.
eapply Sem_e_comp. eapply Sem_e_comp. apply H2. apply H3. apply H6.
- inversion Heval; subst. inversion H2; subst.
eapply Sem_e_comp. apply H3. eapply Sem_e_comp. apply H6. apply H4.
Qed.

224
code/dawn/DawnEval.v Normal file
View File

@@ -0,0 +1,224 @@
Require Import Coq.Lists.List.
Require Import Dawn.
Require Import Coq.Program.Equality.
Inductive step_result :=
| err
| middle (e : expr) (s : list expr)
| final (s : list expr).
Fixpoint eval_step (s : list expr) (e : expr) : step_result :=
match e, s with
| e_int swap, v' :: v :: vs => final (v :: v' :: vs)
| e_int clone, v :: vs => final (v :: v :: vs)
| e_int drop, v :: vs => final vs
| e_int quote, v :: vs => final (e_quote v :: vs)
| e_int compose, (e_quote v2) :: (e_quote v1) :: vs => final (e_quote (e_comp v1 v2) :: vs)
| e_int apply, (e_quote v1) :: vs => middle v1 vs
| e_quote e', vs => final (e_quote e' :: vs)
| e_comp e1 e2, vs =>
match eval_step vs e1 with
| final vs' => middle e2 vs'
| middle e1' vs' => middle (e_comp e1' e2) vs'
| err => err
end
| _, _ => err
end.
Definition strip_value_list (vs : value_stack) := @map value expr (@projT1 expr IsValue) vs.
Theorem eval_step_correct : forall (e : expr) (vs vs' : value_stack), Sem_expr vs e vs' ->
(eval_step (strip_value_list vs) e = final (strip_value_list vs')) \/
(exists (ei : expr) (vsi : value_stack),
eval_step (strip_value_list vs) e = middle ei (strip_value_list vsi) /\
Sem_expr vsi ei vs').
Proof.
intros e vs vs' Hsem.
(* Proceed by induction on the semantics. *)
induction Hsem.
- inversion H; (* The expression is just an intrnsic. *)
(* Dismiss all the straightforward "final" cases,
of which most intrinsics are *)
try (left; reflexivity).
(* We are in an intermediate / middle case. *)
right.
(* The semantics guarantee that the expression in the
quote evaluates to the final state. *)
exists e, vs0. auto.
- (* The expression is a quote. This is yet another final case. *)
left; reflexivity.
- (* The composition is never a final step, since we have to evaluate both
branches to "finish up". *)
destruct IHHsem1; right.
+ (* If the left branch finihed, only the right branch needs to be evaluted. *)
simpl. rewrite H. exists e2, vs2. auto.
+ (* Otherwise, the left branch has an intermediate evaluation, guaranteed
by induction to be consitent. *)
destruct H as [ei [vsi [Heval Hsem']]].
(* We compose the remaining part of the left branch with the right branch. *)
exists (e_comp ei e2), vsi. simpl.
(* The evaluation is trivially to a "middle" state. *)
rewrite Heval. split. auto.
eapply Sem_e_comp. apply Hsem'. apply Hsem2.
Qed.
Inductive eval_chain (s : list expr) (e : expr) (s' : list expr) : Prop :=
| chain_final (P : eval_step s e = final s')
| chain_middle (ei : expr) (si : list expr)
(P : eval_step s e = middle ei si) (rest : eval_chain si ei s').
Lemma eval_chain_merge : forall (e1 e2 : expr) (s s' s'' : list expr),
eval_chain s e1 s' -> eval_chain s' e2 s'' -> eval_chain s (e_comp e1 e2) s''.
Proof.
intros e1 e2 s s' s'' ch1 ch2.
induction ch1;
eapply chain_middle; simpl; try (rewrite P); auto.
Qed.
Lemma eval_chain_split : forall (e1 e2 : expr) (s s'' : list expr),
eval_chain s (e_comp e1 e2) s'' -> exists s', (eval_chain s e1 s') /\ (eval_chain s' e2 s'').
Proof.
intros e1 e2 s ss'' ch.
dependent induction ch.
- simpl in P. destruct (eval_step s e1); inversion P.
- simpl in P. destruct (eval_step s e1) eqn:Hval; try (inversion P).
+ injection P as Hinj; subst. specialize (IHch e e2 H0) as [s'0 [ch1 ch2]].
eexists. split.
* eapply chain_middle. apply Hval. apply ch1.
* apply ch2.
+ subst. eexists. split.
* eapply chain_final. apply Hval.
* apply ch.
Qed.
Theorem eval_step_preservation : forall (e : expr) (vs : value_stack) (s' : list expr),
((exists e', eval_step (strip_value_list vs) e = middle e' s') ->
exists vs', (strip_value_list vs') = s') /\
((eval_step (strip_value_list vs) e = final s') ->
exists vs', (strip_value_list vs') = s').
Proof.
intros e vs. induction e; split.
- intros [e' Hst].
destruct i; try (destruct vs as [|v1 [|v2 vs]]; inversion Hst; fail).
* destruct vs as [|v1 [|v2 vs]]; inversion Hst;
destruct (projT1 v1); inversion H0.
destruct (projT1 v2); inversion H1.
* destruct vs as [|v1 vs]; inversion Hst.
destruct (projT1 v1); inversion H0.
eauto.
- intros Hst.
destruct i; try (destruct vs as [|v1 [|v2 vs]]; inversion Hst; fail).
* destruct vs as [|v1 [|v2 vs]]; inversion Hst. exists (v2 :: v1 :: vs). reflexivity.
* destruct vs as [|v1 vs]; inversion Hst. exists (v1 :: v1 :: vs). reflexivity.
* destruct vs as [|v1 vs]; inversion Hst. exists vs. reflexivity.
* destruct vs as [|v1 vs]; inversion Hst. exists (v_quote (projT1 v1) :: vs). reflexivity.
* destruct vs as [|v1 [|v2 vs]]; inversion Hst.
+ destruct (projT1 v1); inversion H0.
+ destruct (projT1 v1); destruct (projT1 v2); inversion H0.
exists (v_quote (e_comp e0 e) :: vs). reflexivity.
* destruct vs as [|v1 vs]; inversion Hst;
destruct (projT1 v1); inversion H0.
- intros [e' Hst]. inversion Hst.
- intros Hst. inversion Hst. exists (v_quote e :: vs). reflexivity.
- intros [e' Hst]. inversion Hst.
destruct (eval_step (strip_value_list vs) e1) eqn:Hst1.
* inversion H0.
* specialize (IHe1 s') as [IHe1 _]. apply IHe1. exists e.
injection H0 as Hinj; subst; auto.
* specialize (IHe1 s') as [_ IHe1]. apply IHe1.
injection H0 as Hinj; subst; auto.
- intros Hst.
inversion Hst. destruct (eval_step (strip_value_list vs) e1); inversion H0.
Qed.
Theorem val_step_sem : forall (e : expr) (vs vs' : value_stack),
Sem_expr vs e vs' -> eval_chain (strip_value_list vs) e (strip_value_list vs')
with eval_step_int : forall (i : intrinsic) (vs vs' : value_stack),
Sem_int vs i vs' -> eval_chain (strip_value_list vs) (e_int i) (strip_value_list vs').
Proof.
- intros e vs vs' Hsem.
induction Hsem.
+ (* This is an intrinsic, which is handled by the second
theorem, eval_step_int. This lemma is used here. *)
auto.
+ (* A quote doesn't have a next step, and so is final. *)
apply chain_final. auto.
+ (* In compoition, by induction, we know that the two sub-expressions produce
proper evaluation chains. Chains can be composed (via eval_chain_merge). *)
eapply eval_chain_merge with (strip_value_list vs2); auto.
- intros i vs vs' Hsem.
(* The evaluation chain depends on the specific intrinsic in use. *)
inversion Hsem; subst;
(* Most intrinsics produce a final value, and the evaluation chain is trivial. *)
try (apply chain_final; auto; fail).
(* Only apply is non-final. The first step is popping the quote from the stack,
and the rest of the steps are given by the evaluation of the code in the quote. *)
apply chain_middle with e (strip_value_list vs0); auto.
Qed.
Theorem eval_step_sem_back : forall (e : expr) (vs vs' : value_stack),
eval_chain (strip_value_list vs) e (strip_value_list vs') -> Sem_expr vs e vs'
with eval_step_int_back : forall (i : intrinsic) (vs vs' : value_stack),
eval_chain (strip_value_list vs) (e_int i) (strip_value_list vs') -> Sem_int vs i vs'.
Proof.
(* Thoughts: prove a "step value stack preservation" property. A step suggested by
eval_step will ensure the stack contains values if it contained values initially.
Done! eval_step_preservation.
From there, by induction, a chain preserves value stacks. *)
Admitted.
Theorem eval_step_sem_not : forall (e : expr) (vs : value_stack),
~ (exists vs', Sem_expr vs e vs') -> ~(exists vs', eval_chain (strip_value_list vs) e (strip_value_list vs'))
with eval_step_int_not : forall (i : intrinsic) (vs : value_stack),
~ (exists vs', Sem_int vs i vs') -> ~(exists vs', eval_chain (strip_value_list vs) (e_int i) (strip_value_list vs')).
Proof.
(*
- intros e vs Hnsem [vs' Hev].
destruct e.
+ specialize (eval_step_sem_not _ _ Hnsem). apply eval_step_sem_not.
eexists. apply Hev.
+ apply Hnsem. eexists. apply Sem_e_quote.
+ inversion Hev.
* simpl in P. destruct (eval_step (strip_value_list vs) e1); inversion P.
* specialize (eval_chain_split e1 e2 (strip_value_list vs) (strip_value_list vs') Hev) as [s' [ch1 ch2]].
assert (Hnboth : ~ (exists vsi, Sem_expr vs e1 vsi /\ Sem_expr vsi e2 vs')).
{ intros [vsi [H1 H2]]. apply Hnsem. exists vs'. eapply Sem_e_comp. apply H1. apply H2. }
assert (Hncomp : ~ (exists vsi, Sem_expr vs e1 vsi) \/ exists vsi, Sem_expr vs e1 vsi /\ ~(Sem_expr vsi e2 vs')).
{
- intros i vs Hnint [vs' Hev]. destruct i.
+ destruct vs as [|v [|v' vs]]; inversion Hev; simpl in P; inversion P.
apply Hnint. eexists. apply Sem_swap.
+ destruct vs as [|v vs]; inversion Hev; simpl in P; inversion P.
apply Hnint. eexists. apply Sem_clone.
+ destruct vs as [|v vs]; inversion Hev; simpl in P; inversion P.
apply Hnint. eexists. apply Sem_drop.
+ destruct vs as [|v vs]; inversion Hev; simpl in P; inversion P.
apply Hnint. eexists. apply Sem_quote.
+ destruct vs as [|v [|v' vs]];
try (destruct v; destruct x);
try (destruct v'; destruct x0);
simpl in Hev; inversion Hev; simpl in P; inversion P; subst.
destruct i; destruct i0.
apply Hnint. eexists. apply Sem_compose.
+ destruct vs as [|v vs].
* simpl in Hev; inversion Hev; simpl in P; inversion P.
* (* Monkey at keyboard mode engaged here. *)
destruct v eqn:Hv. destruct i eqn:Hi.
simpl in Hev; inversion Hev; simpl in P; inversion P.
injection P as Heq. subst.
assert (Hna : ~(exists vs', Sem_expr vs ei vs')).
{ intros [vs'0 Hsem]. apply Hnint. eexists. apply Sem_apply. apply Hsem. }
specialize (eval_step_sem_not _ _ Hna). apply eval_step_sem_not. eexists.
apply rest.
*)
Admitted.
Require Extraction.
Require Import ExtrHaskellBasic.
Extraction Language Haskell.
Extraction "UccGen.hs" eval_step true false or.

View File

@@ -0,0 +1,23 @@
class EventEmitter {
constructor() {
this.handlers = {}
}
emit(event) {
this.handlers[event]?.forEach(h => h());
}
addHandler(event, handler) {
if(!this.handlers[event]) {
this.handlers[event] = [handler];
} else {
this.handlers[event].push(handler);
}
}
}
const emitter = new EventEmitter();
emitter.addHandler("start", () => console.log("Started!"));
emitter.addHandler("end", () => console.log("Ended!"));
emitter.emit("end");
emitter.emit("start");

View File

@@ -0,0 +1,23 @@
class EventEmitter {
constructor() {
this.handlers = {}
}
emit(event, value) {
this.handlers[event]?.forEach(h => h(value));
}
addHandler(event, handler) {
if(!this.handlers[event]) {
this.handlers[event] = [handler];
} else {
this.handlers[event].push(handler);
}
}
}
const emitter = new EventEmitter();
emitter.addHandler("numberChange", n => console.log("New number value is: ", n));
emitter.addHandler("stringChange", s => console.log("New string value is: ", s));
emitter.emit("numberChange", 1);
emitter.emit("stringChange", "3");

View File

@@ -0,0 +1,27 @@
class EventEmitter<T> {
private handlers: { [eventName in keyof T]?: ((value: T[eventName]) => void)[] }
constructor() {
this.handlers = {}
}
emit<K extends keyof T>(event: K, value: T[K]): void {
this.handlers[event]?.forEach(h => h(value));
}
addHandler<K extends keyof T>(event: K, handler: (value: T[K]) => void): void {
if(!this.handlers[event]) {
this.handlers[event] = [handler];
} else {
this.handlers[event].push(handler);
}
}
}
const emitter = new EventEmitter<{ numberChange: number, stringChange: string }>();
emitter.addHandler("numberChange", n => console.log("New number value is: ", n));
emitter.addHandler("stringChange", s => console.log("New string value is: ", s));
emitter.emit("numberChange", 1);
emitter.emit("stringChange", "3");
emitter.emit("numberChange", "1");
emitter.emit("stringChange", 3);

View File

@@ -1,7 +1,7 @@
--- ---
title: A Language for an Assignment - Homework 1 title: A Language for an Assignment - Homework 1
date: 2019-12-27T23:27:09-08:00 date: 2019-12-27T23:27:09-08:00
tags: ["Haskell", "Python", "Algorithms"] tags: ["Haskell", "Python", "Algorithms", "Programming Languages"]
--- ---
On a rainy Oregon day, I was walking between classes with a group of friends. On a rainy Oregon day, I was walking between classes with a group of friends.

View File

@@ -1,7 +1,7 @@
--- ---
title: A Language for an Assignment - Homework 2 title: A Language for an Assignment - Homework 2
date: 2019-12-30T20:05:10-08:00 date: 2019-12-30T20:05:10-08:00
tags: ["Haskell", "Python", "Algorithms"] tags: ["Haskell", "Python", "Algorithms", "Programming Languages"]
--- ---
After the madness of the After the madness of the

View File

@@ -1,7 +1,7 @@
--- ---
title: A Language for an Assignment - Homework 3 title: A Language for an Assignment - Homework 3
date: 2020-01-02T22:17:43-08:00 date: 2020-01-02T22:17:43-08:00
tags: ["Haskell", "Python", "Algorithms"] tags: ["Haskell", "Python", "Algorithms", "Programming Languages"]
--- ---
It rained in Sunriver on New Year's Eve, and it continued to rain It rained in Sunriver on New Year's Eve, and it continued to rain

376
content/blog/coq_dawn.md Normal file
View File

@@ -0,0 +1,376 @@
---
title: "Formalizing Dawn in Coq"
date: 2021-11-20T19:04:57-08:00
tags: ["Coq", "Dawn", "Programming Languages"]
---
The [_Foundations of Dawn_](https://www.dawn-lang.org/posts/foundations-ucc/) article came up
on [Lobsters](https://lobste.rs/s/clatuv/foundations_dawn_untyped_concatenative) recently.
In this article, the author of Dawn defines a core calculus for the language, and provides its
semantics. The core calculus is called the _untyped concatenative calculus_, or UCC.
The definitions in the semantics seemed so clean and straightforward that I wanted to try my hand at
translating them into machine-checked code. I am most familiar with [Coq](https://coq.inria.fr/),
and that's what I reached for when making this attempt.
### Defining the Syntax
#### Expressions and Intrinsics
This is mostly the easy part. A UCC expression is one of three things:
* An "intrinsic", written \\(i\\), which is akin to a built-in function or command.
* A "quote", written \\([e]\\), which takes a UCC expression \\(e\\) and moves it onto the stack (UCC is stack-based).
* A composition of several expressions, written \\(e_1\\ e_2\\ \\ldots\\ e_n\\), which effectively evaluates them in order.
This is straightforward to define in Coq, but I'm going to make a little simplifying change.
Instead of making "composition of \\(n\\) expressions" a core language feature, I'll only
allow "composition of \\(e_1\\) and \\(e_2\\)", written \\(e_1\\ e_2\\). This change does not
in any way reduce the power of the language; we can still
{{< sidenote "right" "assoc-note" "write \(e_1\ e_2\ \ldots\ e_n\) as \((e_1\ e_2)\ \ldots\ e_n\)." >}}
The same expression can, of course, be written as \(e_1\ \ldots\ (e_{n-1}\ e_n)\).
So, which way should we <em>really</em> use when translating the many-expression composition
from the Dawn article into the two-expression composition I am using here? Well, the answer is,
it doesn't matter; expression composition is <em>associative</em>, so both ways effectively mean
the same thing.<br>
<br>
This is quite similar to what we do in algebra: the regular old addition operator, \(+\) is formally
only defined for pairs of numbers, like \(a+b\). However, no one really bats an eye when we
write \(1+2+3\), because we can just insert parentheses any way we like, and get the same result:
\((1+2)+3\) is the same as \(1+(2+3)\).
{{< /sidenote >}}
With that in mind, we can translate each of the three types of expressions in UCC into cases
of an inductive data type in Coq.
{{< codelines "Coq" "dawn/Dawn.v" 12 15 >}}
Why do we need `e_int`? We do because a token like \\(\\text{swap}\\) can be viewed
as belonging to the set of intrinsics \\(i\\), or the set of expressions, \\(e\\). While writing
down the rules in mathematical notation, what exactly the token means is inferred from context - clearly
\\(\\text{swap}\\ \\text{drop}\\) is an expression built from two other expressions. In statically-typed
functional languages like Coq or Haskell, however, the same expression can't belong to two different,
arbitrary types. Thus, to turn an intrinsic into an expression, we need to wrap it up in a constructor,
which we called `e_int` here. Other than that, `e_quote` accepts as argument another expression, `e` (the
thing being quoted), and `e_comp` accepts two expressions, `e1` and `e2` (the two sub-expressions being composed).
The definition for intrinsics themselves is even simpler:
{{< codelines "Coq" "dawn/Dawn.v" 4 10 >}}
We simply define a constructor for each of the six intrinsics. Since none of the intrinsic
names are reserved in Coq, we can just call our constructors exactly the same as their names
in the written formalization.
#### Values and Value Stacks
Values are up next. My initial thought was to define a value much like
I defined an intrinsic expression: by wrapping an expression in a constructor for a new data
type. Something like:
```Coq
Inductive value :=
| v_quot (e : expr).
```
Then, `v_quot (e_int swap)` would be the Coq translation of the expression \\([\\text{swap}]\\).
However, I didn't decide on this approach for two reasons:
* There are now two ways to write a quoted expression: either `v_quote e` to represent
a quoted expression that is a value, or `e_quote e` to represent a quoted expression
that is just an expression. In the extreme case, the value \\([[e]]\\) would
be represented by `v_quote (e_quote e)` - two different constructors for the same concept,
in the same expression!
* When formalizing the lambda calculus,
[Programming Language Foundations](https://softwarefoundations.cis.upenn.edu/plf-current/Stlc.html)
uses an inductively-defined property to indicate values. In the simply typed lambda calculus,
much like in UCC, values are a subset of expressions.
I took instead the approach from Programming Language Foundations: a value is merely an expression
for which some predicate, `IsValue`, holds. We will define this such that `IsValue (e_quote e)` is provable,
but also such that here is no way to prove `IsValue (e_int swap)`, since _that_ expression is not
a value. But what does "provable" mean, here?
By the [Curry-Howard correspondence](https://en.wikipedia.org/wiki/Curry%E2%80%93Howard_correspondence),
a predicate is just a function that takes _something_ and returns a type. Thus, if \\(\\text{Even}\\)
is a predicate, then \\(\\text{Even}\\ 3\\) is actually a type. Since \\(\\text{Even}\\) takes
numbers in, it is a predicate on numbers. Our \\(\\text{IsValue}\\) predicate will be a predicate
on expressions, instead. In Coq, we can write this as:
{{< codelines "Coq" "dawn/Dawn.v" 19 19 >}}
You might be thinking,
> Huh, `Prop`? But you just said that predicates return types!
This is a good observation; In Coq, `Prop` is a special sort of type that corresponds to logical
propositions. It's special for a few reasons, but those reasons are beyond the scope of this post;
for our purposes, it's sufficient to think of `IsValue e` as a type.
Alright, so what good is this new `IsValue e` type? Well, we will define `IsValue` such that
this type is only _inhabited_ if `e` is a value according to the UCC specification. A type
is inhabited if and only if we can find a value of that type. For instance, the type of natural
numbers, `nat`, is inhabited, because any number, like `0`, has this type. Uninhabited types
are harder to come by, but take as an example the type `3 = 4`, the type of proofs that three is equal
to four. Three is _not_ equal to four, so we can never find a proof of equality, and thus, `3 = 4` is
uninhabited. As I said, `IsValue e` will only be inhabited if `e` is a value per the formal
specification of UCC; specifically, this means that `e` is a quoted expression, like `e_quote e'`.
To this end, we define `IsValue` as follows:
{{< codelines "Coq" "dawn/Dawn.v" 19 20 >}}
Now, `IsValue` is a new data type with only only constructor, `ValQuote`. For any expression `e`,
this constructor creates a value of type `IsValue (e_quote e)`. Two things are true here:
* Since `Val_quote` accepts any expression `e` to be put inside `e_quote`, we can use
`Val_quote` to create an `IsValue` instance for any quoted expression.
* Because `Val_quote` is the _only_ constructor, and because it always returns `IsValue (e_quote e)`,
there's no way to get `IsValue (e_int i)`, or anything else.
Thus, `IsValue e` is inhabited if and only if `e` is a UCC value, as we intended.
Just one more thing. A value is just an expression, but Coq only knows about this as long
as there's an `IsValue` instance around to vouch for it. To be able to reason about values, then,
we will need both the expression and its `IsValue` proof. Thus, we define the type `value` to mean
a pair of two things: an expression `v` and a proof that it's a value, `IsValue v`:
{{< codelines "Coq" "dawn/Dawn.v" 22 22 >}}
A value stack is just a list of values:
{{< codelines "Coq" "dawn/Dawn.v" 23 23 >}}
### Semantics
Remember our `IsValue` predicate? Well, it's not just any predicate, it's a _unary_ predicate.
_Unary_ means that it's a predicate that only takes one argument, an expression in our case. However,
this is far from the only type of predicate. Here are some examples:
* Equality, `=`, is a binary predicate in Coq. It takes two arguments, say `x` and `y`, and builds
a type `x = y` that is only inhabited if `x` and `y` are equal.
* The mathematical "less than" relation is also a binary predicate, and it's called `le` in Coq.
It takes two numbers `n` and `m` and returns a type `le n m` that is only inhabited if `n` is less
than or equal to `m`.
* The evaluation relation in UCC is a ternary predicate. It takes two stacks, `vs` and `vs'`,
and an expression, `e`, and creates a type that's inhabited if and only if evaluating
`e` starting at a stack `vs` results in the stack `vs'`.
Binary predicates are just functions of two inputs that return types. For instance, here's what Coq has
to say about the type of `eq`:
```
eq : ?A -> ?A -> Prop
```
By a similar logic, ternary predicates, much like UCC's evaluation relation, are functions
of three inputs. We can thus write the type of our evaluation relation as follows:
{{< codelines "Coq" "dawn/Dawn.v" 35 35 >}}
We define the constructors just like we did in our `IsValue` predicate. For each evaluation
rule in UCC, such as:
{{< latex >}}
\langle V, v, v'\rangle\ \text{swap}\ \rightarrow\ \langle V, v', v \rangle
{{< /latex >}}
We introduce a constructor. For the `swap` rule mentioned above, the constructor looks like this:
{{< codelines "Coq" "dawn/Dawn.v" 28 28 >}}
Although the stacks are written in reverse order (which is just a consequence of Coq's list notation),
I hope that the correspondence is fairly clear. If it's not, try reading this rule out loud:
> The rule `Sem_swap` says that for every two values `v` and `v'`, and for any stack `vs`,
evaluating `swap` in the original stack `v' :: v :: vs`, aka \\(\\langle V, v, v'\\rangle\\),
results in a final stack `v :: v' :: vs`, aka \\(\\langle V, v', v\\rangle\\).
With that in mind, here's a definition of a predicate `Sem_int`, the evaluation predicate
for intrinsics:
{{< codelines "Coq" "dawn/Dawn.v" 27 33 >}}
Hey, what's all this with `v_quote` and `projT1`? It's just a little bit of bookkeeping.
Given a value -- a pair of an expression `e` and a proof `IsValue e` -- the function `projT1`
just returns the expression `e`. That is, it's basically a way of converting a value back into
an expression. The function `v_quote` takes us in the other direction: given an expression \\(e\\),
it constructs a quoted expression \\([e]\\), and combines it with a proof that the newly constructed
quote is a value.
The above two function in combination help us define the `quote` intrinsic, which
wraps a value on the stack in an additional layer of quotes. When we create a new quote, we
need to push it onto the value stack, so it needs to be a value; we thus use `v_quote`. However,
`v_quote` needs an expression to wrap in a quote, so we use `projT1` to extract the expression from
the value on top of the stack.
In addition to intrinsics, we also define the evaluation relation for actual expressions.
{{< codelines "Coq" "dawn/Dawn.v" 35 39 >}}
Here, we may as well go through the three constructors to explain what they mean:
* `Sem_e_int` says that if the expression being evaluated is an intrinsic, and if the
intrinsic has an effect on the stack as described by `Sem_int` above, then the effect
of the expression itself is the same.
* `Sem_e_quote` says that if the expression is a quote, then a corresponding quoted
value is placed on top of the stack.
* `Sem_e_comp` says that if one expression `e1` changes the stack from `vs1` to `vs2`,
and if another expression `e2` takes this new stack `vs2` and changes it into `vs3`,
then running the two expressions one after another (i.e. composing them) means starting
at stack `vs1` and ending in stack `vs3`.
### \\(\\text{true}\\), \\(\\text{false}\\), \\(\\text{or}\\) and Proofs
Now it's time for some fun! The UCC language specification starts by defining two values:
true and false. Why don't we do the same thing?
|UCC Spec| Coq encoding |
|---|----|
|\\(\\text{false}\\)=\\([\\text{drop}]\\)| {{< codelines "Coq" "dawn/Dawn.v" 41 42 >}}
|\\(\\text{true}\\)=\\([\\text{swap} \\ \\text{drop}]\\)| {{< codelines "Coq" "dawn/Dawn.v" 44 45 >}}
Let's try prove that these two work as intended.
{{< codelines "Coq" "dawn/Dawn.v" 47 53 >}}
This is the first real proof in this article. Rather than getting into the technical details,
I invite you to take a look at the "shape" of the proof:
* After the initial use of `intros`, which brings the variables `v`, `v`, and `vs` into
scope, we start by applying `Sem_e_comp`. Intuitively, this makes sense - at the top level,
our expression, \\(\\text{false}\\ \\text{apply}\\),
is a composition of two other expressions, \\(\\text{false}\\) and \\(\\text{apply}\\).
Because of this, we need to use the rule from our semantics that corresponds to composition.
* The composition rule requires that we describe the individual effects on the stack of the
two constituent expressions (recall that the first expression takes us from the initial stack `v1`
to some intermediate stack `v2`, and the second expression takes us from that stack `v2` to the
final stack `v3`). Thus, we have two "bullet points":
* The first expression, \\(\\text{false}\\), is just a quoted expression. Thus, the rule
`Sem_e_quote` applies, and the contents of the quote are puhsed onto the stack.
* The second expression, \\(\\text{apply}\\), is an intrinsic, so we need to use the rule
`Sem_e_int`, which handles the intrinsic case. This, in turn, requires that we show
the effect of the intrinsic itself; the `apply` intrinsic evaluates the quoted expression
on the stack.
The quoted expression contains the body of false, or \\(\\text{drop}\\). This is
once again an intrinsic, so we use `Sem_e_int`; the intrinsic in question is \\(\\text{drop}\\),
so the `Sem_drop` rule takes care of that.
Following these steps, we arrive at the fact that evaluating `false` on the stack simply drops the top
element, as specified. The proof for \\(\\text{true}\\) is very similar in spirit:
{{< codelines "Coq" "dawn/Dawn.v" 55 63 >}}
We can also formalize the \\(\\text{or}\\) operator:
|UCC Spec| Coq encoding |
|---|----|
|\\(\\text{or}\\)=\\(\\text{clone}\\ \\text{apply}\\)| {{< codelines "Coq" "dawn/Dawn.v" 65 65 >}}
We can write two top-level proofs about how this works: the first says that \\(\\text{or}\\),
when the first argument is \\(\\text{false}\\), just returns the second argument (this is in agreement
with the truth table, since \\(\\text{false}\\) is the identity element of \\(\\text{or}\\)).
The proof proceeds much like before:
{{< codelines "Coq" "dawn/Dawn.v" 67 73 >}}
To shorten the proof a little bit, I used the `Proof with` construct from Coq, which runs
an additional _tactic_ (like `apply`) whenever `...` is used.
Because of this, in this proof writing `apply Sem_apply...` is the same
as `apply Sem_apply. apply Sem_e_int`. Since the `Sem_e_int` rule is used a lot, this makes for a
very convenient shorthand.
Similarly, we prove that \\(\\text{or}\\) applied to \\(\\text{true}\\) always returns \\(\\text{true}\\).
{{< codelines "Coq" "dawn/Dawn.v" 75 83 >}}
Finally, the specific facts (like \\(\\text{false}\\ \\text{or}\\ \\text{false}\\) evaluating to \\(\\text{false}\\))
can be expressed using our two new proofs, `or_false_v` and `or_true`.
{{< codelines "Coq" "dawn/Dawn.v" 85 88 >}}
### Derived Expressions
#### Quotes
The UCC specification defines \\(\\text{quote}_n\\) to make it more convenient to quote
multiple terms. For example, \\(\\text{quote}_2\\) composes and quotes the first two values
on the stack. This is defined in terms of other UCC expressions as follows:
{{< latex >}}
\text{quote}_n = \text{quote}_{n-1}\ \text{swap}\ \text{quote}\ \text{swap}\ \text{compose}
{{< /latex >}}
We can write this in Coq as follows:
{{< codelines "Coq" "dawn/Dawn.v" 90 94 >}}
This definition diverges slightly from the one given in the UCC specification; particularly,
UCC's spec mentions that \\(\\text{quote}_n\\) is only defined for \\(n \\geq 1\\).However,
this means that in our code, we'd have to somehow handle the error that would arise if the
term \\(\\text{quote}\_0\\) is used. Instead, I defined `quote_n n` to simply mean
\\(\\text{quote}\_{n+1}\\); thus, in Coq, no matter what `n` we use, we will have a valid
expression, since `quote_n 0` will simply correspond to \\(\\text{quote}_1 = \\text{quote}\\).
We can now attempt to prove that this definition is correct by ensuring that the examples given
in the specification are valid. We may thus write,
{{< codelines "Coq" "dawn/Dawn.v" 96 106 >}}
We used a new tactic here, `repeat`, but overall, the structure of the proof is pretty straightforward:
the definition of `quote_n` consists of many intrinsics, and we apply the corresponding rules
one-by-one until we arrive at the final stack. Writing this proof was kind of boring, since
I just had to see which intrinsic is being used in each step, and then write a line of `apply`
code to handle that intrinsic. This gets worse for \\(\\text{quote}_3\\):
{{< codelines "Coq" "dawn/Dawn.v" 108 122 >}}
It's so long! Instead, I decided to try out Coq's `Ltac2` mechanism to teach Coq how
to write proofs like this itself. Here's what I came up with:
{{< codelines "Coq" "dawn/Dawn.v" 124 136 >}}
You don't have to understand the details, but in brief, this checks what kind of proof
we're asking Coq to do (for instance, if we're trying to prove that a \\(\\text{swap}\\)
instruction has a particular effect), and tries to apply a corresponding semantic rule.
Thus, it will try `Sem_swap` if the expression is \\(\\text{swap}\\),
`Sem_clone` if the expression is \\(\\text{clone}\\), and so on. Then, the two proofs become:
{{< codelines "Coq" "dawn/Dawn.v" 138 144 >}}
#### Rotations
There's a little trick to formalizing rotations. Values have an important property:
when a value is run against a stack, all it does is place itself on a stack. We can state
this as follows:
{{< latex >}}
\langle V \rangle\ v = \langle V\ v \rangle
{{< /latex >}}
Or, in Coq,
{{< codelines "Coq" "dawn/Dawn.v" 148 149 >}}
This is the trick to how \\(\\text{rotate}_n\\) works: it creates a quote of \\(n\\) reordered and composed
values on the stack, and then evaluates that quote. Since evaluating each value
just places it on the stack, these values end up back on the stack, in the same order that they
were in the quote. When writing the proof, `solve_basic ()` gets us almost all the way to the
end (evaluating a list of values against a stack). Then, we simply apply the composition
rule over and over, following it up with `eval_value` to prove that the each value is just being
placed back on the stack.
{{< codelines "Coq" "dawn/Dawn.v" 156 168 >}}
### `e_comp` is Associative
When composing three expressions, which way of inserting parentheses is correct?
Is it \\((e_1\\ e_2)\\ e_3\\)? Or is it \\(e_1\\ (e_2\\ e_3)\\)? Well, both!
Expression composition is associative, which means that the order of the parentheses
doesn't matter. We state this in the following theorem, which says that the two
ways of writing the composition, if they evaluate to anything, evaluate to the same thing.
{{< codelines "Coq" "dawn/Dawn.v" 170 171 >}}
### Conclusion
That's all I've got in me for today. However, we got pretty far! The UCC specification
says:
> One of my long term goals for UCC is to democratize formal software verification in order to make it much more feasible and realistic to write perfect software.
I think that UCC is definitely getting there: formally defining the semantics outlined
on the page was quite straightforward. We can now have complete confidence in the behavior
of \\(\\text{true}\\), \\(\\text{false}\\), \\(\\text{or}\\), \\(\\text{quote}_n\\) and
\\(\\text{rotate}_n\\). The proof of associativity is also enough to possibly argue for simplifying
the core calculus' syntax even more. All of this we got from an official source, with only
a little bit of tweaking to get from the written description of the language to code! I'm
looking forward to reading the next post about the _multistack_ concatenative calculus.

View File

@@ -0,0 +1,83 @@
---
title: "I Don't Like Coq's Documentation"
date: 2021-11-24T21:48:59-08:00
draft: true
tags: ["Coq"]
---
Recently, I wrote an article on [Formalizing Dawn's Core Calculus in Coq]({{< relref "./coq_dawn.md" >}}).
One of the proofs (specifically, correctness of \\(\\text{quote}_3\\)) was the best candidate I've ever
encountered for proof automation. I knew that proof automation was possible from the second book of Software
Foundations, [Programming Language Foundations](https://softwarefoundations.cis.upenn.edu/plf-current/index.html).
I went there to learn more, and started my little journey in picking up Coq's `Ltac2`.
Before I go any further, let me say that I'd self-describe as an "advanced beginner" in Coq, maybe "intermediate"
on a good day. I am certainly far from a master. I will also say that I am quite young, and thus possibly spoiled
by the explosion of well-documented languages, tools, and libraries. I don't frequently check `man` pages, and I don't
often read straight up technical manuals. Maybe the fault lies with me.
Nevertheles, I feel like I am where I am in the process of learning Coq
in part because of the state of its learning resources.
As a case study, let's take my attempt to automate away a pretty simple proof.
### Grammars instead of Examples
I did not specifically remember the chapter of Software Foundation in which tactics were introduced.
Instead of skimming through chapters until I found it, I tried to look up "coq custom tactic". The
first thing that comes up is the page for `Ltac`.
After a brief explanation of what `Ltac` is, the documentation jumps straight into the grammar of the entire
tactic language. Here' a screenshot of what that looks like:
{{< figure src="ltac_grammar.png" caption="The grammar of Ltac from the Coq page." class="large" alt="A grammar for the `Ltac` language within Coq. The grammar is in BackusNaur form, and has 9 nonterminals." >}}
Good old Backus-Naur form. Now, my two main areas of interest are programming language theory and compilers, and so I'm no stranger to BNF. In fact, the first
time I saw such a documentation page (most pages describing Coq language feature have some production rules), I didn't even consciously process that I was looking
at grammar rules. However, and despite CompCert (a compiler) being probably the most well known example of a project in Coq, I don't think that Coq is made _just_
for people familiar with PLT or compilers. I don't think it should be expected for the average newcomer to Coq (or the average beginner-ish person like me) to know how to read production rules, or to know
what terminals and nonterminals are. Logical Founadtions sure managed to explain Gallina without resorting to BNFs.
And even if I, as a newcomer, know what BNF is, and how to read the rules, there's another layer to this specification:
the precedence of various operators is encoded in the BNF rules. This is a pretty common pattern
for writing down programming language grammars; for each level of operator precedence, there's another
nonterminal. We have `ltac_expr4` for sequencing (the least binding operator), and `ltac_expr3` for "level 3 tactics",
`ltac_expr2` for addition, logical "or", and "level 2 tactics". The creators of this documentation page clearly knew
what they were getting at here, and I've seen this pattern enough times to recognize it right away. But if you
_haven't_ seen this pattern before (and why should you have?), you'll need to take some time to decipher the rules.
That's time that you'd rather spend trying to write your tactic.
The page could just as well have mentioned the types of constructs in `Ltac`, and given a table of their relative precedence.
This could be an opportunity to give an example of what a program in `Ltac` looks like. Instead, despite having seen all of these nonterminals,
I still don't have an image in my mind's eye of what the language looks like. And better yet, I think that the grammar's incorrect:
```
ltac_expr4 ::= ltac_expr3 ; ltac_expr3|binder_tactic
```
The way this is written, there's no way to sequence (using a semicolon) more than two things. The semicolon
only occurs on level four, and both nonterminals in this rule are level three. However, Coq is perfectly happy
to accept the following:
```Coq
Ltac test := auto; auto; auto.
```
In the `Ltac2` grammar, this is written the way I'd expect:
```
ltac2_expr ::= ltac2_expr5 ; ltac2_expr
```
Let's do a quick recap. We have an encoding that requires a degree of experience with grammars and
programming languages to be useful to the reader, _and_ this encoding
{{< sidenote "right" "errors-note" "leaves room for errors," >}}
Here's a question to ponder: how come this error has gone unnoticed? Surely
people used this page to learn <code>Ltac</code>. I believe that part of the reason is
pattern matching: an experienced reader will recognize the "precedence trick", and quickly
scan the grammar levels to estblish precedence. The people writing and proofreading this
documentation likely read it this way, too.
{{< /sidenote >}}
errors that _actually appear in practice_.
{{< sidenote "left" "achilles-note" "We have a very precise, yet incorrect definition." >}}
Amusingly, I think this is very close to what is considered the achilles heel of formal verification:
software that precisely adheres to an incorrect or incomplete specification.
{{< /sidenote >}}
Somehow, "a sequence of statements separated by a semicolon" seems like a simpler explanation.

Binary file not shown.

After

Width:  |  Height:  |  Size: 188 KiB

View File

@@ -1,7 +1,7 @@
--- ---
title: "A Typesafe Representation of an Imperative Language" title: "A Typesafe Representation of an Imperative Language"
date: 2020-11-02T01:07:21-08:00 date: 2020-11-02T01:07:21-08:00
tags: ["Idris"] tags: ["Idris", "Programming Languages"]
--- ---
A recent homework assignment for my university's programming languages A recent homework assignment for my university's programming languages

View File

@@ -1,7 +1,7 @@
--- ---
title: Meaningfully Typechecking a Language in Idris title: Meaningfully Typechecking a Language in Idris
date: 2020-02-27T21:58:55-08:00 date: 2020-02-27T21:58:55-08:00
tags: ["Haskell", "Idris"] tags: ["Haskell", "Idris", "Programming Languages"]
--- ---
This term, I'm a TA for Oregon State University's Programming Languages course. This term, I'm a TA for Oregon State University's Programming Languages course.

View File

@@ -1,7 +1,7 @@
--- ---
title: Meaningfully Typechecking a Language in Idris, Revisited title: Meaningfully Typechecking a Language in Idris, Revisited
date: 2020-07-22T14:37:35-07:00 date: 2020-07-22T14:37:35-07:00
tags: ["Idris"] tags: ["Idris", "Programming Languages"]
favorite: true favorite: true
--- ---

View File

@@ -1,7 +1,7 @@
--- ---
title: Meaningfully Typechecking a Language in Idris, With Tuples title: Meaningfully Typechecking a Language in Idris, With Tuples
date: 2020-08-12T15:48:04-07:00 date: 2020-08-12T15:48:04-07:00
tags: ["Idris"] tags: ["Idris", "Programming Languages"]
--- ---
Some time ago, I wrote a post titled Some time ago, I wrote a post titled

View File

@@ -0,0 +1,120 @@
---
title: "Type-Safe Event Emitter in TypeScript"
date: 2021-09-04T17:18:49-07:00
tags: ["TypeScript"]
---
I've been playing around with TypeScript recently, and enjoying it too.
Nearly all of my compile-time type safety desires have been accomodated
by the language, and in a rather intuitive and clean way. Today, I'm going
to share a little trick I've discovered which allows me to do something that
I suspect would normally require [dependent types](https://en.wikipedia.org/wiki/Dependent_type).
### The Problem
Suppose you want to write a class that emits events. Clients can then install handlers,
functions that are notified whenever an event is emitted. Easy enough; in JavaScript,
this would look something like the following:
{{< codelines "JavaScript" "typescript-emitter/js1.js" 1 17 >}}
We can even write some code to test that this works (just to ease my nerves):
{{< codelines "JavaScript" "typescript-emitter/js1.js" 19 23 >}}
As expected, we get:
```
Ended!
Started!
```
As you probably guessed, we're going to build on this problem a little bit.
In certain situations, you don't just care that an event occured; you also
care about additional event data. For instance, when a number changes, you
may want to know the number's new value. In JavaScript, this is a trivial change:
{{< codelines "JavaScript" "typescript-emitter/js2.js" 1 17 "hl_lines = 6-8" >}}
That's literally it. Once again, let's ensure that this works by sending two new events:
`stringChange` and `numberChange`.
{{< codelines "JavaScript" "typescript-emitter/js2.js" 19 23 >}}
The result of this code is once again unsurprising:
```
New number value is: 1
New string value is: 3
```
But now, how would one go about encoding this in TypeScript? In particular, what is the
type of a handler? We could, of course, give each handler the type `(value: any) => void`.
This, however, makes handlers unsafe. We could very easily write:
```TypeScript
emitter.addHandler("numberChanged", (value: string) => {
console.log("String length is", value.length);
});
emitted.emit("numberChanged", 1);
```
Which would print out:
```
String length is undefined
```
No, I don't like this. TypeScript is supposed to be all about adding type safety to our code,
and this is not at all type safe. We could do all sorts of weird things! There is a way,
however, to make this use case work.
### The Solution
Let me show you what I came up with:
{{< codelines "TypeScript" "typescript-emitter/ts.ts" 1 19 "hl_lines=1 2 8 12">}}
The important changes are on lines 1, 2, 8, and 12 (highlighted in the above code block).
Let's go through each one of them step-by-step.
* __Line 1__: Parameterize the `EventEmitter` by some type `T`. We will use this type `T`
to specify the exact kind of events that our `EventEmitter` will be able to emit and handle.
Specifically, this type will be in the form `{ event: EventValueType }`. For example,
for a `mouseClick` event, we may write `{ mouseClick: { x: number, y: number }}`.
* __Line 2__: Add a proper type to `handlers`. This requires several ingredients of its own:
* We use [index signatures](https://www.typescriptlang.org/docs/handbook/2/objects.html#index-signatures)
to limit the possible names to which handlers can be assigned. We limit these names
to the keys of our type `T`; in the preceding example, `keyof T` would be `"mouseClick"`.
* We also limit the values: `T[eventName]` retrieves the type of the value associated with
key `eventName`. In the mouse example, this type would be `{ x: number, y: number }`. We require
that a key can only be associated with an array of functions to void, each of which accepts
`T[K]` as first argument. This is precisely what we want; for example, `mouseClick` would map to
an array of functions that accept the mouse click location.
* __Line 8__: We restrict the type of `emit` to only accept values that correspond to the keys
of the type `T`. We can't simply write `event: keyof T`, because this would not give us enough
information to retrieve the type of `value`: if `event` is just `keyof T`,
then `value` can be any of the values of `T`. Instead, we use generics; this way, when the
function is called with `"mouseClick"`, the type of `K` is inferred to also be `"mouseClick"`, which
gives TypeScript enough information to narrow the type of `value`.
* __Line 12__: We use the exact same trick here as we did on line 8.
Let's give this a spin with our `numberChange`/`stringChange` example from earlier:
{{< codelines "TypeScript" "typescript-emitter/ts.ts" 21 27 >}}
The function calls on lines 24 and 25 are correct, but the subsequent two (on lines 26 and 27)
are not, as they attempt to emit the _opposite_ type of the one they're supposed to. And indeed,
TypeScript complains about only these two lines:
```
code/typescript-emitter/ts.ts:26:30 - error TS2345: Argument of type 'string' is not assignable to parameter of type 'number'.
26 emitter.emit("numberChange", "1");
~~~
code/typescript-emitter/ts.ts:27:30 - error TS2345: Argument of type 'number' is not assignable to parameter of type 'string'.
27 emitter.emit("stringChange", 3);
~
Found 2 errors.
```
And there you have it! This approach is now also in use in [Hydrogen](https://github.com/vector-im/hydrogen-web),
a lightweight chat client for the [Matrix](https://matrix.org/) protocol. In particular, check out [`EventEmitter.ts`](https://github.com/vector-im/hydrogen-web/blob/master/src/utils/EventEmitter.ts).

Binary file not shown.

Binary file not shown.