LogrelCC.F_mu_ref_cc.rules_unary

From iris.base_logic Require Export gen_heap.
From LogrelCC.program_logic Require Export weakestpre.
From LogrelCC.program_logic Require Import CC_ectx_lifting
     CC_ectxi_language CC_ectx_lifting.
From LogrelCC.F_mu_ref_cc Require Export lang.
From iris.proofmode Require Import tactics.
From stdpp Require Import fin_maps.
Set Default Proof Using "Type".
Import uPred.

Basic rules for language operations.

Class heapG Σ := HeapG {
  heapG_invG : invG Σ;
  heapG_gen_heapG :> gen_heapG loc val Σ
}.

Instance heapG_irisG `{heapG Σ} : irisG lang Σ := {
  iris_invG := heapG_invG;
  state_interp := gen_heap_ctx
}.
Global Opaque iris_invG.

Override the notations so that scopes and coercions work out
Notation "l ↦ᵢ{ q } v" := (mapsto (L:=loc) (V:=val) l q v%V)
  (at level 20, q at level 50, format "l ↦ᵢ{ q } v") : bi_scope.
Notation "l ↦ᵢ v" :=
  (mapsto (L:=loc) (V:=val) l 1 v%V) (at level 20) : bi_scope.
Notation "l ↦ᵢ{ q } -" := ( v, l ↦ᵢ{q} v)%I
  (at level 20, q at level 50, format "l ↦ᵢ{ q } -") : bi_scope.
Notation "l ↦ᵢ -" := (l ↦ᵢ{1} -)%I (at level 20) : bi_scope.

The tactic inv_head_step performs inversion on hypotheses of the shape head_step. The tactic will discharge head-reductions starting from values, and simplifies hypothesis related to conversions from and to values, and finite map operations. This tactic is slightly ad-hoc and tuned for proving our lifting lemmas.
Ltac inv_head_step :=
  repeat match goal with
  | _ => progress simplify_map_eq/= (* simplify memory stuff *)
  | H : IntoVal _ _ |- _ => apply of_to_val in H
  | H : AsVal _ |- _ => destruct H
  | H : context [to_val (of_val _)] |- _ => rewrite to_of_val in H
  | H : to_val _ = Some _ |- _ => progress rewrite (of_to_val _ _ H)
  | H : of_val ?v _ = _ |- _ =>
     is_var v; destruct v; first[discriminate H|injection H as H]
  | H : head_step _ ?e _ _ _ _ _ |- _ =>
     try (is_var e; fail 1); (* inversion yields many goals if e is a variable
     and can thus better be avoided. *)

     inversion H; subst; clear H
  end.

Local Hint Extern 0 (atomic _) => solve_atomic.
Local Hint Extern 0 (head_reducible _ _) => eexists _, _, _, _, _; simpl.

Local Hint Constructors head_step.
Local Hint Resolve alloc_fresh.
Local Hint Resolve to_of_val.

Section lifting.
Context `{heapG Σ}.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val iProp Σ.
Implicit Types efs : list expr.
Implicit Types σ : state.

Hint Extern 1 =>
  match goal with
  | |- σ, head_nonthrow_reducible _ _ σ => repeat econstructor
  end.

Hint Extern 1 =>
  match goal with
  | _ : head_step _ _ _ _ _ _ _ |- _ => inv_head_step
  end.

Base axioms for core primitives of the language: Stateless reductions

Lemma wp_rec K E e1 e1' e2 Φ :
  e1 = (Rec e1') AsVal e2
   WP fill K e1'.[e1, e2/] @ E {{ Φ }} WP fill K (App e1 e2) @ E {{ Φ }}.
Proof.
  intros -> [].
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (App _ _)) /=; eauto.
Qed.

Lemma wp_Lam K E e1 e1' e2 Φ :
  e1 = (Lam e1') AsVal e2
   WP fill K e1'.[e2/] @ E {{ Φ }} WP fill K (App e1 e2) @ E {{ Φ }}.
Proof.
  intros -> [].
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (App _ _)) /=;
    eauto.
Qed.

Lemma wp_LetIn K E e1 e2 Φ : AsVal e1
   WP fill K e2.[e1/] @ E {{ Φ }} WP fill K (LetIn e1 e2) @ E {{ Φ }}.
Proof.
  intros [].
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (LetIn _ _)) /=;
    eauto.
Qed.

Lemma wp_Seq K E e1 e2 Φ : AsVal e1
   WP fill K e2 @ E {{ Φ }} WP fill K (Seq e1 e2) @ E {{ Φ }}.
Proof.
  intros [].
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (Seq _ _)) /=;
    eauto.
Qed.

Lemma wp_bin_op K E op e1 e2 v1 v2 w Φ `{!IntoVal e1 v1, !IntoVal e2 v2} :
  binop_eval op v1 v2 = Some w
   WP fill K (of_val w) @ E {{ Φ }}
   WP fill K (BinOp op e1 e2) @ E {{ Φ }}.
Proof.
  intros ?.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (BinOp op _ _)) /=;
    eauto.
Qed.

Lemma wp_if_true K E e1 e2 Φ :
   WP fill K e1 @ E {{ Φ }} WP fill K (If (#♭ true) e1 e2) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (If _ _ _)) /=;
    eauto.
Qed.

Lemma wp_if_false K E e1 e2 Φ :
   WP fill K e2 @ E {{ Φ }} WP fill K (If (#♭ false) e1 e2) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (If _ _ _)) /=;
    eauto.
Qed.

Lemma wp_fst K E e1 e2 v1 Φ `{!IntoVal e1 v1, !AsVal e2} :
   WP fill K e1 @ E {{ Φ }}
   WP fill K (Fst (Pair e1 e2)) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (Fst _)) /=;
    inv_head_step; eauto.
Qed.

Lemma wp_snd K E e1 e2 v2 Φ `{!AsVal e1, !IntoVal e2 v2} :
   WP fill K e2 @ E {{ Φ }}
   WP fill K (Snd (Pair e1 e2)) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (Snd _)) /=;
    inv_head_step; eauto.
Qed.

Lemma wp_case_injl K E e0 e1 e2 Φ `{!AsVal e0} :
   WP fill K e1.[e0/] @ E {{ Φ }}
   WP fill K (Case (InjL e0) e1 e2) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (Case _ _ _)) /=;
    inv_head_step; eauto.
Qed.

Lemma wp_case_injr K E e0 e1 e2 Φ `{!AsVal e0} :
   WP fill K e2.[e0/] @ E {{ Φ }}
   WP fill K (Case (InjR e0) e1 e2) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (Case _ _ _)) /=;
    inv_head_step; eauto.
Qed.

Lemma wp_unfold K E e v Φ `{!IntoVal e v} :
   WP fill K e @ E {{ Φ }}
   WP fill K (Unfold (Fold e)) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (Unfold _)) /=;
    eauto.
Qed.

Lemma wp_tapp K E e Φ :
   WP fill K e @ E {{ Φ }}
   WP fill K (TApp (TLam e)) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (TApp _)) /=;
    eauto.
Qed.

Lemma wp_callcc K E e Φ :
   WP fill K e.[Cont K/] @ E {{ Φ }}
   WP fill K (Callcc e) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_nonthrow_pure_det_head_step_no_fork' K (Callcc _)) /=;
    eauto.
  - intros ?. exists CaptureMode; repeat econstructor.
Qed.

Lemma wp_throw K K' E e v Φ `{!IntoVal e v} :
   WP fill K' e @ E {{ Φ }}
   WP fill K (Throw e (Cont K')) @ E {{ Φ }}.
Proof.
  rewrite -(wp_lift_throw_pure_det_head_step_no_fork' K (Throw _ _)) /=;
    eauto.
  - intros ?. repeat econstructor; eauto.
Qed.

Heap
Lemma wp_alloc K E e v Φ `{!IntoVal e v} :
   ( l, l ↦ᵢ v -∗ WP fill K (Loc l) @ E {{ Φ }})
   WP fill K (Alloc e) @ E {{ Φ }}.
Proof.
  iIntros "HΦ".
  iApply wp_lift_nonthrow_atomic_head_step_no_fork; auto.
  iIntros (σ1) "Hσ !>".
  assert (σ1 !! (fresh (dom (gset loc) σ1)) = None).
  { eapply (@not_elem_of_dom _ (gmap loc) (gset loc));
      first by typeclasses eauto.
    apply is_fresh. }
  iSplit; first (iPureIntro; eauto).
  { repeat econstructor; eauto. }
  iNext; iIntros (rm v2 σ2 efs Hstep); inv_head_step.
  iMod (@gen_heap_alloc with "Hσ") as "[Hσ Hl]"; first done.
  iModIntro; iSplit=> //. iFrame. by iApply "HΦ".
Qed.

Lemma wp_load K E l q v Φ:
   (l ↦ᵢ{q} v -∗ WP fill K (of_val v) @ E {{ Φ }}) l ↦ᵢ{q} v
   WP fill K (Load (Loc l)) @ E {{ Φ }}.
Proof.
  iIntros "[HΦ >Hl]".
  iApply wp_lift_nonthrow_atomic_head_step_no_fork; auto.
  iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
  iSplit.
  { iPureIntro. repeat econstructor; eauto. }
  iNext; iIntros (rm v2 σ2 efs Hstep); inv_head_step.
  iModIntro; iSplit=> //. iFrame. by iApply "HΦ".
Qed.

Lemma wp_store K E l v' e v Φ `{!IntoVal e v} :
   (l ↦ᵢ v -∗ WP fill K Unit @ E {{ Φ }}) l ↦ᵢ v'
   WP fill K (Store (Loc l) e) @ E {{ Φ }}.
Proof.
  iIntros "[HΦ >Hl]".
  iApply wp_lift_nonthrow_atomic_head_step_no_fork; auto.
  iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
  iSplit.
  { iPureIntro. repeat econstructor; eauto. }
  iNext; iIntros (rm v2 σ2 efs Hstep); inv_head_step.
  iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]".
  iModIntro. iSplit=>//. by iApply "HΦ".
Qed.

Lemma wp_alloc' E e v Φ `{!IntoVal e v} :
   ( l, l ↦ᵢ v ={E}=∗ Φ (LocV l)) WP (Alloc e) @ E {{ Φ }}.
Proof.
  iIntros "HΦ".
  iApply (wp_alloc [] with "[HΦ]"); eauto.
  iNext; iIntros (l) "Hl". iApply wp_value_fupd; eauto.
  by iMod ("HΦ" with "Hl").
Qed.

Lemma wp_load' E l q v Φ:
   (l ↦ᵢ{q} v -∗ WP (of_val v) @ E {{ Φ }}) l ↦ᵢ{q} v
   WP (Load (Loc l)) @ E {{ Φ }}.
Proof. iIntros "HΦ". iApply (wp_load [] with "[HΦ]"); eauto. Qed.

Lemma wp_store' E l v' e v Φ `{!IntoVal e v} :
   (l ↦ᵢ v -∗ WP Unit @ E {{ Φ }}) l ↦ᵢ v'
   WP Store (Loc l) e @ E {{ Φ }}.
Proof. iIntros "HΦ". iApply (wp_store [] with "[HΦ]"); eauto. Qed.

Lemma wp_atomic_under_ectx E1 E2 K e Φ :
  is_atomic e
  (|={E1,E2}=> WP e @ E2 {{ v, |={E2,E1}=> WP fill K (of_val v) @ E1 {{ Φ }} }})
     WP fill K e @ E1 {{ Φ }}.
Proof.
  iIntros (Ha) "H".
  iApply wp_atomic_under_ectx;
    eauto using is_atomic_normal, is_atomic_sub_values, is_atomic_correct.
Qed.

End lifting.