📝 Add a few comments for paper proofs

amethyst
Shad Amethyst 5 months ago
parent 82042039b9
commit 02b9ce9abb

@ -338,6 +338,14 @@ Proof.
by apply of_to_val.
Qed.
(*
By induction on the transitive closure,
we need to show that if `e = v`, then `v | v` (trivial), and if `e >- e'` and `e' | v`,
then `e | v`.
To prove this, we do an induction on the rule used for the single small step.
Each of the cases are individually quite easy: one just needs to provide which big step rule should be used,
and the induction hypothesis can be plugged into the conditions of the big step.
*)
Lemma steps_big_step e (v: val):
rtc step e v big_step e v.
Proof.
@ -345,8 +353,7 @@ Proof.
remember (of_val v) eqn:H_val in H.
induction H; subst.
apply big_step_vals.
remember ((IHrtc (eq_refl (of_val v)))) as IH.
clear HeqIH.
have IH := ((IHrtc (eq_refl (of_val v)))).
clear IHrtc.
(* This is ridiculous. *)
revert H0.

@ -107,7 +107,10 @@ Qed.
(** Now the other direction. *)
(* You may find it helpful to introduce intermediate lemmas. *)
(*
To prove this, we first need to invert the contextual step.
By induction on the context, we can reconstruct the small-step rule used.
*)
Lemma contextual_step_step e1 e2:
contextual_step e1 e2 step e1 e2.
Proof.
@ -320,6 +323,12 @@ Inductive src_typed : typing_context → src_expr → type → Prop :=
where "Γ '⊢S' E : A" := (src_typed Γ E%E A%ty) : FType_scope.
#[export] Hint Constructors src_typed : core.
(*
This is trivial enough for `eauto` to be able to solve all the cases for us.
Doing an induction on the source typing rule used, we can infer the structure of `E`,
the hypothesis needed, and then unfold `erase` to reveal in each case an equivalent
syntactical typing rule.
*)
Lemma type_erasure_correctness Γ E A:
(Γ S E : A)%ty (Γ erase E : A)%ty.
Proof.
@ -329,6 +338,14 @@ Qed.
(** ** Exercise 4: Unique Typing *)
(*
This is a bit of a boring lemma.
We do induction on the first typing and inversion on the second typing.
- for integers, this is trivial
- for variables, since the Γ is shared, by injection we can see that A = B since Γ[x] = Some A = Some B
- for functions of type `C -> A2` =? `C -> B2`, we need to use the induction hypothesis, which asks us to prove that `Γ |- body : B2` to get that `A2 = B2`
- for function applications, the induction hypothesis tells us that if `Γ |- e_fn : B1 -> B2`, then `A1 -> A2` = `B1 -> B2`, allowing us to show that `A2 = B2`.
*)
Lemma src_typing_unique Γ E A B:
(Γ S E : A)%ty (Γ S E : B)%ty A = B.
Proof.

@ -287,6 +287,18 @@ Proof.
Qed.
(*
First, unfold the definition of `Γ |=` everywhere.
The closedness of our term can easily be solved by `naive_solver`.
We can then use the same θ and H_ctx in all hypotheses,
and unroll all definitions of `E[A]`, yielding `v_lhs` and `v_rhs` to which resp. `e_lhs` and `e_rhs` step.
Then, pick `(v_lhs, v_rhs)` as the value this should step to.
We then just need to show that `(e_lhs, e_rhs)` step towards `(v_lhs, v_rhs)`,
and that `(v_lhs, v_rhs) V[A × B]`.
*)
Lemma compat_pair Γ e_lhs e_rhs A B:
Γ e_lhs: A
Γ e_rhs: B
@ -323,6 +335,10 @@ Proof.
eauto. (* Trivial from then on *)
Qed.
(*
Proving this is very similar to the above, except that we need to work on two goals at once,
and we apply the definition of `V[A × B]` to get the necessary stuff to prove both statements.
*)
Lemma compat_fst_snd {Γ e_pair A B}:
Γ e_pair: A × B
Γ (Fst e_pair): A Γ (Snd e_pair): B.

@ -342,6 +342,14 @@ Proof.
intros H1 H2; by eapply H2.
Qed.
(*
By induction on the base step,
we can invert the the typing rules used for `|- e : A` and its hypotheses.
- in the `base_step_case_l` case, we can also invert the `InjL` typing rule
that appears after the first inversion. We then get that `e1 : A' -> A` and `e2 : B' -> A`,
and `v : A'`.
We now just reconstruct the typing of `e'` by using the application constructor.
*)
Lemma typed_preservation_base_step e e' A:
e : A
base_step e e'
@ -366,7 +374,7 @@ Proof.
assumption.
- eapply case_inversion in Hty as (A' & B' & Hty_inj & Hty_lhs & Hty_rhs).
eapply injl_inversion in Hty_inj as (A'' & B'' & Heq & Hty_inj).
injection Heq as -> ->.
injection Heq as <- <-.
eapply typed_app.
exact Hty_lhs.
exact Hty_inj.

@ -171,5 +171,6 @@ Qed.
In the lecture, we also saw `A[id] = (. #0 -> #1)[id] = . (#0 -> #1)[id] = . (#0 -> (id(0))) = . #0 -> #0`
To fix this, we would have to carefully tweak each value in our substitution map to increment all the free type variables
To fix this, we would have to carefully tweak each value in our substitution map to increment all the free type variables.
This is what `ren` does.
*)

@ -531,6 +531,18 @@ Section boring_lemmas.
End boring_lemmas.
(** Bind lemma *)
(*
From the definition of `E[A]` and `E[B]`,
we get that ` e', n < k, K[e] \^n e' -> (k - n, e') V[B]`
- We know that ` j, e \^j e''` and `K[e''] \^(n-j) e'` using a lemma.
- From the expansion of `E[A]`, we can plug `e''` and `j` into `e E[A]`,
yielding `(k-j, e'') V[A]`
- We can use the second hypothesis with `k-j` and the value associated to `e''` as input, yielding `(k-j, K[e'']) E[B]`
- Plugging in the expansion of `E[B]` into `K[e''] \^(n-j) e'`,
we get that `(k-j-(n-j), e') V[B]`
- `k-j-(n-j) = k-n`, allowing us to prove our goal
*)
Lemma bind K e k δ A B :
A δ k e
( j v, j k 𝒱 A δ j v B δ j (fill K (of_val v)))

Loading…
Cancel
Save