Library refinement

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

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).

Notation bres := (bres sw) (only parsing).
Notation bbehavior := (smallstep.behavior xevent bres) (only parsing).

Section Refinement.

Variable R : (cbehavior -> Prop) -> (cbehavior -> Prop) -> Prop.

Definition oR (o1 o2: Map funcname (args -> heap -> cbehavior -> Prop)): Prop :=
  dom o1 = dom o2 /\
  forall f f1, lookup o1 f f1 ->
    forall f2, lookup o2 f f2 ->
      forall a h, R (f1 a h) (f2 a h).

Definition never_empty (o: Map funcname (args -> heap -> cbehavior -> Prop)): Prop :=
  (forall f f1, lookup o f f1 -> forall a h, exists beh, f1 a h beh).

Lemma never_empty_merge: forall o1 o2, disjoint o1 o2 ->
  never_empty o1 -> never_empty o2 -> never_empty (merge o1 o2).
Proof.
  unfold never_empty.
  intros.
  destruct (mergelookup H2); eauto.
Qed.

Section Congruence.

Definition congruence :=
  forall o1, never_empty o1 -> forall o2,
    oR o1 o2 -> oR (create_map (dom o1) (bsemantics o1)) (create_map (dom o2) (bsemantics o2)).

Definition compositionality :=
  forall o o1, disjoint o o1 -> never_empty o -> never_empty o1 ->
    forall o2, oR o1 o2 ->
      oR (link o o1) (link o o2).

Hypothesis refines_refl: reflexive _ R.

Lemma oR_merge : forall ol1 ol2, oR ol1 ol2 ->
  forall ol, disjoint ol ol1 ->
    oR (merge ol ol1) (merge ol ol2).
Proof.
  destruct 1.
  intros.
  generalize (disjoint_same_dom H1 H).
  intro.
  split.
   rewrite dom_merge.
   rewrite dom_merge.
   rewrite H.
   reflexivity.
  intros.
  destruct (mergelookup H3).
   generalize (lookup_merge_lookup_1 ol2 H5).
   intro.
   generalize (lookup_func H4 H6).
   intro; subst.
   apply refines_refl.
  generalize (in_dom_dom (lookup_in_dom H5)).
  rewrite H.
  intro.
  generalize (dom_in_dom H6).
  destruct 1.
  generalize (lookup_disj_merge_lookup_2 H7 H2).
  intro.
  generalize (lookup_func H4 H8).
  intro; subst.
  eapply H0.
  eassumption.
  assumption.
Qed.

CPP 2015 submission: Theorem 2.


Theorem congruence_is_compositionality : congruence <-> compositionality.
Proof.
  unfold congruence, compositionality, link.
  split.
   intros.
   eapply H.
   eapply never_empty_merge.
   assumption.
   assumption.
   assumption.
   apply oR_merge.
   assumption.
   assumption.
  intros.
  rewrite <- (merge_empty o1).
  rewrite <- (merge_empty o2).
  eapply H.
  apply disjoint_empty.
  unfold never_empty.
  intros; discriminate.
  assumption.
  assumption.
Qed.

Lemma oR_refl: reflexive _ oR.
Proof.
  split.
   reflexivity.
  intros until 2.
  generalize (lookup_func H H0).
  intro; subst.
  intros.
  eapply refines_refl.
Qed.

End Congruence.

Section Trans.

Hypothesis refines_trans: transitive _ R.

Lemma oR_trans: transitive _ oR.
Proof.
  unfold transitive.
  destruct 1.
  destruct 1.
  split.
   rewrite H.
   assumption.
  intros.
  generalize (in_dom_dom (lookup_in_dom H3)).
  rewrite H.
  intro.
  generalize (dom_in_dom H5).
  destruct 1.
  eauto.
Qed.

End Trans.

CPP 2015 submission: Definition 14. Refinement relation


Record refinement : Prop := Refinement {
  refines_refl: reflexive _ R;
  refines_trans: transitive _ R;
  refines_congruence: congruence
}.

A reformulation of Theorem 2.

Lemma refines_compositionality: refinement -> compositionality.
Proof. exact (fun REF => let (H, _) := congruence_is_compositionality (refines_refl REF) in H (refines_congruence REF)). Qed.

Hypothesis Hr: refinement.

CPP 2015 submission: Definition 15. Compiler correctness


Section Compilers.

Variables lang1 lang2: language sw.

Definition correct_compiler (c: module lang1 -> option (module lang2)) : Prop :=
  forall m1 m2, c m1 = Some m2 ->
    oR (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)).

Here is the Coq translation of: forall C1, C2, ..., Cn correct compilers, if Ci mi = Some m'i, then m'i refines mi

Inductive compiler_source_target_list : list (module lang1 -> option (module lang2)) -> list (module lang1) -> list (module lang2) -> Prop :=
| compiler_source_target_list_nil: compiler_source_target_list nil nil nil
| compiler_source_target_list_cons: forall c m1 m2,
  c m1 = Some m2 ->
  forall lc lm1 lm2,
    compiler_source_target_list lc lm1 lm2 ->
    compiler_source_target_list (c::lc) (m1::lm1) (m2::lm2).

Inductive all_refine: list (module lang1) -> list (module lang2) -> Prop :=
| all_refine_nil: all_refine nil nil
| all_refine_cons: forall m1 m2,
  oR (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)) ->
  forall l1 l2, all_refine l1 l2 ->
    all_refine (m1::l1) (m2::l2).

Lemma all_refine_correct_compilers: forall lc,
  (forall c, In c lc -> correct_compiler c) ->
  forall lm1 lm2, compiler_source_target_list lc lm1 lm2 ->
    all_refine lm1 lm2.
Proof.
  induction lc; inversion 2; subst.
   constructor.
  constructor.
  eapply H.
  left.
  reflexivity.
  assumption.
  eapply IHlc.
  intros.
  eapply H.
  right.
  assumption.
  assumption.
Qed.

Now let us state Theorem 3 only using program refinement. Let us first do it at the level of the compositional semantics

Lemma all_refine_same_domains: forall lm1 lm2,
  all_refine lm1 lm2 ->
  same_domains_embedded lm1 lm2.
Proof.
  induction 1.
   constructor.
  destruct H.
  rewrite create_map_dom in H.
  rewrite create_map_dom in H.
  constructor.
  assumption.
  assumption.
Qed.

Lemma all_refine_merge: forall lm1 lm2,
  all_refine lm1 lm2 ->
  disjoint_embedded lm1 ->
  oR (create_map (dom (merge_embedded lm1)) (csemantics (merge_embedded lm1))) (create_map (dom (merge_embedded lm2)) (csemantics (merge_embedded lm2))).
Proof.
  intros.
  rewrite <- csemantics_bsemantics_n.
  2: assumption.
  rewrite <- csemantics_bsemantics_n.
  revert lm2 H H0.
  induction lm1; inversion 1; inversion 1; subst; intros.
   simpl.
   split.
   reflexivity.
   intros; discriminate.
  simpl.
  generalize (all_refine_same_domains H).
  inversion 1; subst.
  destruct (disjoint_embedded_same_domains H5 H0).
  inversion H1; subst.
  eapply oR_trans.
  eapply refines_trans.
  assumption.
  eapply refines_compositionality.
  assumption.
  rewrite csemantics_bsemantics_n.
  intro.
  intros.
  generalize (in_dom_dom H6).
  rewrite create_map_dom.
  intro.
  generalize (in_dom_dom H10).
  rewrite create_map_dom.
  intro.
  eapply H8.
  eapply dom_in_dom.
  eassumption.
  eapply dom_in_dom.
  assumption.
  assumption.
  intro.
  intros until 1.
  generalize (create_map_lookup_elim H6).
  intro; subst.
  generalize (in_dom_dom (lookup_in_dom H6)).
  rewrite create_map_dom.
  intro.
  destruct (dom_in_dom H10).
  intros.
  case_eq (callee_init a0 h x).
  intros.
  destruct (exists_behavior (@config_final _ _) (cstep a) (Config h0 l nil)).
  esplit.
  econstructor.
  eassumption.
  eassumption.
  eassumption.
  rewrite csemantics_bsemantics_n.
  intro.
  intros until 1.
  generalize (create_map_lookup_elim H6).
  intro; subst.
  generalize (in_dom_dom (lookup_in_dom H6)).
  rewrite create_map_dom.
  intro.
  destruct (dom_in_dom H10).
  intros.
  case_eq (callee_init a0 h x).
  intros.
  destruct (exists_behavior (@config_final _ _) (cstep (merge_embedded lm1)) (Config h0 l nil)).
  esplit.
  econstructor.
  eassumption.
  eassumption.
  eassumption.
  assumption.
 eapply IHlm1.
 eassumption.
 assumption.
 rewrite link_sym.
 rewrite (link_sym (psi1 := create_map (dom m2) (csemantics m2))).
  eapply refines_compositionality.
  assumption.
  rewrite csemantics_bsemantics_n.
  intro.
  intros.
  generalize (in_dom_dom H6).
  rewrite create_map_dom.
  intro.
  generalize (in_dom_dom H10).
  rewrite create_map_dom.
  rewrite H7.
  intro.
  eapply H12.
  eapply dom_in_dom.
  eassumption.
  eapply dom_in_dom.
  assumption.
  assumption.
  rewrite csemantics_bsemantics_n.
  intro.
  intros until 1.
  generalize (create_map_lookup_elim H6).
  intro; subst.
  generalize (in_dom_dom (lookup_in_dom H6)).
  rewrite create_map_dom.
  intro.
  destruct (dom_in_dom H10).
  intros.
  case_eq (callee_init a0 h x).
  intros.
  destruct (exists_behavior (@config_final _ _) (cstep (merge_embedded l2)) (Config h0 l nil)).
  esplit.
  econstructor.
  eassumption.
  eassumption.
  eassumption.
  assumption.
  intro.
  intros until 1.
  generalize (create_map_lookup_elim H6).
  intro; subst.
  generalize (in_dom_dom (lookup_in_dom H6)).
  rewrite create_map_dom.
  intro.
  destruct (dom_in_dom H10).
  intros.
  case_eq (callee_init a0 h x).
  intros.
  destruct (exists_behavior (@config_final _ _) (cstep a) (Config h0 l nil)).
  esplit.
  econstructor.
  eassumption.
  eassumption.
  eassumption.
  assumption.
  rewrite csemantics_bsemantics_n.
  intro.
  intros until 1.
  generalize (in_dom_dom H6).
  rewrite create_map_dom.
  intros.
  generalize (in_dom_dom H14).
  rewrite create_map_dom.
  intros; eauto using dom_in_dom.
  assumption.
  rewrite csemantics_bsemantics_n.
  intro.
  intros until 1.
  generalize (in_dom_dom H6).
  rewrite create_map_dom.
  intros.
  generalize (in_dom_dom H14).
  rewrite create_map_dom.
  intros.
  eapply H12.
  eapply dom_in_dom.
  rewrite <- H7.
  eassumption.
  apply dom_in_dom.
  assumption.
  assumption.
 destruct (disjoint_embedded_same_domains H0 (all_refine_same_domains H)).
 assumption.
Qed.

Lemma all_refine_merge_semantics : forall lm1 lm2,
  all_refine lm1 lm2 ->
  disjoint_embedded lm1 ->
  oR (create_map (dom (merge_embedded lm1)) (lift_sem (semantics (merge_embedded lm1)))) (create_map (dom (merge_embedded lm2)) (lift_sem (semantics (merge_embedded lm2)))).
Proof.
  intros.
  generalize (create_map_ext (fun f (Hf : _ (merge_embedded lm1) f = _) => semantics_bsemantics_stuck_closure (dom_in_dom Hf))).
  generalize (create_map_ext (fun f (Hf : _ (merge_embedded lm2) f = _) => semantics_bsemantics_stuck_closure (dom_in_dom Hf))).
  intros.
  rewrite <- H1.
  rewrite <- H2.
 generalize (disjoint_sym (stuck_closure_disjoint (m := merge_embedded lm1))).
 intro.
 refine (let J1 := _ in let J2 := _ in _ (refines_compositionality Hr H3 J1 J2 (all_refine_merge H H0))).
 intro.
 unfold lookup, lookup_strong, stuck_closure.
 intro.
 destruct (in_dom_dec (merge_embedded lm1) f).
  discriminate.
 unfold be_stuck.
 injection 1; intros; subst.
 esplit.
 reflexivity.
 intro.
 intros.
 generalize (create_map_lookup_elim H4).
 intro; subst.
 generalize (in_dom_dom (lookup_in_dom H4)).
 rewrite create_map_dom.
 intro.
 destruct (dom_in_dom H5).
 case_eq (callee_init a h x).
 intros.
 destruct (exists_behavior (@config_final _ _) (cstep (merge_embedded lm1)) (Config h0 l nil)).
 esplit.
 econstructor.
 eassumption.
 eassumption.
 eassumption.
 unfold link at 1 2.
 simpl.
 destruct 1.
 generalize (all_refine_same_domains H).
 intro.
 destruct (disjoint_embedded_same_domains H0 H6).
 split.
  rewrite create_map_dom.
  rewrite create_map_dom.
  assumption.
 intros.
 generalize (create_map_lookup_elim H9).
 intro; subst.
 generalize (create_map_lookup_elim H10).
 intro; subst.
 rewrite disj_merge_sym.
 rewrite (disj_merge_sym (M2 := stuck_closure (merge_embedded lm2))).
 eapply H5.
 eapply create_map_lookup_intro.
 eapply in_dom_dom.
 destruct (in_dom_dec (stuck_closure (merge_embedded lm1)) f).
  eapply indom_merge_indom_1.
  assumption.
  eapply indom_merge_indom_2.
  eapply dom_in_dom.
  rewrite create_map_dom.
  eapply in_dom_dom.
  eapply not_in_stuck_closure_dom.
  assumption.
 rewrite <- psistuck_stuck_closure.
 rewrite H8.
 rewrite psistuck_stuck_closure.
 eapply create_map_lookup_intro.
 eapply in_dom_dom.
 destruct (in_dom_dec (stuck_closure (merge_embedded lm2)) f).
  eapply indom_merge_indom_1.
  assumption.
  eapply indom_merge_indom_2.
  eapply dom_in_dom.
  rewrite create_map_dom.
  eapply in_dom_dom.
  eapply not_in_stuck_closure_dom.
  assumption.
 eapply stuck_closure_disjoint.
 eapply stuck_closure_disjoint.
Qed.

CPP 2015 submission: Theorem 3.


Corollary correct_compiler_semantics : forall lc, (forall c, In c lc -> correct_compiler c) ->
  forall lm1 lm2,
    compiler_source_target_list lc lm1 lm2 ->
    disjoint_embedded lm1 ->
    oR (create_map (dom (merge_embedded lm1)) (lift_sem (semantics (merge_embedded lm1)))) (create_map (dom (merge_embedded lm2)) (lift_sem (semantics (merge_embedded lm2)))).
Proof.
  intros; eauto using all_refine_merge_semantics, all_refine_correct_compilers.
Qed.

End Compilers.

End Refinement.

End Semworld.

Summarize the final statement and used axioms.
Goal True.
 idtac "Theorem:".
 Check congruence_is_compositionality.
 Print Assumptions congruence_is_compositionality.
 idtac "".
 idtac "Theorem:".
 Check correct_compiler_semantics.
 Print Assumptions correct_compiler_semantics.
 idtac "".
 trivial.
Save.