LogrelCC.coop_logrel.translation
From LogrelCC.cooperative.cooplang Require Export lang typing.
From LogrelCC.F_mu_ref_cc Require Export lang typing.
Import coop_notations.
Import coop_type_notations.
Fixpoint Translate_coop (e : coopexpr) (boundvars : nat) {struct e} : expr :=
match e with
| coop_Var n => Var n
| coop_Rec e => Rec (Translate_coop e (2 + boundvars))
| coop_Lam e => Lam (Translate_coop e (1 + boundvars))
| coop_LetIn e e' =>
LetIn (Translate_coop e boundvars) (Translate_coop e' (1 + boundvars))
| coop_Seq e e' =>
Seq (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_App e e' =>
App (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_Unit => Unit
| coop_Nat n => Nat n
| coop_Bool b => Bool b
| coop_BinOp o e e' =>
BinOp o (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_If e1 e2 e3 =>
If (Translate_coop e1 boundvars)
(Translate_coop e2 boundvars) (Translate_coop e3 boundvars)
| coop_Pair e1 e2 =>
Pair (Translate_coop e1 boundvars) (Translate_coop e2 boundvars)
| coop_Fst e => Fst (Translate_coop e boundvars)
| coop_Snd e => Snd (Translate_coop e boundvars)
| coop_InjL e => InjL (Translate_coop e boundvars)
| coop_InjR e => InjR (Translate_coop e boundvars)
| coop_Case e1 e2 e3 =>
Case (Translate_coop e1 boundvars)
(Translate_coop e2 (1 + boundvars)) (Translate_coop e3 (1 + boundvars))
| coop_Fold e => Fold (Translate_coop e boundvars)
| coop_Unfold e => Unfold (Translate_coop e boundvars)
| coop_TLam e => TLam (Translate_coop e boundvars)
| coop_TApp e => TApp (Translate_coop e boundvars)
| coop_CoFork e =>
(App (Var boundvars)
(Lam (Translate_coop e boundvars).[ren (+1)]))
| coop_CoYield => App (Var (1 + boundvars)) Unit
| coop_Loc l => Loc l
| coop_Alloc e => Alloc (Translate_coop e boundvars)
| coop_Load e => Load (Translate_coop e boundvars)
| coop_Store e e' =>
Store (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_Cont K => Cont (map (λ Ki, Translate_coop_ectxi Ki boundvars) K)
| coop_Callcc e => Callcc (Translate_coop e (1 + boundvars))
| coop_Throw e e' =>
Throw (Translate_coop e boundvars) (Translate_coop e' boundvars)
end
with Translate_coop_ectxi K (boundvars : nat) {struct K} : ectx_item :=
match K with
| coop_AppLCtx e => AppLCtx (Translate_coop e boundvars)
| coop_AppRCtx v => AppRCtx (Translate_coop_val v boundvars)
| coop_LetInCtx e => LetInCtx (Translate_coop e boundvars)
| coop_SeqCtx e => SeqCtx (Translate_coop e boundvars)
| coop_TAppCtx => TAppCtx
| coop_PairLCtx e => PairLCtx (Translate_coop e boundvars)
| coop_PairRCtx v => PairRCtx (Translate_coop_val v boundvars)
| coop_BinOpLCtx op e =>
BinOpLCtx op (Translate_coop e boundvars)
| coop_BinOpRCtx op v =>
BinOpRCtx op (Translate_coop_val v boundvars)
| coop_FstCtx => FstCtx
| coop_SndCtx => SndCtx
| coop_InjLCtx => InjLCtx
| coop_InjRCtx => InjRCtx
| coop_CaseCtx e2 e3 =>
CaseCtx (Translate_coop e2 (1 + boundvars))
(Translate_coop e3 (1 + boundvars))
| coop_IfCtx e2 e3 =>
IfCtx (Translate_coop e2 boundvars)
(Translate_coop e3 boundvars)
| coop_FoldCtx => FoldCtx
| coop_UnfoldCtx => UnfoldCtx
| coop_AllocCtx => AllocCtx
| coop_LoadCtx => LoadCtx
| coop_StoreLCtx e => StoreLCtx (Translate_coop e boundvars)
| coop_StoreRCtx v => StoreRCtx (Translate_coop_val v boundvars)
| coop_ThrowLCtx e => ThrowLCtx (Translate_coop e boundvars)
| coop_ThrowRCtx v => ThrowRCtx (Translate_coop_val v boundvars)
end
with Translate_coop_val (v : coopval) (boundvars : nat) {struct v} : val :=
match v with
| coop_RecV e => RecV (Translate_coop e (2 + boundvars))
| coop_LamV e => LamV (Translate_coop e (1 + boundvars))
| coop_TLamV e => TLamV (Translate_coop e boundvars)
| coop_UnitV => UnitV
| coop_NatV n => NatV n
| coop_BoolV b => BoolV b
| coop_PairV v1 v2 =>
PairV (Translate_coop_val v1 boundvars) (Translate_coop_val v2 boundvars)
| coop_InjLV v => InjLV (Translate_coop_val v boundvars)
| coop_InjRV v => InjRV (Translate_coop_val v boundvars)
| coop_FoldV v => FoldV (Translate_coop_val v boundvars)
| coop_LocV l => LocV l
| coop_ContV K => ContV (map (λ Ki, Translate_coop_ectxi Ki boundvars) K)
end.
Lemma Translate_coop_typed Γ e τ :
coop_typed Γ e τ →
typed
(Γ ++ [(TArrow (TArrow TUnit TUnit) TUnit); (TArrow TUnit TUnit)])
(Translate_coop e (length Γ)) τ.
Proof.
induction 1; try by simpl; econstructor; eauto.
- constructor.
rewrite lookup_app_l; eauto using lookup_lt_is_Some_1.
- simpl. econstructor; eauto.
rewrite fmap_app; simpl.
rewrite fmap_length in IHtyped.
auto.
- simpl. repeat econstructor.
+ rewrite lookup_app_r; eauto.
replace ((length Γ) - length Γ) with 0 by lia.
eauto.
+ by apply (context_weakening [_]).
- simpl. repeat econstructor.
rewrite lookup_app_r; eauto.
replace (S (length Γ) - length Γ) with 1 by lia.
eauto.
Qed.
Lemma typed_Translate_n_closed Γ τ e:
coop_typed Γ e τ →
(∀ f, (Translate_coop e (length Γ)).[upn (length (Γ ++
[(TArrow (TArrow TUnit TUnit) TUnit);
(TArrow TUnit TUnit)])) f]
= Translate_coop e (length Γ)).
Proof. intros Htp. by eapply typed_n_closed, Translate_coop_typed. Qed.
From LogrelCC.F_mu_ref_cc Require Export lang typing.
Import coop_notations.
Import coop_type_notations.
Fixpoint Translate_coop (e : coopexpr) (boundvars : nat) {struct e} : expr :=
match e with
| coop_Var n => Var n
| coop_Rec e => Rec (Translate_coop e (2 + boundvars))
| coop_Lam e => Lam (Translate_coop e (1 + boundvars))
| coop_LetIn e e' =>
LetIn (Translate_coop e boundvars) (Translate_coop e' (1 + boundvars))
| coop_Seq e e' =>
Seq (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_App e e' =>
App (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_Unit => Unit
| coop_Nat n => Nat n
| coop_Bool b => Bool b
| coop_BinOp o e e' =>
BinOp o (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_If e1 e2 e3 =>
If (Translate_coop e1 boundvars)
(Translate_coop e2 boundvars) (Translate_coop e3 boundvars)
| coop_Pair e1 e2 =>
Pair (Translate_coop e1 boundvars) (Translate_coop e2 boundvars)
| coop_Fst e => Fst (Translate_coop e boundvars)
| coop_Snd e => Snd (Translate_coop e boundvars)
| coop_InjL e => InjL (Translate_coop e boundvars)
| coop_InjR e => InjR (Translate_coop e boundvars)
| coop_Case e1 e2 e3 =>
Case (Translate_coop e1 boundvars)
(Translate_coop e2 (1 + boundvars)) (Translate_coop e3 (1 + boundvars))
| coop_Fold e => Fold (Translate_coop e boundvars)
| coop_Unfold e => Unfold (Translate_coop e boundvars)
| coop_TLam e => TLam (Translate_coop e boundvars)
| coop_TApp e => TApp (Translate_coop e boundvars)
| coop_CoFork e =>
(App (Var boundvars)
(Lam (Translate_coop e boundvars).[ren (+1)]))
| coop_CoYield => App (Var (1 + boundvars)) Unit
| coop_Loc l => Loc l
| coop_Alloc e => Alloc (Translate_coop e boundvars)
| coop_Load e => Load (Translate_coop e boundvars)
| coop_Store e e' =>
Store (Translate_coop e boundvars) (Translate_coop e' boundvars)
| coop_Cont K => Cont (map (λ Ki, Translate_coop_ectxi Ki boundvars) K)
| coop_Callcc e => Callcc (Translate_coop e (1 + boundvars))
| coop_Throw e e' =>
Throw (Translate_coop e boundvars) (Translate_coop e' boundvars)
end
with Translate_coop_ectxi K (boundvars : nat) {struct K} : ectx_item :=
match K with
| coop_AppLCtx e => AppLCtx (Translate_coop e boundvars)
| coop_AppRCtx v => AppRCtx (Translate_coop_val v boundvars)
| coop_LetInCtx e => LetInCtx (Translate_coop e boundvars)
| coop_SeqCtx e => SeqCtx (Translate_coop e boundvars)
| coop_TAppCtx => TAppCtx
| coop_PairLCtx e => PairLCtx (Translate_coop e boundvars)
| coop_PairRCtx v => PairRCtx (Translate_coop_val v boundvars)
| coop_BinOpLCtx op e =>
BinOpLCtx op (Translate_coop e boundvars)
| coop_BinOpRCtx op v =>
BinOpRCtx op (Translate_coop_val v boundvars)
| coop_FstCtx => FstCtx
| coop_SndCtx => SndCtx
| coop_InjLCtx => InjLCtx
| coop_InjRCtx => InjRCtx
| coop_CaseCtx e2 e3 =>
CaseCtx (Translate_coop e2 (1 + boundvars))
(Translate_coop e3 (1 + boundvars))
| coop_IfCtx e2 e3 =>
IfCtx (Translate_coop e2 boundvars)
(Translate_coop e3 boundvars)
| coop_FoldCtx => FoldCtx
| coop_UnfoldCtx => UnfoldCtx
| coop_AllocCtx => AllocCtx
| coop_LoadCtx => LoadCtx
| coop_StoreLCtx e => StoreLCtx (Translate_coop e boundvars)
| coop_StoreRCtx v => StoreRCtx (Translate_coop_val v boundvars)
| coop_ThrowLCtx e => ThrowLCtx (Translate_coop e boundvars)
| coop_ThrowRCtx v => ThrowRCtx (Translate_coop_val v boundvars)
end
with Translate_coop_val (v : coopval) (boundvars : nat) {struct v} : val :=
match v with
| coop_RecV e => RecV (Translate_coop e (2 + boundvars))
| coop_LamV e => LamV (Translate_coop e (1 + boundvars))
| coop_TLamV e => TLamV (Translate_coop e boundvars)
| coop_UnitV => UnitV
| coop_NatV n => NatV n
| coop_BoolV b => BoolV b
| coop_PairV v1 v2 =>
PairV (Translate_coop_val v1 boundvars) (Translate_coop_val v2 boundvars)
| coop_InjLV v => InjLV (Translate_coop_val v boundvars)
| coop_InjRV v => InjRV (Translate_coop_val v boundvars)
| coop_FoldV v => FoldV (Translate_coop_val v boundvars)
| coop_LocV l => LocV l
| coop_ContV K => ContV (map (λ Ki, Translate_coop_ectxi Ki boundvars) K)
end.
Lemma Translate_coop_typed Γ e τ :
coop_typed Γ e τ →
typed
(Γ ++ [(TArrow (TArrow TUnit TUnit) TUnit); (TArrow TUnit TUnit)])
(Translate_coop e (length Γ)) τ.
Proof.
induction 1; try by simpl; econstructor; eauto.
- constructor.
rewrite lookup_app_l; eauto using lookup_lt_is_Some_1.
- simpl. econstructor; eauto.
rewrite fmap_app; simpl.
rewrite fmap_length in IHtyped.
auto.
- simpl. repeat econstructor.
+ rewrite lookup_app_r; eauto.
replace ((length Γ) - length Γ) with 0 by lia.
eauto.
+ by apply (context_weakening [_]).
- simpl. repeat econstructor.
rewrite lookup_app_r; eauto.
replace (S (length Γ) - length Γ) with 1 by lia.
eauto.
Qed.
Lemma typed_Translate_n_closed Γ τ e:
coop_typed Γ e τ →
(∀ f, (Translate_coop e (length Γ)).[upn (length (Γ ++
[(TArrow (TArrow TUnit TUnit) TUnit);
(TArrow TUnit TUnit)])) f]
= Translate_coop e (length Γ)).
Proof. intros Htp. by eapply typed_n_closed, Translate_coop_typed. Qed.