LogrelCC.cooperative.cooplang.lang
Require Import Logic.FunctionalExtensionality.
From LogrelCC Require Export lang_base.
From LogrelCC Require Export prelude.
From LogrelCC.cooperative.program_logic Require Export
coop_language coop_ectx_language coop_ectxi_language.
From stdpp Require Import gmap.
Module lang.
Inductive coopexpr :=
| Var (x : var)
| Rec (e : {bind 2 of coopexpr})
| Lam (e : {bind coopexpr})
| LetIn (e1 : coopexpr) (e2 : {bind coopexpr})
| Seq (e1 e2 : coopexpr)
| App (e1 e2 : coopexpr)
(* Base Types *)
| Unit
| Nat (n : nat)
| Bool (b : bool)
| BinOp (op : binop) (e1 e2 : coopexpr)
(* If then else *)
| If (e0 e1 e2 : coopexpr)
(* Products *)
| Pair (e1 e2 : coopexpr)
| Fst (e : coopexpr)
| Snd (e : coopexpr)
(* Sums *)
| InjL (e : coopexpr)
| InjR (e : coopexpr)
| Case (e0 : coopexpr) (e1 : {bind coopexpr}) (e2 : {bind coopexpr})
(* Recursive Types *)
| Fold (e : coopexpr)
| Unfold (e : coopexpr)
(* Polymorphic Types *)
| TLam (e : coopexpr)
| TApp (e : coopexpr)
(* Cooperative Concurrency *)
| CoFork (e : coopexpr)
| CoYield
(* Reference Types *)
| Loc (l : loc)
| Alloc (e : coopexpr)
| Load (e : coopexpr)
| Store (e1 : coopexpr) (e2 : coopexpr)
(* Continuations *)
| Cont (K : list ectx_item)
| Callcc (e : {bind coopexpr})
| Throw (e1 : coopexpr) (e2 : coopexpr)
with ectx_item :=
From LogrelCC Require Export lang_base.
From LogrelCC Require Export prelude.
From LogrelCC.cooperative.program_logic Require Export
coop_language coop_ectx_language coop_ectxi_language.
From stdpp Require Import gmap.
Module lang.
Inductive coopexpr :=
| Var (x : var)
| Rec (e : {bind 2 of coopexpr})
| Lam (e : {bind coopexpr})
| LetIn (e1 : coopexpr) (e2 : {bind coopexpr})
| Seq (e1 e2 : coopexpr)
| App (e1 e2 : coopexpr)
(* Base Types *)
| Unit
| Nat (n : nat)
| Bool (b : bool)
| BinOp (op : binop) (e1 e2 : coopexpr)
(* If then else *)
| If (e0 e1 e2 : coopexpr)
(* Products *)
| Pair (e1 e2 : coopexpr)
| Fst (e : coopexpr)
| Snd (e : coopexpr)
(* Sums *)
| InjL (e : coopexpr)
| InjR (e : coopexpr)
| Case (e0 : coopexpr) (e1 : {bind coopexpr}) (e2 : {bind coopexpr})
(* Recursive Types *)
| Fold (e : coopexpr)
| Unfold (e : coopexpr)
(* Polymorphic Types *)
| TLam (e : coopexpr)
| TApp (e : coopexpr)
(* Cooperative Concurrency *)
| CoFork (e : coopexpr)
| CoYield
(* Reference Types *)
| Loc (l : loc)
| Alloc (e : coopexpr)
| Load (e : coopexpr)
| Store (e1 : coopexpr) (e2 : coopexpr)
(* Continuations *)
| Cont (K : list ectx_item)
| Callcc (e : {bind coopexpr})
| Throw (e1 : coopexpr) (e2 : coopexpr)
with ectx_item :=
Evaluation contexts
| AppLCtx (e2 : coopexpr)
| AppRCtx (v1 : coopval)
| LetInCtx (e2 : coopexpr)
| SeqCtx (e2 : coopexpr)
| TAppCtx
| PairLCtx (e2 : coopexpr)
| PairRCtx (v1 : coopval)
| BinOpLCtx (op : binop) (e2 : coopexpr)
| BinOpRCtx (op : binop) (v1 : coopval)
| FstCtx
| SndCtx
| InjLCtx
| InjRCtx
| CaseCtx (e1 : {bind coopexpr}) (e2 : {bind coopexpr})
| IfCtx (e1 : {bind coopexpr}) (e2 : {bind coopexpr})
| FoldCtx
| UnfoldCtx
| AllocCtx
| LoadCtx
| StoreLCtx (e2 : coopexpr)
| StoreRCtx (v1 : coopval)
| ThrowLCtx (e2 : coopexpr)
| ThrowRCtx (v1 : coopval)
with coopval := (* Values *)
| RecV (e : {bind 2 of coopexpr})
| LamV (e : {bind 1 of coopexpr})
| TLamV (e : {bind 1 of coopexpr})
| UnitV
| NatV (n : nat)
| BoolV (b : bool)
| PairV (v1 v2 : coopval)
| InjLV (v : coopval)
| InjRV (v : coopval)
| FoldV (v : coopval)
| LocV (l : loc)
| ContV (K : list ectx_item).
Notation coopectx := (list ectx_item).
Fixpoint of_val (v : coopval) : coopexpr :=
match v with
| RecV e => Rec e
| LamV e => Lam e
| TLamV e => TLam e
| UnitV => Unit
| NatV v => Nat v
| BoolV v => Bool v
| PairV v1 v2 => Pair (of_val v1) (of_val v2)
| InjLV v => InjL (of_val v)
| InjRV v => InjR (of_val v)
| FoldV v => Fold (of_val v)
| LocV l => Loc l
| ContV K => Cont K
end.
Fixpoint to_val (e : coopexpr) : option coopval :=
match e with
| Rec e => Some (RecV e)
| Lam e => Some (LamV e)
| TLam e => Some (TLamV e)
| Unit => Some UnitV
| Nat n => Some (NatV n)
| Bool b => Some (BoolV b)
| Pair e1 e2 => v1 ← to_val e1; v2 ← to_val e2; Some (PairV v1 v2)
| InjL e => InjLV <$> to_val e
| InjR e => InjRV <$> to_val e
| Fold e => v ← to_val e; Some (FoldV v)
| Loc l => Some (LocV l)
| Cont K => Some (ContV K)
| _ => None
end.
Definition fill_item (Ki : ectx_item) (e : coopexpr) : coopexpr :=
match Ki with
| AppLCtx e2 => App e e2
| AppRCtx v1 => App (of_val v1) e
| LetInCtx e2 => LetIn e e2
| SeqCtx e2 => Seq e e2
| TAppCtx => TApp e
| PairLCtx e2 => Pair e e2
| PairRCtx v1 => Pair (of_val v1) e
| BinOpLCtx op e2 => BinOp op e e2
| BinOpRCtx op v1 => BinOp op (of_val v1) e
| FstCtx => Fst e
| SndCtx => Snd e
| InjLCtx => InjL e
| InjRCtx => InjR e
| CaseCtx e1 e2 => Case e e1 e2
| IfCtx e1 e2 => If e e1 e2
| FoldCtx => Fold e
| UnfoldCtx => Unfold e
| AllocCtx => Alloc e
| LoadCtx => Load e
| StoreLCtx e2 => Store e e2
| StoreRCtx v1 => Store (of_val v1) e
| ThrowLCtx e2 => Throw e e2
| ThrowRCtx v1 => Throw (of_val v1) e
end.
(* Notation for bool and nat *)
Notation "#♭ b" := (Bool b) (at level 20).
Notation "#n n" := (Nat n) (at level 20).
(* Notation for bool and nat *)
Notation "'#♭v' b" := (BoolV b) (at level 20).
Notation "'#nv' n" := (NatV n) (at level 20).
Lemma coopexpr_rect' (P : coopexpr → Type) (Q : coopval → Type) :
(∀ x : var, P (Var x))
→ (∀ e : {bind 2 of coopexpr}, P e → P (Rec e))
→ (∀ e : {bind coopexpr}, P e → P (Lam e))
→ (∀ e1 : coopexpr, P e1 → (∀ (e2 : {bind 1 of coopexpr}), P e2 → P (LetIn e1 e2)))
→ (∀ e1 : coopexpr, P e1 → (∀ (e2 : coopexpr), P e2 → P (Seq e1 e2)))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (App e1 e2))
→ P Unit
→ (∀ n : nat, P (#n n))
→ (∀ b : bool, P (#♭ b))
→ (∀ (op : binop) (e1 : coopexpr), P e1 → ∀ e2 : coopexpr, P e2 → P (BinOp op e1 e2))
→ (∀ e0 : coopexpr, P e0 → ∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (If e0 e1 e2))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (Pair e1 e2))
→ (∀ e : coopexpr, P e → P (Fst e))
→ (∀ e : coopexpr, P e → P (Snd e))
→ (∀ e : coopexpr, P e → P (InjL e))
→ (∀ e : coopexpr, P e → P (InjR e))
→ (∀ e0 : coopexpr, P e0 → ∀ e1 : {bind coopexpr}, P e1
→ ∀ e2 : {bind coopexpr}, P e2 → P (Case e0 e1 e2))
→ (∀ e : coopexpr, P e → P (Fold e))
→ (∀ e : coopexpr, P e → P (Unfold e))
→ (∀ e : coopexpr, P e → P (TLam e))
→ (∀ e : coopexpr, P e → P (TApp e))
→ (∀ e : coopexpr, P e → P (CoFork e))
→ P CoYield
→ (∀ l : loc, P (Loc l))
→ (∀ e : coopexpr, P e → P (Alloc e))
→ (∀ e : coopexpr, P e → P (Load e))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (Store e1 e2))
→ P (Cont [])
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (AppLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (AppRCtx v1 :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (LetInCtx e2 :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (SeqCtx e2 :: K)))
→ (∀ K, P (Cont K) → P (Cont (TAppCtx :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (PairLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (PairRCtx v1 :: K)))
→ (∀ K op e2, P (Cont K) → P e2 → P (Cont (BinOpLCtx op e2 :: K)))
→ (∀ K op v1, P (Cont K) → Q v1 → P (Cont (BinOpRCtx op v1 :: K)))
→ (∀ K, P (Cont K) → P (Cont (FstCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (SndCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (InjLCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (InjRCtx :: K)))
→ (∀ K e1 e2, P (Cont K) → P e1 → P e2 → P (Cont (CaseCtx e1 e2 :: K)))
→ (∀ K e1 e2, P (Cont K) → P e1 → P e2 → P (Cont (IfCtx e1 e2 :: K)))
→ (∀ K, P (Cont K) → P (Cont (FoldCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (UnfoldCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (AllocCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (LoadCtx :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (StoreLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (StoreRCtx v1 :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (ThrowLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (ThrowRCtx v1 :: K)))
→ (∀ e, P e → Q (RecV e))
→ (∀ e, P e → Q (LamV e))
→ (∀ e, P e → Q (TLamV e))
→ (Q UnitV)
→ (∀ b, Q (NatV b))
→ (∀ b, Q (BoolV b))
→ (∀ v1 v2, Q v1 → Q v2 → Q (PairV v1 v2))
→ (∀ v, Q v → Q (InjLV v))
→ (∀ v, Q v → Q (InjRV v))
→ (∀ v, Q v → Q (FoldV v))
→ (∀ l, Q (LocV l))
→ (∀ K, P (Cont K) → Q (ContV K))
→ Q (ContV [])
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (AppLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (AppRCtx v1 :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (LetInCtx e2 :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (SeqCtx e2 :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (TAppCtx :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (PairLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (PairRCtx v1 :: K)))
→ (∀ K op e2, Q (ContV K) → P e2 → Q (ContV (BinOpLCtx op e2 :: K)))
→ (∀ K op v1, Q (ContV K) → Q v1 → Q (ContV (BinOpRCtx op v1 :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (FstCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (SndCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (InjLCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (InjRCtx :: K)))
→ (∀ K e1 e2, Q (ContV K) → P e1 → P e2 → Q (ContV (CaseCtx e1 e2 :: K)))
→ (∀ K e1 e2, Q (ContV K) → P e1 → P e2 → Q (ContV (IfCtx e1 e2 :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (FoldCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (UnfoldCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (AllocCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (LoadCtx :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (StoreLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (StoreRCtx v1 :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (ThrowLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (ThrowRCtx v1 :: K)))
→ (∀ e : {bind coopexpr}, P e → P (Callcc e))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (Throw e1 e2))
→ ∀ e : coopexpr, P e.
Proof.
intros Hvar Hrec HLam HLetIn HSeq Happ Hunit Hnat Hbool Hbinop Hif Hpair Hfst Hsnd
Hinjl Hinjr Hcase Hfold Hunfold HTLam HTapp Hfork Hyield Hloc Halloc Hload
Hstore
HKnil HAppLCtx HAppRCtx HLetInCtx HSeqCtx HTAppCtx HPairLCtx HPairRCtx
HBinOpLCtx HBinOpRCtx HFstCtx HSndCtx HInjLCtx HInjRCtx HCaseCtx HIfCtx
HFoldCtx HUnfoldCtx HAllocCtx HLoadCtx HStoreLCtx HStoreRCtx
HThrowLCtx HThrowRCtx
HRecV HLamV HTLamV HUnitV HNatV HBoolV HPairV HInjLV HInjRV HFoldV HLocV
HContV
HKnilV HAppLCtxV HAppRCtxV HLetInCtxV HSeqCtxV HTAppCtxV
HPairLCtxV HPairRCtxV HBinOpLCtxV
HBinOpRCtxV HFstCtxV HSndCtxV HInjLCtxV HInjRCtxV HCaseCtxV HIfCtxV
HFoldCtxV HUnfoldCtxV HAllocCtxV HLoadCtxV HStoreLCtxV HStoreRCtxV
HThrowLCtxV HThrowRCtxV
Hcallcc Hthrow e.
refine (
(fix fx (e : coopexpr) {struct e} :=
match e as u return (* F u → *) P u with
| Var x => Hvar _
| Rec e => Hrec _ (fx e)
| Lam e => HLam _ (fx e)
| LetIn e1 e2 => HLetIn _ (fx e1) _ (fx e2)
| Seq e1 e2 => HSeq _ (fx e1) _ (fx e2)
| App e1 e2 => Happ _ (fx e1) _ (fx e2)
| Unit => Hunit
| Nat n => Hnat _
| Bool b => Hbool _
| BinOp op e1 e2 => Hbinop _ _ (fx e1) _ (fx e2)
| If e0 e1 e2 => Hif _ (fx e0) _ (fx e1) _ (fx e2)
| Pair e1 e2 => Hpair _ (fx e1) _ (fx e2)
| Fst e => Hfst _ (fx e)
| Snd e => Hsnd _ (fx e)
| InjL e => Hinjl _ (fx e)
| InjR e => Hinjr _ (fx e)
| Case e0 e1 e2 => Hcase _ (fx e0) _ (fx e1) _ (fx e2)
| Fold e => Hfold _ (fx e)
| Unfold e => Hunfold _ (fx e)
| TLam e => HTLam _ (fx e)
| TApp e => HTapp _ (fx e)
| CoFork e => Hfork _ (fx e)
| CoYield => Hyield
| Loc l => Hloc _
| Alloc e => Halloc _ (fx e)
| Load e => Hload _ (fx e)
| Store e1 e2 => Hstore _ (fx e1) _ (fx e2)
| Cont K =>
(fix gx (K : coopectx) {struct K} :=
let HX :=
fix hx (v : coopval) {struct v} :=
match v as u return Q u with
| RecV e => HRecV _ (fx e)
| LamV e => HLamV _ (fx e)
| TLamV e => HTLamV _ (fx e)
| UnitV => HUnitV
| NatV n => HNatV n
| BoolV b => HBoolV b
| PairV v1 v2 => HPairV _ _ (hx v1) (hx v2)
| InjLV v => HInjLV _ (hx v)
| InjRV v => HInjRV _ (hx v)
| FoldV v => HFoldV _ (hx v)
| LocV l => HLocV l
| ContV K'' =>
(fix gx' (K : coopectx) {struct K} :=
match K as u return Q (ContV u) with
| nil => HKnilV
| (AppLCtx e2) :: K' => HAppLCtxV _ _ (gx' K') (fx e2)
| (AppRCtx v1) :: K' => HAppRCtxV _ _ (gx' K') (hx v1)
| (LetInCtx e2) :: K' => HLetInCtxV _ _ (gx' K') (fx e2)
| (SeqCtx e2) :: K' => HSeqCtxV _ _ (gx' K') (fx e2)
| (TAppCtx) :: K' => HTAppCtxV _ (gx' K')
| (PairLCtx e2) :: K' => HPairLCtxV _ _ (gx' K') (fx e2)
| (PairRCtx v1) :: K' => HPairRCtxV _ _ (gx' K') (hx v1)
| (BinOpLCtx op e2) :: K' => HBinOpLCtxV _ _ _ (gx' K') (fx e2)
| (BinOpRCtx op v1) :: K' => HBinOpRCtxV _ _ _ (gx' K') (hx v1)
| (FstCtx) :: K' => HFstCtxV _ (gx' K')
| (SndCtx) :: K' => HSndCtxV _ (gx' K')
| (InjLCtx) :: K' => HInjLCtxV _ (gx' K')
| (InjRCtx) :: K' => HInjRCtxV _ (gx' K')
| (CaseCtx e1 e2) :: K' =>
HCaseCtxV _ _ _ (gx' K') (fx e1) (fx e2)
| (IfCtx e1 e2) :: K' => HIfCtxV _ _ _ (gx' K') (fx e1) (fx e2)
| (FoldCtx) :: K' => HFoldCtxV _ (gx' K')
| (UnfoldCtx) :: K' => HUnfoldCtxV _ (gx' K')
| (AllocCtx) :: K' => HAllocCtxV _ (gx' K')
| (LoadCtx) :: K' => HLoadCtxV _ (gx' K')
| (StoreLCtx e2) :: K' => HStoreLCtxV _ _ (gx' K') (fx e2)
| (StoreRCtx v1) :: K' => HStoreRCtxV _ _ (gx' K') (hx v1)
| (ThrowLCtx e2) :: K' => HThrowLCtxV _ _ (gx' K') (fx e2)
| (ThrowRCtx v1) :: K' => HThrowRCtxV _ _ (gx' K') (hx v1)
end) K''
end
in
match K as u return P (Cont u) with
| nil => HKnil
| (AppLCtx e2) :: K' => HAppLCtx _ _ (gx K') (fx e2)
| (AppRCtx v1) :: K' => HAppRCtx _ _ (gx K') (HX v1)
| (LetInCtx e2) :: K' => HLetInCtx _ _ (gx K') (fx e2)
| (SeqCtx e2) :: K' => HSeqCtx _ _ (gx K') (fx e2)
| (TAppCtx) :: K' => HTAppCtx _ (gx K')
| (PairLCtx e2) :: K' => HPairLCtx _ _ (gx K') (fx e2)
| (PairRCtx v1) :: K' => HPairRCtx _ _ (gx K') (HX v1)
| (BinOpLCtx op e2) :: K' => HBinOpLCtx _ _ _ (gx K') (fx e2)
| (BinOpRCtx op v1) :: K' => HBinOpRCtx _ _ _ (gx K') (HX v1)
| (FstCtx) :: K' => HFstCtx _ (gx K')
| (SndCtx) :: K' => HSndCtx _ (gx K')
| (InjLCtx) :: K' => HInjLCtx _ (gx K')
| (InjRCtx) :: K' => HInjRCtx _ (gx K')
| (CaseCtx e1 e2) :: K' => HCaseCtx _ _ _ (gx K') (fx e1) (fx e2)
| (IfCtx e1 e2) :: K' => HIfCtx _ _ _ (gx K') (fx e1) (fx e2)
| (FoldCtx) :: K' => HFoldCtx _ (gx K')
| (UnfoldCtx) :: K' => HUnfoldCtx _ (gx K')
| (AllocCtx) :: K' => HAllocCtx _ (gx K')
| (LoadCtx) :: K' => HLoadCtx _ (gx K')
| (StoreLCtx e2) :: K' => HStoreLCtx _ _ (gx K') (fx e2)
| (StoreRCtx v1) :: K' => HStoreRCtx _ _ (gx K') (HX v1)
| (ThrowLCtx e2) :: K' => HThrowLCtx _ _ (gx K') (fx e2)
| (ThrowRCtx v1) :: K' => HThrowRCtx _ _ (gx K') (HX v1)
end) K
| Callcc e => Hcallcc _ (fx e)
| Throw e1 e2 => Hthrow _ (fx e1) _ (fx e2)
end) e).
Qed.
Instance Ids_coopexpr : Ids coopexpr. Proof. derive. Defined.
Fixpoint coopexpr_rename (sb : var → var) (e : coopexpr) : coopexpr :=
let a := coopexpr_rename : Rename coopexpr in
match e with
| Var x => Var (sb x)
| Rec e => Rec (rename (iterate upren 2 sb) e)
| Lam e => Lam (rename (upren sb) e)
| LetIn e1 e2 => LetIn (rename sb e1) (rename (upren sb) e2)
| Seq e1 e2 => Seq (rename sb e1) (rename sb e2)
| App e1 e2 => App (rename sb e1) (rename sb e2)
| Unit => Unit
| Nat n => Nat n
| Bool b => Bool b
| BinOp op e1 e2 => BinOp op (rename sb e1) (rename sb e2)
| If e0 e1 e2 =>
If (rename sb e0) (rename sb e1) (rename sb e2)
| Pair e1 e2 => Pair (rename sb e1) (rename sb e2)
| Fst e => Fst (rename sb e)
| Snd e => Snd (rename sb e)
| InjL e => InjL (rename sb e)
| InjR e => InjR (rename sb e)
| Case e0 e1 e2 => Case (rename sb e0)
(rename (upren sb) e1)
(rename (upren sb) e2)
| Fold e => Fold (rename sb e)
| Unfold e => Unfold (rename sb e)
| TLam e => TLam (rename sb e)
| TApp e => TApp (rename sb e)
| CoFork e => CoFork (rename sb e)
| CoYield => CoYield
| Loc l => Loc l
| Alloc e => Alloc (rename sb e)
| Load e => Load (rename sb e)
| Store e1 e2 => Store (rename sb e1) (rename sb e2)
| Cont K => Cont (map (ectx_item_rename sb) K)
| Callcc e => Callcc (rename (upren sb) e)
| Throw e1 e2 => Throw (rename sb e1) (rename sb e2)
end
with coopval_rename (sb : var → var) (v : coopval) : coopval :=
let a := coopexpr_rename : Rename coopexpr in
match v with
| RecV e => RecV (rename (iterate upren 2 sb) e)
| LamV e => LamV (rename (upren sb) e)
| TLamV e => TLamV (rename sb e)
| UnitV => UnitV
| NatV n => NatV n
| BoolV b => BoolV b
| PairV v1 v2 => PairV (coopval_rename sb v1) (coopval_rename sb v2)
| InjLV v => InjLV (coopval_rename sb v)
| InjRV v => InjRV (coopval_rename sb v)
| FoldV v => FoldV (coopval_rename sb v)
| LocV l => LocV l
| ContV K => ContV (map (ectx_item_rename sb) K)
end
with ectx_item_rename (sb : var → var) (K : ectx_item) : ectx_item :=
let a := coopexpr_rename : Rename coopexpr in
match K with
| AppLCtx e2 => AppLCtx (rename sb e2)
| AppRCtx v1 => AppRCtx (coopval_rename sb v1)
| LetInCtx e2 => LetInCtx (rename (upren sb) e2)
| SeqCtx e2 => SeqCtx (rename sb e2)
| TAppCtx => TAppCtx
| PairLCtx e2 => PairLCtx (rename sb e2)
| PairRCtx v1 => PairRCtx (coopval_rename sb v1)
| BinOpLCtx op e2 => BinOpLCtx op (rename sb e2)
| BinOpRCtx op v1 => BinOpRCtx op (coopval_rename sb v1)
| FstCtx => FstCtx
| SndCtx => SndCtx
| InjLCtx => InjLCtx
| InjRCtx => InjRCtx
| CaseCtx e1 e2 => CaseCtx (rename (upren sb ) e1)
(rename (upren sb) e2)
| IfCtx e1 e2 => IfCtx (rename sb e1) (rename sb e2)
| FoldCtx => FoldCtx
| UnfoldCtx => UnfoldCtx
| AllocCtx => AllocCtx
| LoadCtx => LoadCtx
| StoreLCtx e2 => StoreLCtx (rename sb e2)
| StoreRCtx v1 => StoreRCtx (coopval_rename sb v1)
| ThrowLCtx e2 => ThrowLCtx (rename sb e2)
| ThrowRCtx v1 => ThrowRCtx (coopval_rename sb v1)
end.
Instance Rename_coopexpr : Rename coopexpr := coopexpr_rename.
Fixpoint coopexpr_subst (sb : var → coopexpr) (e : coopexpr) : coopexpr :=
let a := coopexpr_subst : Subst coopexpr in
match e with
| Var x => sb x
| Rec e => Rec (subst (upn 2 sb) e)
| Lam e => Lam (subst (up sb) e)
| LetIn e1 e2 => LetIn (subst sb e1) (subst (up sb) e2)
| Seq e1 e2 => Seq (subst sb e1) (subst sb e2)
| App e1 e2 => App (subst sb e1) (subst sb e2)
| Unit => Unit
| Nat n => Nat n
| Bool b => Bool b
| BinOp op e1 e2 => BinOp op (subst sb e1) (subst sb e2)
| If e0 e1 e2 =>
If (subst sb e0) (subst sb e1) (subst sb e2)
| Pair e1 e2 => Pair (subst sb e1) (subst sb e2)
| Fst e => Fst (subst sb e)
| Snd e => Snd (subst sb e)
| InjL e => InjL (subst sb e)
| InjR e => InjR (subst sb e)
| Case e0 e1 e2 => Case (subst sb e0)
(subst (up sb) e1)
(subst (up sb) e2)
| Fold e => Fold (subst sb e)
| Unfold e => Unfold (subst sb e)
| TLam e => TLam (subst sb e)
| TApp e => TApp (subst sb e)
| CoFork e => CoFork (subst sb e)
| CoYield => CoYield
| Loc l => Loc l
| Alloc e => Alloc (subst sb e)
| Load e => Load (subst sb e)
| Store e1 e2 => Store (subst sb e1) (subst sb e2)
| Cont K => Cont (map (ectx_item_subst sb) K)
| Callcc e => Callcc (subst (up sb) e)
| Throw e1 e2 => Throw (subst sb e1) (subst sb e2)
end
with coopval_subst (sb : var → coopexpr) (v : coopval) : coopval :=
let a := coopexpr_subst : Subst coopexpr in
match v with
| RecV e => RecV (subst (upn 2 sb) e)
| LamV e => LamV (subst (up sb) e)
| TLamV e => TLamV (subst sb e)
| UnitV => UnitV
| NatV n => NatV n
| BoolV b => BoolV b
| PairV v1 v2 => PairV (coopval_subst sb v1) (coopval_subst sb v2)
| InjLV v => InjLV (coopval_subst sb v)
| InjRV v => InjRV (coopval_subst sb v)
| FoldV v => FoldV (coopval_subst sb v)
| LocV l => LocV l
| ContV K => ContV (map (ectx_item_subst sb) K)
end
with ectx_item_subst (sb : var → coopexpr) (K : ectx_item) : ectx_item :=
let a := coopexpr_subst : Subst coopexpr in
match K with
| AppLCtx e2 => AppLCtx (subst sb e2)
| AppRCtx v1 => AppRCtx (coopval_subst sb v1)
| LetInCtx e2 => LetInCtx (subst (up sb) e2)
| SeqCtx e2 => SeqCtx (subst sb e2)
| TAppCtx => TAppCtx
| PairLCtx e2 => PairLCtx (subst sb e2)
| PairRCtx v1 => PairRCtx (coopval_subst sb v1)
| BinOpLCtx op e2 => BinOpLCtx op (subst sb e2)
| BinOpRCtx op v1 => BinOpRCtx op (coopval_subst sb v1)
| FstCtx => FstCtx
| SndCtx => SndCtx
| InjLCtx => InjLCtx
| InjRCtx => InjRCtx
| CaseCtx e1 e2 =>
CaseCtx (subst (up sb) e1) (subst (up sb) e2)
| IfCtx e1 e2 => IfCtx (subst sb e1) (subst sb e2)
| FoldCtx => FoldCtx
| UnfoldCtx => UnfoldCtx
| AllocCtx => AllocCtx
| LoadCtx => LoadCtx
| StoreLCtx e2 => StoreLCtx (subst sb e2)
| StoreRCtx v1 => StoreRCtx (coopval_subst sb v1)
| ThrowLCtx e2 => ThrowLCtx (subst sb e2)
| ThrowRCtx v1 => ThrowRCtx (coopval_subst sb v1)
end.
Instance Subst_coopexpr : Subst coopexpr := coopexpr_subst.
Lemma helper1 : ∀ xi : var → var, up (ren xi) = ren (upren xi).
Proof. intros sb; extensionality z; induction z; simpl; auto. Qed.
Lemma helper2 : up ids = ids.
Proof. extensionality z; induction z; simpl; auto. Qed.
Lemma helper3:
∀ (xi : var → var) (sigma : var → coopexpr) (s : coopexpr),
(rename xi s).[sigma] = s.[xi >>> sigma].
Proof.
intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ (f : var → var) (g : var → coopexpr),
(rename f e).[g] = e.[f >>> g])
(λ v, ∀ (f : var → var) (g : var → coopexpr),
coopval_subst g (coopval_rename f v) = coopval_subst (f >>> g) v));
intros; simpl in *;
repeat match goal with
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), Cont _ = Cont _ |- _ =>
specialize (H f g); inversion H
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), ContV _ = ContV _ |- _ =>
specialize (H f g); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 -?up_comp_ren_subst; auto.
Qed.
Lemma helper4 : ∀ (xi : var → var) (s : coopexpr), rename xi s = s.[ren xi].
Proof.
intros f e; revert f.
apply (coopexpr_rect'
(λ e, ∀ f : var → var, rename f e = e.[ren f])
(λ v, ∀ f : var → var, coopval_rename f v = coopval_subst (ren f) v));
intros; simpl in *;
repeat match goal with
| f : var → var, H : ∀ f : var → var, Cont _ = Cont _ |- _ =>
specialize (H f); inversion H
| f : var → var, H : ∀ f : var → var, ContV _ = ContV _ |- _ =>
specialize (H f); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 ?up_upren_n_internal ?up_upren_internal;
auto using helper1.
Qed.
Lemma helper5 :
∀ (xi : var → var) (sigma : var → coopexpr) (s : coopexpr),
(rename xi s).[sigma] = s.[xi >>> sigma].
Proof.
intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ (f : var → var) (g : var → coopexpr),
(rename f e).[g] = e.[f >>> g])
(λ v, ∀ (f : var → var) (g : var → coopexpr),
coopval_subst g (coopval_rename f v) = coopval_subst (f >>> g) v));
intros; simpl in *;
repeat match goal with
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), Cont _ = Cont _ |- _ =>
specialize (H f g); inversion H
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), ContV _ = ContV _ |- _ =>
specialize (H f g); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 -?up_comp_ren_subst; auto.
Qed.
Lemma helper6 :
∀ (sigma : var → coopexpr) (xi : var → var) (s : coopexpr),
rename xi s.[sigma] = s.[sigma >>> rename xi].
Proof.
intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ (f : var → coopexpr) (g : var → var),
rename g e.[f] = e.[f >>> rename g])
(λ v, ∀ (f : var → coopexpr) (g : var → var),
coopval_rename g (coopval_subst f v) = coopval_subst (f >>> rename g) v));
intros; simpl in *;
repeat match goal with
| f : var → var, g : var → coopexpr,
H : ∀ (g : var → coopexpr) (f : var → var), Cont _ = Cont _ |- _ =>
specialize (H g f); inversion H
| f : var → var, g : var → coopexpr,
H : ∀ (g : var → coopexpr) (f : var → var), ContV _ = ContV _ |- _ =>
specialize (H g f); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 ?up_comp_subst_ren_n_internal
?up_comp_subst_ren_internal;
try apply up_comp_subst_ren_internal; auto using helper4, helper5.
Qed.
Instance SubstLemmas_coopexpr : SubstLemmas coopexpr.
Proof.
econstructor; eauto using helper4.
- apply (coopexpr_rect'
(λ e, e.[ids] = e)
(λ v, coopval_subst ids v = v));
intros; simpl in *;
repeat match goal with
| H : Cont _ = Cont _ |- _ => inversion H; clear H
| H : ContV _ = ContV _ |- _ => inversion H; clear H
end;
rewrite ?up_id_internal ?up_id_n_internal ?H ?H0 ?H1 ?H2 ?H3;
auto using helper2.
- intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ f g : var → coopexpr, e.[f].[g] = e.[f >> g])
(λ v, ∀ f g : var → coopexpr, coopval_subst g (coopval_subst f v)
= coopval_subst (f >> g) v));
intros; simpl in *;
repeat match goal with
| f : var → coopexpr, g : var → coopexpr,
H : ∀ (f : var → coopexpr) (g : var → coopexpr), Cont _ = Cont _ |- _ =>
let H' := fresh in let H'' := fresh in
pose proof (H f g) as H'; inversion H';
pose proof (H g f) as H''; inversion H''; clear H
| f : var → coopexpr, g : var → coopexpr,
H : ∀ (f : var → coopexpr) (g : var → coopexpr), ContV _ = ContV _ |- _ =>
let H' := fresh in let H'' := fresh in
pose proof (H f g) as H'; inversion H';
pose proof (H g f) as H''; inversion H''; clear H
end;
rewrite ?H ?H0 ?H1 ?H2 ?H3 ?up_comp_internal
?up_comp_n_internal; try apply up_comp_internal;
auto using helper5, helper6.
Qed.
Instance coopexpr_dec_eq : EqDecision coopexpr.
Proof.
intros e.
apply (coopexpr_rect' (λ e, ∀ e' : coopexpr, Decision (e = e'))
(λ w, ∀ v, Decision (w = v))); intros;
match goal with
|- Decision (?A = ?B) => destruct B
end;
try match goal with
|- Decision (Cont ?A = Cont ?B) => destruct B as [|[]]
end;
try match goal with
| |- Decision (ContV (_ :: _) = ContV ?B) => destruct B as [|[]]
| |- Decision (ContV [] = ContV ?B) => destruct B as [|[]]
end; try (right; inversion 1; fail).
all: try match goal with
| |- Decision (?A = ?A) => left; trivial
| |- Decision (?A ?B = ?A ?C) => destruct (decide (B = C));
[| right; inversion 1; tauto]; subst; left; eauto
| |- Decision (?A ?B1 ?B2 = ?A ?C1 ?C2) =>
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (?A ?B1 ?B2 ?B3 = ?A ?C1 ?C2 ?C3) =>
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
destruct (decide (B3 = C3)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (?A ?B1 ?B2 ?B3 = ?A ?C1 ?C2 ?C3) =>
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
destruct (decide (B3 = C3)); [| right; inversion 1; tauto];
subst; left; eauto
end.
all: try match goal with
| |- Decision (Cont (?A :: ?K1) = Cont (?A :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto]; subst; left; eauto
| |- Decision (Cont (?A ?B :: ?K1) = Cont (?A ?C :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B = C)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (Cont (?A ?B1 ?B2 :: ?K1) = Cont (?A ?C1 ?C2 :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
subst; left; eauto
end.
all: try match goal with
| |- Decision (ContV (?A :: ?K1) = ContV (?A :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (ContV K1 = ContV K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto]; subst; left; eauto
| |- Decision (ContV (?A ?B :: ?K1) = ContV (?A ?C :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (ContV K1 = ContV K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B = C)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (ContV (?A ?B1 ?B2 :: ?K1) = ContV (?A ?C1 ?C2 :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (ContV K1 = ContV K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
subst; left; eauto
end.
match goal with
| |- Decision (ContV ?K1 = ContV ?K2) =>
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
subst; left; eauto
end.
Qed.
| AppRCtx (v1 : coopval)
| LetInCtx (e2 : coopexpr)
| SeqCtx (e2 : coopexpr)
| TAppCtx
| PairLCtx (e2 : coopexpr)
| PairRCtx (v1 : coopval)
| BinOpLCtx (op : binop) (e2 : coopexpr)
| BinOpRCtx (op : binop) (v1 : coopval)
| FstCtx
| SndCtx
| InjLCtx
| InjRCtx
| CaseCtx (e1 : {bind coopexpr}) (e2 : {bind coopexpr})
| IfCtx (e1 : {bind coopexpr}) (e2 : {bind coopexpr})
| FoldCtx
| UnfoldCtx
| AllocCtx
| LoadCtx
| StoreLCtx (e2 : coopexpr)
| StoreRCtx (v1 : coopval)
| ThrowLCtx (e2 : coopexpr)
| ThrowRCtx (v1 : coopval)
with coopval := (* Values *)
| RecV (e : {bind 2 of coopexpr})
| LamV (e : {bind 1 of coopexpr})
| TLamV (e : {bind 1 of coopexpr})
| UnitV
| NatV (n : nat)
| BoolV (b : bool)
| PairV (v1 v2 : coopval)
| InjLV (v : coopval)
| InjRV (v : coopval)
| FoldV (v : coopval)
| LocV (l : loc)
| ContV (K : list ectx_item).
Notation coopectx := (list ectx_item).
Fixpoint of_val (v : coopval) : coopexpr :=
match v with
| RecV e => Rec e
| LamV e => Lam e
| TLamV e => TLam e
| UnitV => Unit
| NatV v => Nat v
| BoolV v => Bool v
| PairV v1 v2 => Pair (of_val v1) (of_val v2)
| InjLV v => InjL (of_val v)
| InjRV v => InjR (of_val v)
| FoldV v => Fold (of_val v)
| LocV l => Loc l
| ContV K => Cont K
end.
Fixpoint to_val (e : coopexpr) : option coopval :=
match e with
| Rec e => Some (RecV e)
| Lam e => Some (LamV e)
| TLam e => Some (TLamV e)
| Unit => Some UnitV
| Nat n => Some (NatV n)
| Bool b => Some (BoolV b)
| Pair e1 e2 => v1 ← to_val e1; v2 ← to_val e2; Some (PairV v1 v2)
| InjL e => InjLV <$> to_val e
| InjR e => InjRV <$> to_val e
| Fold e => v ← to_val e; Some (FoldV v)
| Loc l => Some (LocV l)
| Cont K => Some (ContV K)
| _ => None
end.
Definition fill_item (Ki : ectx_item) (e : coopexpr) : coopexpr :=
match Ki with
| AppLCtx e2 => App e e2
| AppRCtx v1 => App (of_val v1) e
| LetInCtx e2 => LetIn e e2
| SeqCtx e2 => Seq e e2
| TAppCtx => TApp e
| PairLCtx e2 => Pair e e2
| PairRCtx v1 => Pair (of_val v1) e
| BinOpLCtx op e2 => BinOp op e e2
| BinOpRCtx op v1 => BinOp op (of_val v1) e
| FstCtx => Fst e
| SndCtx => Snd e
| InjLCtx => InjL e
| InjRCtx => InjR e
| CaseCtx e1 e2 => Case e e1 e2
| IfCtx e1 e2 => If e e1 e2
| FoldCtx => Fold e
| UnfoldCtx => Unfold e
| AllocCtx => Alloc e
| LoadCtx => Load e
| StoreLCtx e2 => Store e e2
| StoreRCtx v1 => Store (of_val v1) e
| ThrowLCtx e2 => Throw e e2
| ThrowRCtx v1 => Throw (of_val v1) e
end.
(* Notation for bool and nat *)
Notation "#♭ b" := (Bool b) (at level 20).
Notation "#n n" := (Nat n) (at level 20).
(* Notation for bool and nat *)
Notation "'#♭v' b" := (BoolV b) (at level 20).
Notation "'#nv' n" := (NatV n) (at level 20).
Lemma coopexpr_rect' (P : coopexpr → Type) (Q : coopval → Type) :
(∀ x : var, P (Var x))
→ (∀ e : {bind 2 of coopexpr}, P e → P (Rec e))
→ (∀ e : {bind coopexpr}, P e → P (Lam e))
→ (∀ e1 : coopexpr, P e1 → (∀ (e2 : {bind 1 of coopexpr}), P e2 → P (LetIn e1 e2)))
→ (∀ e1 : coopexpr, P e1 → (∀ (e2 : coopexpr), P e2 → P (Seq e1 e2)))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (App e1 e2))
→ P Unit
→ (∀ n : nat, P (#n n))
→ (∀ b : bool, P (#♭ b))
→ (∀ (op : binop) (e1 : coopexpr), P e1 → ∀ e2 : coopexpr, P e2 → P (BinOp op e1 e2))
→ (∀ e0 : coopexpr, P e0 → ∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (If e0 e1 e2))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (Pair e1 e2))
→ (∀ e : coopexpr, P e → P (Fst e))
→ (∀ e : coopexpr, P e → P (Snd e))
→ (∀ e : coopexpr, P e → P (InjL e))
→ (∀ e : coopexpr, P e → P (InjR e))
→ (∀ e0 : coopexpr, P e0 → ∀ e1 : {bind coopexpr}, P e1
→ ∀ e2 : {bind coopexpr}, P e2 → P (Case e0 e1 e2))
→ (∀ e : coopexpr, P e → P (Fold e))
→ (∀ e : coopexpr, P e → P (Unfold e))
→ (∀ e : coopexpr, P e → P (TLam e))
→ (∀ e : coopexpr, P e → P (TApp e))
→ (∀ e : coopexpr, P e → P (CoFork e))
→ P CoYield
→ (∀ l : loc, P (Loc l))
→ (∀ e : coopexpr, P e → P (Alloc e))
→ (∀ e : coopexpr, P e → P (Load e))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (Store e1 e2))
→ P (Cont [])
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (AppLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (AppRCtx v1 :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (LetInCtx e2 :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (SeqCtx e2 :: K)))
→ (∀ K, P (Cont K) → P (Cont (TAppCtx :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (PairLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (PairRCtx v1 :: K)))
→ (∀ K op e2, P (Cont K) → P e2 → P (Cont (BinOpLCtx op e2 :: K)))
→ (∀ K op v1, P (Cont K) → Q v1 → P (Cont (BinOpRCtx op v1 :: K)))
→ (∀ K, P (Cont K) → P (Cont (FstCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (SndCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (InjLCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (InjRCtx :: K)))
→ (∀ K e1 e2, P (Cont K) → P e1 → P e2 → P (Cont (CaseCtx e1 e2 :: K)))
→ (∀ K e1 e2, P (Cont K) → P e1 → P e2 → P (Cont (IfCtx e1 e2 :: K)))
→ (∀ K, P (Cont K) → P (Cont (FoldCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (UnfoldCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (AllocCtx :: K)))
→ (∀ K, P (Cont K) → P (Cont (LoadCtx :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (StoreLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (StoreRCtx v1 :: K)))
→ (∀ K e2, P (Cont K) → P e2 → P (Cont (ThrowLCtx e2 :: K)))
→ (∀ K v1, P (Cont K) → Q v1 → P (Cont (ThrowRCtx v1 :: K)))
→ (∀ e, P e → Q (RecV e))
→ (∀ e, P e → Q (LamV e))
→ (∀ e, P e → Q (TLamV e))
→ (Q UnitV)
→ (∀ b, Q (NatV b))
→ (∀ b, Q (BoolV b))
→ (∀ v1 v2, Q v1 → Q v2 → Q (PairV v1 v2))
→ (∀ v, Q v → Q (InjLV v))
→ (∀ v, Q v → Q (InjRV v))
→ (∀ v, Q v → Q (FoldV v))
→ (∀ l, Q (LocV l))
→ (∀ K, P (Cont K) → Q (ContV K))
→ Q (ContV [])
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (AppLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (AppRCtx v1 :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (LetInCtx e2 :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (SeqCtx e2 :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (TAppCtx :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (PairLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (PairRCtx v1 :: K)))
→ (∀ K op e2, Q (ContV K) → P e2 → Q (ContV (BinOpLCtx op e2 :: K)))
→ (∀ K op v1, Q (ContV K) → Q v1 → Q (ContV (BinOpRCtx op v1 :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (FstCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (SndCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (InjLCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (InjRCtx :: K)))
→ (∀ K e1 e2, Q (ContV K) → P e1 → P e2 → Q (ContV (CaseCtx e1 e2 :: K)))
→ (∀ K e1 e2, Q (ContV K) → P e1 → P e2 → Q (ContV (IfCtx e1 e2 :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (FoldCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (UnfoldCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (AllocCtx :: K)))
→ (∀ K, Q (ContV K) → Q (ContV (LoadCtx :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (StoreLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (StoreRCtx v1 :: K)))
→ (∀ K e2, Q (ContV K) → P e2 → Q (ContV (ThrowLCtx e2 :: K)))
→ (∀ K v1, Q (ContV K) → Q v1 → Q (ContV (ThrowRCtx v1 :: K)))
→ (∀ e : {bind coopexpr}, P e → P (Callcc e))
→ (∀ e1 : coopexpr, P e1 → ∀ e2 : coopexpr, P e2 → P (Throw e1 e2))
→ ∀ e : coopexpr, P e.
Proof.
intros Hvar Hrec HLam HLetIn HSeq Happ Hunit Hnat Hbool Hbinop Hif Hpair Hfst Hsnd
Hinjl Hinjr Hcase Hfold Hunfold HTLam HTapp Hfork Hyield Hloc Halloc Hload
Hstore
HKnil HAppLCtx HAppRCtx HLetInCtx HSeqCtx HTAppCtx HPairLCtx HPairRCtx
HBinOpLCtx HBinOpRCtx HFstCtx HSndCtx HInjLCtx HInjRCtx HCaseCtx HIfCtx
HFoldCtx HUnfoldCtx HAllocCtx HLoadCtx HStoreLCtx HStoreRCtx
HThrowLCtx HThrowRCtx
HRecV HLamV HTLamV HUnitV HNatV HBoolV HPairV HInjLV HInjRV HFoldV HLocV
HContV
HKnilV HAppLCtxV HAppRCtxV HLetInCtxV HSeqCtxV HTAppCtxV
HPairLCtxV HPairRCtxV HBinOpLCtxV
HBinOpRCtxV HFstCtxV HSndCtxV HInjLCtxV HInjRCtxV HCaseCtxV HIfCtxV
HFoldCtxV HUnfoldCtxV HAllocCtxV HLoadCtxV HStoreLCtxV HStoreRCtxV
HThrowLCtxV HThrowRCtxV
Hcallcc Hthrow e.
refine (
(fix fx (e : coopexpr) {struct e} :=
match e as u return (* F u → *) P u with
| Var x => Hvar _
| Rec e => Hrec _ (fx e)
| Lam e => HLam _ (fx e)
| LetIn e1 e2 => HLetIn _ (fx e1) _ (fx e2)
| Seq e1 e2 => HSeq _ (fx e1) _ (fx e2)
| App e1 e2 => Happ _ (fx e1) _ (fx e2)
| Unit => Hunit
| Nat n => Hnat _
| Bool b => Hbool _
| BinOp op e1 e2 => Hbinop _ _ (fx e1) _ (fx e2)
| If e0 e1 e2 => Hif _ (fx e0) _ (fx e1) _ (fx e2)
| Pair e1 e2 => Hpair _ (fx e1) _ (fx e2)
| Fst e => Hfst _ (fx e)
| Snd e => Hsnd _ (fx e)
| InjL e => Hinjl _ (fx e)
| InjR e => Hinjr _ (fx e)
| Case e0 e1 e2 => Hcase _ (fx e0) _ (fx e1) _ (fx e2)
| Fold e => Hfold _ (fx e)
| Unfold e => Hunfold _ (fx e)
| TLam e => HTLam _ (fx e)
| TApp e => HTapp _ (fx e)
| CoFork e => Hfork _ (fx e)
| CoYield => Hyield
| Loc l => Hloc _
| Alloc e => Halloc _ (fx e)
| Load e => Hload _ (fx e)
| Store e1 e2 => Hstore _ (fx e1) _ (fx e2)
| Cont K =>
(fix gx (K : coopectx) {struct K} :=
let HX :=
fix hx (v : coopval) {struct v} :=
match v as u return Q u with
| RecV e => HRecV _ (fx e)
| LamV e => HLamV _ (fx e)
| TLamV e => HTLamV _ (fx e)
| UnitV => HUnitV
| NatV n => HNatV n
| BoolV b => HBoolV b
| PairV v1 v2 => HPairV _ _ (hx v1) (hx v2)
| InjLV v => HInjLV _ (hx v)
| InjRV v => HInjRV _ (hx v)
| FoldV v => HFoldV _ (hx v)
| LocV l => HLocV l
| ContV K'' =>
(fix gx' (K : coopectx) {struct K} :=
match K as u return Q (ContV u) with
| nil => HKnilV
| (AppLCtx e2) :: K' => HAppLCtxV _ _ (gx' K') (fx e2)
| (AppRCtx v1) :: K' => HAppRCtxV _ _ (gx' K') (hx v1)
| (LetInCtx e2) :: K' => HLetInCtxV _ _ (gx' K') (fx e2)
| (SeqCtx e2) :: K' => HSeqCtxV _ _ (gx' K') (fx e2)
| (TAppCtx) :: K' => HTAppCtxV _ (gx' K')
| (PairLCtx e2) :: K' => HPairLCtxV _ _ (gx' K') (fx e2)
| (PairRCtx v1) :: K' => HPairRCtxV _ _ (gx' K') (hx v1)
| (BinOpLCtx op e2) :: K' => HBinOpLCtxV _ _ _ (gx' K') (fx e2)
| (BinOpRCtx op v1) :: K' => HBinOpRCtxV _ _ _ (gx' K') (hx v1)
| (FstCtx) :: K' => HFstCtxV _ (gx' K')
| (SndCtx) :: K' => HSndCtxV _ (gx' K')
| (InjLCtx) :: K' => HInjLCtxV _ (gx' K')
| (InjRCtx) :: K' => HInjRCtxV _ (gx' K')
| (CaseCtx e1 e2) :: K' =>
HCaseCtxV _ _ _ (gx' K') (fx e1) (fx e2)
| (IfCtx e1 e2) :: K' => HIfCtxV _ _ _ (gx' K') (fx e1) (fx e2)
| (FoldCtx) :: K' => HFoldCtxV _ (gx' K')
| (UnfoldCtx) :: K' => HUnfoldCtxV _ (gx' K')
| (AllocCtx) :: K' => HAllocCtxV _ (gx' K')
| (LoadCtx) :: K' => HLoadCtxV _ (gx' K')
| (StoreLCtx e2) :: K' => HStoreLCtxV _ _ (gx' K') (fx e2)
| (StoreRCtx v1) :: K' => HStoreRCtxV _ _ (gx' K') (hx v1)
| (ThrowLCtx e2) :: K' => HThrowLCtxV _ _ (gx' K') (fx e2)
| (ThrowRCtx v1) :: K' => HThrowRCtxV _ _ (gx' K') (hx v1)
end) K''
end
in
match K as u return P (Cont u) with
| nil => HKnil
| (AppLCtx e2) :: K' => HAppLCtx _ _ (gx K') (fx e2)
| (AppRCtx v1) :: K' => HAppRCtx _ _ (gx K') (HX v1)
| (LetInCtx e2) :: K' => HLetInCtx _ _ (gx K') (fx e2)
| (SeqCtx e2) :: K' => HSeqCtx _ _ (gx K') (fx e2)
| (TAppCtx) :: K' => HTAppCtx _ (gx K')
| (PairLCtx e2) :: K' => HPairLCtx _ _ (gx K') (fx e2)
| (PairRCtx v1) :: K' => HPairRCtx _ _ (gx K') (HX v1)
| (BinOpLCtx op e2) :: K' => HBinOpLCtx _ _ _ (gx K') (fx e2)
| (BinOpRCtx op v1) :: K' => HBinOpRCtx _ _ _ (gx K') (HX v1)
| (FstCtx) :: K' => HFstCtx _ (gx K')
| (SndCtx) :: K' => HSndCtx _ (gx K')
| (InjLCtx) :: K' => HInjLCtx _ (gx K')
| (InjRCtx) :: K' => HInjRCtx _ (gx K')
| (CaseCtx e1 e2) :: K' => HCaseCtx _ _ _ (gx K') (fx e1) (fx e2)
| (IfCtx e1 e2) :: K' => HIfCtx _ _ _ (gx K') (fx e1) (fx e2)
| (FoldCtx) :: K' => HFoldCtx _ (gx K')
| (UnfoldCtx) :: K' => HUnfoldCtx _ (gx K')
| (AllocCtx) :: K' => HAllocCtx _ (gx K')
| (LoadCtx) :: K' => HLoadCtx _ (gx K')
| (StoreLCtx e2) :: K' => HStoreLCtx _ _ (gx K') (fx e2)
| (StoreRCtx v1) :: K' => HStoreRCtx _ _ (gx K') (HX v1)
| (ThrowLCtx e2) :: K' => HThrowLCtx _ _ (gx K') (fx e2)
| (ThrowRCtx v1) :: K' => HThrowRCtx _ _ (gx K') (HX v1)
end) K
| Callcc e => Hcallcc _ (fx e)
| Throw e1 e2 => Hthrow _ (fx e1) _ (fx e2)
end) e).
Qed.
Instance Ids_coopexpr : Ids coopexpr. Proof. derive. Defined.
Fixpoint coopexpr_rename (sb : var → var) (e : coopexpr) : coopexpr :=
let a := coopexpr_rename : Rename coopexpr in
match e with
| Var x => Var (sb x)
| Rec e => Rec (rename (iterate upren 2 sb) e)
| Lam e => Lam (rename (upren sb) e)
| LetIn e1 e2 => LetIn (rename sb e1) (rename (upren sb) e2)
| Seq e1 e2 => Seq (rename sb e1) (rename sb e2)
| App e1 e2 => App (rename sb e1) (rename sb e2)
| Unit => Unit
| Nat n => Nat n
| Bool b => Bool b
| BinOp op e1 e2 => BinOp op (rename sb e1) (rename sb e2)
| If e0 e1 e2 =>
If (rename sb e0) (rename sb e1) (rename sb e2)
| Pair e1 e2 => Pair (rename sb e1) (rename sb e2)
| Fst e => Fst (rename sb e)
| Snd e => Snd (rename sb e)
| InjL e => InjL (rename sb e)
| InjR e => InjR (rename sb e)
| Case e0 e1 e2 => Case (rename sb e0)
(rename (upren sb) e1)
(rename (upren sb) e2)
| Fold e => Fold (rename sb e)
| Unfold e => Unfold (rename sb e)
| TLam e => TLam (rename sb e)
| TApp e => TApp (rename sb e)
| CoFork e => CoFork (rename sb e)
| CoYield => CoYield
| Loc l => Loc l
| Alloc e => Alloc (rename sb e)
| Load e => Load (rename sb e)
| Store e1 e2 => Store (rename sb e1) (rename sb e2)
| Cont K => Cont (map (ectx_item_rename sb) K)
| Callcc e => Callcc (rename (upren sb) e)
| Throw e1 e2 => Throw (rename sb e1) (rename sb e2)
end
with coopval_rename (sb : var → var) (v : coopval) : coopval :=
let a := coopexpr_rename : Rename coopexpr in
match v with
| RecV e => RecV (rename (iterate upren 2 sb) e)
| LamV e => LamV (rename (upren sb) e)
| TLamV e => TLamV (rename sb e)
| UnitV => UnitV
| NatV n => NatV n
| BoolV b => BoolV b
| PairV v1 v2 => PairV (coopval_rename sb v1) (coopval_rename sb v2)
| InjLV v => InjLV (coopval_rename sb v)
| InjRV v => InjRV (coopval_rename sb v)
| FoldV v => FoldV (coopval_rename sb v)
| LocV l => LocV l
| ContV K => ContV (map (ectx_item_rename sb) K)
end
with ectx_item_rename (sb : var → var) (K : ectx_item) : ectx_item :=
let a := coopexpr_rename : Rename coopexpr in
match K with
| AppLCtx e2 => AppLCtx (rename sb e2)
| AppRCtx v1 => AppRCtx (coopval_rename sb v1)
| LetInCtx e2 => LetInCtx (rename (upren sb) e2)
| SeqCtx e2 => SeqCtx (rename sb e2)
| TAppCtx => TAppCtx
| PairLCtx e2 => PairLCtx (rename sb e2)
| PairRCtx v1 => PairRCtx (coopval_rename sb v1)
| BinOpLCtx op e2 => BinOpLCtx op (rename sb e2)
| BinOpRCtx op v1 => BinOpRCtx op (coopval_rename sb v1)
| FstCtx => FstCtx
| SndCtx => SndCtx
| InjLCtx => InjLCtx
| InjRCtx => InjRCtx
| CaseCtx e1 e2 => CaseCtx (rename (upren sb ) e1)
(rename (upren sb) e2)
| IfCtx e1 e2 => IfCtx (rename sb e1) (rename sb e2)
| FoldCtx => FoldCtx
| UnfoldCtx => UnfoldCtx
| AllocCtx => AllocCtx
| LoadCtx => LoadCtx
| StoreLCtx e2 => StoreLCtx (rename sb e2)
| StoreRCtx v1 => StoreRCtx (coopval_rename sb v1)
| ThrowLCtx e2 => ThrowLCtx (rename sb e2)
| ThrowRCtx v1 => ThrowRCtx (coopval_rename sb v1)
end.
Instance Rename_coopexpr : Rename coopexpr := coopexpr_rename.
Fixpoint coopexpr_subst (sb : var → coopexpr) (e : coopexpr) : coopexpr :=
let a := coopexpr_subst : Subst coopexpr in
match e with
| Var x => sb x
| Rec e => Rec (subst (upn 2 sb) e)
| Lam e => Lam (subst (up sb) e)
| LetIn e1 e2 => LetIn (subst sb e1) (subst (up sb) e2)
| Seq e1 e2 => Seq (subst sb e1) (subst sb e2)
| App e1 e2 => App (subst sb e1) (subst sb e2)
| Unit => Unit
| Nat n => Nat n
| Bool b => Bool b
| BinOp op e1 e2 => BinOp op (subst sb e1) (subst sb e2)
| If e0 e1 e2 =>
If (subst sb e0) (subst sb e1) (subst sb e2)
| Pair e1 e2 => Pair (subst sb e1) (subst sb e2)
| Fst e => Fst (subst sb e)
| Snd e => Snd (subst sb e)
| InjL e => InjL (subst sb e)
| InjR e => InjR (subst sb e)
| Case e0 e1 e2 => Case (subst sb e0)
(subst (up sb) e1)
(subst (up sb) e2)
| Fold e => Fold (subst sb e)
| Unfold e => Unfold (subst sb e)
| TLam e => TLam (subst sb e)
| TApp e => TApp (subst sb e)
| CoFork e => CoFork (subst sb e)
| CoYield => CoYield
| Loc l => Loc l
| Alloc e => Alloc (subst sb e)
| Load e => Load (subst sb e)
| Store e1 e2 => Store (subst sb e1) (subst sb e2)
| Cont K => Cont (map (ectx_item_subst sb) K)
| Callcc e => Callcc (subst (up sb) e)
| Throw e1 e2 => Throw (subst sb e1) (subst sb e2)
end
with coopval_subst (sb : var → coopexpr) (v : coopval) : coopval :=
let a := coopexpr_subst : Subst coopexpr in
match v with
| RecV e => RecV (subst (upn 2 sb) e)
| LamV e => LamV (subst (up sb) e)
| TLamV e => TLamV (subst sb e)
| UnitV => UnitV
| NatV n => NatV n
| BoolV b => BoolV b
| PairV v1 v2 => PairV (coopval_subst sb v1) (coopval_subst sb v2)
| InjLV v => InjLV (coopval_subst sb v)
| InjRV v => InjRV (coopval_subst sb v)
| FoldV v => FoldV (coopval_subst sb v)
| LocV l => LocV l
| ContV K => ContV (map (ectx_item_subst sb) K)
end
with ectx_item_subst (sb : var → coopexpr) (K : ectx_item) : ectx_item :=
let a := coopexpr_subst : Subst coopexpr in
match K with
| AppLCtx e2 => AppLCtx (subst sb e2)
| AppRCtx v1 => AppRCtx (coopval_subst sb v1)
| LetInCtx e2 => LetInCtx (subst (up sb) e2)
| SeqCtx e2 => SeqCtx (subst sb e2)
| TAppCtx => TAppCtx
| PairLCtx e2 => PairLCtx (subst sb e2)
| PairRCtx v1 => PairRCtx (coopval_subst sb v1)
| BinOpLCtx op e2 => BinOpLCtx op (subst sb e2)
| BinOpRCtx op v1 => BinOpRCtx op (coopval_subst sb v1)
| FstCtx => FstCtx
| SndCtx => SndCtx
| InjLCtx => InjLCtx
| InjRCtx => InjRCtx
| CaseCtx e1 e2 =>
CaseCtx (subst (up sb) e1) (subst (up sb) e2)
| IfCtx e1 e2 => IfCtx (subst sb e1) (subst sb e2)
| FoldCtx => FoldCtx
| UnfoldCtx => UnfoldCtx
| AllocCtx => AllocCtx
| LoadCtx => LoadCtx
| StoreLCtx e2 => StoreLCtx (subst sb e2)
| StoreRCtx v1 => StoreRCtx (coopval_subst sb v1)
| ThrowLCtx e2 => ThrowLCtx (subst sb e2)
| ThrowRCtx v1 => ThrowRCtx (coopval_subst sb v1)
end.
Instance Subst_coopexpr : Subst coopexpr := coopexpr_subst.
Lemma helper1 : ∀ xi : var → var, up (ren xi) = ren (upren xi).
Proof. intros sb; extensionality z; induction z; simpl; auto. Qed.
Lemma helper2 : up ids = ids.
Proof. extensionality z; induction z; simpl; auto. Qed.
Lemma helper3:
∀ (xi : var → var) (sigma : var → coopexpr) (s : coopexpr),
(rename xi s).[sigma] = s.[xi >>> sigma].
Proof.
intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ (f : var → var) (g : var → coopexpr),
(rename f e).[g] = e.[f >>> g])
(λ v, ∀ (f : var → var) (g : var → coopexpr),
coopval_subst g (coopval_rename f v) = coopval_subst (f >>> g) v));
intros; simpl in *;
repeat match goal with
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), Cont _ = Cont _ |- _ =>
specialize (H f g); inversion H
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), ContV _ = ContV _ |- _ =>
specialize (H f g); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 -?up_comp_ren_subst; auto.
Qed.
Lemma helper4 : ∀ (xi : var → var) (s : coopexpr), rename xi s = s.[ren xi].
Proof.
intros f e; revert f.
apply (coopexpr_rect'
(λ e, ∀ f : var → var, rename f e = e.[ren f])
(λ v, ∀ f : var → var, coopval_rename f v = coopval_subst (ren f) v));
intros; simpl in *;
repeat match goal with
| f : var → var, H : ∀ f : var → var, Cont _ = Cont _ |- _ =>
specialize (H f); inversion H
| f : var → var, H : ∀ f : var → var, ContV _ = ContV _ |- _ =>
specialize (H f); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 ?up_upren_n_internal ?up_upren_internal;
auto using helper1.
Qed.
Lemma helper5 :
∀ (xi : var → var) (sigma : var → coopexpr) (s : coopexpr),
(rename xi s).[sigma] = s.[xi >>> sigma].
Proof.
intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ (f : var → var) (g : var → coopexpr),
(rename f e).[g] = e.[f >>> g])
(λ v, ∀ (f : var → var) (g : var → coopexpr),
coopval_subst g (coopval_rename f v) = coopval_subst (f >>> g) v));
intros; simpl in *;
repeat match goal with
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), Cont _ = Cont _ |- _ =>
specialize (H f g); inversion H
| f : var → var, g : var → coopexpr,
H : ∀ (f : var → var) (g : var → coopexpr), ContV _ = ContV _ |- _ =>
specialize (H f g); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 -?up_comp_ren_subst; auto.
Qed.
Lemma helper6 :
∀ (sigma : var → coopexpr) (xi : var → var) (s : coopexpr),
rename xi s.[sigma] = s.[sigma >>> rename xi].
Proof.
intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ (f : var → coopexpr) (g : var → var),
rename g e.[f] = e.[f >>> rename g])
(λ v, ∀ (f : var → coopexpr) (g : var → var),
coopval_rename g (coopval_subst f v) = coopval_subst (f >>> rename g) v));
intros; simpl in *;
repeat match goal with
| f : var → var, g : var → coopexpr,
H : ∀ (g : var → coopexpr) (f : var → var), Cont _ = Cont _ |- _ =>
specialize (H g f); inversion H
| f : var → var, g : var → coopexpr,
H : ∀ (g : var → coopexpr) (f : var → var), ContV _ = ContV _ |- _ =>
specialize (H g f); inversion H
end;
simpl in *; rewrite ?H ?H0 ?H1 ?up_comp_subst_ren_n_internal
?up_comp_subst_ren_internal;
try apply up_comp_subst_ren_internal; auto using helper4, helper5.
Qed.
Instance SubstLemmas_coopexpr : SubstLemmas coopexpr.
Proof.
econstructor; eauto using helper4.
- apply (coopexpr_rect'
(λ e, e.[ids] = e)
(λ v, coopval_subst ids v = v));
intros; simpl in *;
repeat match goal with
| H : Cont _ = Cont _ |- _ => inversion H; clear H
| H : ContV _ = ContV _ |- _ => inversion H; clear H
end;
rewrite ?up_id_internal ?up_id_n_internal ?H ?H0 ?H1 ?H2 ?H3;
auto using helper2.
- intros f g e; revert f g.
apply (coopexpr_rect'
(λ e, ∀ f g : var → coopexpr, e.[f].[g] = e.[f >> g])
(λ v, ∀ f g : var → coopexpr, coopval_subst g (coopval_subst f v)
= coopval_subst (f >> g) v));
intros; simpl in *;
repeat match goal with
| f : var → coopexpr, g : var → coopexpr,
H : ∀ (f : var → coopexpr) (g : var → coopexpr), Cont _ = Cont _ |- _ =>
let H' := fresh in let H'' := fresh in
pose proof (H f g) as H'; inversion H';
pose proof (H g f) as H''; inversion H''; clear H
| f : var → coopexpr, g : var → coopexpr,
H : ∀ (f : var → coopexpr) (g : var → coopexpr), ContV _ = ContV _ |- _ =>
let H' := fresh in let H'' := fresh in
pose proof (H f g) as H'; inversion H';
pose proof (H g f) as H''; inversion H''; clear H
end;
rewrite ?H ?H0 ?H1 ?H2 ?H3 ?up_comp_internal
?up_comp_n_internal; try apply up_comp_internal;
auto using helper5, helper6.
Qed.
Instance coopexpr_dec_eq : EqDecision coopexpr.
Proof.
intros e.
apply (coopexpr_rect' (λ e, ∀ e' : coopexpr, Decision (e = e'))
(λ w, ∀ v, Decision (w = v))); intros;
match goal with
|- Decision (?A = ?B) => destruct B
end;
try match goal with
|- Decision (Cont ?A = Cont ?B) => destruct B as [|[]]
end;
try match goal with
| |- Decision (ContV (_ :: _) = ContV ?B) => destruct B as [|[]]
| |- Decision (ContV [] = ContV ?B) => destruct B as [|[]]
end; try (right; inversion 1; fail).
all: try match goal with
| |- Decision (?A = ?A) => left; trivial
| |- Decision (?A ?B = ?A ?C) => destruct (decide (B = C));
[| right; inversion 1; tauto]; subst; left; eauto
| |- Decision (?A ?B1 ?B2 = ?A ?C1 ?C2) =>
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (?A ?B1 ?B2 ?B3 = ?A ?C1 ?C2 ?C3) =>
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
destruct (decide (B3 = C3)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (?A ?B1 ?B2 ?B3 = ?A ?C1 ?C2 ?C3) =>
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
destruct (decide (B3 = C3)); [| right; inversion 1; tauto];
subst; left; eauto
end.
all: try match goal with
| |- Decision (Cont (?A :: ?K1) = Cont (?A :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto]; subst; left; eauto
| |- Decision (Cont (?A ?B :: ?K1) = Cont (?A ?C :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B = C)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (Cont (?A ?B1 ?B2 :: ?K1) = Cont (?A ?C1 ?C2 :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
subst; left; eauto
end.
all: try match goal with
| |- Decision (ContV (?A :: ?K1) = ContV (?A :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (ContV K1 = ContV K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto]; subst; left; eauto
| |- Decision (ContV (?A ?B :: ?K1) = ContV (?A ?C :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (ContV K1 = ContV K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B = C)); [| right; inversion 1; tauto];
subst; left; eauto
| |- Decision (ContV (?A ?B1 ?B2 :: ?K1) = ContV (?A ?C1 ?C2 :: ?K2)) =>
let Hf := fresh "H" in
destruct (decide (ContV K1 = ContV K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
destruct (decide (B1 = C1)); [| right; inversion 1; tauto];
destruct (decide (B2 = C2)); [| right; inversion 1; tauto];
subst; left; eauto
end.
match goal with
| |- Decision (ContV ?K1 = ContV ?K2) =>
destruct (decide (Cont K1 = Cont K2)) as [Hf|];
[inversion Hf| right; inversion 1; subst; tauto];
subst; left; eauto
end.
Qed.
Basic properties about the language
Lemma to_of_val v : to_val (of_val v) = Some v.
Proof. by induction v; simplify_option_eq. Qed.
Lemma of_to_val e v : to_val e = Some v → of_val v = e.
Proof.
revert v; induction e; intros; simplify_option_eq; auto with f_equal.
Qed.
Instance of_val_inj : Inj (=) (=) of_val.
Proof. by intros ?? Hv; apply (inj Some); rewrite -!to_of_val Hv. Qed.
Lemma of_val_dec_eq v1 v2 :
Decision (of_val v1 = of_val v2) → Decision (v1 = v2).
Proof.
destruct 1 as [Hv%of_val_inj|Hv]; [left| right => Hw; rewrite Hw in Hv]; done.
Qed.
Instance coopval_dec_eq : EqDecision coopval.
Proof. intros v v'; apply of_val_dec_eq, coopexpr_dec_eq. Qed.
Instance ectx_item_dec_eq : EqDecision ectx_item.
Proof.
intros K K'.
rewrite /Decision; destruct (decide (Cont [K] = Cont [K'])) as [He|Hne].
- inversion He; subst; auto.
- by right; intros HnK; subst.
Qed.
Fixpoint binop_eval (op : binop) (v v' : coopval) : option coopval :=
match v with
| (#nv a) =>
match v' with
| (#nv b) =>
match op with
| Add => Some (#nv (a + b))
| Sub => Some (#nv (a - b))
| Eq => Some (#♭v (bool_decide (a = b)))
| Le => Some (#♭v (bool_decide (a ≤ b)))
| Lt => Some (#♭v (bool_decide (a < b)))
end
| _ => None
end
| _ => None
end.
Instance coopval_inh : Inhabited coopval := populate UnitV.
Definition state : Type := gmap loc coopval.
Inductive head_step :
list ectx_item → coopexpr → state → coopexpr → state → option coopexpr → RedMode → Prop :=
(* rec-β *)
| RecBetaS K e1 e2 v2 σ :
to_val e2 = Some v2 →
head_step K (App (Rec e1) e2) σ e1.[(Rec e1), e2/] σ None NormalMode
(* Lam-β *)
| LamBetaS K e1 e2 v2 σ :
to_val e2 = Some v2 →
head_step K (App (Lam e1) e2) σ e1.[e2/] σ None NormalMode
(* zeta *)
| ZetaS K e1 e2 v2 σ :
to_val e2 = Some v2 →
head_step K (LetIn e2 e1) σ e1.[e2/] σ None NormalMode
(* Seq_β *)
| SeqBetaS K e1 e2 v1 σ :
to_val e1 = Some v1 →
head_step K (Seq e1 e2) σ e2 σ None NormalMode
(* Products *)
| FstS K e1 v1 e2 v2 σ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
head_step K (Fst (Pair e1 e2)) σ e1 σ None NormalMode
| SndS K e1 v1 e2 v2 σ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
head_step K (Snd (Pair e1 e2)) σ e2 σ None NormalMode
(* Sums *)
| CaseLS K e0 v0 e1 e2 σ :
to_val e0 = Some v0 →
head_step K (Case (InjL e0) e1 e2) σ e1.[e0/] σ None NormalMode
| CaseRS K e0 v0 e1 e2 σ :
to_val e0 = Some v0 →
head_step K (Case (InjR e0) e1 e2) σ e2.[e0/] σ None NormalMode
(* nat bin op *)
| BinOpS K op e1 e2 v1 v2 w σ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
binop_eval op v1 v2 = Some w →
head_step K (BinOp op e1 e2) σ (of_val w) σ None NormalMode
(* If then else *)
| IfFalse K e1 e2 σ :
head_step K (If (#♭ false) e1 e2) σ e2 σ None NormalMode
| IfTrue K e1 e2 σ :
head_step K (If (#♭ true) e1 e2) σ e1 σ None NormalMode
(* Recursive Types *)
| Unfold_Fold K e v σ :
to_val e = Some v →
head_step K (Unfold (Fold e)) σ e σ None NormalMode
(* Polymorphic Types *)
| TBeta K e σ :
head_step K (TApp (TLam e)) σ e σ None NormalMode
(* Cooperative Concurrency *)
| ForkS K e σ:
head_step K (CoFork e) σ Unit σ (Some e) NormalMode
| YieldS K σ:
head_step K CoYield σ Unit σ None YieldMode
(* Reference Types *)
| AllocS K e v σ l :
to_val e = Some v → σ !! l = None →
head_step K (Alloc e) σ (Loc l) (<[l:=v]>σ) None NormalMode
| LoadS K l v σ :
σ !! l = Some v →
head_step K (Load (Loc l)) σ (of_val v) σ None NormalMode
| StoreS K l e v σ :
to_val e = Some v → is_Some (σ !! l) →
head_step K (Store (Loc l) e) σ Unit (<[l:=v]>σ) None NormalMode
(* Continuations *)
| CallccS K e σ: head_step K (Callcc e) σ e.[Cont K/] σ None CaptureMode
| ThrowS K K' e v σ:
to_val e = Some v →
head_step K (Throw e (Cont K')) σ
(foldl (flip fill_item) e K') σ None ThrowMode.
Lemma fill_item_val Ki e :
is_Some (to_val (fill_item Ki e)) → is_Some (to_val e).
Proof. intros [v ?]. destruct Ki; simplify_option_eq; eauto. Qed.
Instance fill_item_inj Ki : Inj (=) (=) (fill_item Ki).
Proof. destruct Ki; intros ???; simplify_eq; auto with f_equal. Qed.
Lemma val_stuck K e1 σ1 e2 σ2 ef rm :
head_step K e1 σ1 e2 σ2 ef rm → to_val e1 = None.
Proof. destruct 1; naive_solver. Qed.
Lemma head_ctx_step_val K Ki e σ1 e2 σ2 ef rm :
head_step K (fill_item Ki e) σ1 e2 σ2 ef rm → is_Some (to_val e).
Proof. destruct Ki; inversion_clear 1; simplify_option_eq; eauto. Qed.
Lemma fill_item_no_val_inj Ki1 Ki2 e1 e2 :
to_val e1 = None → to_val e2 = None →
fill_item Ki1 e1 = fill_item Ki2 e2 → Ki1 = Ki2.
Proof.
destruct Ki1, Ki2; intros; try discriminate; simplify_eq;
repeat match goal with
| H : to_val (of_val _) = None |- _ => by rewrite to_of_val in H
end; auto.
Qed.
Lemma red_mode_det K e1 σ1 e2 σ2 efs rm :
head_step K e1 σ1 e2 σ2 efs rm →
∀ K' σ1' e2' σ2' efs' rm',
head_step K' e1 σ1' e2' σ2' efs' rm' → rm' = rm.
Proof. by destruct 1; intros ??????; inversion 1. Qed.
Definition capture K e :=
match e with
| Callcc e' => Some e'.[Cont K/]
| _ => None
end.
Lemma ectxi_capture_captures K e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs CaptureMode →
capture K e1 = Some e2.
Proof. inversion 1; eauto. Qed.
Lemma ectxi_normal_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs NormalMode →
head_step K' e1 σ1 e2 σ2 efs NormalMode.
Proof. inversion 1; econstructor; eauto. Qed.
Lemma ectxi_capture_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs CaptureMode →
∃ e2', capture K' e1 = Some e2' ∧
head_step K' e1 σ1 e2' σ2 efs CaptureMode.
Proof. inversion 1; subst; simpl; eauto using head_step. Qed.
Lemma ectxi_throw_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs ThrowMode →
head_step K' e1 σ1 e2 σ2 efs ThrowMode.
Proof. inversion 1; econstructor; eauto. Qed.
Lemma ectxi_yield_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs YieldMode →
head_step K' e1 σ1 e2 σ2 efs YieldMode.
Proof. inversion 1; econstructor; eauto. Qed.
Lemma alloc_fresh K e v σ :
let l := fresh (dom (gset _) σ) in
to_val e = Some v →
head_step K (Alloc e) σ (Loc l) (<[l:=v]>σ) None NormalMode.
Proof. by intros; apply AllocS, (not_elem_of_dom (D:=gset loc)), is_fresh. Qed.
Canonical Structure stateC := leibnizC state.
Canonical Structure coopvalC := leibnizC coopval.
Canonical Structure coopexprC := leibnizC coopexpr.
End lang.
Proof. by induction v; simplify_option_eq. Qed.
Lemma of_to_val e v : to_val e = Some v → of_val v = e.
Proof.
revert v; induction e; intros; simplify_option_eq; auto with f_equal.
Qed.
Instance of_val_inj : Inj (=) (=) of_val.
Proof. by intros ?? Hv; apply (inj Some); rewrite -!to_of_val Hv. Qed.
Lemma of_val_dec_eq v1 v2 :
Decision (of_val v1 = of_val v2) → Decision (v1 = v2).
Proof.
destruct 1 as [Hv%of_val_inj|Hv]; [left| right => Hw; rewrite Hw in Hv]; done.
Qed.
Instance coopval_dec_eq : EqDecision coopval.
Proof. intros v v'; apply of_val_dec_eq, coopexpr_dec_eq. Qed.
Instance ectx_item_dec_eq : EqDecision ectx_item.
Proof.
intros K K'.
rewrite /Decision; destruct (decide (Cont [K] = Cont [K'])) as [He|Hne].
- inversion He; subst; auto.
- by right; intros HnK; subst.
Qed.
Fixpoint binop_eval (op : binop) (v v' : coopval) : option coopval :=
match v with
| (#nv a) =>
match v' with
| (#nv b) =>
match op with
| Add => Some (#nv (a + b))
| Sub => Some (#nv (a - b))
| Eq => Some (#♭v (bool_decide (a = b)))
| Le => Some (#♭v (bool_decide (a ≤ b)))
| Lt => Some (#♭v (bool_decide (a < b)))
end
| _ => None
end
| _ => None
end.
Instance coopval_inh : Inhabited coopval := populate UnitV.
Definition state : Type := gmap loc coopval.
Inductive head_step :
list ectx_item → coopexpr → state → coopexpr → state → option coopexpr → RedMode → Prop :=
(* rec-β *)
| RecBetaS K e1 e2 v2 σ :
to_val e2 = Some v2 →
head_step K (App (Rec e1) e2) σ e1.[(Rec e1), e2/] σ None NormalMode
(* Lam-β *)
| LamBetaS K e1 e2 v2 σ :
to_val e2 = Some v2 →
head_step K (App (Lam e1) e2) σ e1.[e2/] σ None NormalMode
(* zeta *)
| ZetaS K e1 e2 v2 σ :
to_val e2 = Some v2 →
head_step K (LetIn e2 e1) σ e1.[e2/] σ None NormalMode
(* Seq_β *)
| SeqBetaS K e1 e2 v1 σ :
to_val e1 = Some v1 →
head_step K (Seq e1 e2) σ e2 σ None NormalMode
(* Products *)
| FstS K e1 v1 e2 v2 σ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
head_step K (Fst (Pair e1 e2)) σ e1 σ None NormalMode
| SndS K e1 v1 e2 v2 σ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
head_step K (Snd (Pair e1 e2)) σ e2 σ None NormalMode
(* Sums *)
| CaseLS K e0 v0 e1 e2 σ :
to_val e0 = Some v0 →
head_step K (Case (InjL e0) e1 e2) σ e1.[e0/] σ None NormalMode
| CaseRS K e0 v0 e1 e2 σ :
to_val e0 = Some v0 →
head_step K (Case (InjR e0) e1 e2) σ e2.[e0/] σ None NormalMode
(* nat bin op *)
| BinOpS K op e1 e2 v1 v2 w σ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
binop_eval op v1 v2 = Some w →
head_step K (BinOp op e1 e2) σ (of_val w) σ None NormalMode
(* If then else *)
| IfFalse K e1 e2 σ :
head_step K (If (#♭ false) e1 e2) σ e2 σ None NormalMode
| IfTrue K e1 e2 σ :
head_step K (If (#♭ true) e1 e2) σ e1 σ None NormalMode
(* Recursive Types *)
| Unfold_Fold K e v σ :
to_val e = Some v →
head_step K (Unfold (Fold e)) σ e σ None NormalMode
(* Polymorphic Types *)
| TBeta K e σ :
head_step K (TApp (TLam e)) σ e σ None NormalMode
(* Cooperative Concurrency *)
| ForkS K e σ:
head_step K (CoFork e) σ Unit σ (Some e) NormalMode
| YieldS K σ:
head_step K CoYield σ Unit σ None YieldMode
(* Reference Types *)
| AllocS K e v σ l :
to_val e = Some v → σ !! l = None →
head_step K (Alloc e) σ (Loc l) (<[l:=v]>σ) None NormalMode
| LoadS K l v σ :
σ !! l = Some v →
head_step K (Load (Loc l)) σ (of_val v) σ None NormalMode
| StoreS K l e v σ :
to_val e = Some v → is_Some (σ !! l) →
head_step K (Store (Loc l) e) σ Unit (<[l:=v]>σ) None NormalMode
(* Continuations *)
| CallccS K e σ: head_step K (Callcc e) σ e.[Cont K/] σ None CaptureMode
| ThrowS K K' e v σ:
to_val e = Some v →
head_step K (Throw e (Cont K')) σ
(foldl (flip fill_item) e K') σ None ThrowMode.
Lemma fill_item_val Ki e :
is_Some (to_val (fill_item Ki e)) → is_Some (to_val e).
Proof. intros [v ?]. destruct Ki; simplify_option_eq; eauto. Qed.
Instance fill_item_inj Ki : Inj (=) (=) (fill_item Ki).
Proof. destruct Ki; intros ???; simplify_eq; auto with f_equal. Qed.
Lemma val_stuck K e1 σ1 e2 σ2 ef rm :
head_step K e1 σ1 e2 σ2 ef rm → to_val e1 = None.
Proof. destruct 1; naive_solver. Qed.
Lemma head_ctx_step_val K Ki e σ1 e2 σ2 ef rm :
head_step K (fill_item Ki e) σ1 e2 σ2 ef rm → is_Some (to_val e).
Proof. destruct Ki; inversion_clear 1; simplify_option_eq; eauto. Qed.
Lemma fill_item_no_val_inj Ki1 Ki2 e1 e2 :
to_val e1 = None → to_val e2 = None →
fill_item Ki1 e1 = fill_item Ki2 e2 → Ki1 = Ki2.
Proof.
destruct Ki1, Ki2; intros; try discriminate; simplify_eq;
repeat match goal with
| H : to_val (of_val _) = None |- _ => by rewrite to_of_val in H
end; auto.
Qed.
Lemma red_mode_det K e1 σ1 e2 σ2 efs rm :
head_step K e1 σ1 e2 σ2 efs rm →
∀ K' σ1' e2' σ2' efs' rm',
head_step K' e1 σ1' e2' σ2' efs' rm' → rm' = rm.
Proof. by destruct 1; intros ??????; inversion 1. Qed.
Definition capture K e :=
match e with
| Callcc e' => Some e'.[Cont K/]
| _ => None
end.
Lemma ectxi_capture_captures K e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs CaptureMode →
capture K e1 = Some e2.
Proof. inversion 1; eauto. Qed.
Lemma ectxi_normal_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs NormalMode →
head_step K' e1 σ1 e2 σ2 efs NormalMode.
Proof. inversion 1; econstructor; eauto. Qed.
Lemma ectxi_capture_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs CaptureMode →
∃ e2', capture K' e1 = Some e2' ∧
head_step K' e1 σ1 e2' σ2 efs CaptureMode.
Proof. inversion 1; subst; simpl; eauto using head_step. Qed.
Lemma ectxi_throw_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs ThrowMode →
head_step K' e1 σ1 e2 σ2 efs ThrowMode.
Proof. inversion 1; econstructor; eauto. Qed.
Lemma ectxi_yield_reduciblity K K' e1 σ1 e2 σ2 efs :
head_step K e1 σ1 e2 σ2 efs YieldMode →
head_step K' e1 σ1 e2 σ2 efs YieldMode.
Proof. inversion 1; econstructor; eauto. Qed.
Lemma alloc_fresh K e v σ :
let l := fresh (dom (gset _) σ) in
to_val e = Some v →
head_step K (Alloc e) σ (Loc l) (<[l:=v]>σ) None NormalMode.
Proof. by intros; apply AllocS, (not_elem_of_dom (D:=gset loc)), is_fresh. Qed.
Canonical Structure stateC := leibnizC state.
Canonical Structure coopvalC := leibnizC coopval.
Canonical Structure coopexprC := leibnizC coopexpr.
End lang.
Language
Program Instance heap_ectxi_lang :
CCEctxiLanguage
(lang.coopexpr) lang.coopval lang.ectx_item lang.state := {|
of_val := lang.of_val; to_val := lang.to_val;
fill_item := lang.fill_item; head_step := lang.head_step;
capture := lang.capture
|}.
Solve Obligations with simpl; eauto using lang.to_of_val, lang.of_to_val,
lang.val_stuck, lang.fill_item_val, lang.fill_item_no_val_inj,
lang.head_ctx_step_val, lang.red_mode_det, lang.ectxi_capture_captures,
lang.ectxi_normal_reduciblity, lang.ectxi_throw_reduciblity,
lang.ectxi_capture_reduciblity, lang.ectxi_yield_reduciblity.
Canonical Structure lang := CC_ectx_lang (lang.coopexpr).
Module coop_notations.
Notation coop_fill_item := (@fill_item _ _ _ _ heap_ectxi_lang).
Notation coop_fill := (@fill _ _ _ _ heap_ectxi_lang).
Notation coop_of_val := (@of_val _ _ _ _ heap_ectxi_lang).
Notation coop_to_val := (@to_val _ _ _ _ heap_ectxi_lang).
Notation coop_Var := lang.Var.
Notation coop_Rec := lang.Rec.
Notation coop_Lam := lang.Lam.
Notation coop_LetIn := lang.LetIn.
Notation coop_Seq := lang.Seq.
Notation coop_App := lang.App.
Notation coop_Unit := lang.Unit.
Notation coop_Nat := lang.Nat.
Notation coop_Bool := lang.Bool.
Notation coop_BinOp := lang.BinOp.
Notation coop_If := lang.If.
Notation coop_Pair := lang.Pair.
Notation coop_Fst := lang.Fst.
Notation coop_Snd := lang.Snd.
Notation coop_InjL := lang.InjL.
Notation coop_InjR := lang.InjR.
Notation coop_Case := lang.Case.
Notation coop_Fold := lang.Fold.
Notation coop_Unfold := lang.Unfold.
Notation coop_TLam := lang.TLam.
Notation coop_TApp := lang.TApp.
Notation coop_CoFork := lang.CoFork.
Notation coop_CoYield := lang.CoYield.
Notation coop_Loc := lang.Loc.
Notation coop_Alloc := lang.Alloc.
Notation coop_Load := lang.Load.
Notation coop_Store := lang.Store.
Notation coop_Cont := lang.Cont.
Notation coop_Callcc := lang.Callcc.
Notation coop_Throw := lang.Throw.
Notation coop_AppLCtx := lang.AppLCtx.
Notation coop_AppRCtx := lang.AppRCtx.
Notation coop_LetInCtx := lang.LetInCtx.
Notation coop_SeqCtx := lang.SeqCtx.
Notation coop_TAppCtx := lang.TAppCtx.
Notation coop_PairLCtx := lang.PairLCtx.
Notation coop_PairRCtx := lang.PairRCtx.
Notation coop_BinOpLCtx := lang.BinOpLCtx.
Notation coop_BinOpRCtx := lang.BinOpRCtx.
Notation coop_FstCtx := lang.FstCtx.
Notation coop_SndCtx := lang.SndCtx.
Notation coop_InjLCtx := lang.InjLCtx.
Notation coop_InjRCtx := lang.InjRCtx.
Notation coop_CaseCtx := lang.CaseCtx.
Notation coop_IfCtx := lang.IfCtx.
Notation coop_FoldCtx := lang.FoldCtx.
Notation coop_UnfoldCtx := lang.UnfoldCtx.
Notation coop_AllocCtx := lang.AllocCtx.
Notation coop_LoadCtx := lang.LoadCtx.
Notation coop_StoreLCtx := lang.StoreLCtx.
Notation coop_StoreRCtx := lang.StoreRCtx.
Notation coop_ThrowLCtx := lang.ThrowLCtx.
Notation coop_ThrowRCtx := lang.ThrowRCtx.
Notation coop_RecV := lang.RecV.
Notation coop_LamV := lang.LamV.
Notation coop_TLamV := lang.TLamV.
Notation coop_UnitV := lang.UnitV.
Notation coop_NatV := lang.NatV.
Notation coop_BoolV := lang.BoolV.
Notation coop_PairV := lang.PairV.
Notation coop_InjLV := lang.InjLV.
Notation coop_InjRV := lang.InjRV.
Notation coop_FoldV := lang.FoldV.
Notation coop_LocV := lang.LocV.
Notation coop_ContV := lang.ContV.
End coop_notations.
Export lang.
Definition is_atomic (e : coopexpr) : Prop :=
match e with
| Alloc e => is_Some (to_val e)
| Load e => is_Some (to_val e)
| Store e1 e2 => is_Some (to_val e1) ∧ is_Some (to_val e2)
| _ => False
end.
Lemma is_atomic_sub_values e : is_atomic e → sub_values e.
Proof.
destruct e; inversion 1; simpl; apply ectxi_language_sub_values;
intros [] ?; inversion 1; subst; simpl in *; tauto.
Qed.
Lemma is_atomic_correct e : is_atomic e → Atomic StronglyAtomic e.
Proof.
intros ?; apply ectx_language_atomic; simpl.
- destruct 1; simpl; try done; rewrite ?to_of_val; eauto.
- apply ectxi_language_sub_values => Ki e' Heq; subst; simpl in *.
destruct Ki; naive_solver.
Qed.
Lemma is_atomic_normal e : is_atomic e → is_normal e.
Proof. by destruct e; inversion 1; intros ????; inversion 1; simpl; auto. Qed.
Ltac solve_atomic :=
apply is_atomic_correct; simpl; repeat split;
rewrite ?to_of_val; eapply mk_is_Some; fast_done.
Ltac solve_is_atomic :=
simpl; repeat split; rewrite ?to_of_val; eapply mk_is_Some; fast_done.
Hint Extern 0 (language.atomic _) => solve_atomic.
Hint Extern 0 (language.atomic _) => solve_atomic : typeclass_instances.
Hint Extern 0 (is_atomic _) => solve_is_atomic.
Hint Extern 0 (is_atomic _) => solve_is_atomic : typeclass_instances.
(* helper lemmas for autosubst to treat manually defined substitution well. *)
Lemma ectx_item_subst_simpl f K :
Cont (map (ectx_item_subst f) K) = (Cont K).[f].
Proof. trivial. Qed.
Hint Rewrite ectx_item_subst_simpl : autosubst.
Lemma ectx_item_subst_fold f K :
map (ectx_item_subst f) K = map (ectx_item_subst f) K.
Proof. trivial. Qed.
Hint Rewrite ectx_item_subst_fold : autosubst.
Lemma ectx_item_subst_comp K f g :
map (ectx_item_subst f) (map (ectx_item_subst g) K) =
map (ectx_item_subst (g >> f)) K.
Proof.
match goal with
|- ?A = ?B => assert (Cont A = Cont B) as Hcnt by (by asimpl);
by inversion Hcnt
end.
Qed.
Hint Rewrite ectx_item_subst_comp : autosubst.
Lemma ectx_item_subst_ids K :
map (ectx_item_subst ids) K = K.
Proof.
match goal with
|- ?A = ?B => assert (Cont A = Cont B) as Hcnt by (by asimpl);
by inversion Hcnt as [[Hcnt']]; rewrite Hcnt'
end.
Qed.
Hint Rewrite ectx_item_subst_ids : autosubst.
Hint Extern 1 (IntoVal _ _) =>
rewrite /IntoVal /=; repeat rewrite to_of_val /=.
Hint Extern 1 (IntoVal _ _) =>
rewrite /IntoVal /=; repeat rewrite to_of_val /=: typeclass_instances.
Hint Extern 1 (AsVal _) =>
rewrite /AsVal /=; repeat rewrite to_of_val /=.
Hint Extern 1 (AsVal _) =>
rewrite /AsVal /=; repeat rewrite to_of_val /= : typeclass_instances.
CCEctxiLanguage
(lang.coopexpr) lang.coopval lang.ectx_item lang.state := {|
of_val := lang.of_val; to_val := lang.to_val;
fill_item := lang.fill_item; head_step := lang.head_step;
capture := lang.capture
|}.
Solve Obligations with simpl; eauto using lang.to_of_val, lang.of_to_val,
lang.val_stuck, lang.fill_item_val, lang.fill_item_no_val_inj,
lang.head_ctx_step_val, lang.red_mode_det, lang.ectxi_capture_captures,
lang.ectxi_normal_reduciblity, lang.ectxi_throw_reduciblity,
lang.ectxi_capture_reduciblity, lang.ectxi_yield_reduciblity.
Canonical Structure lang := CC_ectx_lang (lang.coopexpr).
Module coop_notations.
Notation coop_fill_item := (@fill_item _ _ _ _ heap_ectxi_lang).
Notation coop_fill := (@fill _ _ _ _ heap_ectxi_lang).
Notation coop_of_val := (@of_val _ _ _ _ heap_ectxi_lang).
Notation coop_to_val := (@to_val _ _ _ _ heap_ectxi_lang).
Notation coop_Var := lang.Var.
Notation coop_Rec := lang.Rec.
Notation coop_Lam := lang.Lam.
Notation coop_LetIn := lang.LetIn.
Notation coop_Seq := lang.Seq.
Notation coop_App := lang.App.
Notation coop_Unit := lang.Unit.
Notation coop_Nat := lang.Nat.
Notation coop_Bool := lang.Bool.
Notation coop_BinOp := lang.BinOp.
Notation coop_If := lang.If.
Notation coop_Pair := lang.Pair.
Notation coop_Fst := lang.Fst.
Notation coop_Snd := lang.Snd.
Notation coop_InjL := lang.InjL.
Notation coop_InjR := lang.InjR.
Notation coop_Case := lang.Case.
Notation coop_Fold := lang.Fold.
Notation coop_Unfold := lang.Unfold.
Notation coop_TLam := lang.TLam.
Notation coop_TApp := lang.TApp.
Notation coop_CoFork := lang.CoFork.
Notation coop_CoYield := lang.CoYield.
Notation coop_Loc := lang.Loc.
Notation coop_Alloc := lang.Alloc.
Notation coop_Load := lang.Load.
Notation coop_Store := lang.Store.
Notation coop_Cont := lang.Cont.
Notation coop_Callcc := lang.Callcc.
Notation coop_Throw := lang.Throw.
Notation coop_AppLCtx := lang.AppLCtx.
Notation coop_AppRCtx := lang.AppRCtx.
Notation coop_LetInCtx := lang.LetInCtx.
Notation coop_SeqCtx := lang.SeqCtx.
Notation coop_TAppCtx := lang.TAppCtx.
Notation coop_PairLCtx := lang.PairLCtx.
Notation coop_PairRCtx := lang.PairRCtx.
Notation coop_BinOpLCtx := lang.BinOpLCtx.
Notation coop_BinOpRCtx := lang.BinOpRCtx.
Notation coop_FstCtx := lang.FstCtx.
Notation coop_SndCtx := lang.SndCtx.
Notation coop_InjLCtx := lang.InjLCtx.
Notation coop_InjRCtx := lang.InjRCtx.
Notation coop_CaseCtx := lang.CaseCtx.
Notation coop_IfCtx := lang.IfCtx.
Notation coop_FoldCtx := lang.FoldCtx.
Notation coop_UnfoldCtx := lang.UnfoldCtx.
Notation coop_AllocCtx := lang.AllocCtx.
Notation coop_LoadCtx := lang.LoadCtx.
Notation coop_StoreLCtx := lang.StoreLCtx.
Notation coop_StoreRCtx := lang.StoreRCtx.
Notation coop_ThrowLCtx := lang.ThrowLCtx.
Notation coop_ThrowRCtx := lang.ThrowRCtx.
Notation coop_RecV := lang.RecV.
Notation coop_LamV := lang.LamV.
Notation coop_TLamV := lang.TLamV.
Notation coop_UnitV := lang.UnitV.
Notation coop_NatV := lang.NatV.
Notation coop_BoolV := lang.BoolV.
Notation coop_PairV := lang.PairV.
Notation coop_InjLV := lang.InjLV.
Notation coop_InjRV := lang.InjRV.
Notation coop_FoldV := lang.FoldV.
Notation coop_LocV := lang.LocV.
Notation coop_ContV := lang.ContV.
End coop_notations.
Export lang.
Definition is_atomic (e : coopexpr) : Prop :=
match e with
| Alloc e => is_Some (to_val e)
| Load e => is_Some (to_val e)
| Store e1 e2 => is_Some (to_val e1) ∧ is_Some (to_val e2)
| _ => False
end.
Lemma is_atomic_sub_values e : is_atomic e → sub_values e.
Proof.
destruct e; inversion 1; simpl; apply ectxi_language_sub_values;
intros [] ?; inversion 1; subst; simpl in *; tauto.
Qed.
Lemma is_atomic_correct e : is_atomic e → Atomic StronglyAtomic e.
Proof.
intros ?; apply ectx_language_atomic; simpl.
- destruct 1; simpl; try done; rewrite ?to_of_val; eauto.
- apply ectxi_language_sub_values => Ki e' Heq; subst; simpl in *.
destruct Ki; naive_solver.
Qed.
Lemma is_atomic_normal e : is_atomic e → is_normal e.
Proof. by destruct e; inversion 1; intros ????; inversion 1; simpl; auto. Qed.
Ltac solve_atomic :=
apply is_atomic_correct; simpl; repeat split;
rewrite ?to_of_val; eapply mk_is_Some; fast_done.
Ltac solve_is_atomic :=
simpl; repeat split; rewrite ?to_of_val; eapply mk_is_Some; fast_done.
Hint Extern 0 (language.atomic _) => solve_atomic.
Hint Extern 0 (language.atomic _) => solve_atomic : typeclass_instances.
Hint Extern 0 (is_atomic _) => solve_is_atomic.
Hint Extern 0 (is_atomic _) => solve_is_atomic : typeclass_instances.
(* helper lemmas for autosubst to treat manually defined substitution well. *)
Lemma ectx_item_subst_simpl f K :
Cont (map (ectx_item_subst f) K) = (Cont K).[f].
Proof. trivial. Qed.
Hint Rewrite ectx_item_subst_simpl : autosubst.
Lemma ectx_item_subst_fold f K :
map (ectx_item_subst f) K = map (ectx_item_subst f) K.
Proof. trivial. Qed.
Hint Rewrite ectx_item_subst_fold : autosubst.
Lemma ectx_item_subst_comp K f g :
map (ectx_item_subst f) (map (ectx_item_subst g) K) =
map (ectx_item_subst (g >> f)) K.
Proof.
match goal with
|- ?A = ?B => assert (Cont A = Cont B) as Hcnt by (by asimpl);
by inversion Hcnt
end.
Qed.
Hint Rewrite ectx_item_subst_comp : autosubst.
Lemma ectx_item_subst_ids K :
map (ectx_item_subst ids) K = K.
Proof.
match goal with
|- ?A = ?B => assert (Cont A = Cont B) as Hcnt by (by asimpl);
by inversion Hcnt as [[Hcnt']]; rewrite Hcnt'
end.
Qed.
Hint Rewrite ectx_item_subst_ids : autosubst.
Hint Extern 1 (IntoVal _ _) =>
rewrite /IntoVal /=; repeat rewrite to_of_val /=.
Hint Extern 1 (IntoVal _ _) =>
rewrite /IntoVal /=; repeat rewrite to_of_val /=: typeclass_instances.
Hint Extern 1 (AsVal _) =>
rewrite /AsVal /=; repeat rewrite to_of_val /=.
Hint Extern 1 (AsVal _) =>
rewrite /AsVal /=; repeat rewrite to_of_val /= : typeclass_instances.