Library resolution

Require Export comp.
Set Implicit Arguments.
Unset Strict Implicit.

CPP 2015 submission: Section V (definitions only)

Section Semworld.

Variable sw: semworld.

Notation funcname := (funcname sw) (only parsing).
Notation event := (event sw) (only parsing).
Notation heap := (heap sw) (only parsing).
Notation args := (args sw) (only parsing).
Notation ret := (ret sw) (only parsing).
Notation xevent := (xevent sw) (only parsing).

Notation trace := (list event) (only parsing).
Notation traceinf := (Stream event) (only parsing).

Notation cbehavior := (behavior xevent (heap * ret)) (only parsing).
Notation behavior := (behavior event (heap * ret)) (only parsing).

CPP 2015 submission: Definition 11. Behavior simulation small-step semantics

Results
Inductive bres : Type :=
| Spurious
| Normal (h: heap) (r: ret)
.

Normal and spurious behaviors
Notation bbehavior := (smallstep.behavior xevent bres) (only parsing).

Function bbehavior_of_cbehavior (b: cbehavior): bbehavior :=
  match b with
    | Terminates l (h,r) => Terminates l (Normal h r)
    | Stuck l => Stuck _ l
    | Diverges l => Diverges _ l
    | Reacts f => Reacts _ f
  end.

Lemma bbehavior_of_cbehavior_cons: forall e b,
  cons_behavior e (bbehavior_of_cbehavior b) = bbehavior_of_cbehavior (cons_behavior e b).
Proof.
  destruct b; simpl; try reflexivity.
  destruct hfin; reflexivity.
Qed.

Lemma bbehavior_of_cbehavior_append: forall le b,
  append1_behavior le (bbehavior_of_cbehavior b) = bbehavior_of_cbehavior (append1_behavior le b).
Proof.
  induction le; simpl.
  trivial.
 intros.
 rewrite IHle.
 eapply bbehavior_of_cbehavior_cons.
Qed.

Lemma bbehavior_of_cbehavior_cons_ex: forall e b b',
  cons_behavior e b = bbehavior_of_cbehavior b'
 ->
 exists b'', b' = cons_behavior e b''.
Proof.
  intros.
  generalize (behavior_elim_ex b').
  inversion 1; subst; simpl in *.
   destruct r. functional inversion H.
   functional inversion H.
   functional inversion H.
  rewrite <- bbehavior_of_cbehavior_cons in H.
  generalize (cons_behavior_inj H).
  injection 1; intros; subst.
  eauto.
Qed.

Lemma bbehavior_of_cbehavior_append_ex: forall e b b',
  append1_behavior e b = bbehavior_of_cbehavior b'
 ->
 exists b'', b' = append1_behavior e b''.
Proof.
  induction e; simpl.
   eauto.
  intros.
  edestruct (bbehavior_of_cbehavior_cons_ex).
  eassumption.
  subst.
  rewrite <- bbehavior_of_cbehavior_cons in H.
  generalize (cons_behavior_inj H).
  injection 1; intros; subst.
  edestruct IHe.
  eassumption.
  subst.
  eauto.
Qed.

Configuration
Inductive bstate: Type :=
| BSpurious
| BBehavior (_: cbehavior) (_: list ((heap * ret) * cbehavior))
.

Section Simulate.

Variable psi: Map funcname (args -> heap -> cbehavior -> Prop).

Transition relation
Inductive bstep : bstate -> list xevent -> bstate -> Prop :=
| bstep_cons e b b'
  (Hb': b' = cons_behavior (Event e) b)
  s:
  bstep (BBehavior b' s) (Event e :: nil) (BBehavior b s)
| bstep_external f (Hf: ~ in_dom psi f)
  a h1 r h2 b b' (Hb': b' = cons_behavior (External f a h1 r h2) b) s:
  bstep (BBehavior b' s) (External f a h1 r h2 :: nil) (BBehavior b s)
| bstep_internal f fb (Hf: lookup psi f fb)
  a h1 bf (Hbf: fb a h1 bf)
  r h2 b b' (Hb': b' = cons_behavior (External f a h1 r h2) b) s:
  bstep (BBehavior b' s) nil (BBehavior bf (((h2,r),b) :: s))
| bstep_terminates hr b' s:
  bstep (BBehavior (Terminates nil hr) ((hr,b') :: s)) nil (BBehavior b' s)
| bstep_spurious hr hr' (Hdi: hr <> hr') b' s:
  bstep (BBehavior (Terminates nil hr) ((hr', b') :: s)) nil BSpurious
| bstep_diverges s:
  bstep (BBehavior (Diverges _ nil) s) nil (BBehavior (Diverges _ nil) s)
.

Final states
Inductive bstate_final: bstate -> bres -> Prop :=
| bstate_final_normal h r:
  bstate_final (BBehavior (Terminates nil (h,r)) nil) (Normal h r)
| bstate_final_spurious:
  bstate_final BSpurious Spurious
.

CPP 2015 submission: Definition 12. Resolution.


Inductive bsemantics (f: funcname) (ar: args) (h: heap): cbehavior -> Prop :=
| bsemantics_intro fb (Hcmd: lookup psi f fb)
  cb (Hcb: fb ar h cb)
  cb' bb' (Hbeh': bb' = bbehavior_of_cbehavior cb')
  (Hbeh: config_behaviors bstep bstate_final (BBehavior cb nil) bb') :
  bsemantics f ar h cb'
  .

End Simulate.

CPP 2015 submission: Definition 13. Linking


Definition link psi2 psi1 :=
  create_map (dom (merge psi2 psi1)) (bsemantics (merge psi2 psi1)).

Lemma link_sym : forall psi1 psi2, disjoint psi1 psi2 -> link psi1 psi2 = link psi2 psi1.
Proof.
  intros.
  unfold link.
  rewrite disj_merge_sym.
  reflexivity.
  assumption.
Qed.

End Semworld.