LogrelCC.examples.refinement.list_basics
From iris.proofmode Require Import tactics.
From LogrelCC Require Export rules_unary rules_binary typing lang cl_rules.
From LogrelCC.examples Require Import list_rev.
From iris.base_logic Require Import invariants.
Definition list_length : expr :=
TLam (Rec (Case (Unfold (Var 1))
(BinOp Add (#n 1) (App (Var 1) (Snd (Var 0)))) (#n 0))).
Definition list_snoc : expr :=
TLam (Rec
(Lam (Case (Unfold (Var 0))
(Fold (InjL (Pair (Fst (Var 0))
(App (App (Var 2) (Var 3)) (Snd (Var 0))))))
(Fold (InjL (Pair (Var 3) (Fold (InjR Unit)))))))).
Definition list_head : expr :=
TLam (Lam (Case (Unfold (Var 0)) (InjR (Fst (Var 0))) (InjL Unit))).
Definition list_tail : expr :=
TLam (Lam (Case (Unfold (Var 0)) (Snd (Var 0)) (Fold (InjR Unit)))).
Definition list_get : expr :=
TLam
(Rec
(Lam
(If
(BinOp Eq (Var 2) (#n 0))
(App (TApp list_head) (Var 0))
((App (App (Var 1)
(BinOp Sub (Var 2) (#n 1)))
(App (TApp list_tail) (Var 0))))))).
Lemma list_length_closed f : list_length.[f] = list_length.
Proof. by asimpl. Qed.
Hint Rewrite list_length_closed : autosubst.
Lemma list_snoc_closed f : list_snoc.[f] = list_snoc.
Proof. by asimpl. Qed.
Hint Rewrite list_snoc_closed : autosubst.
Lemma list_head_closed f : list_head.[f] = list_head.
Proof. by asimpl. Qed.
Hint Rewrite list_head_closed : autosubst.
Lemma list_tail_closed f : list_tail.[f] = list_tail.
Proof. by asimpl. Qed.
Hint Rewrite list_tail_closed : autosubst.
Lemma list_get_closed f : list_get.[f] = list_get.
Proof. by asimpl. Qed.
Hint Rewrite list_get_closed : autosubst.
Definition ListType τ := TRec (TSum (TProd τ.[ren (+1)] (TVar 0)) TUnit).
Lemma list_length_type Γ :
typed Γ list_length (TForall (TArrow (ListType (TVar 0)) TNat)).
Proof.
do 2 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_snoc_type Γ :
typed Γ list_snoc
(TForall (TArrow (TVar 0)
(TArrow (ListType (TVar 0)) (ListType (TVar 0))))).
Proof.
do 3 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_head_type Γ :
typed Γ list_head (TForall (TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0)))).
Proof.
do 2 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_tail_type Γ :
typed Γ list_tail (TForall (TArrow (ListType (TVar 0)) (ListType (TVar 0)))).
Proof.
do 2 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_get_type Γ :
typed Γ list_get
(TForall (TArrow TNat
(TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0))))).
Proof.
repeat econstructor.
- replace (TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0))) with
(TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0))).[(TVar 0)/];
last by asimpl.
econstructor. apply list_head_type.
- replace (TArrow (ListType (TVar 0)) (ListType (TVar 0))) with
(TArrow (ListType (TVar 0)) (ListType (TVar 0))).[(TVar 0)/];
last by asimpl.
econstructor. apply list_tail_type.
Qed.
Section list_basics.
Context `{cfgSG Σ}.
Context `{heapG Σ}.
Lemma clwp_list_length E l :
True ⊢ CLWP App (TApp list_length) (of_val (embed_list l)) @ E
{{v, ⌜v = #nv (length l)⌝ }}.
iIntros "_".
iApply (clwp_bind [AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iLöb as "IH" forall (l).
iApply clwp_rec; eauto using to_of_val. iNext. asimpl.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]); iApply clwp_unfold; eauto; simpl.
iNext; iApply clwp_value; eauto; simpl. iApply clwp_case_injr; eauto.
iNext; simpl. iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto; simpl.
iNext. iApply clwp_value; auto.
simpl; iApply clwp_case_injl; eauto.
iNext. iApply (clwp_bind [BinOpRCtx _ (#nv _)]). asimpl.
iApply (clwp_bind [AppRCtx (RecV _)]).
iApply clwp_snd; eauto using to_of_val.
iNext. iApply clwp_value; eauto using to_of_val.
iApply clwp_wand_r; iSplitL; first by iApply "IH".
iIntros (w Hw); subst. iApply clwp_bin_op; eauto.
iNext. iApply clwp_value; eauto.
Qed.
Lemma clwp_list_snoc E v l :
True ⊢ CLWP App (App (TApp list_snoc) (of_val v)) (of_val (embed_list l))
@E {{w, ⌜w = embed_list (l ++ [v])⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _; AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply (clwp_bind [AppLCtx _]). iApply clwp_rec; eauto using to_of_val.
asimpl. iNext. iApply clwp_value; eauto. simpl.
iLöb as "IH" forall (l).
iApply clwp_Lam; eauto using to_of_val.
iNext. asimpl.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; eauto. simpl.
iApply clwp_case_injr; eauto.
iNext. asimpl.
by iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; auto. simpl.
iApply clwp_case_injl; eauto.
iNext. asimpl.
iApply (clwp_bind [PairLCtx _; InjLCtx; FoldCtx]).
iApply clwp_fst; eauto using to_of_val.
iNext; iApply clwp_value; eauto using to_of_val.
iApply (clwp_bind [AppLCtx _; PairRCtx _; InjLCtx; FoldCtx]).
iApply clwp_rec; eauto using to_of_val.
asimpl. iNext. iApply clwp_value; eauto.
iApply (clwp_bind [AppRCtx _; PairRCtx _; InjLCtx; FoldCtx]).
iApply clwp_snd; eauto using to_of_val.
iApply clwp_value; eauto using to_of_val.
iNext; simpl. iApply (clwp_bind [PairRCtx _; InjLCtx; FoldCtx]).
iApply clwp_wand_r; iSplitL; first by iApply "IH".
iIntros (w Hw); subst. iApply clwp_value; eauto.
Qed.
Lemma clwp_list_head E l :
True ⊢ CLWP App (TApp list_head) (of_val (embed_list l))
@E {{w, ⌜match head l with
None => w = InjLV UnitV
| Some v => w = InjRV v
end⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply clwp_Lam; eauto. asimpl.
iNext.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; eauto. simpl.
iApply clwp_case_injr; eauto.
iNext. asimpl.
iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; auto. iNext.
iApply clwp_value; auto. simpl.
iApply clwp_case_injl; eauto.
iNext. asimpl.
iApply (clwp_bind [InjRCtx]).
iApply clwp_fst; eauto using to_of_val.
iNext; iApply clwp_value; eauto using to_of_val.
iApply clwp_value; eauto.
Qed.
Lemma clwp_list_tail E l :
True ⊢ CLWP App (TApp list_tail) (of_val (embed_list l))
@E {{w, ⌜w = embed_list (tail l)⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply clwp_Lam; eauto. asimpl.
iNext.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; eauto. simpl.
iApply clwp_case_injr; eauto.
iNext. asimpl.
iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; auto. iNext.
iApply clwp_value; auto. simpl.
iApply clwp_case_injl; eauto.
iNext. asimpl.
iApply clwp_snd; eauto.
iNext; iApply clwp_value; eauto.
Qed.
Lemma clwp_list_get E i l :
True ⊢ CLWP App (App (TApp list_get) (#n i)) (of_val (embed_list l))
@E {{w, ⌜match l !! i with
None => w = InjLV UnitV
| Some v => w = InjRV v
end⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _; AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply (clwp_bind [AppLCtx _]). iApply clwp_rec; eauto.
asimpl. iNext. iApply clwp_value; eauto. simpl.
iLöb as "IH" forall (i l).
iApply clwp_Lam; eauto. iNext. asimpl.
destruct i; simpl.
- iApply (clwp_bind [IfCtx _ _]). iApply clwp_bin_op; eauto.
simpl. iApply clwp_value; eauto. iNext. iApply clwp_if_true.
iNext. rewrite-/list_head.
iApply clwp_wand_r; iSplitL; first by iApply (clwp_list_head).
destruct l; simpl; auto.
- iApply (clwp_bind [IfCtx _ _]). iApply clwp_bin_op; eauto.
simpl. iApply clwp_value; eauto. iNext. iApply clwp_if_false.
iNext. rewrite-/list_tail.
iApply (clwp_bind [AppRCtx (RecV _); AppLCtx _]).
iApply clwp_bin_op; eauto.
simpl. replace (i - 0) with i by lia.
iNext; iApply clwp_value; eauto. simpl.
iApply (clwp_bind [AppLCtx _]).
iApply clwp_rec; eauto.
asimpl. iNext. rewrite - /list_head -/list_tail.
iApply clwp_value; eauto. simpl.
destruct l; simpl.
+ iApply (clwp_bind [AppRCtx (LamV _)]).
iApply clwp_wand_r; iSplitL;
first by iApply (clwp_list_tail _ []).
iIntros (w ->); simpl.
change (Fold (InjR Unit)) with (of_val (embed_list [])).
iApply clwp_wand_r; iSplitL;
first iApply "IH"; auto.
+ iApply (clwp_bind [AppRCtx (LamV _)]).
iApply clwp_wand_r; iSplitL;
first by iApply (clwp_list_tail _ (_ :: _)).
iIntros (w ->); simpl. iApply "IH".
Qed.
Lemma steps_list_length E ρ j K l :
nclose specN ⊆ E →
spec_ctx ρ ∗ j ⤇ fill K (App (TApp list_length) (of_val (embed_list l)))
⊢ |={E}=> j ⤇ fill K (#n (length l)).
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj"; auto.
simpl.
iInduction l as [|a l] "IHl" forall (K) "Hctx Hj".
- iMod (step_rec with "[$Hj]") as "Hj"; auto. asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
iMod (step_case_inr with "[$Hj]") as "Hj"; auto.
- iMod (step_rec with "[$Hj]") as "Hj"; auto.
asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto. asimpl.
iMod (step_snd _ _ _ (AppRCtx (RecV _) :: BinOpRCtx _ (#nv _) :: _)
with "[$Hj]") as "Hj"; simpl; eauto using to_of_val.
iMod ("IHl" $! (BinOpRCtx _ (#nv _) :: _) with "[] [Hj]") as "Hj"; eauto.
iMod (step_nat_binop with "[$Hj]") as "Hj"; eauto.
Qed.
Lemma step_list_snoc E ρ j K v l :
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (App (TApp list_snoc) (of_val v)) (of_val (embed_list l)))
⊢ |={E}=> j ⤇ fill K (of_val (embed_list (l ++ [v]))).
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto. simpl. asimpl.
iInduction l as [|a l] "IHl" forall (K) "Hctx Hj".
- iMod (step_Lam with "[$Hj]") as "Hj"; eauto. asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; eauto.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto.
by asimpl.
- iMod (step_Lam with "[$Hj]") as "Hj"; eauto. asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
simpl.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto. simpl.
iMod (step_fst _ _ _ (PairLCtx _ :: InjLCtx :: FoldCtx :: _) with "[$Hj]")
as "Hj"; eauto using to_of_val; simpl.
asimpl.
iMod (step_rec _ _ _ (AppLCtx _ :: PairRCtx _ :: InjLCtx :: FoldCtx :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val.
asimpl.
iMod (step_snd _ _ _
(AppRCtx (LamV _) :: PairRCtx _ :: InjLCtx :: FoldCtx :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val; simpl.
iMod ("IHl" $! (PairRCtx _ :: InjLCtx :: FoldCtx :: _) with "[] [Hj]")
as "Hj"; eauto.
Qed.
Lemma step_list_head E ρ j K l:
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (TApp list_head) (of_val (embed_list l)))
⊢ |={E}=> ∃ w, j ⤇ fill K (of_val w) ∗
⌜match head l with
None => w = InjLV UnitV
| Some v => w = InjRV v
end⌝.
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_Lam with "[$Hj]") as "Hj"; eauto using to_of_val.
asimpl.
destruct l; simpl.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; eauto.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto. simpl.
iModIntro; iExists (InjLV UnitV); auto.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto.
asimpl.
iMod (step_fst _ _ _ (InjRCtx :: _) with "[$Hj]") as "Hj";
eauto using to_of_val.
iExists (InjRV _); eauto.
Qed.
Lemma step_list_tail E ρ j K l :
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (TApp list_tail) (of_val (embed_list l)))
⊢ |={E}=> j ⤇ fill K (of_val (embed_list (tail l))).
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_Lam with "[$Hj]") as "Hj";
auto; simpl.
destruct l; simpl.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_case_inr with "[$Hj]") as "Hj"; auto.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto;
simpl.
iMod (step_snd with "[$Hj]") as "Hj"; eauto.
Qed.
Lemma step_list_get E ρ j K i l :
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (App (TApp list_get) (#n i)) (of_val (embed_list l)))
⊢ |={E}=> ∃ w, j ⤇ fill K (of_val w) ∗
⌜match l !! i with
None => w = InjLV UnitV
| Some v =>w = InjRV v
end⌝.
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto.
simpl. asimpl.
iInduction l as [|a l] "IHl" forall (i K) "Hctx Hj".
- iInduction i as [|i] "IHl" forall (K) "Hctx Hj"; simpl in *.
+ iMod (step_Lam with "[$Hj]") as "Hj"; eauto using to_of_val. asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val; simpl.
iMod (step_if_true with "[$Hj]") as "Hj"; auto.
rewrite -/list_head.
iMod (step_list_head _ _ _ _ [] with "[$Hj]") as "Hj"; auto.
+ iMod (step_Lam with "[$Hj]") as "Hj"; eauto using to_of_val. asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _) with "[$Hj]") as "Hj"; eauto.
iMod (step_if_false with "[$Hj]") as "Hj"; auto.
rewrite -/list_head -/list_tail.
iMod (step_nat_binop _ _ _ (AppRCtx (RecV _) :: AppLCtx _ :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val; simpl.
replace (i - 0) with i by lia.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val. simpl. asimpl.
rewrite -/list_head -/list_tail.
iMod (step_list_tail _ _ _ (AppRCtx (LamV _) :: _) [] with "[$Hj]")
as "Hj"; auto; simpl.
iMod ("IHl" with "[] [Hj]") as "Hj"; auto.
- destruct i as [|i]; simpl.
+ iMod (step_Lam with "[$Hj]") as "Hj"; auto.
asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val; simpl.
iMod (step_if_true with "[$Hj]") as "Hj"; auto.
rewrite -/list_head -/list_tail.
iMod (step_list_head _ _ _ _ (_ :: _) with "[$Hj]")
as "Hj"; auto; simpl.
+ rewrite -/list_head -/list_tail.
iMod (step_Lam with "[$Hj]") as "Hj"; auto. asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _)
with "[$Hj]") as "Hj"; eauto. simpl.
iMod (step_if_false with "[$Hj]") as "Hj"; auto.
iMod (step_nat_binop _ _ _ (AppRCtx (RecV _) :: AppLCtx _ :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val; simpl.
replace (i - 0) with i by lia.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val. simpl. asimpl.
rewrite -/list_head -/list_tail.
iMod (step_list_tail _ _ _ (AppRCtx (LamV _) :: _) (_ :: _)
with "[$Hj]") as "Hj"; auto; simpl.
iMod ("IHl" with "[] [Hj]") as "Hj"; auto.
Qed.
End list_basics.
Typeclasses Opaque list_length.
Global Opaque list_length.
Typeclasses Opaque list_snoc.
Global Opaque list_snoc.
Typeclasses Opaque list_head.
Global Opaque list_head.
Typeclasses Opaque list_tail.
Global Opaque list_tail.
Typeclasses Opaque list_get.
Global Opaque list_get.
From LogrelCC Require Export rules_unary rules_binary typing lang cl_rules.
From LogrelCC.examples Require Import list_rev.
From iris.base_logic Require Import invariants.
Definition list_length : expr :=
TLam (Rec (Case (Unfold (Var 1))
(BinOp Add (#n 1) (App (Var 1) (Snd (Var 0)))) (#n 0))).
Definition list_snoc : expr :=
TLam (Rec
(Lam (Case (Unfold (Var 0))
(Fold (InjL (Pair (Fst (Var 0))
(App (App (Var 2) (Var 3)) (Snd (Var 0))))))
(Fold (InjL (Pair (Var 3) (Fold (InjR Unit)))))))).
Definition list_head : expr :=
TLam (Lam (Case (Unfold (Var 0)) (InjR (Fst (Var 0))) (InjL Unit))).
Definition list_tail : expr :=
TLam (Lam (Case (Unfold (Var 0)) (Snd (Var 0)) (Fold (InjR Unit)))).
Definition list_get : expr :=
TLam
(Rec
(Lam
(If
(BinOp Eq (Var 2) (#n 0))
(App (TApp list_head) (Var 0))
((App (App (Var 1)
(BinOp Sub (Var 2) (#n 1)))
(App (TApp list_tail) (Var 0))))))).
Lemma list_length_closed f : list_length.[f] = list_length.
Proof. by asimpl. Qed.
Hint Rewrite list_length_closed : autosubst.
Lemma list_snoc_closed f : list_snoc.[f] = list_snoc.
Proof. by asimpl. Qed.
Hint Rewrite list_snoc_closed : autosubst.
Lemma list_head_closed f : list_head.[f] = list_head.
Proof. by asimpl. Qed.
Hint Rewrite list_head_closed : autosubst.
Lemma list_tail_closed f : list_tail.[f] = list_tail.
Proof. by asimpl. Qed.
Hint Rewrite list_tail_closed : autosubst.
Lemma list_get_closed f : list_get.[f] = list_get.
Proof. by asimpl. Qed.
Hint Rewrite list_get_closed : autosubst.
Definition ListType τ := TRec (TSum (TProd τ.[ren (+1)] (TVar 0)) TUnit).
Lemma list_length_type Γ :
typed Γ list_length (TForall (TArrow (ListType (TVar 0)) TNat)).
Proof.
do 2 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_snoc_type Γ :
typed Γ list_snoc
(TForall (TArrow (TVar 0)
(TArrow (ListType (TVar 0)) (ListType (TVar 0))))).
Proof.
do 3 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_head_type Γ :
typed Γ list_head (TForall (TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0)))).
Proof.
do 2 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_tail_type Γ :
typed Γ list_tail (TForall (TArrow (ListType (TVar 0)) (ListType (TVar 0)))).
Proof.
do 2 econstructor.
eapply (Case_typed _ _ _ _ (TProd (TVar 0) (ListType (TVar 0))) TUnit);
repeat econstructor.
replace (TSum (TProd (TVar 0) (ListType (TVar 0))) TUnit) with
(TSum (TProd (TVar 1) (TVar 0)) TUnit).[(ListType (TVar 0))/];
last by asimpl.
repeat econstructor.
Qed.
Lemma list_get_type Γ :
typed Γ list_get
(TForall (TArrow TNat
(TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0))))).
Proof.
repeat econstructor.
- replace (TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0))) with
(TArrow (ListType (TVar 0)) (TSum TUnit (TVar 0))).[(TVar 0)/];
last by asimpl.
econstructor. apply list_head_type.
- replace (TArrow (ListType (TVar 0)) (ListType (TVar 0))) with
(TArrow (ListType (TVar 0)) (ListType (TVar 0))).[(TVar 0)/];
last by asimpl.
econstructor. apply list_tail_type.
Qed.
Section list_basics.
Context `{cfgSG Σ}.
Context `{heapG Σ}.
Lemma clwp_list_length E l :
True ⊢ CLWP App (TApp list_length) (of_val (embed_list l)) @ E
{{v, ⌜v = #nv (length l)⌝ }}.
iIntros "_".
iApply (clwp_bind [AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iLöb as "IH" forall (l).
iApply clwp_rec; eauto using to_of_val. iNext. asimpl.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]); iApply clwp_unfold; eauto; simpl.
iNext; iApply clwp_value; eauto; simpl. iApply clwp_case_injr; eauto.
iNext; simpl. iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto; simpl.
iNext. iApply clwp_value; auto.
simpl; iApply clwp_case_injl; eauto.
iNext. iApply (clwp_bind [BinOpRCtx _ (#nv _)]). asimpl.
iApply (clwp_bind [AppRCtx (RecV _)]).
iApply clwp_snd; eauto using to_of_val.
iNext. iApply clwp_value; eauto using to_of_val.
iApply clwp_wand_r; iSplitL; first by iApply "IH".
iIntros (w Hw); subst. iApply clwp_bin_op; eauto.
iNext. iApply clwp_value; eauto.
Qed.
Lemma clwp_list_snoc E v l :
True ⊢ CLWP App (App (TApp list_snoc) (of_val v)) (of_val (embed_list l))
@E {{w, ⌜w = embed_list (l ++ [v])⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _; AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply (clwp_bind [AppLCtx _]). iApply clwp_rec; eauto using to_of_val.
asimpl. iNext. iApply clwp_value; eauto. simpl.
iLöb as "IH" forall (l).
iApply clwp_Lam; eauto using to_of_val.
iNext. asimpl.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; eauto. simpl.
iApply clwp_case_injr; eauto.
iNext. asimpl.
by iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; auto. simpl.
iApply clwp_case_injl; eauto.
iNext. asimpl.
iApply (clwp_bind [PairLCtx _; InjLCtx; FoldCtx]).
iApply clwp_fst; eauto using to_of_val.
iNext; iApply clwp_value; eauto using to_of_val.
iApply (clwp_bind [AppLCtx _; PairRCtx _; InjLCtx; FoldCtx]).
iApply clwp_rec; eauto using to_of_val.
asimpl. iNext. iApply clwp_value; eauto.
iApply (clwp_bind [AppRCtx _; PairRCtx _; InjLCtx; FoldCtx]).
iApply clwp_snd; eauto using to_of_val.
iApply clwp_value; eauto using to_of_val.
iNext; simpl. iApply (clwp_bind [PairRCtx _; InjLCtx; FoldCtx]).
iApply clwp_wand_r; iSplitL; first by iApply "IH".
iIntros (w Hw); subst. iApply clwp_value; eauto.
Qed.
Lemma clwp_list_head E l :
True ⊢ CLWP App (TApp list_head) (of_val (embed_list l))
@E {{w, ⌜match head l with
None => w = InjLV UnitV
| Some v => w = InjRV v
end⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply clwp_Lam; eauto. asimpl.
iNext.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; eauto. simpl.
iApply clwp_case_injr; eauto.
iNext. asimpl.
iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; auto. iNext.
iApply clwp_value; auto. simpl.
iApply clwp_case_injl; eauto.
iNext. asimpl.
iApply (clwp_bind [InjRCtx]).
iApply clwp_fst; eauto using to_of_val.
iNext; iApply clwp_value; eauto using to_of_val.
iApply clwp_value; eauto.
Qed.
Lemma clwp_list_tail E l :
True ⊢ CLWP App (TApp list_tail) (of_val (embed_list l))
@E {{w, ⌜w = embed_list (tail l)⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply clwp_Lam; eauto. asimpl.
iNext.
destruct l; simpl.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; eauto. iNext.
iApply clwp_value; eauto. simpl.
iApply clwp_case_injr; eauto.
iNext. asimpl.
iApply clwp_value; eauto.
- iApply (clwp_bind [CaseCtx _ _]).
iApply clwp_unfold; auto. iNext.
iApply clwp_value; auto. simpl.
iApply clwp_case_injl; eauto.
iNext. asimpl.
iApply clwp_snd; eauto.
iNext; iApply clwp_value; eauto.
Qed.
Lemma clwp_list_get E i l :
True ⊢ CLWP App (App (TApp list_get) (#n i)) (of_val (embed_list l))
@E {{w, ⌜match l !! i with
None => w = InjLV UnitV
| Some v => w = InjRV v
end⌝ }}.
Proof.
iIntros "_".
iApply (clwp_bind [AppLCtx _; AppLCtx _]); iApply clwp_tapp.
iNext; iApply clwp_value; eauto; simpl.
iApply (clwp_bind [AppLCtx _]). iApply clwp_rec; eauto.
asimpl. iNext. iApply clwp_value; eauto. simpl.
iLöb as "IH" forall (i l).
iApply clwp_Lam; eauto. iNext. asimpl.
destruct i; simpl.
- iApply (clwp_bind [IfCtx _ _]). iApply clwp_bin_op; eauto.
simpl. iApply clwp_value; eauto. iNext. iApply clwp_if_true.
iNext. rewrite-/list_head.
iApply clwp_wand_r; iSplitL; first by iApply (clwp_list_head).
destruct l; simpl; auto.
- iApply (clwp_bind [IfCtx _ _]). iApply clwp_bin_op; eauto.
simpl. iApply clwp_value; eauto. iNext. iApply clwp_if_false.
iNext. rewrite-/list_tail.
iApply (clwp_bind [AppRCtx (RecV _); AppLCtx _]).
iApply clwp_bin_op; eauto.
simpl. replace (i - 0) with i by lia.
iNext; iApply clwp_value; eauto. simpl.
iApply (clwp_bind [AppLCtx _]).
iApply clwp_rec; eauto.
asimpl. iNext. rewrite - /list_head -/list_tail.
iApply clwp_value; eauto. simpl.
destruct l; simpl.
+ iApply (clwp_bind [AppRCtx (LamV _)]).
iApply clwp_wand_r; iSplitL;
first by iApply (clwp_list_tail _ []).
iIntros (w ->); simpl.
change (Fold (InjR Unit)) with (of_val (embed_list [])).
iApply clwp_wand_r; iSplitL;
first iApply "IH"; auto.
+ iApply (clwp_bind [AppRCtx (LamV _)]).
iApply clwp_wand_r; iSplitL;
first by iApply (clwp_list_tail _ (_ :: _)).
iIntros (w ->); simpl. iApply "IH".
Qed.
Lemma steps_list_length E ρ j K l :
nclose specN ⊆ E →
spec_ctx ρ ∗ j ⤇ fill K (App (TApp list_length) (of_val (embed_list l)))
⊢ |={E}=> j ⤇ fill K (#n (length l)).
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj"; auto.
simpl.
iInduction l as [|a l] "IHl" forall (K) "Hctx Hj".
- iMod (step_rec with "[$Hj]") as "Hj"; auto. asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
iMod (step_case_inr with "[$Hj]") as "Hj"; auto.
- iMod (step_rec with "[$Hj]") as "Hj"; auto.
asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto. asimpl.
iMod (step_snd _ _ _ (AppRCtx (RecV _) :: BinOpRCtx _ (#nv _) :: _)
with "[$Hj]") as "Hj"; simpl; eauto using to_of_val.
iMod ("IHl" $! (BinOpRCtx _ (#nv _) :: _) with "[] [Hj]") as "Hj"; eauto.
iMod (step_nat_binop with "[$Hj]") as "Hj"; eauto.
Qed.
Lemma step_list_snoc E ρ j K v l :
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (App (TApp list_snoc) (of_val v)) (of_val (embed_list l)))
⊢ |={E}=> j ⤇ fill K (of_val (embed_list (l ++ [v]))).
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto. simpl. asimpl.
iInduction l as [|a l] "IHl" forall (K) "Hctx Hj".
- iMod (step_Lam with "[$Hj]") as "Hj"; eauto. asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; eauto.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto.
by asimpl.
- iMod (step_Lam with "[$Hj]") as "Hj"; eauto. asimpl.
iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
simpl.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto. simpl.
iMod (step_fst _ _ _ (PairLCtx _ :: InjLCtx :: FoldCtx :: _) with "[$Hj]")
as "Hj"; eauto using to_of_val; simpl.
asimpl.
iMod (step_rec _ _ _ (AppLCtx _ :: PairRCtx _ :: InjLCtx :: FoldCtx :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val.
asimpl.
iMod (step_snd _ _ _
(AppRCtx (LamV _) :: PairRCtx _ :: InjLCtx :: FoldCtx :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val; simpl.
iMod ("IHl" $! (PairRCtx _ :: InjLCtx :: FoldCtx :: _) with "[] [Hj]")
as "Hj"; eauto.
Qed.
Lemma step_list_head E ρ j K l:
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (TApp list_head) (of_val (embed_list l)))
⊢ |={E}=> ∃ w, j ⤇ fill K (of_val w) ∗
⌜match head l with
None => w = InjLV UnitV
| Some v => w = InjRV v
end⌝.
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_Lam with "[$Hj]") as "Hj"; eauto using to_of_val.
asimpl.
destruct l; simpl.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; eauto.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto. simpl.
iModIntro; iExists (InjLV UnitV); auto.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj"; auto.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto.
asimpl.
iMod (step_fst _ _ _ (InjRCtx :: _) with "[$Hj]") as "Hj";
eauto using to_of_val.
iExists (InjRV _); eauto.
Qed.
Lemma step_list_tail E ρ j K l :
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (TApp list_tail) (of_val (embed_list l)))
⊢ |={E}=> j ⤇ fill K (of_val (embed_list (tail l))).
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_Lam with "[$Hj]") as "Hj";
auto; simpl.
destruct l; simpl.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_case_inr with "[$Hj]") as "Hj"; auto.
- iMod (step_Fold _ _ _ (CaseCtx _ _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_case_inl with "[$Hj]") as "Hj"; auto;
simpl.
iMod (step_snd with "[$Hj]") as "Hj"; eauto.
Qed.
Lemma step_list_get E ρ j K i l :
nclose specN ⊆ E →
spec_ctx ρ ∗
j ⤇ fill K (App (App (TApp list_get) (#n i)) (of_val (embed_list l)))
⊢ |={E}=> ∃ w, j ⤇ fill K (of_val w) ∗
⌜match l !! i with
None => w = InjLV UnitV
| Some v =>w = InjRV v
end⌝.
Proof.
iIntros (HE) "[#Hctx Hj]".
iMod (step_tlam _ _ _ (AppLCtx _ :: AppLCtx _ :: _) with "[$Hj]") as "Hj";
auto; simpl.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto.
simpl. asimpl.
iInduction l as [|a l] "IHl" forall (i K) "Hctx Hj".
- iInduction i as [|i] "IHl" forall (K) "Hctx Hj"; simpl in *.
+ iMod (step_Lam with "[$Hj]") as "Hj"; eauto using to_of_val. asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val; simpl.
iMod (step_if_true with "[$Hj]") as "Hj"; auto.
rewrite -/list_head.
iMod (step_list_head _ _ _ _ [] with "[$Hj]") as "Hj"; auto.
+ iMod (step_Lam with "[$Hj]") as "Hj"; eauto using to_of_val. asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _) with "[$Hj]") as "Hj"; eauto.
iMod (step_if_false with "[$Hj]") as "Hj"; auto.
rewrite -/list_head -/list_tail.
iMod (step_nat_binop _ _ _ (AppRCtx (RecV _) :: AppLCtx _ :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val; simpl.
replace (i - 0) with i by lia.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val. simpl. asimpl.
rewrite -/list_head -/list_tail.
iMod (step_list_tail _ _ _ (AppRCtx (LamV _) :: _) [] with "[$Hj]")
as "Hj"; auto; simpl.
iMod ("IHl" with "[] [Hj]") as "Hj"; auto.
- destruct i as [|i]; simpl.
+ iMod (step_Lam with "[$Hj]") as "Hj"; auto.
asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val; simpl.
iMod (step_if_true with "[$Hj]") as "Hj"; auto.
rewrite -/list_head -/list_tail.
iMod (step_list_head _ _ _ _ (_ :: _) with "[$Hj]")
as "Hj"; auto; simpl.
+ rewrite -/list_head -/list_tail.
iMod (step_Lam with "[$Hj]") as "Hj"; auto. asimpl.
iMod (step_nat_binop _ _ _ (IfCtx _ _ :: _)
with "[$Hj]") as "Hj"; eauto. simpl.
iMod (step_if_false with "[$Hj]") as "Hj"; auto.
iMod (step_nat_binop _ _ _ (AppRCtx (RecV _) :: AppLCtx _ :: _)
with "[$Hj]") as "Hj"; eauto using to_of_val; simpl.
replace (i - 0) with i by lia.
iMod (step_rec _ _ _ (AppLCtx _ :: _) with "[$Hj]") as "Hj";
eauto using to_of_val. simpl. asimpl.
rewrite -/list_head -/list_tail.
iMod (step_list_tail _ _ _ (AppRCtx (LamV _) :: _) (_ :: _)
with "[$Hj]") as "Hj"; auto; simpl.
iMod ("IHl" with "[] [Hj]") as "Hj"; auto.
Qed.
End list_basics.
Typeclasses Opaque list_length.
Global Opaque list_length.
Typeclasses Opaque list_snoc.
Global Opaque list_snoc.
Typeclasses Opaque list_head.
Global Opaque list_head.
Typeclasses Opaque list_tail.
Global Opaque list_tail.
Typeclasses Opaque list_get.
Global Opaque list_get.