Revising for this week's lectures

This commit is contained in:
Adam Chlipala 2022-03-27 13:40:08 -04:00
parent c5a69b6253
commit 33733a0450
4 changed files with 50 additions and 50 deletions

View file

@ -47,8 +47,8 @@ Qed.
Definition pred_strong1 (n : nat) : n > 0 -> nat := Definition pred_strong1 (n : nat) : n > 0 -> nat :=
match n with match n with
| O => fun pf : 0 > 0 => match zgtz pf with end | O => fun pf : 0 > 0 => match zgtz pf with end
| S n' => fun _ => n' | S n' => fun _ => n'
end. end.
(* We expand the type of [pred] to include a _proof_ that its argument [n] is (* We expand the type of [pred] to include a _proof_ that its argument [n] is
@ -76,8 +76,8 @@ Compute pred_strong1 two_gt0.
Fail Definition pred_strong1' (n : nat) (pf : n > 0) : nat := Fail Definition pred_strong1' (n : nat) (pf : n > 0) : nat :=
match n with match n with
| O => match zgtz pf with end | O => match zgtz pf with end
| S n' => n' | S n' => n'
end. end.
(* The term [zgtz pf] fails to type-check. Somehow the type checker has failed (* The term [zgtz pf] fails to type-check. Somehow the type checker has failed
@ -99,8 +99,8 @@ Fail Definition pred_strong1' (n : nat) (pf : n > 0) : nat :=
Definition pred_strong1' (n : nat) : n > 0 -> nat := Definition pred_strong1' (n : nat) : n > 0 -> nat :=
match n return n > 0 -> nat with match n return n > 0 -> nat with
| O => fun pf : 0 > 0 => match zgtz pf with end | O => fun pf : 0 > 0 => match zgtz pf with end
| S n' => fun _ => n' | S n' => fun _ => n'
end. end.
(* By making explicit the functional relationship between value [n] and the (* By making explicit the functional relationship between value [n] and the
@ -134,8 +134,8 @@ Locate "{ _ : _ | _ }".
Definition pred_strong2 (s : {n : nat | n > 0} ) : nat := Definition pred_strong2 (s : {n : nat | n > 0} ) : nat :=
match s with match s with
| exist O pf => match zgtz pf with end | exist O pf => match zgtz pf with end
| exist (S n') _ => n' | exist (S n') _ => n'
end. end.
(* To build a value of a subset type, we use the [exist] constructor, and the (* To build a value of a subset type, we use the [exist] constructor, and the
@ -165,8 +165,8 @@ Extraction pred_strong2.
Definition pred_strong3 (s : {n : nat | n > 0}) : {m : nat | proj1_sig s = S m} := Definition pred_strong3 (s : {n : nat | n > 0}) : {m : nat | proj1_sig s = S m} :=
match s return {m : nat | proj1_sig s = S m} with match s return {m : nat | proj1_sig s = S m} with
| exist 0 pf => match zgtz pf with end | exist 0 pf => match zgtz pf with end
| exist (S n') pf => exist _ n' (eq_refl _) | exist (S n') pf => exist _ n' (eq_refl _)
end. end.
Compute pred_strong3 (exist _ 2 two_gt0). Compute pred_strong3 (exist _ 2 two_gt0).
@ -197,8 +197,8 @@ Extraction pred_strong3.
Definition pred_strong4 : forall (n : nat), n > 0 -> {m : nat | n = S m}. Definition pred_strong4 : forall (n : nat), n > 0 -> {m : nat | n = S m}.
refine (fun n => refine (fun n =>
match n with match n with
| O => fun _ => False_rec _ _ | O => fun _ => False_rec _ _
| S n' => fun _ => exist _ n' _ | S n' => fun _ => exist _ n' _
end). end).
(* We build [pred_strong4] using tactic-based proving, beginning with a (* We build [pred_strong4] using tactic-based proving, beginning with a
@ -222,8 +222,8 @@ Definition pred_strong4 : forall (n : nat), n > 0 -> {m : nat | n = S m}.
Undo. Undo.
refine (fun n => refine (fun n =>
match n with match n with
| O => fun _ => False_rec _ _ | O => fun _ => False_rec _ _
| S n' => fun _ => exist _ n' _ | S n' => fun _ => exist _ n' _
end); equality || linear_arithmetic. end); equality || linear_arithmetic.
Defined. Defined.
@ -252,8 +252,8 @@ Notation "[ e ]" := (exist _ e _).
Definition pred_strong5 : forall (n : nat), n > 0 -> {m : nat | n = S m}. Definition pred_strong5 : forall (n : nat), n > 0 -> {m : nat | n = S m}.
refine (fun n => refine (fun n =>
match n with match n with
| O => fun _ => ! | O => fun _ => !
| S n' => fun _ => [n'] | S n' => fun _ => [n']
end); equality || linear_arithmetic. end); equality || linear_arithmetic.
Defined. Defined.
@ -297,9 +297,9 @@ Notation "'Reduce' x" := (if x then Yes else No) (at level 50).
Definition eq_nat_dec : forall n m : nat, {n = m} + {n <> m}. Definition eq_nat_dec : forall n m : nat, {n = m} + {n <> m}.
refine (fix f (n m : nat) : {n = m} + {n <> m} := refine (fix f (n m : nat) : {n = m} + {n <> m} :=
match n, m with match n, m with
| O, O => Yes | O, O => Yes
| S n', S m' => Reduce (f n' m') | S n', S m' => Reduce (f n' m')
| _, _ => No | _, _ => No
end); equality. end); equality.
Defined. Defined.
@ -349,8 +349,8 @@ Section In_dec.
Definition In_dec : forall (x : A) (ls : list A), {In x ls} + {~ In x ls}. Definition In_dec : forall (x : A) (ls : list A), {In x ls} + {~ In x ls}.
refine (fix f (x : A) (ls : list A) : {In x ls} + {~ In x ls} := refine (fix f (x : A) (ls : list A) : {In x ls} + {~ In x ls} :=
match ls with match ls with
| nil => No | nil => No
| x' :: ls' => A_eq_dec x x' || f x ls' | x' :: ls' => A_eq_dec x x' || f x ls'
end); simplify; equality. end); simplify; equality.
Defined. Defined.
End In_dec. End In_dec.
@ -391,8 +391,8 @@ Notation "[| x |]" := (Found _ x _).
Definition pred_strong7 : forall n : nat, {{m | n = S m}}. Definition pred_strong7 : forall n : nat, {{m | n = S m}}.
refine (fun n => refine (fun n =>
match n return {{m | n = S m}} with match n return {{m | n = S m}} with
| O => ?? | O => ??
| S n' => [|n'|] | S n' => [|n'|]
end); trivial. end); trivial.
Defined. Defined.
@ -422,8 +422,8 @@ Notation "[|| x ||]" := (inleft _ [x]).
Definition pred_strong8 : forall n : nat, {m : nat | n = S m} + {n = 0}. Definition pred_strong8 : forall n : nat, {m : nat | n = S m} + {n = 0}.
refine (fun n => refine (fun n =>
match n with match n with
| O => !! | O => !!
| S n' => [||n'||] | S n' => [||n'||]
end); trivial. end); trivial.
Defined. Defined.
@ -466,8 +466,8 @@ Defined.
* similarly straightforward version of this function. *) * similarly straightforward version of this function. *)
Notation "x <-- e1 ; e2" := (match e1 with Notation "x <-- e1 ; e2" := (match e1 with
| inright _ => !! | inright _ => !!
| inleft (exist x _) => e2 | inleft (exist x _) => e2
end) end)
(right associativity, at level 60). (right associativity, at level 60).
@ -541,20 +541,20 @@ Local Hint Constructors hasType : core.
Definition typeCheck : forall e : exp, {{t | hasType e t}}. Definition typeCheck : forall e : exp, {{t | hasType e t}}.
refine (fix F (e : exp) : {{t | hasType e t}} := refine (fix F (e : exp) : {{t | hasType e t}} :=
match e return {{t | hasType e t}} with match e return {{t | hasType e t}} with
| Nat _ => [|TNat|] | Nat _ => [|TNat|]
| Plus e1 e2 => | Plus e1 e2 =>
t1 <- F e1; t1 <- F e1;
t2 <- F e2; t2 <- F e2;
eq_type_dec t1 TNat;; eq_type_dec t1 TNat;;
eq_type_dec t2 TNat;; eq_type_dec t2 TNat;;
[|TNat|] [|TNat|]
| Bool _ => [|TBool|] | Bool _ => [|TBool|]
| And e1 e2 => | And e1 e2 =>
t1 <- F e1; t1 <- F e1;
t2 <- F e2; t2 <- F e2;
eq_type_dec t1 TBool;; eq_type_dec t1 TBool;;
eq_type_dec t2 TBool;; eq_type_dec t2 TBool;;
[|TBool|] [|TBool|]
end); subst; eauto. end); subst; eauto.
Defined. Defined.

View file

@ -16,8 +16,8 @@ Set Asymmetric Patterns.
Definition pred (n : nat) : nat := Definition pred (n : nat) : nat :=
match n with match n with
| O => O | O => O
| S n' => n' | S n' => n'
end. end.
Extraction pred. Extraction pred.
@ -124,8 +124,8 @@ Notation "[|| x ||]" := (inleft _ [x]).
(** * Monadic Notations *) (** * Monadic Notations *)
Notation "x <- e1 ; e2" := (match e1 with Notation "x <- e1 ; e2" := (match e1 with
| Unknown => ?? | Unknown => ??
| Found x _ => e2 | Found x _ => e2
end) end)
(right associativity, at level 60). (right associativity, at level 60).
@ -145,8 +145,8 @@ Admitted.
Notation "x <-- e1 ; e2" := (match e1 with Notation "x <-- e1 ; e2" := (match e1 with
| inright _ => !! | inright _ => !!
| inleft (exist x _) => e2 | inleft (exist x _) => e2
end) end)
(right associativity, at level 60). (right associativity, at level 60).

View file

@ -311,7 +311,7 @@ Module References.
end. end.
Qed. Qed.
(* Now, a series of lemmas essentially copied from original type-soundness (* Now, a series of lemmas essentially copied from the original type-soundness
* proof. *) * proof. *)
Lemma weakening_override : forall (G G' : fmap var type) x t, Lemma weakening_override : forall (G G' : fmap var type) x t,

View file

@ -3567,7 +3567,7 @@ We begin by copying over the two basic-step rules from last chapter, threading t
$$\infer{\smallstepo{(h, (\lambda x. \; e) \; v)}{(h, \subst{e}{x}{v})}}{} $$\infer{\smallstepo{(h, (\lambda x. \; e) \; v)}{(h, \subst{e}{x}{v})}}{}
\quad \infer{\smallstepo{(h, n + m)}{(h, n \textbf{+} m)}}{}$$ \quad \infer{\smallstepo{(h, n + m)}{(h, n \textbf{+} m)}}{}$$
To write out the rules that are specific to references, it's helpful to extend our language syntax with a form that will never appear in original programs, but which does show up at intermediate execution steps. To write out the rules that are specific to references, it's helpful to extend our language syntax with a form that will never appear in original programs but which does show up at intermediate execution steps.
In particular, let's add an expression form for \emph{locations}\index{locations}, the runtime values of references, and let's say that locations also count as values. In particular, let's add an expression form for \emph{locations}\index{locations}, the runtime values of references, and let's say that locations also count as values.
$$\begin{array}{rrcl} $$\begin{array}{rrcl}
\textrm{Locations} & \ell &\in& \mathbb N \\ \textrm{Locations} & \ell &\in& \mathbb N \\
@ -3600,7 +3600,7 @@ As a small exercise for the reader, it may be worth using this judgment to deriv
Even fixing the empty heap in the starting state, there is some nondeterminism in which final heap it returns: the possibilities are all the single-location heaps, mapping their single locations to value 1. Even fixing the empty heap in the starting state, there is some nondeterminism in which final heap it returns: the possibilities are all the single-location heaps, mapping their single locations to value 1.
It is natural to allow this nondeterminism in allocation, since typical memory allocators in real systems don't give promises about predictability in the addresses that they return. It is natural to allow this nondeterminism in allocation, since typical memory allocators in real systems don't give promises about predictability in the addresses that they return.
However, we will be able to prove that, for instance, any program returning a number \emph{gives the same answer, independently of nondeterministic choices made by the allocator}. However, we will be able to prove that, for instance, any program returning a number \emph{gives the same answer, independently of nondeterministic choices made by the allocator}.
That property is not true in programming languages like C\index{C programming language} that are not \emph{memory safe}\index{memory safety}, as they allow arithmetic and comparisons on pointers\index{pointers}, the closest C equivalent of our references. That property is not true in programming languages like C\index{C programming language} that are not \emph{memory-safe}\index{memory safety}, as they allow arithmetic and comparisons on pointers\index{pointers}, the closest C equivalent of our references.
\section{Type Soundness} \section{Type Soundness}