RunST.logrel_shared

From iris.prelude Require Import tactics.
From iris.proofmode Require Import tactics.
From iris.algebra Require Import list.
From RunST Require Import lang typing rules regions.
Import uPred.

(* HACK: move somewhere else *)
Ltac auto_equiv ::=
  (* Deal with "pointwise_relation" *)
  repeat lazymatch goal with
  | |- pointwise_relation _ _ _ _ => intros ?
  end;
  (* Normalize away equalities. *)
  repeat match goal with
  | H : _ ≡{_}≡ _ |- _ => apply (timeless_iff _ _) in H
  | _ => progress simplify_eq
  end;
  (* repeatedly apply congruence lemmas and use the equalities in the hypotheses. *)
  try (f_equiv; fast_done || auto_equiv).

Section logrel_shared.
  Context `{heapG Σ, LogRelG Σ}.

  Notation Cx := (prodC ((prodC valC valC) -n> iProp Σ) (leibnizC positive)).
  Notation D := (prodC valC valC -n> iProp Σ).
  Implicit Types τi : D.
  Implicit Types Δ : listC Cx.
  Implicit Types interp : listC Cx D.

  Program Definition ctx_lookup (x : var) : listC Cx -n> D := λne Δ,
    (from_option fst (cconst True)%I (Δ !! x)).
  Solve Obligations with solve_proper_alt.

  Program Definition type_to_reg_def (τ : type) :
    (listC Cx) -n> (leibnizC positive) :=
    λne Δ,
    match τ return leibnizC positive with
    | TVar x => (from_option snd 1%positive (Δ !! x))
    | _ => 1%positive
    end.
  Next Obligation.
  Proof.
    intros [] n d1 d2 Heq; trivial.
    eapply (@from_option_ne
              (prodC D (leibnizC positive)) positive (dist n) snd n snd_ne);
      eauto.
      by rewrite Heq.
  Qed.

  Definition type_to_reg_aux : { x | x = @type_to_reg_def }. by eexists. Qed.
  Definition type_to_reg := proj1_sig type_to_reg_aux.
  Definition type_to_reg_eq : @type_to_reg = @type_to_reg_def :=
    proj2_sig type_to_reg_aux.

  Lemma type_to_reg_var x Δ :
    type_to_reg (TVar x) Δ = (from_option snd 1%positive (Δ !! x)).
  Proof. by rewrite type_to_reg_eq /type_to_reg_def. Qed.

  Program Definition interp_unit : listC Cx -n> D := λne Δ ww,
    (ww.1 = LitV LitUnit ww.2 = LitV LitUnit)%I.
  Solve Obligations with solve_proper_alt.
  Program Definition interp_int : listC Cx -n> D := λne Δ ww,
    ( a : Z, ww.1 = LitV (LitInt a) ww.2 = LitV (LitInt a))%I.
  Solve Obligations with solve_proper.
  Program Definition interp_bool : listC Cx -n> D := λne Δ ww,
    ( b : bool, ww.1 = LitV (LitBool b) ww.2 = LitV (LitBool b))%I.
  Solve Obligations with solve_proper.

  Program Definition interp_prod
      (interp1 interp2 : listC Cx -n> D) : listC Cx -n> D := λne Δ ww,
    ( vv1 vv2, ww = (PairV (vv1.1) (vv2.1), PairV (vv1.2) (vv2.2))
                interp1 Δ vv1 interp2 Δ vv2)%I.
  Solve Obligations with solve_proper.

  Program Definition interp_sum
      (interp1 interp2 : listC Cx -n> D) : listC Cx -n> D := λne Δ ww,
    (( vv, ww = (InjLV (vv.1), InjLV (vv.2)) interp1 Δ vv)
     ( vv, ww = (InjRV (vv.1), InjRV (vv.2)) interp2 Δ vv))%I.
  Solve Obligations with solve_proper.

  Program Definition interp_arrow
          (interp_expr :
             (listC Cx -n> D) (listC Cx) (expr * expr) iProp Σ)
          (H: n, Proper (dist n ==> dist n ==> (=) ==> dist n) interp_expr)
          (interp1 interp2 : listC Cx -n> D) : listC Cx -n> D :=
    λne Δ ww,
    ( vv, interp1 Δ vv
             interp_expr
               interp2 Δ (App (of_val (ww.1)) (of_val (vv.1)),
                          App (of_val (ww.2)) (of_val (vv.2))))%I.
  Solve Obligations with solve_proper.

  Program Definition interp_forall
          (interp_expr :
             (listC Cx -n> D) (listC Cx) (expr * expr) iProp Σ)
          (H: n, Proper (dist n ==> dist n ==> (=) ==> dist n) interp_expr)
      (interp : listC Cx -n> D) : listC Cx -n> D := λne Δ ww,
    ( τi n,
          ( ww, PersistentP (τi ww))
          interp_expr
            interp ((τi, n) :: Δ) (TApp (of_val (ww.1)), TApp (of_val (ww.2))))%I.
  Solve Obligations with solve_proper.

  Program Definition interp_rec1 τ
      (interp : listC Cx -n> D) (Δ : listC Cx) (τi : D) : D := λne ww,
    ( vv, ww = (FoldV (vv.1), FoldV (vv.2))
              interp ((τi, (type_to_reg (TRec τ)) Δ) :: Δ) vv)%I.
  Solve Obligations with solve_proper.

  Global Instance interp_rec1_contractive τ
         (interp : listC Cx -n> D) (Δ : listC Cx) :
    Contractive (interp_rec1 τ interp Δ).
  Proof.
    intros n τi1 τi2 Hτi ww; cbn.
    apply always_ne, exist_ne; intros vv; apply and_ne; trivial.
    apply later_contractive =>i Hi. by rewrite Hτi.
  Qed.

  Program Definition interp_rec τ (interp : listC Cx -n> D) : listC Cx -n> D :=
    λne Δ, fixpoint (interp_rec1 τ interp Δ).
  Next Obligation.
    intros τ interp n Δ1 Δ2 ; apply fixpoint_ne => τi ww. solve_proper.
  Qed.

  Program Definition interp_ref ρ
      (interp : listC Cx -n> D) : listC Cx -n> D :=
    λne Δ ww, STRefREL (type_to_reg ρ Δ) (interp Δ) ww.
  Solve Obligations with solve_proper.
  Next Obligation.
  Proof.
    intros ? ? ? ? ? Heq ?; simpl.
    repeat eapply cofe_mor_car_ne; auto.
    by rewrite Heq.
  Qed.

  Definition interp_env (interp : type -> listC Cx -n> D) (Γ : list type)
      (Δ : listC Cx) (vvs : list (val * val)) : iProp Σ :=
    (length Γ = length vvs [★] zip_with (λ τ, interp τ Δ) Γ vvs)%I.

  Class env_PersistentP (Δ : listC Cx) :=
    ctx_persistentP : Forall (λ τi : Cx, vv, PersistentP ((fst τi) vv)) Δ.
  Global Instance ctx_persistent_nil : env_PersistentP [].
  Proof. by constructor. Qed.
  Global Instance ctx_persistent_cons τi p Δ :
    ( vv, PersistentP (τi vv)) env_PersistentP Δ env_PersistentP ((τi, p) :: Δ).
  Proof. by constructor. Qed.
  Global Instance ctx_persistent_lookup Δ x vv :
    env_PersistentP Δ PersistentP (ctx_lookup x Δ vv).
  Proof. intros ; revert x; induction =>-[|?] /=; apply _. Qed.

  Global Instance interp_env_persistent
         (interp : type -> listC Cx -n> D) Γ Δ vvs :
    ( τ Δ vv,
        env_PersistentP Δ PersistentP (interp τ Δ vv))
    env_PersistentP Δ PersistentP (interp_env interp Γ Δ vvs) := _.

  Lemma type_to_reg_weaken Δ1 Π Δ2 τ :
    (type_to_reg τ.[upn (length Δ1) (ren (+length Π))]) (Δ1 ++ Π ++ Δ2) =
    (type_to_reg τ) (Δ1 ++ Δ2).
  Proof.
    rewrite type_to_reg_eq /type_to_reg_def /=.
    destruct τ; trivial.
    asimpl. rewrite iter_up. destruct lt_dec as [Hl | Hl]; simpl.
    { by rewrite !lookup_app_l. }
    rewrite !lookup_app_r; first do 2 f_equiv; trivial.
    all: revert Hl; move: (length Δ1) (length Π) => ? ? ?; lia.
  Qed.

  Lemma interp_env_length (interp : type -> listC Cx -n> D) Δ Γ vvs :
    interp_env interp Γ Δ vvs length Γ = length vvs.
  Proof. by iIntros "[% ?]". Qed.

  Lemma interp_env_Some_l (interp : type -> listC Cx -n> D) Δ Γ vvs x τ :
    Γ !! x = Some τ
    interp_env interp Γ Δ vvs vv, vvs !! x = Some vv interp τ Δ vv.
  Proof.
    iIntros (?) "[Hlen HΓ]"; iDestruct "Hlen" as %Hlen.
    destruct (lookup_lt_is_Some_2 vvs x) as [v Hv].
    { by rewrite -Hlen; apply lookup_lt_Some with τ. }
    iExists v; iSplit. done. iApply (big_sep_elem_of with "HΓ").
    apply elem_of_list_lookup_2 with x.
    rewrite lookup_zip_with; by simplify_option_eq.
  Qed.

  Lemma interp_env_nil (interp : type -> listC Cx -n> D) Δ :
    True interp_env interp [] Δ [].
  Proof. iIntros ""; iSplit; auto. Qed.

  Lemma interp_env_cons (interp : type -> listC Cx -n> D) Δ Γ vvs τ vv :
    interp_env interp (τ :: Γ) Δ (vv :: vvs) ⊣⊢
    interp τ Δ vv interp_env interp Γ Δ vvs.
  Proof.
    rewrite /interp_env /= (assoc _ (interp _ _ _)) -(comm _ (_ = _)%I) -assoc.
    by apply sep_proper; [apply pure_proper; omega|].
  Qed.

  Lemma type_to_reg_closed τ Δ1 Δ2 τi rn rn':
    ( τ', τ.[upn (length Δ1) (τ' .: ids)] = τ)
    type_to_reg τ (Δ1 ++ ((τi, rn) :: Δ2))
    type_to_reg τ (Δ1 ++ ((τi, rn') :: Δ2)).
  Proof.
    intros Hτc. rewrite type_to_reg_eq /type_to_reg_def /=.
    destruct τ; auto.
    unfold var in *.
    destruct (decide (x < length Δ1)) as [Hl|Hl].
    - rewrite !lookup_app_l; auto.
    - rewrite !lookup_app_r;
        try (revert Hl; case: (length Δ1); intros; omega).
      destruct (decide ((x - length Δ1) = 0)) as [Heq|Heq]; subst.
      + assert (x = length Δ1) by omega; subst.
        specialize (Hτc TUnit). asimpl in Hτc. rewrite iter_up in Hτc.
        destruct lt_dec; first omega.
        rewrite Heq in Hτc. asimpl in Hτc. inversion Hτc.
      + destruct (x - length Δ1) eqn:Heq'; first omega.
        by rewrite Heq'.
  Qed.

End logrel_shared.