Giving up on exercises 6

amethyst
Shad Amethyst 10 months ago
parent 28c21e9e93
commit ad5f832f96

@ -13,6 +13,9 @@ Implicit Types
(A B : type) (A B : type)
. .
Ltac is_val_to_val v H :=
rewrite (is_val_spec _) in H; destruct H as [v H]; apply of_to_val in H; symmetry in H.
(** ** Exercise 5: Keep Rollin' *) (** ** Exercise 5: Keep Rollin' *)
(** This exercise is slightly tricky. (** This exercise is slightly tricky.
We strongly recommend you to first work on the other exercises. We strongly recommend you to first work on the other exercises.
@ -56,9 +59,6 @@ Section recursion_combinator.
tauto. tauto.
Qed. Qed.
Ltac is_val_to_val v H :=
rewrite (is_val_spec _) in H; destruct H as [v H]; apply of_to_val in H; symmetry in H.
(* (*
(Rec t) e = (λx. (unroll (rec' t)) (rec' t) x) e (Rec t) e = (λx. (unroll (rec' t)) (rec' t) x) e
~> (unroll (rec' t)) (rec' t) e ~> (unroll (rec' t)) (rec' t) e
@ -123,7 +123,6 @@ Do not attempt to understand what is going on in the proof, this clearly isn't m
reflexivity. reflexivity.
Qed. Qed.
Search (?A <[?f := ?x]> ?A).
Lemma rec_body_typing n Γ (A B: type) t : Lemma rec_body_typing n Γ (A B: type) t :
Γ !! x = None Γ !! x = None
Γ !! f = None Γ !! f = None
@ -157,9 +156,30 @@ Do not attempt to understand what is going on in the proof, this clearly isn't m
TY n; Γ (Rec t) : (A B). TY n; Γ (Rec t) : (A B).
Proof. Proof.
intros. intros.
econstructor.
(* TODO: exercise *) solve_typing.
Admitted. - unfold rec_body.
eapply typed_unroll'.
eapply (typed_weakening n _ Γ).
apply rec_body_typing; try assumption.
3: exact H3.
all: try eauto.
apply insert_subseteq; assumption.
asimpl.
reflexivity.
- asimpl.
solve_typing.
eapply (typed_weakening n _ Γ).
assumption.
etrans; [ apply insert_subseteq | apply insert_subseteq ].
assumption.
rewrite lookup_insert_ne.
assumption.
symmetry.
assumption.
reflexivity.
- assumption.
Qed.
End recursion_combinator. End recursion_combinator.
@ -194,8 +214,57 @@ Lemma fix_red (f x: string) (e e': expr):
f x f x
rtc contextual_step ((fix: f x := e) e')%V (sub x e' (sub f (fix: f x := e)%V e)). rtc contextual_step ((fix: f x := e) e')%V (sub x e' (sub f (fix: f x := e)%V e)).
Proof. Proof.
(* TODO: exercise *) intros.
Admitted. rewrite Fix_eq.
unfold Fix.
eapply rtc_l.
apply base_contextual_step.
eapply BetaS; [ assumption | reflexivity ].
simpl.
repeat rewrite bnamed_eq.
repeat (rewrite (decide_False (P := x = f)); last auto).
repeat (rewrite (decide_True (P := x = x)); last reflexivity).
is_val_to_val v H1; subst.
eapply rtc_l.
eapply (Ectx_step _ _ [
AppLCtx (RollV (λ: (BNamed f),
App (λ: (BNamed f) (BNamed x), e) (λ: (BNamed x),
App (App (unroll Var f) (Var f)) (Var x)
))
);
AppLCtx v
]); [reflexivity | reflexivity | econstructor; eauto].
eapply rtc_l.
eapply (Ectx_step _ _ [
AppLCtx v
]); [reflexivity | reflexivity | econstructor; eauto].
simpl.
repeat rewrite bnamed_eq.
repeat (rewrite (decide_False (P := f = x)); last auto).
repeat (rewrite (decide_True (P := f = f)); last reflexivity).
eapply rtc_l.
eapply (Ectx_step _ _ [
AppLCtx v
]); [reflexivity | reflexivity | econstructor; eauto].
simpl.
repeat rewrite bnamed_eq.
repeat (rewrite (decide_False (P := f = x)); last auto).
eapply rtc_l.
eapply base_contextual_step.
econstructor; [eauto | reflexivity].
simpl.
reflexivity.
Qed.
Lemma fix_typing n Γ (f x: string) (A B: type) (e: expr): Lemma fix_typing n Γ (f x: string) (A B: type) (e: expr):
@ -205,7 +274,7 @@ Lemma fix_typing n Γ (f x: string) (A B: type) (e: expr):
TY n; <[x := A]> (<[f := (A B)%ty]> Γ) e : B TY n; <[x := A]> (<[f := (A B)%ty]> Γ) e : B
TY n; Γ (fix: f x := e) : (A B). TY n; Γ (fix: f x := e) : (A B).
Proof. Proof.
(* TODO: exercise *) (* Weee, let's do everything from scratch again! *)
Admitted. Admitted.
@ -289,22 +358,299 @@ Definition list_t (A : type) : type :=
Definition mylist_impl : val := PackV ( Definition mylist_impl : val := PackV (
(#0, (LitV LitUnit)), (#0, (LitV LitUnit)),
(λ: "a" "l", (#1 + (Fst "l"), ("a", (Snd "")))), (λ: "a" "l", (#1 + (Fst "l"), ("a", (Snd "l")))),
(Λ, λ: "l" "zero" "cb", (Λ, λ: "l" "zero" "cb",
(* (fix "rec" "l" := *) (* (fix "rec" "l" := *)
if: (Fst "l") = #0 then if: (Fst "l") = #0 then
"zero" "zero"
else else
"cb" (Fst (Snd "l")) (Snd (Snd "l")) "cb" (Fst (Snd "l")) ((Fst "l") - #1, Snd (Snd "l"))
(* ) "l1" *) (* ) "l1" *)
) )
). ).
Inductive list_sem_type {A : type} : nat val Prop :=
| empty k : list_sem_type k (PairV (LitV (LitInt 0)) (LitV LitUnit))
| push k (n : nat) rest head :
is_closed [] (of_val head)
list_sem_type k (PairV (LitV (LitInt n)) rest)
𝒱 A δ_any k head
list_sem_type k (PairV (LitV (LitInt (1 + n))) (PairV head rest)).
(* I'm not gonna reiterate my beliefs on what constitute bad pedagogy. *)
Lemma mylist_impl_sem_typed A : Lemma mylist_impl_sem_typed A :
type_wf 0 A type_wf 0 A
k, 𝒱 (list_t A) δ_any k mylist_impl. k, 𝒱 (list_t A) δ_any k mylist_impl.
Proof. Proof.
(* TODO: exercise *) intros A_wf k.
Admitted. unfold list_t.
simp type_interp.
eexists; split; [ reflexivity | ].
pose_sem_type (list_sem_type (A := A)) as τ.
{
intros l v sem.
induction sem.
eauto.
simpl.
simpl in IHsem.
apply andb_True; split; assumption.
}
{
intros l l' v sem l'_le_l.
induction sem.
econstructor.
econstructor.
assumption.
exact (IHsem l'_le_l).
eapply val_rel_mono.
2: exact H0.
assumption.
}
exists τ.
simp type_interp; eexists; eexists; split; last split; first reflexivity.
simp type_interp; eexists; eexists; split; last split; first reflexivity.
- simp type_interp; simpl.
constructor.
- simp type_interp; eexists; eexists; split; last split; [reflexivity | eauto | ].
intros arg j j_le_k Hsem.
simpl.
eapply sem_val_expr_rel'; eauto.
simp type_interp; eexists; eexists; split; last split; [reflexivity | simpl; apply andb_True; split; eauto | ].
apply is_closed_weaken_nil.
apply val_rel_is_closed in Hsem.
assumption.
intros arg i i_le_k Hsem.
simpl.
rewrite subst_is_closed_nil.
2: apply val_rel_is_closed in Hsem; assumption.
simp type_interp in Hsem; simpl in Hsem.
induction Hsem.
+ eapply expr_det_step_closure.
apply det_step_pair_r.
apply det_step_pair_r.
apply det_step_snd; eauto.
eapply expr_det_step_closure.
apply det_step_pair_l; eauto.
apply det_step_binop_r.
apply det_step_fst; eauto.
eapply expr_det_step_closure.
apply det_step_pair_l; eauto.
eapply sem_val_expr_rel'; eauto.
simpl.
simplify_val.
eauto.
simp type_interp; simpl.
have h : 0%Z = Z.of_nat 0.
reflexivity.
rewrite h.
econstructor.
apply val_rel_is_closed in Hsem; assumption.
econstructor.
rewrite <-sem_val_rel_cons in Hsem.
eapply val_rel_mono.
2: exact Hsem.
lia.
+ eapply expr_det_step_closure.
apply det_step_pair_r.
apply det_step_pair_r.
apply det_step_snd; eauto.
fold of_val.
eapply expr_det_step_closure.
apply det_step_pair_l; eauto.
apply det_step_binop_r.
apply det_step_fst; eauto.
eapply expr_det_step_closure.
apply det_step_pair_l; eauto.
eapply sem_val_expr_rel'; eauto.
simpl.
simplify_val.
eauto.
simp type_interp; simpl.
have h : (1 + Z.of_nat n)%Z = Z.of_nat (1 + n).
lia.
rewrite h.
econstructor.
apply val_rel_is_closed in Hsem; assumption.
rewrite <-h.
econstructor.
apply val_rel_is_closed in Hsem; assumption.
eapply __p2.
exact Hsem.
lia.
(* Little annoying buggers: *)
eapply val_rel_mono.
2: exact H0.
lia.
rewrite <-sem_val_rel_cons in Hsem.
eapply val_rel_mono.
2: exact Hsem.
lia.
- simp type_interp.
eexists; split; last split; [ reflexivity | eauto | ].
intro τ.
eapply sem_val_expr_rel'; eauto.
simp type_interp; eexists; eexists; split; last split; [reflexivity | eauto | ].
intros arg j j_le_k Hsem.
have Hcl := (val_rel_is_closed _ _ _ _ Hsem).
simpl.
eapply sem_val_expr_rel'; eauto.
simp type_interp; eexists; eexists; split; last split; [reflexivity | eauto | ].
simpl.
repeat (apply andb_True; split); eauto; apply is_closed_weaken_nil; assumption.
intros arg i i_le_j Hsem.
simpl.
have Hcl := (val_rel_is_closed _ _ _ _ Hsem).
eapply sem_val_expr_rel'; eauto.
simp type_interp; eexists; eexists; split; last split; [reflexivity | eauto | ].
(* Big closedness detour *)
simpl.
repeat (apply andb_True; split); eauto; apply is_closed_weaken_nil;
try assumption;
rewrite subst_is_closed_nil; assumption.
intros arg l l_le_i Hsem.
simpl.
have Hcl := (val_rel_is_closed _ _ _ _ Hsem).
rewrite subst_is_closed_nil.
2: rewrite subst_is_closed_nil; assumption.
rewrite subst_is_closed_nil.
2: assumption.
simp type_interp in Hsem; simpl in Hsem.
destruct Hsem.
+ eapply expr_det_step_closure.
apply det_step_if.
apply det_step_binop_l; eauto.
apply det_step_fst; eauto.
eapply expr_det_step_closure.
apply det_step_if.
eauto.
simpl.
eapply expr_det_step_closure.
apply det_step_if_true.
rewrite subst_is_closed_nil.
2: assumption.
eapply sem_val_expr_rel'; eauto.
rewrite to_of_val; reflexivity.
simp type_interp; simpl.
eapply sem_type_mono.
simp type_interp in Hsem; simpl in Hsem.
lia.
+ eapply expr_det_step_closure.
apply det_step_if.
apply det_step_binop_l; eauto.
apply det_step_fst; eauto.
eapply expr_det_step_closure.
apply det_step_if.
eauto.
simpl.
have h : bool_decide ((1 + n)%Z = 0%Z) = false.
apply bool_decide_false.
lia.
rewrite h.
eapply expr_det_step_closure.
apply det_step_if_false.
eapply expr_det_step_closure.
apply det_step_app_r.
apply det_step_pair_r.
apply det_step_snd_lift.
apply det_step_snd; eauto.
eapply expr_det_step_closure.
apply det_step_app_r.
apply det_step_pair_r.
apply det_step_snd; eauto.
eapply expr_det_step_closure.
apply det_step_app_r.
apply det_step_pair_l; eauto.
eapply expr_det_step_closure.
apply det_step_app_r.
apply det_step_pair_l; eauto.
have h : (1 + Z.of_nat n - 1)%Z = Z.of_nat n.
lia.
rewrite h.
eapply expr_det_step_closure.
apply det_step_app_l; eauto.
apply det_step_app_r.
apply det_step_fst_lift.
apply det_step_snd; eauto.
eapply expr_det_step_closure.
apply det_step_app_l; eauto.
simp type_interp in Hsem.
destruct Hsem as (x & e & -> & Hcl & Hsem).
eapply expr_det_step_closure.
apply det_step_app_l; eauto.
apply det_step_beta; eauto.
have h : 𝒱 A.[ren (+2)] (τ .: τ .: δ_any) l head.
{
have h : A.[ren (+2)] = A.[ren (+1)].[ren (+1)].
asimpl; reflexivity.
rewrite h.
rewrite <-sem_val_rel_cons.
rewrite <-sem_val_rel_cons.
eapply val_rel_mono.
2: exact H0.
lia.
}
have h : l l.
lia.
specialize (Hsem _ _ h h).
eapply (bind [AppLCtx (LitV (Z.of_nat n), rest)]).
eapply expr_rel_mono.
2: exact Hsem.
lia.
intros j v j_le_funny Hsem.
simpl.
simp type_interp in Hsem.
destruct Hsem as (x' & e' & -> & Hcl & Hsem).
eapply expr_det_step_closure.
apply det_step_beta; eauto.
have h : (#(LitInt (Z.of_nat n)), of_val rest)%E = of_val (PairV (LitV n) rest).
reflexivity.
rewrite h.
apply Hsem.
lia.
simp type_interp; simpl.
eapply __p2.
exact Hsem.
lia.
Qed.

@ -3,12 +3,56 @@ From iris Require Import prelude.
From semantics.lib Require Import maps. From semantics.lib Require Import maps.
From semantics.ts.systemf_mu Require Import lang notation. From semantics.ts.systemf_mu Require Import lang notation.
Print base_reducible_contextual_step_ectx.
Lemma contextual_ectx_step_case' {a e e'} :
contextual_step (fill_item a e) e'
( e'', e' = fill_item a e'' contextual_step e e'') is_val e.
Proof.
intro H.
admit.
Admitted.
Lemma contextual_ectx_step_case K e e' : Lemma contextual_ectx_step_case K e e' :
contextual_step (fill K e) e' contextual_step (fill K e) e'
( e'', e' = fill K e'' contextual_step e e'') is_val e. ( e'', e' = fill K e'' contextual_step e e'') is_val e.
Proof. Proof.
(* TODO: exercise *) revert e e'.
Admitted. induction K.
- intros e e' H.
simpl.
left. eexists. split; [reflexivity | exact H].
- intros e e' H.
simpl in H.
specialize (IHK _ _ H).
induction IHK.
+ destruct H0 as (e'' & -> & Hctx).
induction (contextual_ectx_step_case' Hctx).
destruct H0 as (e''' & -> & Hctx').
left.
eexists.
split.
reflexivity.
exact Hctx'.
right.
exact H0.
+ induction a; simpl in H0; try (exfalso; exact H0); right; tauto.
(* Restart.
revert e e'.
induction K using rev_ind.
- intros e e' H.
simpl.
left. eexists. split; [reflexivity | exact H].
- intros e e' H.
have h : K ++ [x] = comp_ectx [x] K.
reflexivity.
rewrite h.
rewrite h in H. *)
Qed.
(** ** Deterministic reduction *) (** ** Deterministic reduction *)
@ -162,7 +206,17 @@ Lemma red_nsteps_fill K k e e' :
red_nsteps j e e'' red_nsteps j e e''
red_nsteps (k - j) (fill K e'') e'. red_nsteps (k - j) (fill K e'') e'.
Proof. Proof.
(* TODO: exercise *) (* revert e e'.
induction K.
{
admit.
(* trivial *)
}
intros e e' H.
induction a.
specialize (IHK (fill [a] e)).
TODO: exercise *)
Admitted. Admitted.

Loading…
Cancel
Save