Library ccimproves
Require Export refinement.
Set Implicit Arguments.
Unset Strict Implicit.
Set Implicit Arguments.
Unset Strict Implicit.
CompCert improvement relation.
Section Event.
Variable EVENT: Type.
Variable final : Type.
Notation behavior := (behavior EVENT final) (only parsing).
Definition behavior_improves (beh1 beh2: behavior) : Prop :=
beh2 = beh1 \/ exists t, beh2 = Stuck _ t /\ exists beh', beh1 = append1_behavior t beh'.
Definition compcert_refinement (beh1 beh2: behavior -> Prop) :=
forall b1, beh1 b1 -> exists b2, beh2 b2 /\ behavior_improves b1 b2.
Section Backwards.
Section Safe.
Variables (CONFIG: Type) (config_final : CONFIG -> final -> Prop) (STEP: CONFIG -> list EVENT -> CONFIG -> Prop).
Definition safe (s: CONFIG) : Prop :=
forall s',
star STEP s nil s' ->
(exists r, config_final s' r)
\/ (exists t, exists s'', STEP s' t s'').
Lemma star_safe:
forall s s',
star STEP s nil s' -> safe s -> safe s'.
Proof.
intros; red; intros. apply H0. eapply star_trans'; eauto.
Qed.
Definition safe_along_behavior (s: CONFIG) (b: behavior) : Prop :=
forall t1 s' b2, star STEP s t1 s' -> b = append1_behavior t1 b2 ->
(exists r, config_final s' r)
\/ (exists t2, exists s'', STEP s' t2 s'').
Remark safe_along_safe:
forall s b, safe_along_behavior s b -> safe s.
Proof.
intros; red; intros. eapply H; eauto. reflexivity.
Qed.
Remark star_safe_along:
forall s b t1 s' b2,
safe_along_behavior s b ->
star STEP s t1 s' -> b = append1_behavior t1 b2 ->
safe_along_behavior s' b2.
Proof.
intros; red; intros. eapply H. eapply star_trans; eauto.
subst. rewrite append_behavior_app. reflexivity.
Qed.
Remark not_safe_along_behavior:
forall s b,
~ safe_along_behavior s b ->
exists t, exists s',
(exists beh, b = append1_behavior t beh)
/\ star STEP s t s'
/\ (forall t s'', ~(STEP s' t s''))
/\ (forall r, ~(config_final s' r)).
Proof.
intros.
destruct (not_all_ex_not _ _ H) as [t1 A]; clear H.
destruct (not_all_ex_not _ _ A) as [s' B]; clear A.
destruct (not_all_ex_not _ _ B) as [b2 C]; clear B.
destruct (imply_to_and _ _ C) as [D E]; clear C.
destruct (imply_to_and _ _ E) as [F G]; clear E.
destruct (not_or_and _ _ G) as [P Q]; clear G.
exists t1; exists s'.
split. exists b2; auto.
split. auto.
split. intros. intro. elim Q. exists t; exists s''; eauto.
intros; red; intros. elim P. exists r; auto.
Qed.
End Safe.
Variables
(CONFIG1: Type) (config_final1 : CONFIG1 -> final -> Prop) (STEP1: CONFIG1 -> list EVENT -> CONFIG1 -> Prop)
(CONFIG2: Type) (config_final2 : CONFIG2 -> final -> Prop) (STEP2: CONFIG2 -> list EVENT -> CONFIG2 -> Prop).
The general form of a backward simulation.
Record backward_simulation: Type :=
Backward_simulation {
bsim_index: Type;
bsim_order: bsim_index -> bsim_index -> Prop;
bsim_match_states :> bsim_index -> CONFIG1 -> CONFIG2 -> Prop
}.
Variable bs: backward_simulation.
Record backward_simulation_prop : Prop :=
backward_simulation_intro {
bsim_order_wf: well_founded (@bsim_order bs);
bsim_match_final_states:
forall i s1 s2 r,
(@bsim_match_states bs) i s1 s2 -> safe config_final1 STEP1 s1 -> config_final2 s2 r ->
exists s1', star STEP1 s1 nil s1' /\ config_final1 s1' r;
bsim_progress:
forall i s1 s2,
(@bsim_match_states bs) i s1 s2 -> safe config_final1 STEP1 s1 ->
(exists r, config_final2 s2 r) \/
(exists t, exists s2', STEP2 s2 t s2');
bsim_simulation:
forall s2 t s2', STEP2 s2 t s2' ->
forall i s1, (@bsim_match_states bs) i s1 s2 -> safe config_final1 STEP1 s1 ->
exists i', exists s1',
(plus STEP1 s1 t s1' \/ (star STEP1 s1 t s1' /\ bsim_order i' i))
/\ bsim_match_states i' s1' s2'
}.
Hypothesis S: backward_simulation_prop.
Lemma backward_simulation_star:
forall s2 t s2', star STEP2 s2 t s2' ->
forall i s1 b, bs i s1 s2 -> safe_along_behavior config_final1 STEP1 s1 (append1_behavior t b) ->
exists i', exists s1', star STEP1 s1 t s1' /\ bs i' s1' s2'.
Proof.
induction 1 using star_ind3_left; intros.
exists i; exists s1; split; auto. apply star_refl.
refine (_ (@bsim_simulation S _ _ _ Hstep _ _ H0 _)).
intros [i' [s1' [A B]]].
assert (star STEP1 s1 tr1 s1').
destruct A.
eright. assumption.
tauto.
refine (_ (IHstar _ _ _ B _)).
intros [i'' [s2'' [C D]]].
exists i''; exists s2''; split; auto. eapply star_trans; eauto.
eapply star_safe_along. eassumption. eassumption.
rewrite append_behavior_app. reflexivity.
eapply safe_along_safe; eauto.
Qed.
Lemma backward_simulation_forever_silent:
forall i s1 s2,
config_silently_diverges STEP2 s2 -> bs i s1 s2 -> safe config_final1 STEP1 s1 ->
config_silently_diverges STEP1 s1.
Proof.
assert (forall i s1 s2,
config_silently_diverges STEP2 s2 -> bs i s1 s2 -> safe config_final1 STEP1 s1 ->
forever_silent_N STEP1 (@bsim_order bs) i s1).
cofix COINDHYP; intros.
inversion H; subst. destruct (@bsim_simulation S _ _ _ Hstep _ _ H0 H1) as [i' [s2' [A B]]].
destruct A as [C | [C D]].
eapply forever_silent_N_plus; eauto. eapply COINDHYP; eauto.
eapply star_safe; eauto. eright; auto.
eapply forever_silent_N_star; eauto. eapply COINDHYP; eauto.
eapply star_safe; eauto.
intros. eapply forever_silent_N_forever; eauto. apply bsim_order_wf. assumption.
Qed.
Lemma backward_simulation_forever_reactive:
forall i s1 s2 T,
config_reacts STEP2 s2 T -> bs i s1 s2 -> safe_along_behavior config_final1 STEP1 s1 (Reacts _ (func_of_stream T)) ->
config_reacts STEP1 s1 T.
Proof.
cofix COINDHYP; intros. inversion H; subst.
assert (Reacts _ (func_of_stream (appinf (e :: l) il)) = append1_behavior (e :: l) (Reacts final (func_of_stream il))).
clear COINDHYP.
rewrite append_behavior.
simpl.
f_equal.
apply bisim_func_eq.
constructor.
apply bisim_appinf.
apply stream_of_func_of_stream.
rewrite H2 in H1.
destruct (backward_simulation_star Hstep H0 H1) as [i' [s1' [A B]]]; eauto.
econstructor; eauto. eapply COINDHYP; eauto. eapply star_safe_along; eauto.
Qed.
Lemma backward_simulation_state_behaves:
forall i s1 s2 beh2,
bs i s1 s2 -> config_behaviors STEP2 config_final2 s2 beh2 ->
exists beh1, config_behaviors STEP1 config_final1 s1 beh1 /\ behavior_improves beh2 beh1.
Proof.
intros. destruct (classic (safe_along_behavior config_final1 STEP1 s1 beh2)).
exists beh2; split; [idtac|left; auto].
destruct beh2.
assert (Terminates evseq hfin = append1_behavior evseq (Terminates nil hfin)).
rewrite append_behavior. simpl. rewrite <- app_nil_end. reflexivity.
rewrite H2 in H1.
inversion H0; subst.
generalize (backward_simulation_star Hterm H H1).
intros [i' [s1' [ A B ] ] ].
refine (_ (bsim_match_final_states S B _ Hfin)).
2: eapply safe_along_safe; eapply star_safe_along; eauto.
intros [s1'' [C D ] ].
econstructor. eapply star_trans'; eauto. rewrite <- app_nil_end. auto. assumption.
assert (Stuck _ evseq = append1_behavior evseq (Stuck final nil)).
rewrite append_behavior. simpl. rewrite <- app_nil_end. reflexivity.
rewrite H2 in H1.
inversion H0; subst.
generalize (backward_simulation_star Hstar H H1).
intros [i' [s1' [A B] ] ].
refine (_ (bsim_progress S B _)).
2: eapply safe_along_safe; eapply star_safe_along; eauto.
intros [ [ r FIN] | [t' [s2' JSTEP2 ] ] ].
elim (Hstuck_not_final _ FIN).
elim (Hconf_stuck _ _ JSTEP2).
assert (Diverges _ evseq = append1_behavior evseq (Diverges final nil)).
rewrite append_behavior. simpl. rewrite <- app_nil_end. reflexivity.
rewrite H2 in H1.
inversion H0; subst.
generalize (backward_simulation_star Hstar H H1).
intros [i' [s1' [A B] ] ].
econstructor. eauto. eapply backward_simulation_forever_silent; eauto.
eapply safe_along_safe. eapply star_safe_along; eauto.
simpl in *.
eapply backward_simulation_forever_reactive. eassumption. eassumption. rewrite func_of_stream_of_func. assumption.
generalize (not_safe_along_behavior H1).
intros [t [s1' [PREF [STEPS [NOSTEP NOFIN] ] ] ] ].
exists (Stuck _ t); split.
econstructor; eauto.
right. exists t; auto.
Qed.
End Backwards.
End Event.
CPP 2015 submission: Theorem 4. CompCert improvement is a refinement relation.
Congruence of CompCert behavior improvement relation
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 bstate := (bstate sw) (only parsing).
Section BSemantics_Mono.
Variables psi2 psi1: Map funcname (args -> heap -> cbehavior -> Prop).
Hypothesis psi2_improves_psi1: forall f psi2f,
lookup psi2 f psi2f ->
exists psi1f,
lookup psi1 f psi1f /\
forall a h b2, psi2f a h b2 ->
exists b1, psi1f a h b1 /\
behavior_improves b2 b1.
Hypothesis psi2_nonempty: forall f psi1f,
lookup psi1 f psi1f ->
exists psi2f,
lookup psi2 f psi2f /\
forall a h, exists b, psi2f a h b.
Inductive bsemantics_match_stack : list ((heap * ret) * cbehavior) -> list ((heap * ret) * cbehavior) -> Prop :=
| bsemantics_match_stack_nil: bsemantics_match_stack nil nil
| bsemantics_match_stack_cons hr b2 b1 (Himproves: behavior_improves b2 b1) l2 l1 (Hmatch: bsemantics_match_stack l1 l2):
bsemantics_match_stack ((hr, b1) :: l1) ((hr, b2) :: l2).
Inductive bsemantics_match_states (unused : unit) : bstate -> bstate -> Prop :=
| bsemantics_match_states_spurious: bsemantics_match_states unused (@BSpurious _) (@BSpurious _)
| bsemantics_match_states_normal b2 b1 (Himproves: behavior_improves b2 b1) l2 l1 (Hmatch: bsemantics_match_stack l1 l2) :
bsemantics_match_states unused (BBehavior b1 l1) (BBehavior b2 l2).
Definition bsemantics_backward_simulation :=
Backward_simulation (bsim_index := unit) (fun _ _ => False) bsemantics_match_states.
Lemma bsemantics_backward_simulation_prop: backward_simulation_prop (@bstate_final _) (bstep psi1) (@bstate_final _) (bstep psi2) bsemantics_backward_simulation.
Proof.
unfold bsemantics_backward_simulation.
constructor; simpl.
unfold well_founded.
intros.
constructor.
tauto.
inversion 3; subst.
inversion H; subst.
inversion Hmatch; subst.
destruct Himproves.
subst.
esplit.
split.
eleft.
assumption.
destruct H2.
destruct H2.
subst.
unfold safe in H0.
destruct H3.
rewrite append_behavior in H2.
destruct x0; simpl in *; try discriminate.
injection H2; intros; subst.
symmetry in H4.
destruct (app_eq_nil _ _ H4).
subst.
destruct (H0 _ (star_refl _ _)).
destruct H3.
inversion H3.
destruct H3.
destruct H3.
inversion H3; subst.
destruct b; discriminate.
destruct b; discriminate.
destruct b; discriminate.
inversion H; subst.
esplit.
split.
eleft.
assumption.
inversion 1; subst.
unfold safe.
intro.
left.
esplit.
econstructor.
intro.
destruct (H0 _ (star_refl _ _)).
destruct H1.
inversion H1; subst.
inversion Hmatch; subst.
inversion Himproves; subst.
eauto.
destruct H2.
destruct H2.
discriminate.
destruct H1.
destruct H1.
right.
inversion H1; subst.
assert (exists b2', b2 = cons_behavior (Event e) b2').
inversion Himproves; subst; eauto.
destruct H2.
destruct H2.
destruct b; try discriminate.
injection H2; intros; subst.
destruct H3.
subst.
simpl.
eauto.
destruct H2.
subst.
esplit.
esplit.
eapply bstep_cons.
reflexivity.
assert (exists b2', b2 = cons_behavior (External f a h1 r h2) b2').
inversion Himproves; subst; eauto.
destruct H2.
destruct H2.
destruct b; try discriminate.
injection H2; intros; subst.
destruct H3.
subst.
simpl.
eauto.
destruct H2.
subst.
esplit.
esplit.
eapply bstep_external.
2: reflexivity.
intro.
destruct H2.
destruct (psi2_improves_psi1 H2).
destruct H3.
destruct Hf.
eapply lookup_in_dom.
eassumption.
assert (exists b2', b2 = cons_behavior (External f a h1 r h2) b2').
inversion Himproves; subst; eauto.
destruct H2.
destruct H2.
destruct b; try discriminate.
injection H2; intros; subst.
destruct H3.
subst.
simpl.
eauto.
destruct H2.
subst.
destruct (psi2_nonempty Hf).
destruct H2.
destruct (H3 a h1).
esplit.
esplit.
eapply bstep_internal.
eassumption.
eassumption.
reflexivity.
inversion Hmatch; subst.
inversion Himproves; subst.
esplit.
esplit.
eapply bstep_terminates.
destruct H2.
destruct H2.
discriminate.
inversion Hmatch; subst.
inversion Himproves; subst.
esplit.
esplit.
eapply bstep_spurious.
assumption.
destruct H2.
destruct H2.
discriminate.
inversion Himproves; subst.
esplit.
esplit.
eapply bstep_diverges.
destruct H2.
destruct H2.
discriminate.
inversion 1; subst; inversion 1; subst; intros.
assert (exists b'1, b1 = cons_behavior (Event e) b'1 /\ behavior_improves b b'1).
inversion Himproves; subst.
esplit.
split.
reflexivity.
left.
reflexivity.
destruct H2.
destruct H2.
subst.
destruct H3.
destruct x.
destruct (H1 _ (star_refl _ _)).
destruct H3.
inversion H3.
destruct H3.
destruct H3.
inversion H3; subst.
destruct b0; discriminate.
destruct b0; discriminate.
destruct b0; discriminate.
simpl in H2.
generalize (cons_behavior_inj H2).
injection 1; intros; subst.
eexists (Stuck _ _).
split.
simpl.
reflexivity.
right.
esplit.
split.
reflexivity.
eauto.
destruct H2.
destruct H2.
subst.
exists tt.
esplit.
split.
left.
eleft.
eapply bstep_cons.
reflexivity.
econstructor.
assumption.
assumption.
assert (exists b'1, b1 = cons_behavior (External f a h1 r h2) b'1 /\ behavior_improves b b'1 ).
inversion Himproves; subst.
esplit.
split.
reflexivity.
left.
reflexivity.
destruct H2.
destruct H2.
subst.
destruct H3.
destruct x.
destruct (H1 _ (star_refl _ _)).
destruct H3.
inversion H3.
destruct H3.
destruct H3.
inversion H3; subst.
destruct b0; discriminate.
destruct b0; discriminate.
destruct b0; discriminate.
simpl in H2.
generalize (cons_behavior_inj H2).
injection 1; intros; subst.
eexists (Stuck _ _).
split.
simpl.
reflexivity.
right.
esplit.
split.
reflexivity.
eauto.
destruct H2.
destruct H2.
subst.
exists tt.
esplit.
split.
left.
eleft.
eapply bstep_external.
intro.
destruct H2.
destruct (psi2_nonempty H2).
destruct H4.
destruct Hf.
eapply lookup_in_dom.
eassumption.
reflexivity.
econstructor.
assumption.
assumption.
assert (exists b'1, b1 = cons_behavior (External f a h1 r h2) b'1 /\ behavior_improves b b'1).
inversion Himproves; subst.
esplit.
split.
reflexivity.
left.
reflexivity.
destruct H2.
destruct H2.
subst.
destruct H3.
destruct x.
destruct (H1 _ (star_refl _ _)).
destruct H3.
inversion H3.
destruct H3.
destruct H3.
inversion H3; subst.
destruct b0; discriminate.
destruct b0; discriminate.
destruct b0; discriminate.
simpl in H2.
generalize (cons_behavior_inj H2).
injection 1; intros; subst.
eexists (Stuck _ _).
split.
simpl.
reflexivity.
right.
esplit.
split.
reflexivity.
eauto.
destruct H2.
destruct H2.
subst.
exists tt.
destruct (psi2_improves_psi1 Hf).
destruct H2.
destruct (H4 _ _ _ Hbf).
destruct H5.
esplit.
split.
left.
eleft.
eapply bstep_internal.
eassumption.
eassumption.
reflexivity.
econstructor.
assumption.
econstructor.
assumption.
assumption.
inversion Hmatch; subst.
inversion Himproves; subst.
exists tt.
esplit.
split.
left.
eleft.
eapply bstep_terminates.
econstructor.
assumption.
assumption.
destruct H2.
destruct H2.
subst.
destruct H3.
rewrite append_behavior in H2.
destruct x0; try discriminate.
simpl in H2.
injection H2; intros; subst.
symmetry in H4.
destruct (app_eq_nil _ _ H4).
subst.
destruct (H1 _ (star_refl _ _)).
destruct H3.
inversion H3.
destruct H3.
destruct H3.
inversion H3; subst; destruct b; discriminate.
inversion Hmatch; subst.
inversion Himproves; subst.
exists tt.
esplit.
split.
left.
eleft.
eapply bstep_spurious.
assumption.
constructor.
destruct H2.
destruct H2.
subst.
destruct H3.
rewrite append_behavior in H2.
destruct x0; try discriminate.
simpl in H2.
injection H2; intros; subst.
symmetry in H4.
destruct (app_eq_nil _ _ H4).
subst.
destruct (H1 _ (star_refl _ _)).
destruct H3.
inversion H3.
destruct H3.
destruct H3.
inversion H3; subst; destruct b; discriminate.
inversion Himproves; subst.
exists tt.
esplit.
split.
left.
eleft.
eapply bstep_diverges.
constructor.
left.
reflexivity.
assumption.
destruct H2.
destruct H2.
subst.
destruct H3.
rewrite append_behavior in H2.
destruct x0; try discriminate.
simpl in H2.
injection H2; intros; subst.
symmetry in H3.
destruct (app_eq_nil _ _ H3).
subst.
destruct (H1 _ (star_refl _ _)).
destruct H4.
inversion H4.
destruct H4.
destruct H4.
inversion H4; subst; destruct b; discriminate.
Qed.
Theorem bsemantics_mono: forall f a h b2,
bsemantics psi2 f a h b2 ->
exists b1, bsemantics psi1 f a h b1 /\
behavior_improves b2 b1.
Proof.
inversion 1; subst.
destruct (psi2_improves_psi1 Hcmd).
destruct H0.
destruct (H1 _ _ _ Hcb).
destruct H2.
refine (_ (backward_simulation_state_behaves (bsemantics_backward_simulation_prop) (i := tt) (s1 := BBehavior x0 nil) _ Hbeh)).
destruct 1.
destruct H4.
assert (exists b1, x1 = bbehavior_of_cbehavior b1 /\ behavior_improves b2 b1).
inversion H5; subst.
unfold behavior_improves; eauto.
destruct H6.
destruct H6.
destruct H7.
symmetry in H7.
destruct (bbehavior_of_cbehavior_append_ex H7).
subst.
eexists (Stuck _ _).
simpl.
split.
reflexivity.
right.
esplit.
split.
reflexivity.
eauto.
destruct H6.
destruct H6.
subst.
esplit.
split.
econstructor.
eassumption.
eassumption.
3: eassumption.
reflexivity.
assumption.
constructor.
assumption.
constructor.
Qed.
End BSemantics_Mono.
Theorem compcert_refinement_is_refinement: refinement (compcert_refinement (EVENT := xevent) (final := heap * ret)).
Proof.
split; unfold compcert_refinement.
unfold reflexive.
intros.
esplit.
split.
eassumption.
left.
reflexivity.
intro.
intros.
destruct (H _ H1).
destruct H2.
destruct (H0 _ H2).
destruct H4.
esplit.
split.
eassumption.
destruct H3.
subst.
assumption.
destruct H5.
subst.
right.
assumption.
destruct H3.
destruct H3.
destruct H6.
subst.
destruct H5.
destruct H3.
destruct H5.
subst.
right.
esplit.
split.
reflexivity.
rewrite append_behavior in H5.
destruct x4; try discriminate.
injection H5; intros; subst.
exists (append1_behavior evseq x3).
rewrite append_behavior_app.
reflexivity.
intro.
unfold oR.
simpl.
intros.
repeat rewrite create_map_dom in *.
destruct H0.
split.
assumption.
intros.
generalize (create_map_lookup_elim H2).
intro; subst.
generalize (create_map_lookup_elim H3).
intro; subst.
generalize (in_dom_dom (lookup_in_dom H2)).
generalize (in_dom_dom (lookup_in_dom H3)).
repeat rewrite create_map_dom.
intros.
destruct (dom_in_dom H5).
destruct (dom_in_dom H6).
eapply bsemantics_mono.
3: eassumption.
intros.
generalize (in_dom_dom (lookup_in_dom H9)).
rewrite H0.
intro.
destruct (dom_in_dom H10).
esplit.
split.
eassumption.
eauto.
intros.
generalize (in_dom_dom (lookup_in_dom H9)).
rewrite <- H0.
intro.
destruct (dom_in_dom H10).
intros.
esplit.
split.
eassumption.
eauto.
Qed.
End Semworld.
Summarize the final statement and used axioms.
Goal True.
idtac "Theorem:".
Check compcert_refinement_is_refinement.
Print Assumptions compcert_refinement_is_refinement.
idtac "".
trivial.
Save.
idtac "Theorem:".
Check compcert_refinement_is_refinement.
Print Assumptions compcert_refinement_is_refinement.
idtac "".
trivial.
Save.