Library iris.heap_lang.tactics

From iris.heap_lang Require Export lang.
From iris.prelude Require Import fin_maps.
Import heap_lang.

We define an alternative representation of expressions in which the embedding of values and closed expressions is explicit. By reification of expressions into this type we can implementation substitution, closedness checking, atomic checking, and conversion into values, by computation.
Module W.
Inductive expr :=
  | Val (v : val)
  | ClosedExpr (e : heap_lang.expr) `{!Closed [] e}
  
  | Var (x : string)
  | Rec (f x : binder) (e : expr)
  | App (e1 e2 : expr)
  
  | Lit (l : base_lit)
  | UnOp (op : un_op) (e : expr)
  | BinOp (op : bin_op) (e1 e2 : expr)
  | If (e0 e1 e2 : expr)
  
  | Pair (e1 e2 : expr)
  | Fst (e : expr)
  | Snd (e : expr)
  | Letp (x y : binder) (e1 e2 : expr)
  
  | InjL (e : expr)
  | InjR (e : expr)
  | Case (e0 : expr) (e1 : expr) (e2 : expr)
  
  | Fork (e : expr)
  
  | Alloc (e : expr)
  | Load (e : expr)
  | Store (e1 : expr) (e2 : expr)
  | CAS (e0 : expr) (e1 : expr) (e2 : expr)
  | Swap (e1 : expr) (e2 : expr)
  | FAI (e : expr).

Fixpoint to_expr (e : expr) : heap_lang.expr :=
  match e with
  | Val vof_val v
  | ClosedExpr e _e
  | Var xheap_lang.Var x
  | Rec f x eheap_lang.Rec f x (to_expr e)
  | App e1 e2heap_lang.App (to_expr e1) (to_expr e2)
  | Lit lheap_lang.Lit l
  | UnOp op eheap_lang.UnOp op (to_expr e)
  | BinOp op e1 e2heap_lang.BinOp op (to_expr e1) (to_expr e2)
  | If e0 e1 e2heap_lang.If (to_expr e0) (to_expr e1) (to_expr e2)
  | Pair e1 e2heap_lang.Pair (to_expr e1) (to_expr e2)
  | Fst eheap_lang.Fst (to_expr e)
  | Snd eheap_lang.Snd (to_expr e)
  | Letp x y e1 e2heap_lang.Letp x y (to_expr e1) (to_expr e2)
  | InjL eheap_lang.InjL (to_expr e)
  | InjR eheap_lang.InjR (to_expr e)
  | Case e0 e1 e2heap_lang.Case (to_expr e0) (to_expr e1) (to_expr e2)
  | Fork eheap_lang.Fork (to_expr e)
  | Alloc eheap_lang.Alloc (to_expr e)
  | Load eheap_lang.Load (to_expr e)
  | Store e1 e2heap_lang.Store (to_expr e1) (to_expr e2)
  | CAS e0 e1 e2heap_lang.CAS (to_expr e0) (to_expr e1) (to_expr e2)
  | Swap e1 e2heap_lang.Swap (to_expr e1) (to_expr e2)
  | FAI eheap_lang.FAI (to_expr e)
  end.

Ltac of_expr e :=
  lazymatch e with
  | heap_lang.Var ?xconstr:(Var x)
  | heap_lang.Rec ?f ?x ?elet e := of_expr e in constr:(Rec f x e)
  | heap_lang.App ?e1 ?e2
     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(App e1 e2)
  | heap_lang.Lit ?lconstr:(Lit l)
  | heap_lang.UnOp ?op ?elet e := of_expr e in constr:(UnOp op e)
  | heap_lang.BinOp ?op ?e1 ?e2
     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(BinOp op e1 e2)
  | heap_lang.If ?e0 ?e1 ?e2
     let e0 := of_expr e0 in let e1 := of_expr e1 in let e2 := of_expr e2 in
     constr:(If e0 e1 e2)
  | heap_lang.Pair ?e1 ?e2
     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(Pair e1 e2)
  | heap_lang.Fst ?elet e := of_expr e in constr:(Fst e)
  | heap_lang.Snd ?elet e := of_expr e in constr:(Snd e)
  | heap_lang.Letp ?x ?y ?e1 ?e2
     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(Letp x y e1 e2)
  | heap_lang.InjL ?elet e := of_expr e in constr:(InjL e)
  | heap_lang.InjR ?elet e := of_expr e in constr:(InjR e)
  | heap_lang.Case ?e0 ?e1 ?e2
     let e0 := of_expr e0 in let e1 := of_expr e1 in let e2 := of_expr e2 in
     constr:(Case e0 e1 e2)
  | heap_lang.Fork ?elet e := of_expr e in constr:(Fork e)
  | heap_lang.Alloc ?elet e := of_expr e in constr:(Alloc e)
  | heap_lang.Load ?elet e := of_expr e in constr:(Load e)
  | heap_lang.Store ?e1 ?e2
     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(Store e1 e2)
  | heap_lang.CAS ?e0 ?e1 ?e2
     let e0 := of_expr e0 in let e1 := of_expr e1 in let e2 := of_expr e2 in
     constr:(CAS e0 e1 e2)
  | heap_lang.Swap ?e1 ?e2
     let e1 := of_expr e1 in let e2 := of_expr e2 in constr:(Swap e1 e2)
  | heap_lang.FAI ?elet e := of_expr e in constr:(FAI e)
  | to_expr ?ee
  | of_val ?vconstr:(Val v)
  | _constr:(ltac:(
     match goal with H : Closed [] e |- _exact (@ClosedExpr e H) end))
  end.

Fixpoint is_closed (X : list string) (e : expr) : bool :=
  match e with
  | Val _ | ClosedExpr _ _true
  | Var xbool_decide (x X)
  | Rec f x eis_closed (f :b: x :b: X) e
  | Letp x y e1 e2is_closed X e1 && is_closed (x :b: y :b: X) e2
  | Lit _true
  | UnOp _ e | Fst e | Snd e | InjL e | InjR e | Fork e | Alloc e | Load e | FAI e
     is_closed X e
  | App e1 e2 | BinOp _ e1 e2 | Pair e1 e2 | Store e1 e2 | Swap e1 e2
     is_closed X e1 && is_closed X e2
  | If e0 e1 e2 | Case e0 e1 e2 | CAS e0 e1 e2
     is_closed X e0 && is_closed X e1 && is_closed X e2
  end.
Lemma is_closed_correct X e : is_closed X e heap_lang.is_closed X (to_expr e).
Proof.
  revert X.
  induction e; naive_solver eauto using is_closed_of_val, is_closed_weaken_nil.
Qed.

Fixpoint to_val (e : expr) : option val :=
  match e with
  | Val vSome v
  | Rec f x e
     if decide (is_closed (f :b: x :b: []) e) is left H
     then Some (@RecV f x (to_expr e) (is_closed_correct _ _ H)) else None
  | Lit lSome (LitV l)
  | Pair e1 e2v1 to_val e1; v2 to_val e2; Some (PairV v1 v2)
  | InjL eInjLV <$> to_val e
  | InjR eInjRV <$> to_val e
  | _None
  end.
Lemma to_val_Some e v :
  to_val e = Some v heap_lang.to_val (to_expr e) = Some v.
Proof.
  revert v. induction e; intros; simplify_option_eq; rewrite ?to_of_val; auto.
  - do 2 f_equal. apply proof_irrel.
  - exfalso. unfold Closed in *; eauto using is_closed_correct.
Qed.
Lemma to_val_is_Some e :
  is_Some (to_val e) is_Some (heap_lang.to_val (to_expr e)).
Proof. intros [v ?]; v; eauto using to_val_Some. Qed.

Fixpoint subst (x : string) (es : expr) (e : expr) : expr :=
  match e with
  | Val vVal v
  | ClosedExpr e H ⇒ @ClosedExpr e H
  | Var yif decide (x = y) then es else Var y
  | Rec f y e
     Rec f y $ if decide (BNamed x f BNamed x y) then subst x es e else e
  | App e1 e2App (subst x es e1) (subst x es e2)
  | Lit lLit l
  | UnOp op eUnOp op (subst x es e)
  | BinOp op e1 e2BinOp op (subst x es e1) (subst x es e2)
  | If e0 e1 e2If (subst x es e0) (subst x es e1) (subst x es e2)
  | Pair e1 e2Pair (subst x es e1) (subst x es e2)
  | Fst eFst (subst x es e)
  | Snd eSnd (subst x es e)
  | Letp y z e1 e2
    Letp y z (subst x es e1) $ if decide (BNamed x y BNamed x z) then subst x es e2 else e2
  | InjL eInjL (subst x es e)
  | InjR eInjR (subst x es e)
  | Case e0 e1 e2Case (subst x es e0) (subst x es e1) (subst x es e2)
  | Fork eFork (subst x es e)
  | Alloc eAlloc (subst x es e)
  | Load eLoad (subst x es e)
  | Store e1 e2Store (subst x es e1) (subst x es e2)
  | CAS e0 e1 e2CAS (subst x es e0) (subst x es e1) (subst x es e2)
  | Swap e1 e2Swap (subst x es e1) (subst x es e2)
  | FAI eFAI (subst x es e)
  end.
Lemma to_expr_subst x er e :
  to_expr (subst x er e) = heap_lang.subst x (to_expr er) (to_expr e).
Proof.
  induction e; simpl; repeat case_decide;
    f_equal; auto using is_closed_nil_subst, is_closed_of_val, eq_sym.
Qed.

Definition atomic (e : expr) :=
  match e with
  | Alloc ebool_decide (is_Some (to_val e))
  | Load ebool_decide (is_Some (to_val e))
  | Store e1 e2bool_decide (is_Some (to_val e1) is_Some (to_val e2))
  | CAS e0 e1 e2
     bool_decide (is_Some (to_val e0) is_Some (to_val e1) is_Some (to_val e2))
  | Swap e1 e2bool_decide (is_Some (to_val e1) is_Some (to_val e2))
  | FAI ebool_decide (is_Some (to_val e))
  | Fst e | Snd ebool_decide (is_Some (to_val e))
  | Fork _true
  
  | App (Rec _ _ (Lit _)) (Lit _) ⇒ true
  | _false
  end.
Lemma atomic_correct e : atomic e language.atomic (to_expr e).
Proof.
  intros He. apply ectx_language_atomic.
  - intros σ e' σ' ef.
    destruct e; simpl; try done; repeat (case_match; try done);
    inversion 1; rewrite ?to_of_val; eauto. subst.
    unfold subst'; repeat (case_match || contradiction || simplify_eq/=); eauto.
  - intros [|Ki K] e' Hfill Hnotval; [done|exfalso].
    apply (fill_not_val K), eq_None_not_Some in Hnotval. apply Hnotval. simpl.
    revert He. destruct e; simpl; try done; repeat (case_match; try done);
    rewrite ?bool_decide_spec;
    destruct Ki; inversion Hfill; subst; clear Hfill;
    try naive_solver eauto using to_val_is_Some.
    move_ /=; destruct decide as [|Nclosed]; [by eauto | by destruct Nclosed].
Qed.
End W.

Ltac solve_closed :=
  match goal with
  | |- Closed ?X ?e
     let e' := W.of_expr e in change (Closed X (W.to_expr e'));
     apply W.is_closed_correct; vm_compute; exact I
  end.
Hint Extern 0 (Closed _ _) ⇒ solve_closed : typeclass_instances.

Ltac solve_to_val :=
  try match goal with
  | |- context E [language.to_val ?e] ⇒
     let X := context E [to_val e] in change X
  end;
  match goal with
  | |- to_val ?e = Some ?v
     let e' := W.of_expr e in change (to_val (W.to_expr e') = Some v);
     apply W.to_val_Some; simpl; unfold W.to_expr; reflexivity
  | |- is_Some (to_val ?e) ⇒
     let e' := W.of_expr e in change (is_Some (to_val (W.to_expr e')));
     apply W.to_val_is_Some, (bool_decide_unpack _); vm_compute; exact I
  end.

Ltac solve_atomic :=
  match goal with
  | |- language.atomic ?e
     let e' := W.of_expr e in change (language.atomic (W.to_expr e'));
     apply W.atomic_correct; vm_compute; exact I
  end.
Hint Extern 0 (language.atomic _) ⇒ solve_atomic : fsaV.

Substitution
Ltac simpl_subst :=
  simpl;
  repeat match goal with
  | |- context [subst ?x ?er ?e] ⇒
      let er' := W.of_expr er in let e' := W.of_expr e in
      change (subst x er e) with (subst x (W.to_expr er') (W.to_expr e'));
      rewrite <-(W.to_expr_subst x); simpl
  end;
  unfold W.to_expr.
Arguments W.to_expr : simpl never.
Arguments subst : simpl never.

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/=
  | H : to_val _ = Some _ |- _apply of_to_val in H
  | H : context [to_val (of_val _)] |- _rewrite to_of_val in H
  | H : head_step ?e _ _ _ _ |- _
     try (is_var e; fail 1);
     inversion H; subst; clear H
  end.

The tactic reshape_expr e tac decomposes the expression e into an evaluation context K and a subexpression e'. It calls the tactic tac K e' for each possible decomposition until tac succeeds.
Ltac reshape_val e tac :=
  let rec go e :=
  match e with
  | of_val ?vv
  | Rec ?f ?x ?econstr:(RecV f x e)
  | Lit ?lconstr:(LitV l)
  | Pair ?e1 ?e2
    let v1 := go e1 in let v2 := go e2 in constr:(PairV v1 v2)
  | InjL ?elet v := go e in constr:(InjLV v)
  | InjR ?elet v := go e in constr:(InjRV v)
  end in let v := go e in first [tac v | fail 2].

Ltac reshape_expr e tac :=
  let rec go K e :=
  match e with
  | _tac (reverse K) e
  | App ?e1 ?e2reshape_val e1 ltac:(fun v1go (AppRCtx v1 :: K) e2)
  | App ?e1 ?e2go (AppLCtx e2 :: K) e1
  | UnOp ?op ?ego (UnOpCtx op :: K) e
  | BinOp ?op ?e1 ?e2
     reshape_val e1 ltac:(fun v1go (BinOpRCtx op v1 :: K) e2)
  | BinOp ?op ?e1 ?e2go (BinOpLCtx op e2 :: K) e1
  | If ?e0 ?e1 ?e2go (IfCtx e1 e2 :: K) e0
  | Pair ?e1 ?e2reshape_val e1 ltac:(fun v1go (PairRCtx v1 :: K) e2)
  | Pair ?e1 ?e2go (PairLCtx e2 :: K) e1
  | Fst ?ego (FstCtx :: K) e
  | Snd ?ego (SndCtx :: K) e
  | Letp ?x ?y ?e1 ?e2go (LetpCtx x y e2 :: K) e1
  | InjL ?ego (InjLCtx :: K) e
  | InjR ?ego (InjRCtx :: K) e
  | Case ?e0 ?e1 ?e2go (CaseCtx e1 e2 :: K) e0
  | Alloc ?ego (AllocCtx :: K) e
  | Load ?ego (LoadCtx :: K) e
  | Store ?e1 ?e2reshape_val e1 ltac:(fun v1go (StoreRCtx v1 :: K) e2)
  | Store ?e1 ?e2go (StoreLCtx e2 :: K) e1
  | CAS ?e0 ?e1 ?e2reshape_val e0 ltac:(fun v0first
     [ reshape_val e1 ltac:(fun v1go (CasRCtx v0 v1 :: K) e2)
     | go (CasMCtx v0 e2 :: K) e1 ])
  | CAS ?e0 ?e1 ?e2go (CasLCtx e1 e2 :: K) e0
  | Swap ?e1 ?e2reshape_val e1 ltac:(fun v1go (SwapRCtx v1 :: K) e2)
  | Swap ?e1 ?e2go (SwapLCtx e2 :: K) e1
  | FAI ?ego (FaiCtx :: K) e
  end in go (@nil ectx_item) e.

The tactic do_head_step tac solves goals of the shape head_reducible and head_step by performing a reduction step and uses tac to solve any side-conditions generated by individual steps.
Tactic Notation "do_head_step" tactic3(tac) :=
  try match goal with |- head_reducible _ _eexists _, _, _ end;
  simpl;
  match goal with
  | |- head_step ?e1 ?σ1 ?e2 ?σ2 ?ef
     first [apply alloc_fresh|econstructor];
       
       first [rewrite ?to_of_val; reflexivity|simpl_subst; tac; fast_done]
  end.