Skip to content
Snippets Groups Projects
Commit 03200aa5 authored by Ralf Jung's avatar Ralf Jung
Browse files

bring back copy and send typing contexts

For now though, I think I'll leave the other version on paper... there, not being syntax-directed is way less of a problem
parent f9bcc1f9
No related branches found
No related tags found
No related merge requests found
Pipeline #
...@@ -136,18 +136,16 @@ Section typing. ...@@ -136,18 +136,16 @@ Section typing.
+ rewrite /tctx_interp big_sepL_zip_with. done. + rewrite /tctx_interp big_sepL_zip_with. done.
Qed. Qed.
Lemma type_fn {A m} E L E' fb kb (argsb : list binder) e Lemma type_fn {A} E L E' fb kb (argsb : list binder) e
(tys : A vec type (length argsb)) ty (tys : A vec type (length argsb)) ty
(cps : vec path m) (ctyl : vec type m) `{!LstCopy ctyl, !LstSend ctyl} : T `{!CopyC T, !SendC T} :
Closed (fb :b: kb :b: argsb +b+ []) e Closed (fb :b: kb :b: argsb +b+ []) e
( x f k (args : vec val _), ( x f k (args : vec val _),
typed_body (E' x) [] [CCtx_iscont k [] 1 (λ v, [TCtx_hasty (v!!!0) (ty x)])] typed_body (E' x) [] [CCtx_iscont k [] 1 (λ v, [TCtx_hasty (v!!!0) (ty x)])]
(TCtx_hasty f (fn E' tys ty) :: (TCtx_hasty f (fn E' tys ty) ::
zip_with (TCtx_hasty of_val) args (tys x) ++ zip_with (TCtx_hasty of_val) args (tys x) ++ T)
zip_with TCtx_hasty cps ctyl)
(subst' fb f $ subst_v (kb :: argsb) (vmap of_val $ k ::: args) e)) (subst' fb f $ subst_v (kb :: argsb) (vmap of_val $ k ::: args) e))
typed_instruction_ty E L (zip_with TCtx_hasty cps ctyl) typed_instruction_ty E L T (Rec fb (kb :: argsb) e) (fn E' tys ty).
(Rec fb (kb :: argsb) e) (fn E' tys ty).
Proof. Proof.
iIntros (Hc Hbody tid qE) "#HEAP #LFT $ $ $ #HT". iApply wp_value. iIntros (Hc Hbody tid qE) "#HEAP #LFT $ $ $ #HT". iApply wp_value.
{ simpl. rewrite decide_left. done. } { simpl. rewrite decide_left. done. }
...@@ -159,6 +157,6 @@ Section typing. ...@@ -159,6 +157,6 @@ Section typing.
{ by rewrite -vec_to_list_cons -vec_to_list_map -subst_v_eq. } { by rewrite -vec_to_list_cons -vec_to_list_map -subst_v_eq. }
iApply (Hbody with "* HEAP LFT Htl HE HL HC"). iApply (Hbody with "* HEAP LFT Htl HE HL HC").
rewrite tctx_interp_cons tctx_interp_app. iFrame "HT' IH". rewrite tctx_interp_cons tctx_interp_app. iFrame "HT' IH".
iApply tctx_send. by iNext. iApply sendc_change_tid. by iNext.
Qed. Qed.
End typing. End typing.
...@@ -115,28 +115,37 @@ Section type_context. ...@@ -115,28 +115,37 @@ Section type_context.
tctx_interp tid [x] tctx_elt_interp tid x. tctx_interp tid [x] tctx_elt_interp tid x.
Proof. rewrite tctx_interp_cons tctx_interp_nil right_id //. Qed. Proof. rewrite tctx_interp_cons tctx_interp_nil right_id //. Qed.
Global Instance tctx_persistent cps ctyl tid : (** Copy typing contexts *)
LstCopy ctyl PersistentP (tctx_interp tid $ zip_with TCtx_hasty cps ctyl). Class CopyC (T : tctx) :=
copyc_persistent tid : PersistentP (tctx_interp tid T).
Global Existing Instances copyc_persistent.
Global Instance tctx_nil_copy : CopyC [].
Proof. rewrite /CopyC. apply _. Qed.
Global Instance tctx_ty_copy T p ty :
CopyC T Copy ty CopyC (TCtx_hasty p ty :: T).
Proof. Proof.
intros Hcopy. revert cps; induction Hcopy; intros cps;
first by (rewrite zip_with_nil_r tctx_interp_nil; apply _).
destruct cps; first by (rewrite tctx_interp_nil; apply _). simpl.
(* TODO RJ: Should we have instances that PersistentP respects equiv? *) (* TODO RJ: Should we have instances that PersistentP respects equiv? *)
rewrite /PersistentP tctx_interp_cons. intros ???. rewrite /PersistentP tctx_interp_cons.
apply uPred.sep_persistent; by apply _. apply uPred.sep_persistent; by apply _.
Qed. Qed.
Lemma tctx_send cps ctyl tid1 tid2 {Hcopy : LstSend ctyl} : (** Send typing contexts *)
tctx_interp tid1 $ zip_with TCtx_hasty cps ctyl -∗ Class SendC (T : tctx) :=
tctx_interp tid2 $ zip_with TCtx_hasty cps ctyl. sendc_change_tid tid1 tid2 : tctx_interp tid1 T -∗ tctx_interp tid2 T.
Global Instance tctx_nil_send : SendC [].
Proof. done. Qed.
Global Instance tctx_ty_send T p ty :
SendC T Send ty SendC (TCtx_hasty p ty :: T).
Proof. Proof.
revert cps; induction Hcopy; intros cps; iIntros (HT Hty ??). rewrite !tctx_interp_cons.
first by rewrite zip_with_nil_r !tctx_interp_nil. iIntros "[Hty HT]". iSplitR "HT".
destruct cps; first by rewrite !tctx_interp_nil. simpl.
rewrite !tctx_interp_cons. iIntros "[Hty HT]". iSplitR "HT".
- iDestruct "Hty" as (?) "[% Hty]". iExists _. iSplit; first done. - iDestruct "Hty" as (?) "[% Hty]". iExists _. iSplit; first done.
by iApply send_change_tid. by iApply Hty.
- by iApply IHHcopy. - by iApply HT.
Qed. Qed.
(** Type context inclusion *) (** Type context inclusion *)
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment