Library link
Require Export resolution tac.
Set Implicit Arguments.
Unset Strict Implicit.
Set Implicit Arguments.
Unset Strict Implicit.
CPP 2015 submission: Theorem 1.
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 Bstep_properties.
Variable psi: Map funcname (args -> heap -> cbehavior -> Prop).
Lemma bstar_append:
forall l (Hl:
forall f (Hf: in_dom psi f) a h1 r h2, List.In (External f a h1 r h2) l -> False)
cb s,
star (bstep psi) (BBehavior (append1_behavior l cb) s) l (BBehavior cb s).
Proof.
induction l; simpl.
constructor.
intros.
destruct a.
eapply star_trans'.
eright.
eleft.
eapply bstep_cons.
reflexivity.
eapply IHl.
intros.
eauto.
reflexivity.
eapply star_trans'.
eright.
eleft.
eapply bstep_external.
intro.
2: reflexivity.
eauto.
eapply IHl.
intros.
eauto.
reflexivity.
Qed.
Lemma bstep_stack: forall b1 s1 l b2 s2
(Hstep: bstep psi (BBehavior b1 s1) l (BBehavior b2 s2))
s,
bstep psi (BBehavior b1 (s1 ++ s)) l (BBehavior b2 (s2 ++ s)).
Proof.
inversion 1; subst.
constructor. reflexivity.
constructor. assumption. reflexivity.
simpl. econstructor. eassumption. eassumption. reflexivity.
simpl. constructor.
constructor.
Qed.
Corollary bstepn_stack: forall n b1 s1 l b2 s2
(Hstepn: stepn_left (bstep psi) n (BBehavior b1 s1) l (BBehavior b2 s2)) s,
stepn_left (bstep psi) n (BBehavior b1 (s1 ++ s)) l (BBehavior b2 (s2 ++ s)).
Proof.
induction n; inversion 1; subst.
constructor.
destruct conf1.
inversion Hstepn0; subst. inversion Hstep0.
econstructor; eauto using bstep_stack.
Qed.
End Bstep_properties.
Section Language.
Variable lang: language sw.
Notation local := (local lang) (only parsing).
Notation frame := (frame lang) (only parsing).
Section Properties.
Variable m: module lang.
Lemma step_stack: forall h l s tr h' l' s'
(Hstep: step m (Config h l s) tr (Config h' l' s'))
k,
step m (Config h l (s ++ k)) tr (Config h' l' (s' ++ k)).
Proof.
inversion 1; subst.
intros.
eapply step_internal.
eassumption.
trivial.
intros.
simpl.
econstructor.
eassumption.
eassumption.
eassumption.
eassumption.
intros.
simpl.
econstructor.
eassumption.
assumption.
Qed.
Lemma step_stack_recip: forall h l s tr h' l' s',
step m (Config h l s) tr (Config h' l' s') ->
(exists s1, s' = s1 ++ s /\ step m (Config h l nil) tr (Config h' l' s1)) \/ (tr = nil /\ exists r, local_kind l = Some (Return r) /\
exists fr, s = fr :: s' /\
caller_return r h fr = (h', l')
).
Proof.
inversion 1; subst.
left; exists nil.
split.
reflexivity.
econstructor.
eassumption.
reflexivity.
left; exists (fr :: nil).
split.
reflexivity.
econstructor.
eassumption.
eassumption.
eassumption.
assumption.
right.
eauto 6.
Qed.
Lemma cstep_stack_recip: forall h l s tr h' l' s',
cstep m (Config h l s) tr (Config h' l' s') ->
(exists s1, s' = s1 ++ s /\ cstep m (Config h l nil) tr (Config h' l' s1)) \/ (tr = nil /\ exists r, local_kind l = Some (Return r) /\
exists fr, s = fr :: s' /\
caller_return r h fr = (h', l')
).
Proof.
inversion 1; subst.
destruct (step_stack_recip Hstep).
destruct H0.
destruct H0.
left.
esplit.
split.
eassumption.
econstructor.
eassumption.
reflexivity.
destruct H0.
subst; simpl; eauto.
left; exists nil.
split.
reflexivity.
eapply cstep_call_external.
assumption.
assumption.
eassumption.
assumption.
assumption.
Qed.
Lemma cstep_stack: forall h l s tr h' l' s'
(Hstep: cstep m (Config h l s) tr (Config h' l' s'))
k,
cstep m (Config h l (s ++ k)) tr (Config h' l' (s' ++ k)).
Proof.
inversion 1; subst.
intros.
eapply cstep_step.
eapply step_stack.
eassumption.
reflexivity.
intros.
eapply cstep_call_external.
assumption.
assumption.
eassumption.
assumption.
assumption.
Qed.
Corollary cstepn_stack: forall n h l s tr h' l' s'
(Hstep: stepn_left (cstep m) n (Config h l s) tr (Config h' l' s'))
k,
stepn_left (cstep m) n (Config h l (s ++ k)) tr (Config h' l' (s' ++ k)).
Proof.
induction n; inversion 1; subst.
constructor.
intros.
destruct conf1.
econstructor.
eapply cstep_stack.
eassumption.
eapply IHn.
eassumption.
reflexivity.
Qed.
Lemma cstepn_stack_recip: forall n h l s tr h' l' s'
(Hstepn: stepn_left (cstep m) n (Config h l s) tr (Config h' l' s'))
s0 s1 (Hs: s = s1 ++ s0),
(exists s2, s' = s2 ++ s0 /\ stepn_left (cstep m) n (Config h l s1) tr (Config h' l' s2)) \/
exists n1, exists h1, exists l1, exists tr1,
stepn_left (cstep m) n1 (Config h l s1) tr1 (Config h1 l1 nil) /\
exists r, local_kind l1 = Some (Return r) /\
exists fr, exists s'0, s0 = fr :: s'0 /\
exists h2, exists l2, caller_return r h1 fr = (h2, l2) /\
exists n2, exists tr2, stepn_left (cstep m) n2 (Config h2 l2 s'0) tr2 (Config h' l' s') /\
tr = tr1 ++ tr2 /\
n = S (n1 + n2)
.
Proof.
induction n; inversion 1; subst.
intros; subst.
left.
esplit.
split.
reflexivity.
constructor.
clear Hstepn.
destruct conf1.
destruct (cstep_stack_recip Hstep).
destruct H.
destruct H.
subst.
intros; subst.
assert (x ++ s1 ++ s0 = ((x ++ s1) ++ s0)).
rewrite app_ass.
reflexivity.
destruct (IHn _ _ _ _ _ _ _ Hstepn0 _ _ H); clear Hstepn0.
destruct H1.
destruct H1.
left.
esplit.
split.
eassumption.
change s1 with (nil ++ s1).
econstructor.
eapply cstep_stack.
eassumption.
eassumption.
reflexivity.
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H2.
destruct H2.
destruct H3.
destruct H3.
destruct H3.
destruct H4.
destruct H4.
destruct H4.
destruct H5.
destruct H5.
destruct H5.
destruct H6.
right.
change s1 with (nil ++ s1).
esplit.
esplit.
esplit.
esplit.
split.
eright.
eapply cstep_stack.
eassumption.
eassumption.
reflexivity.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
split.
subst.
rewrite app_ass.
reflexivity.
subst.
reflexivity.
destruct H.
destruct H0.
destruct H0.
destruct H1.
destruct H1.
intros.
subst.
destruct s1.
simpl in *.
subst.
right.
esplit.
esplit.
esplit.
esplit.
split.
eleft.
esplit.
split.
eassumption.
esplit.
esplit.
split.
reflexivity.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
split.
reflexivity.
reflexivity.
injection Hs; intros; subst.
assert (cstep m (Config h l (f :: nil)) nil (Config c_heap c_local nil)).
econstructor.
econstructor.
eassumption.
eassumption.
reflexivity.
destruct (IHn _ _ _ _ _ _ _ Hstepn0 _ _ (refl_equal _)).
destruct H1.
destruct H1.
subst.
left.
esplit.
split.
reflexivity.
change (f :: s1) with ((f :: nil) ++ s1).
econstructor.
eapply cstep_stack.
eassumption.
eassumption.
reflexivity.
right.
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H3.
destruct H3.
destruct H4.
destruct H4.
destruct H4.
destruct H5.
destruct H5.
destruct H5.
destruct H6.
destruct H6.
destruct H6.
destruct H7.
change (f :: s1) with ((f :: nil) ++ s1).
esplit.
esplit.
esplit.
esplit.
split.
eright.
eapply cstep_stack.
eassumption.
eassumption.
simpl; reflexivity.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
split.
assumption.
subst.
reflexivity.
Qed.
Lemma cstep_le:
forall h l s tr h' l' s'
(Hstep: cstep m (Config h l s) tr (Config h' l' s')),
heap_le h h'.
Proof.
inversion 1; subst; eauto using step_le, caller_call_le, caller_return_le, heap_le_trans.
Qed.
Corollary cstepn_le:
forall n h l s tr h' l' s'
(Hstep: stepn_left (cstep m) n (Config h l s) tr (Config h' l' s')),
heap_le h h'.
Proof.
induction n; inversion 1; subst; eauto using heap_le_refl.
destruct conf1; eauto using heap_le_trans, cstep_le.
Qed.
Require Export smallstep_guided.
Lemma guided_cstepn_uncons:
forall n h l fr s g tr h' l' sf s' g',
stepn_left (@guided_step _ _) n (Config h l (sf ++ fr :: s), g) tr (Config h' l' s', g') ->
guided_execution_correct (cstep m) (Config h l (sf ++ fr :: s), g) ->
(forall s'', s' <> s'' ++ fr :: s) ->
exists n1, exists h1, exists l1, exists tr1, exists g1,
stepn_left (@guided_step _ _) n1 (Config h l (sf ++ fr :: s), g) tr1 (Config h1 l1 (fr :: s), g1) /\
stepn_left (cstep m) n1 (Config h l sf) tr1 (Config h1 l1 nil) /\
exists h2, exists l2, exists g2,
guided_step (Config h1 l1 (fr :: s), g1) nil (Config h2 l2 s, g2) /\
n1 < n /\
forall n3, n3 <= n1 ->
forall h3 l3 s3 g3 tr3,
stepn_left (@guided_step _ _) n3 (Config h l (sf ++ fr :: s), g) tr3 (Config h3 l3 s3, g3) ->
exists s'3, s3 = s'3 ++ fr :: s
.
Proof.
induction n; inversion 1; subst.
intros.
edestruct H1.
reflexivity.
inversion Hstep; subst.
inversion 1; subst.
destruct s'0.
destruct (cstep_stack_recip Hstep0).
destruct H1.
destruct H1.
subst.
intros.
rewrite <- app_ass in Hstepn.
rewrite <- app_ass in HCOIND.
destruct (IHn _ _ _ _ _ _ _ _ _ _ _ Hstepn HCOIND H1).
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H4.
destruct H5.
destruct H5.
destruct H5.
destruct H5.
destruct H6.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eright.
econstructor.
rewrite <- app_ass.
eassumption.
reflexivity.
split.
change sf with (nil ++ sf).
econstructor.
eapply cstep_stack.
eassumption.
eassumption.
reflexivity.
esplit.
esplit.
esplit.
split.
eassumption.
split.
omega.
intros.
inversion H9; subst.
eauto.
assert (n0 <= x0) by omega.
inversion Hstep1; subst.
rewrite app_ass in H7.
eauto.
destruct H1.
destruct H2.
destruct H2.
destruct H3.
destruct H3.
destruct sf; injection H3; intros; subst; simpl in *.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eleft.
split.
left.
esplit.
esplit.
esplit.
split.
econstructor.
split.
omega.
inversion 2; subst.
exists nil; reflexivity.
exfalso; omega.
destruct (IHn _ _ _ _ _ _ _ _ _ _ _ Hstepn HCOIND H7).
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H5.
destruct H6.
destruct H6.
destruct H6.
destruct H6.
destruct H8.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eright.
econstructor.
eassumption.
simpl; reflexivity.
split.
econstructor.
econstructor.
eapply step_return.
eassumption.
eassumption.
simpl; reflexivity.
eassumption.
reflexivity.
esplit.
esplit.
esplit.
split.
eassumption.
split.
omega.
inversion 2; subst.
exists (x0 :: sf); reflexivity.
inversion Hstep1; subst.
assert (n0 <= x1) by omega.
eauto.
Qed.
Lemma cstep_trace_inv: forall s tr s',
cstep m s tr s' ->
forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False.
Proof.
inversion 1; subst.
clear H Hstep.
intros until h2.
induction tr0; simpl.
tauto.
destruct 1.
discriminate.
contradiction.
inversion 2; subst.
injection H1; intros; subst.
contradiction.
assumption.
Qed.
Corollary cstarn_trace_inv: forall s tr s',
star (cstep m) s tr s' ->
forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False.
Proof.
induction 1 using star_ind3_left.
simpl; tauto.
intros.
destruct (in_app_or _ _ _ H0).
eauto using cstep_trace_inv.
eauto.
Qed.
End Properties.
Section Link.
Variables m1 m2: module lang.
Hypothesis Hdisj: forall f, in_dom m1 f -> in_dom m2 f -> False.
First part of the proof: bsemantics (merge (csemantics m1) (csemantics m2)) is included in csemantics (merge m1 m2)
Lemma step_merge:
forall m (Hm: m = m1 \/ m = m2) s1 l s2
(Hstep: step m s1 l s2),
step (merge m1 m2) s1 l s2.
Proof.
inversion 2; subst.
eapply step_internal. eassumption. reflexivity.
eapply step_call. eassumption. eapply lookup_disj_merge. eassumption. eassumption. assumption. eassumption. assumption.
eapply step_return. eassumption. assumption.
Qed.
Lemma cstep_merge:
forall m (Hm: m = m1 \/ m = m2) s1 l s2
(Hstep: cstep m s1 l s2)
(Hl: forall f a h1 r h2, List.In (External f a h1 r h2) l -> in_dom (merge m1 m2) f -> False),
cstep (merge m1 m2) s1 l s2.
Proof.
inversion 2; intros; subst.
eapply cstep_step.
eapply step_merge.
eassumption.
eassumption.
reflexivity.
eapply cstep_call_external.
assumption.
intro. eapply Hl. 2: eassumption. left. reflexivity.
eassumption.
assumption.
assumption.
Qed.
Lemma cstepn_merge:
forall m (Hm: m = m1 \/ m = m2) n s1 l s2
(Hstep: stepn_left (cstep m) n s1 l s2)
(Hl: forall f a h1 r h2, List.In (External f a h1 r h2) l -> in_dom (merge m1 m2) f -> False),
stepn_left (cstep (merge m1 m2)) n s1 l s2.
Proof.
induction 2.
intros; constructor.
intros.
subst.
econstructor.
eapply cstep_merge.
eassumption.
eassumption.
intros until 1.
eapply Hl.
eapply in_or_app. left. eassumption.
eapply IHHstep.
intros until 1.
eapply Hl.
eapply in_or_app. right. eassumption.
reflexivity.
Qed.
Lemma cstepm_disj: forall m (Hm: m = m1 \/ m = m2) c_heap0 c_local0 c_stack0 tr1 c_heap1 c_local1 c_stack1 (Hstep: cstep (merge m1 m2) (Config c_heap0 c_local0 c_stack0) tr1 (Config c_heap1 c_local1 c_stack1)),
(cstep m (Config c_heap0 c_local0 c_stack0) tr1 (Config c_heap1 c_local1 c_stack1) /\
forall f (Hf: in_dom (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))) f) a h1 r h2, List.In (External f a h1 r h2) tr1 -> False)
\/ tr1 = nil /\ exists f, exists a, local_kind c_local0 = Some (Call f a) /\ ~ in_dom m f /\ exists m', (m' = m1 \/ m' = m2) /\ exists c, lookup m' f c /\ exists h', exists fr, caller_call a c_heap0 c_local0 = (h', fr) /\ callee_init a h' c = (c_heap1, c_local1) /\ c_stack1 = fr :: c_stack0.
Proof.
inversion 2; subst.
inversion Hstep0; subst.
left. split. econstructor. econstructor. eassumption. reflexivity. reflexivity.
clear Hstep1 Hstep0 Hstep.
induction tr0; simpl.
destruct 2; try discriminate; assumption.
intros; assumption.
destruct (in_dom_dec m func).
destruct i.
generalize (lookup_disj_merge Hm H Hdisj).
intros.
generalize (lookup_func H0 Hfunc).
intro; subst.
left.
split.
eapply cstep_step.
eapply step_call.
eassumption.
eassumption.
eassumption.
assumption.
reflexivity.
simpl; intros; assumption.
destruct (mergelookup' Hfunc).
destruct H.
simpl.
eauto 15.
left.
split.
econstructor.
eapply step_return.
eassumption.
assumption.
reflexivity.
simpl; intros; assumption.
left.
split.
eapply cstep_call_external.
assumption.
intro.
eapply Hf.
eapply indom_merge.
eassumption.
assumption.
eassumption.
assumption.
assumption.
inversion 2; subst; try assumption.
injection H0; intros; subst.
apply Hf.
generalize (in_dom_dom Hf0).
rewrite dom_merge.
rewrite create_map_dom.
rewrite create_map_dom.
rewrite <- dom_merge.
intro.
eapply dom_in_dom.
assumption.
Qed.
Inductive stack_invariant : list ((heap * ret) * cbehavior) -> list frame -> Prop :=
| stack_invariant_nil: stack_invariant nil nil
| stack_invariant_cons h r fr h' l'
(Hh: caller_return r h fr = (h', l'))
m (Hm: m = m1 \/ m = m2)
k' b' (Hbeh: config_behaviors' (cstep m) (@config_final _ _) (Config h' l' k') b')
s'' k''
(Hstack: stack_invariant s'' k'')
k (Hk: k = fr :: k' ++ k''):
stack_invariant (((h, r), b') :: s'') k.
Inductive invariant: bstate sw -> config lang -> Prop :=
| invariant_normal
m (Hm: m = m1 \/ m = m2)
b h l k'
(Hbeh: config_behaviors' (cstep m) (@config_final _ _) (Config h l k') b)
s'' k''
(Hstack: stack_invariant s'' k'')
k (Hk: k = k' ++ k''):
invariant (BBehavior b s'') (Config h l k)
| invariant_spurious c:
invariant (@BSpurious _) c.
Lemma invariant_forward: forall bs c
(Hinv: invariant bs c)
tr bs' (Hstep: bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))) bs tr bs'),
exists c', plus (cstep (merge m1 m2)) c tr c' /\
invariant bs' c'.
Proof.
inversion 2; subst; inversion Hinv; subst; clear Hinv.
inversion Hbeh; subst; clear Hbeh.
functional inversion H1.
functional inversion H1.
functional inversion H1.
simpl in Hb'.
generalize (cons_behavior_inj Hb').
injection 1; intros; subst.
inversion Hcons; subst.
destruct tr; try discriminate.
injection Htr'; intros; subst.
inversion Hstep0; subst.
functional inversion Htr'0; subst.
destruct (star_stepn_left Hnil); clear Hnil.
esplit.
split.
eapply stepn_left_plus.
eapply stepn_right_left.
econstructor.
eapply stepn_left_right.
eapply cstepn_stack.
eapply cstepn_merge.
eassumption.
eassumption.
inversion 1.
eapply cstep_stack.
eapply cstep_step.
eapply step_internal.
eassumption.
simpl; reflexivity.
simpl; reflexivity.
reflexivity.
simpl.
econstructor.
eassumption.
eassumption.
eassumption.
reflexivity.
inversion Hbeh; subst; clear Hbeh.
functional inversion H1.
functional inversion H1.
functional inversion H1.
simpl in Hb'.
generalize (cons_behavior_inj Hb'); clear Hb'.
injection 1; intros; subst.
inversion Hcons; subst.
destruct tr; discriminate.
destruct (star_stepn_left Hnil); clear Hnil.
esplit.
split.
eapply stepn_left_plus.
eapply stepn_right_left.
econstructor.
eapply stepn_left_right.
eapply cstepn_stack.
eapply cstepn_merge.
eassumption.
eassumption.
inversion 1.
eapply cstep_call_external.
eassumption.
intro.
eapply Hf.
eapply dom_in_dom.
rewrite dom_merge.
apply or_bapp.
rewrite create_map_dom.
rewrite create_map_dom.
destruct (merge_indom H1); eauto using in_dom_dom.
eassumption.
eassumption.
eassumption.
reflexivity.
econstructor.
eassumption.
simpl; eassumption.
eassumption.
reflexivity.
assert (exists m', csemantics m' f a h1 bf /\ (m' = m1 \/ m' = m2)).
destruct (mergelookup Hf).
generalize (create_map_lookup_elim H).
generalize (lookup_in_dom H).
intros; subst; eauto.
generalize (create_map_lookup_elim H).
generalize (lookup_in_dom H).
intros; subst; eauto.
destruct H.
destruct H.
inversion H; subst.
assert (lookup (merge m1 m2) f cmd).
destruct H0; subst.
eapply lookup_merge_lookup_1.
assumption.
eapply lookup_merge_lookup_2.
intro.
eapply Hdisj.
eassumption.
eapply lookup_in_dom.
eassumption.
assumption.
inversion Hbeh; subst; clear Hbeh.
functional inversion H4.
functional inversion H4.
functional inversion H4.
simpl in Hb'.
generalize (cons_behavior_inj Hb'); clear Hb'.
injection 1; intros; subst.
inversion Hcons; subst.
destruct tr; discriminate.
destruct (star_stepn_left Hnil); clear Hnil.
esplit.
split.
eapply stepn_left_plus.
eapply stepn_right_left.
econstructor.
eapply stepn_left_right.
eapply cstepn_stack.
eapply cstepn_merge.
2: eassumption.
assumption.
inversion 1.
eapply cstep_step.
eapply step_call.
eassumption.
eassumption.
eassumption.
eassumption.
simpl; reflexivity.
reflexivity.
econstructor.
2: eapply config_behaviors_behaviors'.
2: eassumption.
assumption.
econstructor.
eassumption.
2: eassumption.
assumption.
eassumption.
reflexivity.
reflexivity.
inversion Hstack; subst.
inversion Hbeh; subst.
inversion Hconf; subst.
inversion Hfin; subst.
destruct (star_stepn_left Hterm); clear Hterm.
esplit.
split.
eapply stepn_left_plus.
eapply stepn_right_left.
econstructor.
eapply stepn_left_right.
eapply cstepn_stack.
eapply cstepn_merge.
2: eassumption.
assumption.
inversion 1.
simpl.
eapply cstep_step.
eapply step_return.
eassumption.
eassumption.
simpl; reflexivity.
reflexivity.
econstructor.
2: eassumption.
assumption.
eassumption.
reflexivity.
simpl in Hb'.
functional inversion Hb'.
inversion Hstack; subst.
inversion Hbeh; subst.
inversion Hconf; subst.
inversion Hfin; subst.
destruct (star_stepn_left Hterm); clear Hterm.
case_eq (caller_return r0 h1 fr).
intros.
esplit.
split.
eapply stepn_left_plus.
eapply stepn_right_left.
econstructor.
eapply stepn_left_right.
eapply cstepn_stack.
eapply cstepn_merge.
2: eassumption.
assumption.
inversion 1.
simpl.
eapply cstep_step.
eapply step_return.
eassumption.
eassumption.
simpl; reflexivity.
reflexivity.
econstructor.
simpl in Hb'.
functional inversion Hb'.
inversion Hbeh; subst; clear Hbeh.
generalize (config_diverges_nil_silently_diverges Hconf); clear Hconf.
inversion 1; subst.
destruct conf'.
esplit.
split.
eleft.
eapply cstep_stack.
eapply cstep_merge.
2: eassumption.
assumption.
inversion 1.
econstructor.
eassumption.
econstructor.
econstructor.
eleft.
eassumption.
eassumption.
reflexivity.
simpl in Hb'.
functional inversion Hb'.
Qed.
Corollary invariant_forward_plus: forall bs c
(Hinv: invariant bs c)
tr bs' (Hstep: plus (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) bs tr bs'),
exists c', plus (cstep (merge m1 m2)) c tr c' /\
invariant bs' c'.
Proof.
intros.
revert c Hinv.
induction Hstep.
eauto using invariant_forward.
intros.
edestruct IHHstep1.
eassumption.
destruct H.
edestruct IHHstep2.
eassumption.
destruct H1.
esplit.
split.
eapply plus_trans.
eassumption.
eassumption.
assumption.
Qed.
Corollary invariant_forward_star: forall bs c
(Hinv: invariant bs c)
tr bs' (Hstep: star (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) bs tr bs'),
exists c', star (cstep (merge m1 m2)) c tr c' /\
invariant bs' c'.
Proof.
inversion 2; subst.
esplit.
split.
eleft.
assumption.
edestruct invariant_forward_plus.
eassumption.
eassumption.
destruct H.
esplit.
split.
eright.
eassumption.
assumption.
Qed.
Require Export smallstep_exists_beh.
Lemma invariant_forward_close: forall bs c
(Hinv: invariant bs c)
bbeh (Hbeh: config_behaviors' (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))))
(@bstate_final _) bs bbeh)
beh (Hbeh: bbeh = bbehavior_of_cbehavior beh),
config_behaviors_star (cstep (merge m1 m2)) (@config_final _ _) c beh
.
Proof.
cofix COINDHYP.
inversion 2; subst.
clear COINDHYP.
functional inversion 1; subst.
inversion Hconf; subst.
inversion Hfin; subst.
edestruct invariant_forward_star.
eassumption.
eassumption.
destruct H.
inversion H0; subst.
inversion Hstack; subst.
rewrite <- app_nil_end in *.
generalize (config_behaviors_star_config_behaviors (config_behaviors'_config_behaviors_star Hbeh1)); clear Hbeh1.
inversion 1; subst.
inversion Hfin0; subst.
edestruct (star_stepn_left Hterm0); clear Hterm0.
eapply config_behaviors_star_terminates.
econstructor.
eapply star_trans'.
eassumption.
eapply stepn_left_star.
eapply cstepn_merge.
eassumption.
eassumption.
inversion 1.
reflexivity.
assumption.
clear COINDHYP.
functional inversion 1; subst.
inversion Hconf; subst.
edestruct invariant_forward_star.
eassumption.
eassumption.
destruct H.
inversion H0; subst.
generalize (behavior_elim_ex b).
inversion 1; subst.
exfalso.
destruct s''.
destruct r.
edestruct Hstuck_not_final.
econstructor.
destruct p.
destruct p.
destruct r.
destruct (classic ((h1,r) = (h0,r0))).
injection H1; intros; subst.
eapply Hconf_stuck.
eapply bstep_terminates.
eapply Hconf_stuck.
eapply bstep_spurious.
assumption.
generalize (config_behaviors_star_config_behaviors (config_behaviors'_config_behaviors_star Hbeh1)); clear Hbeh1.
inversion 1; subst.
destruct cstuck.
edestruct (star_stepn_left Hstar0); clear Hstar0.
eapply config_behaviors_star_stuck.
econstructor.
eapply star_trans'.
eassumption.
eapply stepn_left_star.
eapply cstepn_stack.
eapply cstepn_merge.
eassumption.
eassumption.
inversion 1.
reflexivity.
destruct k''.
rewrite <- app_nil_end.
assumption.
destruct c_stack; simpl; inversion 1.
inversion 1; subst.
inversion Hstep; subst.
eapply Hconf_stuck0.
eapply cstep_step.
eapply step_internal.
eassumption.
reflexivity.
reflexivity.
destruct (in_dom_dec m func).
destruct i.
case_eq (callee_init ar h' x0).
intros.
eapply Hconf_stuck0.
eapply cstep_step.
eapply step_call.
eassumption.
eassumption.
eassumption.
eassumption.
simpl; reflexivity.
destruct (ret_inh sw).
case_eq (caller_return X0 h' fr).
intros.
eapply Hconf_stuck0.
eapply cstep_call_external.
eassumption.
assumption.
eassumption.
eapply heap_le_refl.
eassumption.
destruct c_stack.
eapply Hstuck_not_final0.
econstructor.
eassumption.
injection H6; intros; subst.
eapply Hconf_stuck0.
eapply cstep_step.
eapply step_return.
eassumption.
eassumption.
simpl; reflexivity.
eapply Hconf_stuck0.
eapply cstep_call_external.
eassumption.
intro.
eapply Hf.
eapply indom_merge.
eassumption.
assumption.
eassumption.
eassumption.
eassumption.
exfalso.
eapply Hconf_stuck.
eapply bstep_diverges.
exfalso.
destruct e.
eapply Hconf_stuck.
eapply bstep_cons.
reflexivity.
destruct (
in_dom_dec
(merge
(create_map (dom m1)
(csemantics m1))
(create_map (dom m2)
(csemantics m2))
)
f
).
destruct i.
assert (exists b, x a h0 b).
destruct (mergelookup' H1).
destruct H2.
assert (exists m, x0 = create_map (dom m) (csemantics m) /\ (m = m1 \/ m = m2)).
destruct H3; subst; eauto.
clear H3.
destruct H4.
destruct H3.
subst.
generalize (create_map_lookup_elim H2).
intro; subst.
generalize (in_dom_dom (lookup_in_dom H2)).
rewrite create_map_dom.
intro.
generalize (dom_in_dom H3).
destruct 1.
case_eq (callee_init a h0 x).
intros.
destruct (exists_behavior (@config_final _ _) (cstep x1) (Config h1 l0 nil)).
esplit.
econstructor.
eassumption.
eassumption.
eassumption.
destruct H2.
eapply Hconf_stuck.
eapply bstep_internal.
eassumption.
eassumption.
reflexivity.
eapply Hconf_stuck.
eapply bstep_external.
eassumption.
reflexivity.
edestruct Hstuck_not_final.
econstructor.
clear COINDHYP.
functional inversion 1; subst.
clear Hbeh Hbeh0.
generalize (config_diverges_nil_silently_diverges Hconf); clear Hconf.
intros.
eapply config_behaviors_star_diverges.
econstructor.
eleft.
eapply config_silently_diverges'_diverges.
revert bs H c Hinv.
cofix COINDHYP.
inversion 1; subst.
intros.
destruct (invariant_forward Hinv Hstep).
destruct H0.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
assumption.
intros.
edestruct (bbehavior_of_cbehavior_append_ex Hbeh0).
subst.
rewrite <- bbehavior_of_cbehavior_append in Hbeh0.
generalize (append_behavior_right_inj Hbeh0).
intro; subst.
destruct (invariant_forward_star Hinv Hnil).
destruct H.
destruct (invariant_forward H0 Hcons).
destruct H1.
econstructor.
eapply star_trans'.
eassumption.
eright.
eassumption.
simpl; reflexivity.
eapply COINDHYP.
eassumption.
eassumption.
reflexivity.
reflexivity.
Qed.
Corollary bsemantics_in_csemantics: forall f a h b,
bsemantics (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))) f a h b -> csemantics (merge m1 m2) f a h b.
Proof.
inversion 1; subst.
destruct (mergelookup' Hcmd); clear Hcmd.
destruct H0.
assert (exists m, x = create_map (dom m) (csemantics m) /\ (m = m1 \/ m = m2)).
destruct H1; subst; eauto.
clear H1.
destruct H2.
destruct H1.
subst.
generalize (create_map_lookup_elim H0); clear H0.
intro; subst.
inversion Hcb; subst.
generalize (lookup_disj_merge H2 Hcmd Hdisj).
intros.
clear Hcb Hcmd.
econstructor.
eassumption.
eassumption.
eapply config_behaviors_star_config_behaviors.
eapply invariant_forward_close.
econstructor.
eassumption.
eapply config_behaviors_behaviors'.
eassumption.
eleft.
reflexivity.
eapply config_behaviors_behaviors'.
eassumption.
reflexivity.
Qed.
Second part of the proof: csemantics (merge m1 m2) is included in bsemantics (merge (csemantics m1) (csemantics m2))
Terminating and stuck behaviors: easy induction
Lemma csemantics_in_bsemantics_terminates_aux: forall tr h r s
(Hterm: config_terminates (cstep (merge m1 m2)) (@config_final _ _) s tr (h,r))
m (Hm: m = m1 \/ m = m2), exists tr',
config_terminates (cstep m) (@config_final _ _) s tr' (h,r) /\
config_terminates (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (@bstate_final _) (BBehavior (Terminates tr' (h,r)) nil) tr (Normal h r).
Proof.
inversion 1; subst.
clear Hterm.
destruct (star_stepn_left Hterm0); clear Hterm0.
revert tr h r s cfin H Hfin.
induction x using (well_founded_induction lt_wf).
inversion 1; subst.
clear H.
intros.
exists nil.
split.
econstructor.
eleft.
assumption.
econstructor.
eleft.
constructor.
clear H0.
destruct s.
intros.
destruct conf1.
destruct (cstepm_disj Hm Hstep).
destruct H0.
assert (n < S n) by omega.
destruct (H _ H2 _ _ _ _ _ Hstepn Hfin _ Hm).
destruct H3.
inversion H3; subst.
inversion H4; subst.
esplit.
split.
econstructor.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
assumption.
replace (Terminates (tr1 ++ x) (h, r)) with (append1_behavior tr1 (Terminates x (h,r))).
econstructor.
eapply star_trans.
eapply bstar_append.
assumption.
eassumption.
assumption.
rewrite append_behavior.
reflexivity.
destruct H0.
subst.
destruct H1.
destruct H0.
destruct H0.
destruct H1.
destruct H2.
destruct H2.
destruct H3.
destruct H3.
destruct H4.
destruct H4.
destruct H4.
destruct H5.
subst.
inversion Hfin; subst.
destruct (cstepn_stack_recip Hstepn (s1 := nil) (refl_equal _)).
exfalso.
destruct H6.
destruct H6.
destruct x5; discriminate.
destruct H6.
destruct H6.
destruct H6.
destruct H6.
destruct H6.
destruct H7.
destruct H7.
destruct H8.
destruct H8.
destruct H8.
destruct H9.
destruct H9.
destruct H9.
destruct H10.
destruct H10.
destruct H10.
destruct H11.
assert (x5 < S n) by omega.
assert (x14 < S n) by omega.
injection H8; intros; subst.
assert (config_final (Config x6 x7 nil) (x6,x9)).
constructor.
assumption.
generalize (cstepn_le H6).
intros.
destruct (H _ H13 _ _ _ _ _ H6 H11 _ H2); clear H13 H6 H11.
destruct H15.
inversion H11; subst; clear H11.
inversion Hfin0; subst; clear Hfin0.
destruct (star_stepn_left Hterm); clear Hterm.
clear Hstepn.
destruct (H _ H14 _ _ _ _ _ H10 Hfin _ Hm); clear H H14 H10 Hfin.
destruct H13.
inversion H; subst; clear H.
inversion H10; subst; clear H10.
inversion Hfin0; subst.
generalize (bstepn_stack H11 (((x6,x9),Terminates x17 (h,r))::nil)); clear H11.
intros.
assert (in_dom x1 x).
esplit.
eassumption.
generalize (in_dom_dom H10); clear H10.
intro.
simpl in H10.
generalize (create_map_dom (dom x1) (csemantics x1)).
intro.
rewrite <- H11 in H10.
clear H11.
generalize (dom_in_dom H10); clear H10.
destruct 1.
generalize (create_map_lookup_elim H10).
intro; subst.
assert (config_behaviors (cstep x1) (@config_final _ _) (Config c_heap0 c_local0 nil) (Terminates x4 (x6,x9))) by assumption.
esplit.
split.
econstructor.
eapply star_trans.
eright.
eleft.
eapply cstep_call_external.
eassumption.
assumption.
eassumption.
eapply heap_le_trans.
eapply callee_init_le. eassumption.
eassumption.
eassumption.
eassumption.
assumption.
econstructor.
eapply star_trans.
eright.
eleft.
eapply bstep_internal.
eapply lookup_disj_merge.
2: eassumption.
destruct H2; subst; eauto.
unfold disjoint.
intros.
generalize (in_dom_dom H13).
rewrite create_map_dom.
intro.
generalize (in_dom_dom H14).
rewrite create_map_dom.
intro.
eapply Hdisj; eauto using dom_in_dom.
econstructor.
eassumption.
eassumption.
eassumption.
simpl.
match goal with |- Terminates (?e :: ?f) ?g = _ => change (Terminates (e :: f) g) with (cons_behavior e (Terminates f g)) end.
reflexivity.
eapply star_trans.
eapply stepn_left_star.
eassumption.
simpl.
eapply star_trans'.
eright.
eleft.
eapply bstep_terminates.
eassumption.
reflexivity.
assumption.
Qed.
Lemma cstep_merge_stuck: forall m (Hm: m = m1 \/ m = m2) cstuck
(Hstuck: forall (tr : list xevent) (conf' : config lang),
cstep (merge (A:=funcname) (B:=code lang) m1 m2) cstuck tr
conf' -> False)
(tr : list xevent) (conf' : config lang),
cstep m cstuck tr conf' -> False.
Proof.
inversion 3; subst.
eapply Hstuck.
eapply cstep_step.
eapply step_merge.
eassumption.
eassumption.
reflexivity.
destruct (in_dom_dec (merge m1 m2) f).
destruct i.
case_eq (callee_init a h1 x).
intros.
eapply Hstuck.
eapply cstep_step.
eapply step_call.
eassumption.
eassumption.
eassumption.
eassumption.
reflexivity.
eapply Hstuck.
eapply cstep_call_external.
eassumption.
assumption.
eassumption.
eassumption.
eassumption.
Qed.
Lemma csemantics_in_bsemantics_stuck_aux: forall tr s
(Hstuck: config_stuck (cstep (merge m1 m2)) (@config_final _ _) s tr)
m (Hm: m = m1 \/ m = m2), exists tr',
config_behaviors (cstep m) (@config_final _ _) s tr' /\
(forall k,
config_stuck (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (@bstate_final _) (BBehavior tr' k) tr) .
Proof.
inversion 1; subst.
clear Hstuck.
destruct (star_stepn_left Hstar); clear Hstar.
revert tr s cstuck Hstuck_not_final Hconf_stuck H.
induction x using (well_founded_induction lt_wf).
inversion 3; subst.
clear H.
intros.
exists (Stuck _ nil).
split.
econstructor.
eleft.
assumption.
eapply cstep_merge_stuck. eassumption. assumption.
intros.
econstructor.
eleft.
inversion 1.
inversion 1; subst;
functional inversion Hb'.
intros.
destruct s.
destruct conf1.
destruct (cstepm_disj Hm Hstep).
destruct H1.
assert (n < S n) by omega.
destruct (H _ H3 _ _ _ Hstuck_not_final Hconf_stuck Hstepn _ Hm).
destruct H4.
exists (append1_behavior tr1 x).
split.
eapply star_config_behaviors.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
intro.
generalize (H5 k).
inversion 1; subst.
econstructor.
eapply star_trans.
eapply bstar_append.
assumption.
eassumption.
assumption.
assumption.
destruct H1.
destruct H2.
destruct H2.
destruct H2.
destruct H3.
destruct H4.
destruct H4.
destruct H5.
destruct H5.
destruct H6.
destruct H6.
destruct H6.
destruct H7.
subst.
simpl in *.
destruct cstuck.
destruct (cstepn_stack_recip Hstepn (s1 := nil) (refl_equal _)).
destruct H1.
destruct H1.
subst.
assert (n < S n) by omega.
refine (_ (H _ H1 _ _ _ _ _ H8 _ H4)).
destruct 1.
destruct H9.
generalize (ret_inh sw).
inversion 1.
case_eq (caller_return X x3 x4).
intros.
destruct (exists_behavior (@config_final _ _) (cstep m) (Config h l (c_stack))).
esplit.
split.
eapply star_config_behaviors.
eright.
eleft.
eapply cstep_call_external.
eassumption.
assumption.
eassumption.
eapply heap_le_refl.
eassumption.
eassumption.
simpl; reflexivity.
intro.
generalize (H10 ((x3, X, x7) :: k)).
inversion 1; subst.
generalize (in_dom_dom (lookup_in_dom H5)).
rewrite <- (create_map_dom (dom x1) (csemantics x1)).
intro.
destruct (dom_in_dom H15).
assert (
create_map (dom x1) (csemantics x1) = create_map (dom m1) (csemantics m1)
\/ create_map (dom x1) (csemantics x1) = create_map (dom m2) (csemantics m2)
).
destruct H4; subst; eauto.
refine (_ (lookup_disj_merge H17 H16 _)).
intro.
generalize (create_map_lookup_elim H16).
intro; subst.
econstructor.
eapply star_trans'.
eright.
eleft.
eapply bstep_internal.
eassumption.
econstructor.
eassumption.
eassumption.
eassumption.
reflexivity.
eassumption.
reflexivity.
assumption.
assumption.
intro. intros.
generalize (in_dom_dom H18).
generalize (in_dom_dom H19).
repeat rewrite create_map_dom.
intros; eauto using dom_in_dom.
inversion 1; subst.
simpl in *.
case_eq (caller_return r c_heap1 x4).
intros.
eapply Hconf_stuck.
eapply cstep_step.
eapply step_return.
eassumption.
eassumption.
reflexivity.
intros.
destruct conf'.
eapply Hconf_stuck.
eapply cstep_stack.
eassumption.
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H1.
destruct H8.
destruct H8.
destruct H9.
destruct H9.
destruct H9.
destruct H10.
destruct H10.
destruct H10.
destruct H11.
destruct H11.
destruct H11.
destruct H12.
injection H9; intros; subst.
assert (x14 < S (S (x5 + x14))) by omega.
destruct (H _ H12 _ _ _ Hstuck_not_final Hconf_stuck H11 _ Hm).
destruct H13.
clear H.
refine (_ (csemantics_in_bsemantics_terminates_aux (s := Config c_heap0 c_local0 nil) _ H4)).
2: econstructor.
2: eapply stepn_left_star.
2: eassumption.
2: econstructor.
2: eassumption.
destruct 1.
destruct H.
inversion H15; subst.
inversion Hfin; subst.
esplit.
split.
eapply star_config_behaviors.
eright.
eleft.
eapply cstep_call_external.
eassumption.
assumption.
eassumption.
2: eassumption.
eapply heap_le_trans.
eapply callee_init_le.
eassumption.
eapply cstepn_le.
eassumption.
eassumption.
simpl; reflexivity.
intro.
generalize (H14 k).
inversion 1; subst.
generalize (in_dom_dom (lookup_in_dom H5)).
rewrite <- (create_map_dom (dom x1) (csemantics x1)).
intro.
destruct (dom_in_dom H17).
assert (
create_map (dom x1) (csemantics x1) = create_map (dom m1) (csemantics m1)
\/ create_map (dom x1) (csemantics x1) = create_map (dom m2) (csemantics m2)
).
destruct H4; subst; eauto.
refine (_ (lookup_disj_merge H19 H18 _)).
intro.
generalize (create_map_lookup_elim H18).
intro; subst.
destruct (star_stepn_left Hterm).
econstructor.
eapply star_trans'.
eright.
eleft.
eapply bstep_internal with (bf := Terminates _ _).
eassumption.
econstructor.
eassumption.
eassumption.
eassumption.
reflexivity.
eapply star_trans'.
change ((x6,x9,x4) :: k) with (nil ++ (x6,x9,x4)::k).
eapply stepn_left_star.
eapply bstepn_stack.
eassumption.
simpl.
eapply star_trans'.
eright.
eleft.
eapply bstep_terminates.
eassumption.
simpl; reflexivity.
reflexivity.
reflexivity.
assumption.
assumption.
intro. intros.
generalize (in_dom_dom H20).
generalize (in_dom_dom H21).
repeat rewrite create_map_dom.
intros; eauto using dom_in_dom.
Qed.
Diverging/reacting behaviors
How to classify diverging/reacting executions following the external function calls.
Section Classify.
Variable m: module lang.
Hypothesis Hm: m = m1 \/ m = m2.
Lemma resolve_no_external_calls: forall n s0 h l s g tr h' l' s' g' (Hstepn: stepn_left (@guided_step _ _) n (Config h l (s ++ s0), g) tr (Config h' l' (s' ++ s0), g')) (Hg_correct: guided_execution_correct (cstep (merge m1 m2)) (Config h l (s ++ s0), g))
(Hstack: forall h' l' s' tr g', star (@guided_step _ _) (Config h l (s ++ s0), g) tr (Config h' l' s', g') ->
exists s'', s' = s'' ++ s0)
(Hnoext: forall n2, n2 < n ->
forall h2 l2 s2 g2 tr2,
stepn_left (@guided_step _ _) n2 (Config h l (s ++ s0), g) tr2 (Config h2 l2 s2, g2) ->
forall f2 a2,
local_kind l2 = Some (Call f2 a2) ->
in_dom (merge m1 m2) f2 ->
in_dom m f2) ,
stepn_left (cstep m) n (Config h l s) tr (Config h' l' s') /\
forall b, forall hs,
star (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (BBehavior (append1_behavior tr b) hs) tr (BBehavior b hs).
Proof.
induction n; inversion 1; subst.
intros.
generalize (app_inv_tail _ _ _ H4).
intro; subst.
split.
left.
intros.
left.
inversion Hstep; subst.
inversion 1; subst.
intros.
destruct s'0.
edestruct Hstack.
eapply stepn_left_star.
eright.
eassumption.
eleft.
reflexivity.
subst.
assert (cstep m (Config h l s) tr1 (Config c_heap c_local x) /\ forall f, in_dom (merge m1 m2) f -> forall a h1 r h2, In (External f a h1 r h2) tr1 -> False).
inversion Hstep0; subst.
inversion Hstep1; subst.
generalize (app_inv_tail _ _ _ H6).
intro; subst.
split.
eapply cstep_step.
eapply step_internal.
eassumption.
reflexivity.
reflexivity.
intros f _.
revert f.
generalize (trace_of_option_event tr0).
clear.
induction l; simpl.
tauto.
destruct 1.
discriminate.
eauto.
assert (O < S n) by omega.
generalize (Hnoext _ H _ _ _ _ _ (@stepn_left_O _ _ _ _) _ _ Hl (lookup_in_dom Hfunc)).
destruct 1.
generalize (lookup_disj_merge Hm H0 Hdisj).
intros.
generalize (lookup_func H1 Hfunc).
intro; subst.
simpl.
change (fr :: s ++ s0) with ((fr :: s) ++ s0) in H6.
generalize (app_inv_tail _ _ _ H6).
intro; subst.
split.
eapply cstep_step.
eapply step_call.
eassumption.
eassumption.
eassumption.
assumption.
reflexivity.
tauto.
simpl.
change (fr :: x ++ s0) with ((fr :: x) ++ s0) in H2.
generalize (app_inv_tail _ _ _ H2).
intro; subst.
split.
eapply cstep_step.
eapply step_return.
eassumption.
assumption.
reflexivity.
tauto.
generalize (app_inv_tail _ _ _ H6).
intros; subst.
split.
eapply cstep_call_external.
assumption.
intro.
apply Hf.
eauto using indom_merge.
eassumption.
eassumption.
assumption.
destruct 2.
injection H0; intros; subst; eauto.
assumption.
destruct H.
edestruct (IHn); clear IHn.
eassumption.
eassumption.
intros.
eapply Hstack.
eapply star_trans'.
2: eassumption.
eright.
eleft.
eassumption.
reflexivity.
intros.
eapply Hnoext; clear Hnoext.
2: eright.
2: eassumption.
2: eassumption.
omega.
reflexivity.
eassumption.
assumption.
split.
econstructor.
eassumption.
eassumption.
reflexivity.
intros.
rewrite append_behavior_app.
eapply star_trans.
eapply bstar_append.
intros.
generalize (in_dom_dom Hf).
rewrite dom_merge.
rewrite create_map_dom.
rewrite create_map_dom.
rewrite <- dom_merge.
intro.
eapply H0.
eapply dom_in_dom.
eassumption.
eassumption.
eapply H2.
Qed.
Lemma resolve_terminating_external_calls: forall n h l s0 s g tr h' l' s' g' (Hstepn: stepn_left (@guided_step _ _) n (Config h l (s ++ s0), g) tr (Config h' l' (s' ++ s0), g')) (Hg_correct: guided_execution_correct (cstep (merge m1 m2)) (Config h l (s ++ s0), g))
(Hstack: forall h2 l2 s2 tr2 g2, star (@guided_step _ _) (Config h l (s ++ s0), g) tr2 (Config h2 l2 s2, g2) ->
exists s'', s2 = s'' ++ s0)
(Hallterm: forall n2, n2 < n ->
forall h2 l2 s2 g2 tr2,
stepn_left (@guided_step _ _) n2 (Config h l (s ++ s0), g) tr2 (Config h2 l2 (s2 ++ s0), g2) ->
forall f2 a2,
local_kind l2 = Some (Call f2 a2) ->
in_dom (merge m1 m2) f2 ->
~ in_dom m f2 ->
exists h3, exists l3, exists g3, exists fr3,
guided_step (Config h2 l2 (s2 ++ s0), g2) nil (Config h3 l3 (fr3 :: s2 ++ s0), g3) /\
exists n4, exists h4, exists l4, exists tr4, exists g4,
stepn_left (@guided_step _ _) (n4) (Config h3 l3 (fr3 :: s2 ++ s0), g3) tr4 (Config h4 l4 (fr3 :: s2 ++ s0), g4) /\
stepn_left (cstep (merge m1 m2)) n4 (Config h3 l3 nil) tr4 (Config h4 l4 nil) /\
exists h5, exists l5, exists g5,
guided_step (Config h4 l4 (fr3 :: s2 ++ s0), g4) nil (Config h5 l5 (s2 ++ s0), g5) /\
n2 + S (S n4) <= n /\
forall n6, n6 <= n4 ->
forall h' l' s' tr' g',
stepn_left (@guided_step _ _) n6 (Config h3 l3 (fr3 :: s2 ++ s0), g3) tr' (Config h' l' s', g') ->
exists s6, s' = s6 ++ fr3 :: s2 ++ s0),
exists tr', star (cstep m) (Config h l s) tr' (Config h' l' s') /\
forall b, forall hs,
star (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (BBehavior (append1_behavior tr' b) hs) tr (BBehavior b hs).
Proof.
induction n using (well_founded_induction lt_wf).
inversion 1; subst.
generalize (app_inv_tail _ _ _ H6).
intros; subst.
clear H Hallterm.
exists nil.
split.
eleft.
intros.
eleft.
inversion Hstep; subst.
inversion 1; subst.
intros.
destruct s'0.
destruct (classic (exists f, exists a,
local_kind l = Some (Call f a) /\
in_dom (merge m1 m2) f /\
~ in_dom m f
)).
destruct H0.
destruct H0.
destruct H0.
destruct H1.
euse Hallterm.
2: eleft.
omega.
eassumption.
assumption.
assumption.
generalize (guided_stepn_inv_correct al0 (guided_step_inv_correct Hg_correct al)).
inversion 1; subst.
inversion al2; subst.
inversion Hstep1; subst.
inversion Hstep2; try subst.
exfalso.
generalize (refl_equal (length (s ++ s0))).
rewrite <- H12 at 1.
simpl.
omega.
exfalso.
generalize (refl_equal (length (s ++ s0))).
rewrite <- H12 at 1.
simpl.
omega.
inversion al; subst.
simpl in *.
inversion Hstep0; subst.
inversion Hstep3; try subst.
exfalso.
generalize (refl_equal (length (s ++ s0))).
rewrite H12 at 1.
simpl; omega.
rewrite H0 in Hl.
injection Hl; intros; subst.
destruct (mergelookup' Hfunc).
destruct H5.
generalize (in_dom_dom (lookup_in_dom H5)).
rewrite <- (create_map_dom (dom x) (csemantics x)).
intro.
generalize (dom_in_dom H7); clear H7.
destruct 1.
assert (
create_map (dom (B:=code lang) x)
(csemantics (lang:=lang) x) = create_map (dom (B:=code lang) m1)
(csemantics (lang:=lang) m1) \/
create_map (dom (B:=code lang) x)
(csemantics (lang:=lang) x) = create_map (dom (B:=code lang) m2)
(csemantics (lang:=lang) m2)
).
destruct H6; subst; auto.
refine (_ (lookup_disj_merge H8 H7 _)).
intro.
generalize (create_map_lookup_elim H7).
intro; subst.
edestruct csemantics_in_bsemantics_terminates_aux.
econstructor.
eapply stepn_left_star.
eassumption.
econstructor.
eassumption.
eassumption.
destruct H9.
inversion H10; subst.
inversion Hfin; subst.
destruct (star_stepn_left Hterm); clear Hterm.
generalize (stepn_right_left (stepn_right_S (stepn_left_right (stepn_left_S al al0 (refl_equal _))) al2 (refl_equal _))).
rewrite <- app_nil_end.
simpl.
intro.
destruct (eq_nat_dec (S (S n4)) (S n0)).
clear H Hallterm.
injection e; intros; subst.
generalize (determ_n (@guided_step_determ _ _) H12 Hstepn).
injection 1; intros; subst.
simpl in *.
generalize (app_inv_tail _ _ _ H14).
intro; subst.
exists (External func ar h'0 r h4 :: nil).
split.
eright.
eleft.
eapply cstep_call_external.
assumption.
assumption.
eassumption.
eapply heap_le_trans.
eapply callee_init_le.
eassumption.
eapply cstepn_le.
eassumption.
assumption.
intros.
eapply star_trans'.
eright.
eleft.
eapply bstep_internal with (bf := Terminates _ _).
eassumption.
econstructor.
eassumption.
eassumption.
simpl.
eassumption.
simpl.
reflexivity.
eapply star_trans'.
change ((h4, r, b) :: hs) with (nil ++ (h4, r, b) :: hs).
eapply stepn_left_star.
eapply bstepn_stack.
eassumption.
simpl.
eright.
eleft.
eapply bstep_terminates.
rewrite <- app_nil_end.
reflexivity.
reflexivity.
assert (S n0 = S (S n4) + (S n0 - S (S n4))) by omega.
destruct (triangle (@guided_step_determ _ _) Hstepn H12 H13).
destruct H14; subst.
edestruct (H (S n0 - S (S n4))); clear H.
omega.
eassumption.
eapply guided_stepn_inv_correct.
eassumption.
eassumption.
intros.
eapply Hstack.
eapply star_trans'.
eright.
eleft.
eassumption.
eapply star_trans'.
eapply stepn_left_star.
eexact al0.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
simpl; reflexivity.
reflexivity.
simpl; reflexivity.
intros.
edestruct (Hallterm (S (S n4) + n2)).
omega.
eapply stepn_left_app.
eassumption.
eassumption.
eassumption.
assumption.
assumption.
destruct H19.
destruct H19.
destruct H19.
destruct H19.
destruct H20.
destruct H20.
destruct H20.
destruct H20.
destruct H20.
destruct H20.
destruct H21.
destruct H22.
destruct H22.
destruct H22.
destruct H22.
destruct H23.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
omega.
assumption.
clear Hallterm.
destruct H15.
exists (External func ar h'0 r h4 :: x4).
split.
eapply star_trans'.
eright.
eleft.
eapply cstep_call_external.
eassumption.
assumption.
eassumption.
eapply heap_le_trans.
eapply callee_init_le.
eassumption.
eapply cstepn_le.
eassumption.
eassumption.
eassumption.
reflexivity.
intros.
eapply star_trans'.
eright.
eleft.
eapply bstep_internal with (bf := Terminates _ _).
eassumption.
econstructor.
eassumption.
eassumption.
simpl.
eassumption.
simpl; reflexivity.
eapply star_trans'.
eapply stepn_left_star.
change ((h4,r,append1_behavior x4 b) :: hs) with (nil ++ (h4,r,append1_behavior x4 b) :: hs).
eapply bstepn_stack.
eassumption.
simpl.
eapply star_trans'.
eright.
eleft.
eapply bstep_terminates.
eapply H15.
simpl; reflexivity.
reflexivity.
reflexivity.
intro.
intros.
generalize (in_dom_dom H9).
rewrite create_map_dom.
intros.
generalize (in_dom_dom H10).
rewrite create_map_dom.
intros.
eapply Hdisj; eauto using dom_in_dom.
rewrite H0 in Hcmd0.
discriminate.
edestruct Hstack.
eright.
eleft.
eassumption.
subst.
assert (cstep m (Config h l s) tr1 (Config c_heap c_local x) /\ forall f, in_dom (merge m1 m2) f -> forall a h1 r h2, In (External f a h1 r h2) tr1 -> False).
inversion Hstep0; subst.
inversion Hstep1; subst.
generalize (app_inv_tail _ _ _ H8).
intro; subst.
split.
eapply cstep_step.
eapply step_internal.
eassumption.
reflexivity.
reflexivity.
intros f _.
revert f.
generalize (trace_of_option_event tr0).
clear.
induction l; simpl.
tauto.
destruct 1.
discriminate.
eauto.
destruct (in_dom_dec m func).
destruct i.
generalize (lookup_disj_merge Hm H1 Hdisj).
intros.
generalize (lookup_func H2 Hfunc).
intro; subst.
simpl.
change (fr :: s ++ s0) with ((fr :: s) ++ s0) in H8.
generalize (app_inv_tail _ _ _ H8).
intro; subst.
split.
eapply cstep_step.
eapply step_call.
eassumption.
eassumption.
eassumption.
assumption.
reflexivity.
tauto.
destruct H0.
eauto 6 using lookup_in_dom.
simpl.
change (fr :: x ++ s0) with ((fr :: x) ++ s0) in H4.
generalize (app_inv_tail _ _ _ H4).
intro; subst.
split.
eapply cstep_step.
eapply step_return.
eassumption.
assumption.
reflexivity.
tauto.
generalize (app_inv_tail _ _ _ H8).
intro; subst.
split.
eapply cstep_call_external.
assumption.
intro.
apply Hf.
eauto using indom_merge.
eassumption.
eassumption.
assumption.
destruct 2.
injection H2; intros; subst; eauto.
assumption.
destruct H1.
edestruct (H n0); clear H.
omega.
eassumption.
eassumption.
intros.
eapply Hstack.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
intros.
edestruct (Hallterm (S n2)); clear Hallterm.
omega.
econstructor.
eassumption.
eassumption.
reflexivity.
eassumption.
assumption.
assumption.
destruct H7.
destruct H7.
destruct H7.
destruct H7.
destruct H8.
destruct H8.
destruct H8.
destruct H8.
destruct H8.
destruct H8.
destruct H9.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H11.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
omega.
assumption.
clear Hallterm.
destruct H3.
exists (tr1 ++ x0).
split.
eapply star_trans.
eright.
eleft.
eassumption.
assumption.
intros.
rewrite append_behavior_app.
eapply star_trans.
eapply bstar_append.
intros.
generalize (in_dom_dom Hf).
rewrite dom_merge.
rewrite create_map_dom.
rewrite create_map_dom.
rewrite <- dom_merge.
intro.
eapply H2.
eapply dom_in_dom.
eassumption.
eassumption.
eapply H3.
Qed.
Variable h: heap.
Variable l: local.
Variable s: list frame.
Variable g: Stream (list xevent * config lang).
Hypothesis Hg_correct: guided_execution_correct (cstep (merge m1 m2)) (Config h l s, g).
Definition case1_diverging_external :=
exists h1, exists l1, exists s1, exists g1, exists tr1,
star (@guided_step _ _) (Config h l s, g) tr1 (Config h1 l1 (s1), g1) /\
exists f, exists a,
local_kind l1 = Some (Call f a) /\
in_dom (merge m1 m2) f /\
~ in_dom m f /\
exists fr,
forall h' l' s' tr g', plus (@guided_step _ _) (Config h1 l1 (s1), g1) tr (Config h' l' s', g') ->
exists s'', s' = s'' ++ fr :: s1.
Definition case2_all_externals_return :=
forall h1 l1 s1 g1 tr1,
star (@guided_step _ _) (Config h l s, g) tr1 (Config h1 l1 (s1), g1) ->
forall f a,
local_kind l1 = Some (Call f a) ->
in_dom (merge m1 m2) f ->
~ in_dom m f ->
exists h2, exists l2, exists tr2, exists g2,
plus (@guided_step _ _) (Config h1 l1 (s1), g1) tr2 (Config h2 l2 (s1), g2).
Definition case2_subcase1_finitely_many_externals :=
exists n,
forall h1 l1 n1 s1 tr1 g1, stepn_left (@guided_step _ _) n1 (Config h l s, g) tr1 (Config h1 l1 (s1), g1) ->
forall f a,
local_kind l1 = Some (Call f a) ->
in_dom (merge m1 m2) f ->
(~ in_dom m f) ->
n1 < n.
Definition case2_subcase2_infinitely_many_externals :=
forall h1 l1 s1 tr1 g1, star (@guided_step _ _) (Config h l s, g) tr1 (Config h1 l1 (s1), g1) ->
exists h2, exists l2, exists s2, exists g2, exists tr2,
star (@guided_step _ _) (Config h1 l1 (s1), g1) tr2 (Config h2 l2 (s2), g2) /\
exists f, exists a,
local_kind l2 = Some (Call f a) /\
in_dom (merge m1 m2) f /\
(~ in_dom m f).
Lemma classify_external_divergence: case1_diverging_external \/
case2_all_externals_return
.
Proof.
destruct (classic case1_diverging_external).
auto.
right.
generalize (not_ex_all_not _ _ H); clear H.
intros.
refine (_ (fun h1 => not_ex_all_not _ _ (H h1))).
clear H.
intros.
refine (_ (fun h1 l1 => not_ex_all_not _ _ (x h1 l1))).
clear x.
intro.
refine (_ (fun h1 l1 s1 => not_ex_all_not _ _ (x0 h1 l1 s1))).
clear x0.
intro.
refine (_ (fun h1 l1 s1 g1 => not_ex_all_not _ _ (x h1 l1 s1 g1))).
clear x.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 => or_to_imply _ _ (not_and_or _ _ (x0 h1 l1 s1 g1 tr1)))).
clear x0.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar => not_ex_all_not _ _ (x h1 l1 s1 g1 tr1 Hstar))).
clear x.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f => not_ex_all_not _ _ (x0 h1 l1 s1 g1 tr1 Hstar f))).
clear x0.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a => or_to_imply _ _ (not_and_or _ _ (x h1 l1 s1 g1 tr1 Hstar f a)))).
clear x.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind => or_to_imply _ _ (not_and_or _ _ (x0 h1 l1 s1 g1 tr1 Hstar f a Hkind)))).
clear x0.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom => or_to_imply _ _ (not_and_or _ _ (x h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom)))).
clear x.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom => not_ex_all_not _ _ (x0 h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom))).
clear x0.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr => not_all_ex_not _ _ (x h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr))).
clear x.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr =>
match x0 h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr with
| ex_intro h2 H =>
ex_intro _ h2 (not_all_ex_not _ _ H)
end
)).
clear x0.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr =>
match x h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr with
| ex_intro h2 (ex_intro l2 H) =>
ex_intro _ h2 (ex_intro _ l2 (not_all_ex_not _ _ H))
end
)).
clear x.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr =>
match x0 h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr with
| ex_intro h2 (ex_intro l2 (ex_intro s2 H)) =>
ex_intro _ h2 (ex_intro _ l2 (ex_intro _ s2 (not_all_ex_not _ _ H)))
end
)).
clear x0.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr =>
match x h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr with
| ex_intro h2 (ex_intro l2 (ex_intro s2 (ex_intro tr2 H))) =>
ex_intro _ h2 (ex_intro _ l2 (ex_intro _ s2 (ex_intro _ tr2 (not_all_ex_not _ _ H))))
end
)).
clear x.
intros.
refine (_ (fun h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr =>
match x0 h1 l1 s1 g1 tr1 Hstar f a Hkind Hdom Hndom Hfr with
| ex_intro h2 (ex_intro l2 (ex_intro s2 (ex_intro tr2 (ex_intro q2 H)))) =>
ex_intro _ h2 (ex_intro _ l2 (ex_intro _ s2 (ex_intro _ tr2 (ex_intro _ q2 (match imply_to_and _ _ H with
| conj H1 H2 => conj H1 (not_ex_all_not _ _ H2)
end
)))))
end
)).
clear x0.
intros.
intro.
intros.
generalize (x _ _ _ _ _ H _ _ H0 H1 H2).
clear x.
intros.
destruct (star_stepn_left H).
generalize (guided_stepn_inv_correct H4 Hg_correct).
inversion 1; subst.
destruct s'.
inversion Hstep; subst.
inversion Hstep0; subst.
generalize (eval_internal_kind Hstep1).
rewrite H0.
discriminate.
rewrite H0 in Hl.
injection Hl.
intros; subst.
destruct (H3 fr); clear H3.
destruct H6.
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct (plus_stepn_left H3); clear H3.
inversion H7; subst.
simpl in Hstep1.
inversion Hstep1; subst.
simpl in *.
destruct (guided_cstepn_uncons (sf := nil) Hstepn HCOIND H6) .
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H8.
destruct H9.
destruct H9.
destruct H9.
destruct H9.
destruct H10.
simpl in *.
esplit.
esplit.
esplit.
esplit.
eapply stepn_left_plus.
econstructor.
econstructor.
eapply stepn_left_app.
eassumption.
eright.
eassumption.
eleft.
simpl; reflexivity.
simpl; rewrite <- app_nil_end; reflexivity.
rewrite Hcmd in H0.
discriminate.
rewrite Hcmd in H0.
injection H0; intros; subst.
contradiction.
Qed.
Lemma classify_external_finiteness: case2_subcase1_finitely_many_externals \/
case2_subcase2_infinitely_many_externals.
Proof.
destruct (classic case2_subcase1_finitely_many_externals).
auto.
right.
generalize (not_ex_all_not _ _ H).
clear H.
intro.
refine (_ (fun n => match not_all_ex_not _ _ (H n) with
| ex_intro h1 Hh1 =>
ex_intro _ h1
match not_all_ex_not _ _ Hh1 with
| ex_intro l1 Hl1 =>
ex_intro _ l1
match not_all_ex_not _ _ Hl1 with
| ex_intro n1 Hn1 =>
ex_intro _ n1
match not_all_ex_not _ _ Hn1 with
| ex_intro s1 Hs1 =>
ex_intro _ s1
match not_all_ex_not _ _ Hs1 with
| ex_intro tr1 Htr1 =>
ex_intro _ tr1
match not_all_ex_not _ _ Htr1 with
| ex_intro g1 Hg1 =>
ex_intro _ g1
match imply_to_and _ _ Hg1 with
| conj Hstepn Hfa =>
conj Hstepn
match not_all_ex_not _ _ Hfa with
| ex_intro f Hf =>
ex_intro _ f
match not_all_ex_not _ _ Hf with
| ex_intro a Ha =>
ex_intro _ a
match imply_to_and _ _ Ha with
| conj Hkind Hdoms =>
conj Hkind
match imply_to_and _ _ Hdoms with
| conj Hdom Hdomlt =>
conj Hdom (imply_to_and _ _ Hdomlt)
end
end
end
end
end
end
end
end
end
end
end
)).
clear H.
intro H.
intro.
intros.
destruct (star_stepn_left H0).
clear H0.
destruct (H x).
destruct H0.
destruct H0.
destruct H0.
destruct H0.
destruct H0.
destruct H0.
destruct H2.
destruct H2.
destruct H2.
destruct H3.
destruct H4.
assert (x <= x2) by omega.
clear H5.
assert (x2 = x + (x2 - x)) by omega.
destruct (triangle (@guided_step_determ _ _) H0 H1 H5).
revert H5 H7.
generalize (x2 - x).
destruct 2.
subst.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eapply stepn_left_star.
eassumption.
eauto.
Qed.
Definition case1_diverging_external' :=
exists n1, exists h1, exists l1, exists s1, exists g1, exists tr1,
stepn_left (@guided_step _ _) n1 (Config h l s, g) tr1 (Config h1 l1 (s1), g1) /\
exists f, exists a,
local_kind l1 = Some (Call f a) /\
in_dom (merge m1 m2) f /\
~ in_dom m f /\ (
exists fr,
forall h' l' s' tr g', plus (@guided_step _ _) (Config h1 l1 (s1 ), g1) tr (Config h' l' s', g') ->
exists s'', s' = s'' ++ fr :: s1
) /\ forall n2, n2 < n1 ->
forall h2 l2 s2 g2 tr2,
stepn_left (@guided_step _ _) n2 (Config h l s, g) tr2 (Config h2 l2 (s2 ), g2) ->
forall f2 a2,
local_kind l2 = Some (Call f2 a2) ->
in_dom (merge m1 m2) f2 ->
~ in_dom m f2 ->
exists h3, exists l3, exists g3, exists fr3,
guided_step (Config h2 l2 (s2 ), g2) nil (Config h3 l3 (fr3 :: s2 ), g3) /\
exists n4, exists h4, exists l4, exists tr4, exists g4,
stepn_left (@guided_step _ _) (n4) (Config h3 l3 (fr3 :: s2 ), g3) tr4 (Config h4 l4 (fr3 :: s2 ), g4) /\
stepn_left (cstep (merge m1 m2)) n4 (Config h3 l3 nil) tr4 (Config h4 l4 nil) /\
exists h5, exists l5, exists g5,
guided_step (Config h4 l4 (fr3 :: s2 ), g4) nil (Config h5 l5 (s2 ), g5) /\
n2 + S (S n4) <= n1 /\
forall n6, n6 <= n4 ->
forall h' l' s' tr' g',
stepn_left (@guided_step _ _) n6 (Config h3 l3 (fr3 :: s2 ), g3) tr' (Config h' l' s', g') ->
exists s6, s' = s6 ++ fr3 :: s2 .
Lemma case1_diverging_external_strong: case1_diverging_external -> case1_diverging_external'.
Proof.
destruct 1.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H.
destruct (star_stepn_left H).
clear H.
generalize (conj H1 H0).
clear H1 H0.
intro.
pattern x3 in H.
generalize (ex_intro _ x3 H).
clear H x3.
intro.
pattern x2 in H.
generalize (ex_intro _ x2 H).
clear x2 H.
intro.
pattern x1 in H.
generalize (ex_intro _ x1 H).
clear x1 H.
intro.
pattern x0 in H.
generalize (ex_intro _ x0 H).
clear x0 H.
intro.
pattern x in H.
generalize (ex_intro _ x H).
clear x H.
intro.
pattern x4 in H.
generalize (ex_intro _ x4 H).
clear x4 H.
intro.
generalize (exists_minimal _ H).
clear H.
destruct 1.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H1.
destruct H1.
destruct H1.
destruct H2.
destruct H3.
destruct H4.
generalize (fun n' Hn' => not_ex_all_not _ _ (H0 n' Hn')).
clear H0.
intro.
generalize (fun n' Hn' h => not_ex_all_not _ _ (H0 n' Hn' h)).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 => not_ex_all_not _ _ (H0 n' Hn' h1 l1)).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 => not_ex_all_not _ _ (H0 n' Hn' h1 l1 s1)).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 => not_ex_all_not _ _ (H0 n' Hn' h1 l1 s1 g1)).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 => or_to_imply _ _ (not_and_or _ _ (H0 n' Hn' h1 l1 s1 g1 tr1))).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 Hstepn1 => not_ex_all_not _ _ (H0 n' Hn' h1 l1 s1 g1 tr1 Hstepn1)).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 => not_ex_all_not _ _ (H0 n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1)).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 => or_to_imply _ _ (not_and_or _ _ (H0 n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1))).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1 => or_to_imply _ _ (not_and_or _ _ (H0 n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1))).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1 Hdom1 => or_to_imply _ _ (not_and_or _ _ (H0 n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1 Hdom1))).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1 Hdom1 Hndom1 => not_ex_all_not _ _ (H0 n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1 Hdom1 Hndom1)).
clear H0.
intro.
generalize (fun n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1 Hdom1 Hndom1 fr =>
match not_all_ex_not _ _ (H0 n' Hn' h1 l1 s1 g1 tr1 Hstepn1 f1 a1 Hkind1 Hdom1 Hndom1 fr) with
| ex_intro h' Hh' =>
ex_intro _ h'
match not_all_ex_not _ _ Hh' with
| ex_intro l' Hl' =>
ex_intro _ l'
match not_all_ex_not _ _ Hl' with
| ex_intro s' Hs' =>
ex_intro _ s'
match not_all_ex_not _ _ Hs' with
| ex_intro tr' Htr' =>
ex_intro _ tr'
match not_all_ex_not _ _ Htr' with
| ex_intro g' Hg' =>
ex_intro _ g'
match imply_to_and _ _ Hg' with
| conj Hplus Hnex =>
conj Hplus (not_ex_all_not _ _ Hnex)
end
end
end
end
end
end
).
clear H0.
intros.
esplit.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eassumption.
split.
assumption.
split.
assumption.
split.
esplit.
eassumption.
intros.
generalize (guided_stepn_inv_correct H6 Hg_correct).
inversion 1; subst.
inversion Hstep; subst.
inversion Hstep0; subst.
generalize (eval_internal_kind Hstep1).
rewrite H7.
discriminate.
rewrite H7 in Hl.
injection Hl; intros; subst.
destruct (H0 _ H5 _ _ _ _ _ H6 _ _ H7 H8 H9 fr).
destruct H11.
destruct H11.
destruct H11.
destruct H11.
destruct H11.
simpl in *.
destruct (plus_stepn_left H11); clear H11.
inversion H13; intros; subst.
inversion Hstep1; subst.
simpl in *.
change (fr :: s2 ) with (nil ++ fr :: s2 ) in Hstepn.
destruct (guided_cstepn_uncons Hstepn HCOIND H12).
destruct H11.
destruct H11.
destruct H11.
destruct H11.
destruct H11.
destruct H14.
destruct H15.
destruct H15.
destruct H15.
destruct H15.
destruct H16.
simpl in *.
generalize (stepn_left_app H6 (stepn_left_app (stepn_left_S Hstep1 H11 (refl_equal _)) (stepn_left_S H15 (@stepn_left_O _ _ _ _) (refl_equal _)))).
intros.
simpl in *.
rewrite <- app_nil_end in H18.
destruct (le_lt_dec (n2 + (S x11 + 1)) x).
assert (x = (n2 + (S x11 + 1)) + (x - (n2 + (S x11 + 1)))) by omega.
destruct (triangle (@guided_step_determ _ _) H H18 H19).
destruct H20.
revert H19 H20.
generalize (x - (n2 + (S x11 + 1))).
intros; subst.
esplit.
esplit.
esplit.
esplit.
split.
econstructor.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
omega.
intros.
eapply H17.
eassumption.
eassumption.
exfalso.
assert (n2 + (S x11 + 1) = x + (n2 + (S x11 + 1) - x)) by omega.
revert H19.
generalize (n2 + (S x11 + 1) - x).
destruct n.
intro; omega.
intros.
destruct (triangle (@guided_step_determ _ _) H18 H H19).
destruct H20.
generalize (stepn_left_plus H20).
intros.
destruct (H4 _ _ _ _ _ H22).
assert (n2 + (x - n2) = x) by omega.
revert H24.
generalize (x - n2).
destruct n0.
intro; omega.
intro; subst.
destruct (triangle (@guided_step_determ _ _) H H6 (refl_equal _)).
destruct H23.
subst.
inversion H23; subst.
inversion Hstep2; subst.
simpl in *.
rewrite app_ass in H21.
generalize (app_inv_head _ _ _ H21).
intro; subst.
assert (n0 <= x11) by omega.
destruct (H17 _ H24 _ _ _ _ _ Hstepn0).
generalize (refl_equal (length (x2 ))).
rewrite H25 at 1.
repeat (rewrite app_length; simpl).
omega.
rewrite H7 in Hcmd.
discriminate.
rewrite H7 in Hcmd.
injection Hcmd; intros; subst.
contradiction.
Qed.
Definition case2_subcase1_finitely_many_externals' :=
exists n,
forall h1 l1 n1 s1 tr1 g1, stepn_left (@guided_step _ _) n1 (Config h l s, g) tr1 (Config h1 l1 (s1 ), g1) ->
forall f a,
local_kind l1 = Some (Call f a) ->
in_dom (merge m1 m2) f ->
(~ in_dom m f) ->
exists h3, exists l3, exists g3, exists fr3,
guided_step (Config h1 l1 (s1 ), g1) nil (Config h3 l3 (fr3 :: s1 ), g3) /\
exists n4, exists h4, exists l4, exists tr4, exists g4,
stepn_left (@guided_step _ _) (n4) (Config h3 l3 (fr3 :: s1 ), g3) tr4 (Config h4 l4 (fr3 :: s1 ), g4) /\
stepn_left (cstep (merge m1 m2)) n4 (Config h3 l3 nil) tr4 (Config h4 l4 nil) /\
exists h5, exists l5, exists g5,
guided_step (Config h4 l4 (fr3 :: s1 ), g4) nil (Config h5 l5 (s1 ), g5) /\
n1 + S (S (n4)) <= n /\
forall n6, n6 <= n4 ->
forall h' l' s' tr' g',
stepn_left (@guided_step _ _) n6 (Config h3 l3 (fr3 :: s1 ), g3) tr' (Config h' l' s', g') ->
exists s6, s' = s6 ++ fr3 :: s1 .
Lemma case2_subcase1_finitely_many_externals_strong:
case2_all_externals_return ->
case2_subcase1_finitely_many_externals ->
case2_subcase1_finitely_many_externals'.
Proof.
destruct 2.
cut (
forall delta n0
(Hn0: n0 + delta = x)
h0 l0 s0 tr0 g0
(Hstepn0: stepn_left (@guided_step _ _) n0 (Config h l s, g) tr0 (Config h0 l0 (s0 ), g0)),
exists n,
forall h1 l1 n1 s1 tr1 g1, stepn_left (@guided_step _ _) n1 (Config h0 l0 (s0 ), g0) tr1 (Config h1 l1 (s1 ), g1) ->
forall f a,
local_kind l1 = Some (Call f a) ->
in_dom (merge m1 m2) f ->
(~ in_dom m f) ->
exists h3, exists l3, exists g3, exists fr3,
guided_step (Config h1 l1 (s1 ), g1) nil (Config h3 l3 (fr3 :: s1 ), g3) /\
exists n4, exists h4, exists l4, exists tr4, exists g4,
stepn_left (@guided_step _ _) (n4) (Config h3 l3 (fr3 :: s1 ), g3) tr4 (Config h4 l4 (fr3 :: s1 ), g4) /\
stepn_left (cstep (merge m1 m2)) n4 (Config h3 l3 nil) tr4 (Config h4 l4 nil) /\
exists h5, exists l5, exists g5,
guided_step (Config h4 l4 (fr3 :: s1 ), g4) nil (Config h5 l5 (s1 ), g5) /\
n1 + S (S (n4)) <= n /\
forall n6, n6 <= n4 ->
forall h' l' s' tr' g',
stepn_left (@guided_step _ _) n6 (Config h3 l3 (fr3 :: s1 ), g3) tr' (Config h' l' s', g') ->
exists s6, s' = s6 ++ fr3 :: s1
).
intros.
refine (H1 x 0 _ _ _ _ _ _ _).
abstract omega.
simpl. eleft.
induction delta using (well_founded_induction lt_wf).
intros; subst.
destruct (classic (exists n'0, exists h'0, exists l'0, exists s'0, exists tr'0, exists g'0,
stepn_left (@guided_step _ _) n'0 (Config h0 l0 (s0 ), g0) tr'0 (Config h'0 l'0 (s'0 ), g'0) /\
exists f, exists a,
local_kind l'0 = Some (Call f a) /\
in_dom (merge m1 m2) f /\
~ in_dom m f
)).
generalize (exists_minimal _ H2).
clear H2.
destruct 1.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H4.
destruct H4.
destruct H4.
destruct H5.
generalize (stepn_left_app Hstepn0 H2).
intros.
generalize (stepn_left_star H7).
intro.
destruct (H _ _ _ _ _ H8 _ _ H4 H5 H6).
destruct H9.
destruct H9.
destruct H9.
destruct (plus_stepn_left H9); clear H9.
inversion H10; subst.
generalize (guided_stepn_inv_correct H7 Hg_correct).
inversion 1; subst.
inversion Hstep; subst.
inversion Hstep0; subst.
inversion Hstep1; subst.
generalize (eval_internal_kind Hstep2).
rewrite H4.
discriminate.
rewrite H4 in Hl.
injection Hl; intros; subst.
change (fr :: x2 ) with (nil ++ fr :: x2 ) in Hstepn.
destruct (guided_cstepn_uncons Hstepn HCOIND).
intros.
intro.
generalize (refl_equal (length x2)).
rewrite H11 at 1.
rewrite app_length; simpl.
intro.
abstract omega.
destruct H11.
destruct H11.
destruct H11.
destruct H11.
destruct H11.
destruct H12.
destruct H13.
destruct H13.
destruct H13.
destruct H13.
destruct H14.
destruct (le_lt_dec delta (x + 1 + x4 + 1)).
clear H1.
exists (x + 1 + x4 + 1).
intros.
destruct (le_lt_dec x n1).
simpl in *.
generalize (guided_stepn_inv_correct H1 (guided_stepn_inv_correct Hstepn0 Hg_correct)).
inversion 1; subst.
inversion Hstep2; subst.
inversion Hstep3; subst.
generalize (eval_internal_kind Hstep4).
rewrite H16.
discriminate.
rewrite H16 in Hl0.
injection Hl0; intros; subst.
destruct (H _ _ _ _ _ (stepn_left_star (stepn_left_app Hstepn0 H1)) _ _ H16 H17 H18).
destruct H20.
destruct H20.
destruct H20.
destruct (plus_stepn_left H20); clear H20.
inversion H21; subst.
inversion Hstep4; subst.
change (fr0 :: s1 ) with (nil ++ fr0 :: s1 ) in Hstepn1.
destruct (guided_cstepn_uncons Hstepn1 HCOIND0).
intros.
intro.
generalize (refl_equal (length s1)).
rewrite H20 at 1.
rewrite app_length.
simpl.
abstract omega.
destruct H20.
destruct H20.
destruct H20.
destruct H20.
destruct H20.
destruct H22.
destruct H23.
destruct H23.
destruct H23.
destruct H23.
destruct H24.
generalize (H0 _ _ _ _ _ _ (stepn_left_app Hstepn0 H1) _ _ H16 H17 H18).
intros.
destruct (eq_nat_dec n1 x).
subst.
generalize (determ_n (@guided_step_determ _ _) H1 H2).
injection 1; intros; subst.
simpl in *.
destruct (eq_nat_dec x4 x18). subst.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
eassumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
abstract omega.
intros; eauto.
exfalso.
destruct (le_lt_dec x4 x18).
assert (S x4 <= x18) by abstract omega.
generalize (H25 _ H28 _ _ _ _ _ (stepn_right_left (stepn_right_S (stepn_left_right H11) H13 (refl_equal _)))).
destruct 1.
generalize (refl_equal (length x2)).
rewrite H29 at 1.
rewrite app_length.
simpl.
abstract omega.
assert (S x18 <= x4) by abstract omega.
generalize (H15 _ H28 _ _ _ _ _ (stepn_right_left (stepn_right_S (stepn_left_right H20) H23 (refl_equal _)))).
destruct 1.
generalize (refl_equal (length x2)).
rewrite H29 at 1.
rewrite app_length.
simpl.
abstract omega.
assert (n1 = x + 1 + (n1 - (x + 1))) by abstract omega.
revert H27.
generalize (n1 - (x + 1)).
intros.
generalize (triangle (@guided_step_determ _ _) H1 (stepn_left_app H2 (stepn_left_S Hstep (@stepn_left_O _ _ _ _) (refl_equal _))) H27).
subst.
destruct 1.
destruct H27.
repeat rewrite <- app_nil_end in *.
subst.
simpl in *.
assert (n2 <= x4) by abstract omega.
destruct (H15 _ H28 _ _ _ _ _ H27).
simpl in *.
destruct (eq_nat_dec n2 x4).
exfalso.
subst.
generalize (determ_n (@guided_step_determ _ _) H11 H27).
injection 1; intros; subst.
change (fr :: x2 ) with (nil ++ fr :: x2 ) in H31 at 1.
generalize (app_inv_tail _ _ _ H31).
intro; subst; simpl in *.
generalize (refl_equal (length (x2 ))).
inversion H13.
rewrite H38 at 1.
simpl.
abstract omega.
assert (x4 = S n2 + (x4 - (S n2))) by abstract omega.
generalize (triangle (@guided_step_determ _ _) H11 (stepn_right_left (stepn_right_S (stepn_left_right H27) Hstep4 (refl_equal _))) H30).
revert H30.
generalize (x4 - S n2).
destruct 2; subst.
destruct H31.
rewrite <- app_nil_end in *.
subst.
simpl in *.
destruct (le_lt_dec n3 x18).
exfalso. generalize (H25 _ l4 _ _ _ _ _ H29).
generalize (fr :: x2).
destruct 1.
generalize (refl_equal (length l5)).
rewrite H30 at 1.
rewrite app_length.
simpl.
rewrite app_length.
abstract omega.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
eassumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
abstract omega.
intros; eauto.
rewrite H16 in Hcmd.
discriminate.
rewrite H16 in Hcmd.
injection Hcmd; intros; subst.
contradiction.
edestruct H3; eauto 11.
assert (delta - (x + 1 + x4 + 1) < delta) by abstract omega.
assert ((n0 + (x + 1 + x4 + 1)) + (delta - (x + 1 + x4 + 1)) = n0 + delta) by abstract omega.
generalize (stepn_left_app (stepn_left_app (stepn_left_app H2 (stepn_left_S Hstep (@stepn_left_O _ _ _ _) (refl_equal _))) H11) (stepn_left_S H13 (@stepn_left_O _ _ _ _) (refl_equal _))).
intro.
generalize (stepn_left_app Hstepn0 H18).
simpl in *; repeat rewrite <- app_nil_end in *; simpl in *.
intros.
destruct (H1 _ H16 _ H17 _ _ _ _ _ H19); clear H1.
exists (x + 1 + x4 + 1 + x16).
intros. destruct (le_lt_dec (x + 1 + x4 + 1) n1).
assert (n1 = (x + 1 + x4 + 1) + (n1 - (x + 1 + x4 + 1))) by abstract omega.
generalize (triangle (@guided_step_determ _ _) H1 H18 H24).
revert H24.
generalize (n1 - (x + 1 + x4 + 1)).
destruct 2; subst.
destruct H25; subst.
destruct (H20 _ _ _ _ _ _ H24 _ _ H21 H22 H23); clear H20.
destruct H25.
destruct H20.
destruct H20.
destruct H20.
destruct H25.
destruct H25.
destruct H25.
destruct H25.
destruct H25.
destruct H25.
destruct H26.
destruct H27.
destruct H27.
destruct H27.
destruct H27.
destruct H28.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
abstract omega.
assumption.
clear H20.
destruct (le_lt_dec (S n1) x).
exfalso.
assert (n1 < x) by abstract omega.
edestruct H3; eauto 11.
clear H3.
destruct (eq_nat_dec x n1).
subst.
generalize (determ_n (@guided_step_determ _ _) H2 H1).
injection 1; intros; subst.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
abstract omega.
intros; eauto.
assert (n1 = S x + (n1 - (S x))) by abstract omega.
generalize (triangle (@guided_step_determ _ _) H1 (stepn_right_left (stepn_right_S (stepn_left_right H2) Hstep (refl_equal _))) H3).
revert H3.
generalize (n1 - (S x)).
rewrite <- app_nil_end.
destruct 2; subst.
destruct H20; subst.
assert (n2 <= x4) by abstract omega.
destruct (H15 _ H20 _ _ _ _ _ H3).
subst; simpl in *; simpl in *.
destruct (eq_nat_dec n2 x4).
exfalso.
subst.
generalize (determ_n (@guided_step_determ _ _) H3 H11).
injection 1; intros; subst.
change (fr :: x2 ) with (nil ++ fr :: x2 ) in H26 at 2.
generalize (app_inv_tail _ _ _ H26).
intro; subst.
simpl in *.
generalize (guided_stepn_inv_correct H3 HCOIND).
inversion 1; subst.
inversion H13; subst.
inversion Hstep2; subst.
inversion Hstep3.
generalize (refl_equal (length (x2 ))).
rewrite <- H34 at 1.
simpl.
abstract omega.
generalize (refl_equal (length (x2 ))).
rewrite <- H34 at 1.
simpl.
abstract omega.
rewrite H21 in Hcmd.
discriminate.
assert (x4 = n2 + (x4 - n2)) by abstract omega.
generalize (triangle (@guided_step_determ _ _) H11 H3 H24).
revert H24.
generalize (x4 - n2).
destruct 2; subst.
destruct H25; subst.
destruct n3.
exfalso.
abstract omega. inversion H24; subst.
generalize (guided_stepn_inv_correct H3 HCOIND).
inversion 1; subst.
inversion Hstep2; subst.
inversion Hstep3; subst.
inversion Hstep4; subst.
generalize (eval_internal_kind Hstep5).
rewrite H21.
discriminate.
rewrite Hl0 in H21.
injection H21; intros; subst.
change (fr0 :: x18 ++ fr :: x2 ) with (nil ++ fr0 :: x18 ++ fr :: x2 ) in Hstepn1.
destruct (guided_cstepn_uncons Hstepn1 HCOIND0).
generalize (fr :: x2 ).
intros.
intro.
generalize (refl_equal (length l5)).
rewrite H26 at 1.
repeat (rewrite app_length; simpl).
abstract omega.
destruct H26.
destruct H26.
destruct H26.
destruct H26.
destruct H26.
destruct H27.
destruct H28.
destruct H28.
destruct H28.
destruct H28.
destruct H29.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
abstract omega.
intros; eauto.
rewrite H21 in Hcmd.
discriminate.
rewrite H21 in Hcmd.
injection Hcmd; intros; subst.
contradiction.
rewrite H4 in Hcmd.
discriminate.
rewrite H4 in Hcmd.
injection Hcmd; intros; subst.
contradiction.
clear H0 H1.
exists O.
intros.
destruct H2; eauto 12.
Qed.
End Classify.
Case 1: call a diverging external function
Inductive invarI (m: module lang) : config lang * Stream (list xevent * config lang) -> bstate sw -> Prop :=
| invarI_intro l f a (Hl: local_kind l = Some (Call f a))
s0 h s g (Hcorrect: guided_execution_correct (cstep (merge m1 m2)) (Config h l (s ++ s0), g))
h0 fr (Hcall: caller_call a h l = (h0, fr))
h1 l1 g1 (Hstep: guided_step (Config h l (s ++ s0), g) nil (Config h1 l1 (fr :: s ++ s0), g1))
(Hforallstep: forall h' l' s' tr' g', star (@guided_step _ _) (Config h1 l1 (fr :: s ++ s0), g1) tr' (Config h' l' s', g') -> exists s'', s' = s'' ++ fr :: s ++ s0)
r h1 b b' (Hb' : b' = cons_behavior (External f a h0 r h1) b)
(Hm: m = m1 \/ m = m2)
(Hfndom: ~ in_dom m f)
foo
(Hfoo: cstep m (Config h l s) (External f a h0 r h1 :: nil) foo)
s' (Hs': s' = s ++ s0)
bs
:
invarI m (Config h l s', g) (BBehavior b' bs).
Inductive invarJ (s0: list frame) (m: module lang) : config lang * Stream (list xevent * config lang) -> bstate sw -> Prop :=
| invarJ_intro h l s g
(Hcorrect: guided_execution_correct (cstep (merge m1 m2)) (Config h l (s ++ s0), g))
(Hforallstep: forall h' l' s' tr' g', star (@guided_step _ _) (Config h l (s ++ s0), g) tr' (Config h' l' s', g') -> exists s'', s' = s'' ++ s0)
(Hm: m = m1 \/ m = m2)
b (Hbconf: config_behaviors (@guided_step _ _) (fun _ _ => False) (Config h l (s ++ s0), g) b)
(Hnoext: forall h' l' s' tr' g', star (@guided_step _ _) (Config h l (s ++ s0), g) tr' (Config h' l' s', g') -> forall f a, local_kind l' = Some (Call f a) -> in_dom (merge m1 m2) f -> in_dom m f)
s' (Hs' : s' = s ++ s0)
bs
:
invarJ s0 m (Config h l s', g) (BBehavior b bs).
Lemma invarJ_sim: forall s0 m,
simult (@guided_step _ _) (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (invarJ s0 m).
Proof.
intros.
intro.
inversion 1; subst.
generalize (config_behaviors_behaviors' Hbconf).
inversion 1; subst.
inversion Hbconf; subst; contradiction.
inversion Hbconf; subst.
destruct cstuck.
destruct s1.
destruct p.
edestruct Hconf_stuck.
econstructor.
generalize (config_diverges_nil_silently_diverges Hbconf).
inversion 1; subst.
inversion Hcorrect; subst.
inversion Hstep; subst.
destruct s'.
edestruct Hforallstep.
eright.
eleft.
eassumption.
subst.
esplit.
esplit.
split.
eleft.
eassumption.
esplit.
split.
eleft.
eapply bstep_diverges.
econstructor.
eassumption.
intros.
eapply Hforallstep.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
simpl; reflexivity.
assumption.
econstructor.
eleft.
eassumption.
intros.
eapply Hnoext.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
simpl; reflexivity.
eassumption.
assumption.
reflexivity.
destruct (star_stepn_left Hnil).
generalize (stepn_right_left (stepn_right_S (stepn_left_right H1) Hcons (refl_equal _))).
intro.
simpl in H2.
destruct c2.
generalize (guided_stepn_inv_correct H2 Hcorrect).
intro.
generalize (stepn_left_plus H2).
intro.
generalize (stepn_left_star H2).
intro.
generalize (cstarn_trace_inv (stepn_left_star (guided_stepn_inv H2 Hcorrect))).
intro.
refine (_ (bstar_append (psi := merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))) (l := e :: l0) _ b0 bs)).
inversion 1; subst.
destruct c.
destruct (Hforallstep _ _ _ _ _ H5).
subst.
esplit.
esplit.
split.
eassumption.
esplit.
split.
eassumption.
econstructor.
eassumption.
intros; eauto using star_trans'.
assumption.
eapply config_behaviors_star_config_behaviors.
eapply config_behaviors'_config_behaviors_star.
eassumption.
intros; eauto using star_trans'.
reflexivity.
intros.
generalize (in_dom_dom Hf).
rewrite dom_merge.
rewrite create_map_dom.
rewrite create_map_dom.
rewrite <- dom_merge.
intro.
generalize (dom_in_dom H8).
eauto.
Qed.
Inductive invarJ' (s0 : list frame) (m: module lang) : config lang * Stream (list xevent * config lang) -> config lang -> Prop :=
| invarJ'_intro h l s g b bs (Hinvar: invarJ s0 m (Config h l (s ++ s0), g) (BBehavior b bs))
s' (Hs': s' = s ++ s0) :
invarJ' s0 m (Config h l s', g) (Config h l s).
Lemma invarJ'_sim: forall s0 m,
simult (@guided_step _ _) (cstep m) (invarJ' s0 m).
Proof.
intros.
intro.
intros.
inversion H; subst.
destruct (invarJ_sim Hinvar).
destruct H0.
destruct H0.
destruct H1.
destruct H1.
destruct x0.
destruct c.
inversion Hinvar; subst.
generalize (app_inv_tail _ _ _ Hs').
intro; subst.
destruct (plus_stepn_left H0).
destruct (Hforallstep _ _ _ _ _ (stepn_left_star H3)).
subst.
edestruct resolve_no_external_calls.
eassumption.
eassumption.
eassumption.
assumption.
intros; eauto using stepn_left_star.
destruct x1.
inversion H2.
esplit.
esplit.
split.
eassumption.
esplit.
split.
eapply stepn_left_plus.
eassumption.
econstructor.
eassumption.
reflexivity.
Qed.
Case 3: infinitely many external functions that all terminate.
CoInductive invarK (s0: list frame) (m: module lang) : config lang * Stream (list xevent * config lang) -> bstate sw -> Prop :=
| invarK_intro
h l s g
(Hcorrect: guided_execution_correct (cstep (merge m1 m2)) (Config h l (s ++ s0), g))
(Hforallstep: forall h' l' s' tr' g', star (@guided_step _ _) (Config h l (s ++ s0), g) tr' (Config h' l' s', g') -> exists s'', s' = s'' ++ s0)
n1 h1 l1 s1 tr1 g1
(Hstepn: stepn_left (@guided_step _ _) n1 (Config h l (s ++ s0), g) tr1 (Config h1 l1 s1, g1))
(Hm: m = m1 \/ m = m2)
(Hnoext: forall n', n' < n1 -> forall h' l' s' tr' g', stepn_left (@guided_step _ _) n' (Config h l (s ++ s0), g) tr' (Config h' l' s', g') -> forall f a, local_kind l' = Some (Call f a) -> in_dom (merge m1 m2) f -> in_dom m f)
f a (Hf: local_kind l1 = Some (Call f a)) (Hdom: in_dom (merge m1 m2) f) (Hndom: ~ in_dom m f)
h2 fr (Hfr: caller_call a h1 l1 = (h2, fr))
h3 l3 g3
(Hcall: guided_step (Config h1 l1 s1, g1) nil (Config h3 l3 (fr :: s1), g3))
n4 h4 l4 tr4 g4
(Hfuncg: stepn_left (@guided_step _ _) (n4) (Config h3 l3 (fr :: s1), g3) tr4 (Config h4 l4 (fr :: s1), g4))
(Hfuncc: stepn_left (cstep (merge m1 m2)) n4 (Config h3 l3 nil) tr4 (Config h4 l4 nil))
r (Hretk: local_kind l4 = Some (Return r))
h5 l5 g5
(Hret: guided_step (Config h4 l4 (fr :: s1), g4) nil (Config h5 l5 s1, g5))
b
bs
(HCOIND: invarK s0 m (Config h5 l5 s1, g5) (BBehavior b bs))
b' (Hb': b' = append1_behavior (tr1 ++ External f a h2 r h4 :: nil) b)
s' (Hs': s' = s ++ s0)
:
invarK s0 m (Config h l s', g) (BBehavior b' bs)
.
Lemma invarK_cbehavior:
forall s0 m h l s g b' bs
(HinvarK: invarK s0 m (Config h l (s ++ s0), g) (BBehavior (Reacts _ (func_of_stream b')) bs))
b'' (Hsim: bisim b' b''),
config_reacts (cstep m) (Config h l s) b''.
Proof.
cofix COINDHYP.
inversion 1; subst.
generalize (app_inv_tail _ _ _ Hs').
intro; subst.
rewrite append_behavior in Hb'.
destruct b; try discriminate.
injection Hb'.
intro.
generalize (func_eq_bisim H).
intro.
apply bisim_sym in H0.
destruct (bisim_appinf_strong H0).
subst.
generalize (appinf_bisim H0).
intro.
case_eq (tr1 ++ External f a h2 r h4 :: nil).
destruct tr1; discriminate.
intros.
destruct (Hforallstep _ _ _ _ _ (stepn_left_star Hstepn)).
subst.
edestruct resolve_no_external_calls.
eassumption.
eassumption.
assumption.
assumption.
clear COINDHYP; intros; eauto.
generalize (guided_stepn_inv_correct Hstepn Hcorrect).
inversion 1; subst.
inversion Hcall; subst.
inversion Hstep; subst.
inversion Hstep0; try subst.
clear COINDHYP; exfalso.
generalize (refl_equal (length (x1 ++ s0))).
rewrite H13 at 1.
simpl; omega.
rewrite Hf in Hl.
injection Hl; intros; subst.
rewrite Hframe in Hfr.
injection Hfr; intros; subst.
generalize (guided_stepn_inv_correct Hfuncg HCOIND0).
inversion 1; subst.
inversion Hret; subst.
inversion Hstep1; subst.
inversion Hstep2; try subst.
clear COINDHYP; exfalso.
generalize (refl_equal (length (x1 ++ s0))).
rewrite <- H14 at 1.
simpl; omega.
clear COINDHYP; exfalso.
generalize (refl_equal (length (x1 ++ s0))).
rewrite <- H14 at 1.
simpl; omega.
rewrite Hcmd in Hretk.
injection Hretk; intros; subst.
rewrite <- (func_of_stream_of_func evinfseq) in HCOIND.
destruct (bisim_appinf_strong Hsim).
subst.
generalize (appinf_bisim Hsim).
intro.
econstructor.
eapply star_trans'.
eapply stepn_left_star.
eassumption.
eright.
eleft.
eapply cstep_call_external.
eassumption.
assumption.
eassumption.
eapply heap_le_trans.
eapply callee_init_le.
eassumption.
eapply cstepn_le.
eassumption.
eassumption.
symmetry.
eassumption.
eapply COINDHYP; clear COINDHYP.
eexact HCOIND.
2: reflexivity.
eapply bisim_trans.
eassumption.
assumption.
clear COINDHYP; exfalso.
generalize (refl_equal (length s)).
rewrite H13 at 1.
simpl; omega.
Qed.
Lemma invarK_sim: forall s0 m,
simult (@guided_step _ _) (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (invarK s0 m).
Proof.
intros.
intro.
inversion 1; subst.
edestruct Hforallstep.
eapply stepn_left_star.
eexact Hstepn.
subst.
edestruct resolve_no_external_calls.
eassumption.
eassumption.
assumption.
intros; eauto.
intros; eauto.
generalize (guided_stepn_inv_correct Hstepn Hcorrect).
inversion 1; subst.
inversion Hcall; subst.
inversion Hstep; subst.
inversion Hstep0; try subst.
exfalso. generalize (refl_equal (length (x ++ s0))). rewrite H10 at 1. simpl. omega.
rewrite Hf in Hl.
injection Hl; intros; subst.
rewrite Hframe in Hfr.
injection Hfr; intros; subst.
generalize (guided_stepn_inv_correct Hfuncg HCOIND0).
inversion 1; subst.
inversion Hret; subst.
inversion Hstep1; subst.
inversion Hstep2; try subst.
exfalso. generalize (refl_equal (length (x ++ s0))). rewrite <- H11 at 1. simpl. omega.
exfalso. generalize (refl_equal (length (x ++ s0))). rewrite <- H11 at 1. simpl. omega.
rewrite Hretk in Hcmd.
injection Hcmd; intros; subst.
destruct (mergelookup' Hfunc).
destruct H4.
generalize (in_dom_dom (lookup_in_dom H4)).
generalize (create_map_dom (dom x0) (csemantics x0)).
intro Y.
pattern (dom (B := code lang) x0) at 1.
rewrite <- Y at 1.
intro.
generalize (dom_in_dom H6).
destruct 1.
generalize (create_map_lookup_elim H7).
intro; subst.
assert (create_map (dom x0) (csemantics x0) = create_map (dom m1) (csemantics m1) \/
create_map (dom x0) (csemantics x0) = create_map (dom m2) (csemantics m2)).
destruct H5; subst; eauto.
refine (_ (lookup_disj_merge H8 H7 _)).
intro.
edestruct csemantics_in_bsemantics_terminates_aux.
2 : eexact H5.
econstructor.
eapply stepn_left_star.
eapply Hfuncc.
econstructor.
eassumption.
destruct H9.
inversion H10; subst.
inversion Hfin; subst.
destruct (star_stepn_left Hterm).
esplit.
esplit.
split.
eapply star_plus_right.
eapply stepn_left_star.
eassumption.
eapply star_plus_left.
eleft.
eassumption.
eapply star_trans'.
eapply stepn_left_star.
eassumption.
eright.
eleft.
eassumption.
rewrite <- app_nil_end.
reflexivity.
simpl; reflexivity.
reflexivity.
rewrite append_behavior_app.
esplit.
split.
eapply star_plus_right.
eapply H1.
simpl.
eapply star_plus_left.
eleft.
eapply bstep_internal with (bf := Terminates _ _).
eassumption.
econstructor.
eassumption.
eassumption.
eassumption.
reflexivity.
change ((h4, r0, b) :: bs) with (nil ++ (h4, r0, b) :: bs).
eapply star_trans'.
eapply stepn_left_star.
eapply bstepn_stack.
eassumption.
simpl.
eright.
eleft.
eapply bstep_terminates.
rewrite <- app_nil_end.
reflexivity.
simpl; reflexivity.
reflexivity.
assumption.
intro.
intros.
generalize (in_dom_dom H9).
generalize (in_dom_dom H10).
rewrite create_map_dom.
rewrite create_map_dom.
eauto using dom_in_dom.
exfalso.
generalize (refl_equal (length s1)).
rewrite H10 at 1.
simpl.
omega.
Qed.
How to build the reactive behavior retrieving the infinitely many external function calls.
Section Build_invarK.
Variable m: module lang.
Hypothesis Hm: m = m1 \/ m = m2.
Variable s0 : list frame.
Definition build_invarK_inv h l s g : Prop :=
(
case2_all_externals_return m h l (s ++ s0) g /\
case2_subcase2_infinitely_many_externals m h l (s ++ s0) g /\
guided_execution_correct (cstep (merge m1 m2)) (Config h l (s ++ s0), g) /\
(forall h' l' s' tr' g', star (@guided_step _ _) (Config h l (s ++ s0), g) tr' (Config h' l' s', g') -> exists s'', s' = s'' ++ s0)
).
Definition build_invarK_prop h l s g tr1 f a h2 r h4 h5 l5 s1 g5 : Prop :=
exists n1, exists h1, exists l1, exists g1,
stepn_left (@guided_step _ _) n1 (Config h l (s ++ s0), g) tr1 (Config h1 l1 (s1 ++ s0), g1) /\
(forall n', n' < n1 -> forall h' l' s' tr' g', stepn_left (@guided_step _ _) n' (Config h l (s ++ s0), g) tr' (Config h' l' s', g') -> forall f a, local_kind l' = Some (Call f a) -> in_dom (merge m1 m2) f -> in_dom m f) /\
local_kind l1 = Some (Call f a) /\
in_dom (merge m1 m2) f /\
~ in_dom m f /\
exists fr, caller_call a h1 l1 = (h2, fr) /\
exists h3, exists l3, exists g3,
guided_step (Config h1 l1 (s1 ++ s0), g1) nil (Config h3 l3 (fr :: (s1 ++ s0)), g3) /\
exists n4, exists l4, exists tr4, exists g4,
stepn_left (@guided_step _ _) (n4) (Config h3 l3 (fr :: (s1 ++ s0)), g3) tr4 (Config h4 l4 (fr :: (s1 ++ s0)), g4) /\
stepn_left (cstep (merge m1 m2)) n4 (Config h3 l3 nil) tr4 (Config h4 l4 nil) /\
local_kind l4 = Some (Return r) /\
guided_step (Config h4 l4 (fr :: (s1 ++ s0)), g4) nil (Config h5 l5 (s1 ++ s0), g5).
Lemma build_invarK_loop_body: forall h l s g,
build_invarK_inv h l s g ->
exists tr1, exists f, exists a, exists h2, exists r, exists h4, exists h5, exists l5, exists s1, exists g5,
tr1 ++ External f a h2 r h4 :: nil <> nil /\
build_invarK_inv h5 l5 s1 g5 /\
build_invarK_prop h l s g tr1 f a h2 r h4 h5 l5 s1 g5.
Proof.
destruct 1.
destruct H0.
destruct H1.
unfold case2_all_externals_return in H.
unfold case2_subcase2_infinitely_many_externals in H0.
refine (_
(exists_minimal _ match (H0 _ _ _ _ _ (star_refl _ _)) with
| ex_intro h2 (ex_intro l2 (ex_intro s2 (ex_intro g2 (ex_intro tr2 (conj Hstar Hf))))) =>
match star_stepn_left Hstar with
| ex_intro n Hstepn =>
ex_intro _ n (ex_intro _ h2 (ex_intro _ l2 (ex_intro _ s2 (ex_intro _ g2 (ex_intro _ tr2 (conj Hstepn Hf))))))
end
end
)).
destruct 1.
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H3.
destruct H5.
destruct H5.
destruct H5.
destruct H6.
destruct (H2 _ _ _ _ _ (stepn_left_star H3)).
subst.
destruct (H _ _ _ _ _ (stepn_left_star H3) _ _ H5 H6 H7).
destruct H8.
destruct H8.
destruct H8.
destruct (plus_stepn_left H8); clear H8.
inversion H9; subst.
generalize (guided_stepn_inv_correct H3 H1).
inversion 1; subst.
inversion Hstep; subst.
inversion Hstep0; subst.
inversion Hstep1; subst.
generalize (eval_internal_kind Hstep2).
rewrite H5.
discriminate.
rewrite H5 in Hl.
injection Hl; intros; subst.
change (fr :: x7 ++ s0) with (nil ++ fr :: x7 ++ s0) in Hstepn.
destruct (guided_cstepn_uncons Hstepn HCOIND).
generalize (x7 ++ s0).
intros.
intro.
generalize (refl_equal (length l0)).
rewrite H10 at 1.
rewrite app_length.
simpl.
omega.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H11.
destruct H12.
destruct H12.
destruct H12.
destruct H12.
destruct H13.
generalize (guided_stepn_inv_correct H10 HCOIND).
inversion 1; subst.
inversion H12; subst.
inversion Hstep2; subst.
inversion Hstep3; subst.
generalize (refl_equal (length (x7 ++ s0))).
rewrite <- H23 at 1.
generalize (x7 ++ s0).
simpl.
intros; exfalso; omega.
generalize (refl_equal (length (x7 ++ s0))).
rewrite <- H23 at 1.
generalize (x7 ++ s0).
simpl.
intros; exfalso; omega.
generalize (stepn_left_app H3 (stepn_left_S Hstep (stepn_left_app H10 (stepn_left_S H12 (stepn_left_O _ _) (refl_equal _))) (refl_equal _))).
rewrite <- app_nil_end in *.
simpl in *.
intro.
exists x4.
exists func.
exists ar.
exists h'.
exists r.
exists x5.
exists x13.
exists x14.
exists x7.
exists x15.
split.
destruct x4; discriminate.
split.
esplit.
intro.
intros.
eapply H.
eapply star_trans'.
eapply stepn_left_star.
eexact H16.
eassumption.
reflexivity.
eassumption.
assumption.
assumption.
split.
intro.
intros.
eapply H0.
eapply star_trans'.
eapply stepn_left_star.
eexact H16.
eassumption.
reflexivity.
split.
assumption.
intros.
eapply H2.
eapply star_trans'.
eapply stepn_left_star.
eexact H16.
eassumption.
reflexivity.
unfold build_invarK_prop.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
intros.
apply NNPP.
intro.
edestruct H4.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
eauto 9.
split.
assumption.
split.
assumption.
split.
assumption.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
split.
assumption.
assumption.
rewrite H5 in Hcmd.
discriminate.
rewrite H5 in Hcmd.
injection Hcmd; intros; subst.
contradiction.
Qed.
Lemma build_invarK_loop_body_type: forall h l s g,
build_invarK_inv h l s g ->
{ tr1 : _ & { f : _ & { a : _ & { h2 : _ & {r : _ & { h4 : _ & { h5 : _ & { l5 : _ & { s1 : _ & { g5 |
tr1 ++ External f a h2 r h4 :: nil <> nil /\
build_invarK_inv h5 l5 s1 g5 /\
build_invarK_prop h l s g tr1 f a h2 r h4 h5 l5 s1 g5
}}}}}}}}}}.
Proof.
intros.
destruct (constructive_indefinite_description _ (build_invarK_loop_body H)).
destruct (constructive_indefinite_description _ e); clear e.
destruct (constructive_indefinite_description _ e0); clear e0.
destruct (constructive_indefinite_description _ e); clear e.
destruct (constructive_indefinite_description _ e0); clear e0.
destruct (constructive_indefinite_description _ e); clear e.
destruct (constructive_indefinite_description _ e0); clear e0.
destruct (constructive_indefinite_description _ e); clear e.
destruct (constructive_indefinite_description _ e0); clear e0.
destruct (constructive_indefinite_description _ e); clear e.
eauto 11.
Qed.
CoFixpoint build_invarK_traceinf' (h: heap) (l: local) (s: list frame) (g: Stream (list xevent * config lang))
(Hg: build_invarK_inv h l s g) : traceinf' xevent :=
match build_invarK_loop_body_type Hg with
| existT tr1 (existT f (existT a (existT h2 (existT r (existT h4 (existT h5 (existT l5 (existT s1 (exist g5 (conj Hnonempty (conj Hg' _))))))))))) =>
Econsinf' _ (build_invarK_traceinf' Hg') Hnonempty
end.
Lemma build_invarK_traceinf'_invarK :
forall h l s g (Hg: build_invarK_inv h l s g) b
(Hbisim: bisim (traceinf_of_traceinf' (build_invarK_traceinf' Hg)) b) bs,
invarK s0 m (Config h l (s ++ s0), g) (BBehavior (Reacts _ (func_of_stream b)) bs).
Proof.
cofix COINDHYP.
intros.
rewrite (unroll_traceinf' (build_invarK_traceinf' Hg)) in Hbisim.
simpl in Hbisim.
destruct (build_invarK_loop_body_type Hg).
destruct s1.
destruct s1.
destruct s1.
destruct s1.
destruct s1.
destruct s1.
destruct s1.
destruct s1.
destruct s1.
destruct a.
destruct a.
rewrite traceinf_traceinf'_app in Hbisim.
destruct (bisim_appinf_strong Hbisim).
subst.
generalize (appinf_bisim Hbisim).
intro.
inversion Hg; subst.
destruct H1.
destruct H2.
destruct b1.
destruct H4.
destruct H4.
destruct H4.
destruct H4.
destruct H5.
destruct H6.
destruct H7.
destruct H8.
destruct H9.
destruct H9.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H11.
destruct H11.
destruct H11.
destruct H11.
destruct H11.
destruct H12.
destruct H13.
econstructor.
eassumption.
assumption.
eassumption.
assumption.
assumption.
eassumption.
assumption.
assumption.
eassumption.
eassumption.
eassumption.
assumption.
eassumption.
eassumption.
eapply COINDHYP.
eassumption.
rewrite append_behavior.
simpl.
cut (
(func_of_stream
(appinf (x ++ External (sw:=sw) x0 x1 x2 x3 x4 :: nil) x9)) =
(func_of_stream
(appinf (x ++ External (sw:=sw) x0 x1 x2 x3 x4 :: nil)
(stream_of_func (func_of_stream x9))))
).
simpl.
intros.
rewrite <- H15.
reflexivity.
apply bisim_func_eq.
eapply bisim_appinf.
apply stream_of_func_of_stream.
reflexivity.
Qed.
End Build_invarK.
Lemma instantiate_invariant : forall h l s s0 g
(Hcorrect: guided_execution_correct (cstep (merge m1 m2)) (Config h l (s ++ s0), g))
(Hforallstep: forall h' l' s' tr' g', star (@guided_step _ _) (Config h l (s ++ s0), g) tr' (Config h' l' s', g') -> exists s'', s' = s'' ++ s0)
m (Hm: m = m1 \/ m = m2),
exists tr', exists s', exists g', exists l', exists h',
star (@guided_step _ _) (Config h l (s ++ s0), g) tr' (Config h' l' (s' ++ s0), g') /\
exists tr,
star (cstep m) (Config h l s) tr (Config h' l' s') /\
exists b,
config_behaviors (cstep m) (@config_final _ _) (Config h' l' s') b /\
forall bs,
star (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (BBehavior (append1_behavior tr b) bs) tr' (BBehavior b bs) /\
(invarI m (Config h' l' (s' ++ s0), g') (BBehavior b bs) \/
invarJ s0 m (Config h' l' (s' ++ s0), g') (BBehavior b bs) \/
invarK s0 m (Config h' l' (s' ++ s0), g') (BBehavior b bs)).
Proof.
intros.
destruct (classify_external_divergence m Hcorrect).
generalize (case1_diverging_external_strong Hm Hcorrect H).
clear H.
intro.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H.
destruct H0.
destruct H0.
destruct H0.
destruct H1.
destruct H2.
destruct H3.
destruct H3.
destruct (Hforallstep _ _ _ _ _ (stepn_left_star H)).
subst.
edestruct resolve_terminating_external_calls.
eassumption.
eassumption.
eassumption.
assumption.
intros.
edestruct H4.
eassumption.
eassumption.
eassumption.
assumption.
assumption.
clear H4.
destruct H10.
destruct H4.
destruct H4.
destruct H4.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H11.
destruct H12.
destruct H12.
destruct H12.
destruct H12.
destruct H13.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eassumption.
split.
assumption.
esplit.
esplit.
esplit.
split.
eassumption.
split.
omega.
assumption.
destruct H5.
generalize (guided_stepn_inv_correct H Hcorrect).
inversion 1; subst.
inversion Hstep; subst.
inversion Hstep0; subst.
exfalso. generalize (eval_internal_kind Hstep1). rewrite H0. discriminate.
rewrite H0 in Hl. injection Hl; intros; subst.
generalize (ret_inh sw).
inversion 1.
case_eq (caller_return X h' fr).
intros.
assert (cstep m (Config x0 x1 x8) (External func ar h' X h' :: nil) (Config h0 l0 x8)).
eapply cstep_call_external.
assumption.
assumption.
eassumption.
eapply heap_le_refl.
assumption.
destruct (exists_behavior (@config_final _ _) (cstep m) (Config h0 l0 x8)).
assert (x7 = fr).
edestruct H3.
eleft.
econstructor.
change (fr :: x8 ++ s0) with ((fr :: nil) ++ (x8 ++ s0)) in H12.
change (x7 :: x8 ++ s0) with ((x7 :: nil) ++ (x8 ++ s0)) in H12.
rewrite <- (app_ass x5) in H12.
generalize (app_inv_tail _ _ _ H12).
destruct x5.
injection 1; intros; subst; trivial.
simpl. destruct x5; discriminate.
subst.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eapply stepn_left_star.
eassumption.
esplit.
split.
eassumption.
esplit.
split.
eapply star_config_behaviors.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
intros.
split.
eapply H6.
left.
econstructor.
eassumption.
econstructor.
eassumption.
assumption.
eassumption.
econstructor.
intros.
eapply H3.
eapply star_plus_left.
eleft.
econstructor.
eassumption.
simpl; reflexivity.
simpl; reflexivity.
assumption.
assumption.
eassumption.
reflexivity.
rewrite H0 in Hcmd. discriminate.
rewrite H0 in Hcmd. injection Hcmd; intros; subst. contradiction.
destruct (classify_external_finiteness Hm h l (s ++ s0) g).
generalize (case2_subcase1_finitely_many_externals_strong Hm Hcorrect H H0).
clear H H0.
destruct 1.
destruct (guided_stepn_exists x (Config h l (s ++ s0), g)).
destruct H0.
destruct x1.
destruct c.
destruct (Hforallstep _ _ _ _ _ (stepn_left_star H0)).
subst.
edestruct resolve_terminating_external_calls.
eassumption.
eassumption.
eassumption.
assumption.
intros.
eauto.
destruct H1.
assert (forall h1 l1 s2 tr1 g1, star (@guided_step _ _) (Config c_heap c_local (x1 ++ s0), s1) tr1 (Config h1 l1 s2, g1) -> forall f a, local_kind l1 = Some (Call f a) -> in_dom (merge m1 m2) f -> in_dom m f).
intros.
destruct (star_stepn_left H3).
destruct (in_dom_dec m f).
assumption.
exfalso.
destruct (H _ _ _ _ _ _ (stepn_left_app H0 H6) _ _ H4 H5 n).
destruct H7.
destruct H7.
destruct H7.
destruct H7.
destruct H8.
destruct H8.
destruct H8.
destruct H8.
destruct H8.
destruct H8.
destruct H9.
destruct H10.
destruct H10.
destruct H10.
destruct H10.
destruct H11.
omega.
destruct (exists_behavior (final := heap * ret) (fun _ _ => False) (@guided_step _ _) (Config c_heap c_local (x1 ++ s0), s1)).
assert (forall hs, invarJ s0 m (Config c_heap c_local (x1 ++ s0), s1) (BBehavior x3 hs)).
intros.
econstructor.
eapply guided_stepn_inv_correct.
eassumption.
assumption.
intros; eauto using star_trans', stepn_left_star.
assumption.
assumption.
assumption.
reflexivity.
assert (invarJ' s0 m (Config c_heap c_local (x1 ++ s0), s1) (Config c_heap c_local x1)).
eapply invarJ'_intro with (bs := nil).
eapply H5.
reflexivity.
generalize (simult_guided_behaviors (@invarJ'_sim s0 m) H4 H6 (@config_final _ _)).
intros.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eapply stepn_left_star.
eassumption.
esplit.
split.
eassumption.
esplit.
split.
eassumption.
intros.
split.
eauto.
right.
left.
eauto.
assert (build_invarK_inv m s0 h l s g).
split.
assumption.
split.
assumption.
split.
assumption.
assumption.
generalize (build_invarK_traceinf'_invarK (Hm := Hm) (Hg := H1) (bisim_refl _)).
intro.
generalize (invarK_cbehavior (H2 nil) (bisim_refl _)).
intro.
esplit.
esplit.
esplit.
esplit.
esplit.
split.
eleft.
esplit.
split.
eleft.
eexists (Reacts _ _).
split.
simpl.
eapply config_reacts_bisim.
eassumption.
eapply stream_of_func_of_stream.
intros.
split.
eleft.
right.
right.
eapply H2.
Qed.
Inductive invarIJK (s: config lang * Stream (list xevent * config lang)) (s': bstate sw): Prop :=
| invarIJK_I m (HI: invarI m s s')
| invarIJK_J s0 m (HJ: invarJ s0 m s s')
| invarIJK_K s0 m (HK: invarK s0 m s s')
.
Lemma invarIJK_sim:
simult (@guided_step _ _) (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) invarIJK.
Proof.
intro.
inversion 1; subst.
inversion HI; subst.
inversion Hcorrect; subst.
inversion Hstep; subst.
inversion Hstep0; subst.
inversion Hstep1; try subst.
generalize (eval_internal_kind Hstep2). rewrite Hl. discriminate.
rewrite Hl0 in Hl. injection Hl; intros; subst.
rewrite Hcall in Hframe.
injection Hframe; intros; subst.
destruct (mergelookup' Hfunc).
destruct H0.
inversion Hfoo; subst.
destruct tr; discriminate.
rewrite Hfr in Hcall.
injection Hcall; intros; subst.
change (fr :: s ++ s0) with (nil ++ (fr :: s ++ s0)) in HCOIND.
destruct (instantiate_invariant HCOIND Hforallstep H1).
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H3.
destruct H3.
destruct H4.
destruct H4.
destruct (H5 (((h2, r), b) :: bs)); clear H5.
generalize (in_dom_dom (lookup_in_dom H0)).
rewrite <- (create_map_dom (dom x) (csemantics x)).
intro.
destruct (dom_in_dom H5).
assert (
create_map (dom x) (csemantics x) = create_map (dom m1) (csemantics m1)
\/ create_map (dom x) (csemantics x) = create_map (dom m2) (csemantics m2)
).
destruct H1; subst; eauto.
refine (_ (lookup_disj_merge H9 H8 _)).
intro.
generalize (create_map_lookup_elim H8).
intro; subst.
esplit.
esplit.
split.
eapply star_plus_left.
eleft.
eassumption.
eassumption.
simpl; reflexivity.
esplit.
split.
eapply star_plus_left.
eleft.
eapply bstep_internal.
eassumption.
econstructor.
eassumption.
eassumption.
eapply star_config_behaviors.
eassumption.
eassumption.
reflexivity.
reflexivity.
eassumption.
reflexivity.
destruct H7.
eapply invarIJK_I. eassumption.
destruct H7.
eapply invarIJK_J. eassumption.
eapply invarIJK_K. eassumption.
intro. intros.
generalize (in_dom_dom H10).
generalize (in_dom_dom H11).
rewrite create_map_dom.
rewrite create_map_dom.
intros; eauto using dom_in_dom.
rewrite Hl in Hcmd.
discriminate.
destruct (invarJ_sim HJ).
destruct H0.
destruct H0.
destruct H1.
destruct H1.
eauto 7 using invarIJK_J.
destruct (invarK_sim HK).
destruct H0.
destruct H0.
destruct H1.
destruct H1.
eauto 7 using invarIJK_K.
Qed.
Corollary csemantics_in_bsemantics_diverges_or_reacts: forall b,
((exists f, b = Diverges _ f) \/ exists f, b = Reacts _ f) -> forall
f a h,
csemantics (merge m1 m2) f a h b ->
bsemantics (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))) f a h b.
Proof.
intros.
inversion H0; subst.
destruct (mergelookup' Hcmd).
destruct H1.
assert (exists t, guided_execution_correct (cstep (merge m1 m2)) (Config h' l' nil, t) /\ config_behaviors (@guided_step _ _) (fun _ _ => False) (Config h' l' nil, t) b).
destruct H; destruct H; subst; simpl; eauto using config_diverges_guided_execution, config_reacts_guided_step.
destruct H3.
destruct H3.
change nil with (nil ++ nil (A := frame)) in H3.
edestruct instantiate_invariant.
eassumption.
intros.
esplit.
rewrite <- app_nil_end.
reflexivity.
eassumption.
destruct H5.
destruct H5.
destruct H5.
destruct H5.
destruct H5.
destruct H6.
destruct H6.
destruct H7.
destruct H7.
destruct (H8 nil); clear H8.
repeat rewrite <- app_nil_end in *.
destruct (exists_behavior (final := heap * ret) (fun _ _ => False) (@guided_step _ _) (Config x5 x4 x2, x3)).
assert (b = append1_behavior x1 x8).
generalize (star_config_behaviors H5 H8 (refl_equal _)).
intro.
destruct (guided_behavior_inv H8).
destruct H12.
subst.
rewrite append_behavior in *.
simpl in *.
destruct H; destruct H; subst; simpl in *.
f_equal; eauto using config_diverges_determ, guided_step_determ.
edestruct diverges_not_reacts; eauto using guided_step_determ.
destruct H12.
subst.
rewrite append_behavior in *.
simpl in *.
destruct H; destruct H; subst; simpl in *.
edestruct diverges_not_reacts; eauto using guided_step_determ.
f_equal.
rewrite <- (func_of_stream_of_func x8).
apply bisim_func_eq.
eapply bisim_trans.
eapply config_reacts_determ.
2: eassumption.
eapply guided_step_determ.
eexact H11.
apply bisim_sym.
apply stream_of_func_of_stream.
subst.
assert (invarIJK (Config x5 x4 x2, x3) (BBehavior x7 nil)).
destruct H10.
eapply invarIJK_I.
eassumption.
destruct H10.
eapply invarIJK_J.
eassumption.
eapply invarIJK_K.
eassumption.
assert (config_behaviors (bstep (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2)))) (@bstate_final _) (BBehavior x7 nil) (bbehavior_of_cbehavior x8)).
destruct (guided_behavior_inv H8).
destruct H12; subst; simpl in *.
eapply config_diverges_simult.
2: eapply invarIJK_sim.
apply guided_step_determ.
eassumption.
assumption.
destruct H12; subst; simpl in *.
eapply config_reacts_simult.
2: eapply invarIJK_sim.
apply guided_step_determ.
eassumption.
assumption.
generalize (star_config_behaviors H9 H12 (refl_equal _)).
rewrite bbehavior_of_cbehavior_append.
intro.
generalize (star_config_behaviors H6 H7 (refl_equal _)).
intro.
generalize (in_dom_dom (lookup_in_dom H1)).
rewrite <- (create_map_dom (dom x) (csemantics x)).
intro.
destruct (dom_in_dom H15).
assert (
create_map (dom x) (csemantics x) = create_map (dom m1) (csemantics m1)
\/ create_map (dom x) (csemantics x) = create_map (dom m2) (csemantics m2)
).
destruct H2; subst; eauto.
refine (_ (lookup_disj_merge H17 H16 _)).
intro.
generalize (create_map_lookup_elim H16).
intro; subst.
econstructor.
eassumption.
econstructor.
eassumption.
eassumption.
eassumption.
reflexivity.
assumption.
intro. intros.
generalize (in_dom_dom H18).
generalize (in_dom_dom H19).
rewrite create_map_dom.
rewrite create_map_dom.
intros; eauto using dom_in_dom.
Qed.
Corollary csemantics_in_bsemantics: forall f a h b,
csemantics (merge m1 m2) f a h b -> bsemantics (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))) f a h b.
Proof.
inversion 1; subst.
destruct (mergelookup' Hcmd).
destruct H0.
generalize (in_dom_dom (lookup_in_dom H0)).
rewrite <- (create_map_dom (dom x) (csemantics x)).
intro.
destruct (dom_in_dom H2).
assert (
create_map (dom x) (csemantics x) = create_map (dom m1) (csemantics m1)
\/ create_map (dom x) (csemantics x) = create_map (dom m2) (csemantics m2)
).
destruct H1; subst; eauto.
refine (_ (lookup_disj_merge H4 H3 _)).
intro.
generalize (create_map_lookup_elim H3).
intro; subst.
destruct b; simpl in *.
inversion Hbeh; subst.
inversion Hfin; subst.
destruct (csemantics_in_bsemantics_terminates_aux Hbeh H1).
destruct H5.
eapply bsemantics_intro with (cb := Terminates _ _).
eassumption.
econstructor.
eassumption.
eassumption.
eassumption.
reflexivity.
assumption.
inversion Hbeh; subst.
destruct (csemantics_in_bsemantics_stuck_aux Hbeh H1).
destruct H5.
eapply bsemantics_intro.
eassumption.
econstructor.
eassumption.
eassumption.
eassumption.
reflexivity.
simpl.
eapply H6.
eapply csemantics_in_bsemantics_diverges_or_reacts.
eauto.
assumption.
eapply csemantics_in_bsemantics_diverges_or_reacts.
eauto.
assumption.
intro.
intros.
generalize (in_dom_dom H5).
generalize (in_dom_dom H6).
rewrite create_map_dom.
rewrite create_map_dom.
eauto using dom_in_dom.
Qed.
Corollary csemantics_bsemantics:
csemantics (merge m1 m2) = bsemantics (merge (create_map (dom m1) (csemantics m1)) (create_map (dom m2) (csemantics m2))).
Proof.
apply functional_extensionality.
intro f.
apply functional_extensionality.
intro a.
apply functional_extensionality.
intro h.
apply Extensionality_Ensembles.
split; unfold Included, Ensembles.In; intros; eauto using csemantics_in_bsemantics, bsemantics_in_csemantics.
Qed.
End Link.
The following generalizes the result to n modules.
Lemma csemantics_bsemantics_n: forall l: list (module lang),
disjoint_embedded l ->
fold_right (@link _) (@empty_map _ _) (map (fun m => create_map (dom m) (csemantics m)) l) = create_map (dom (merge_embedded l)) (csemantics (merge_embedded l)).
Proof.
induction 1.
simpl.
reflexivity.
simpl.
unfold link at 1.
simpl.
rewrite IHdisjoint_embedded.
rewrite csemantics_bsemantics.
rewrite dom_merge.
rewrite create_map_dom.
rewrite create_map_dom.
rewrite dom_merge.
reflexivity.
assumption.
Qed.
End Language.
End Semworld.
Summarize the final statement and used axioms.
Goal True.
idtac "Theorem:".
Check csemantics_bsemantics.
Print Assumptions csemantics_bsemantics.
idtac "".
trivial.
Save.
idtac "Theorem:".
Check csemantics_bsemantics.
Print Assumptions csemantics_bsemantics.
idtac "".
trivial.
Save.