Library smallstep
Tools for small-step and big-step operational semantics
This module defines generic operations and theorems over
the one-step transition relations that are used to specify
operational semantics in small-step style. Then, it defines big-stepping
Require Export List stream Omega MinMax.
Set Implicit Arguments.
Unset Strict Implicit.
Section Event.
The observable behaviour of programs is stated in terms of input/output events, which can also be thought of as system calls to the operating system.
EVENT is the set of events
CPP 2015 submission: Definition 1. Small-step semantics
Variable EVENT: Type.
Section Smallstep.
CONFIG is the set of configurations (or states); and STEP is the transition relation.
Variable (CONFIG: Type) (STEP: CONFIG -> list EVENT -> CONFIG -> Prop).
Transitive and reflexive-transitive closure of the transition relation
Inductive plus : CONFIG -> list EVENT -> CONFIG -> Prop :=
| plus_step: forall conf tr conf'
(Hstep: STEP conf tr conf'),
plus conf tr conf'
| plus_trans: forall conf tr1 conf1
(Hplus1: plus conf tr1 conf1)
tr2 conf2
(Hplus2: plus conf1 tr2 conf2),
plus conf (tr1 ++ tr2) conf2
.
Inductive star : CONFIG -> list EVENT -> CONFIG -> Prop :=
| star_refl : forall conf,
star conf nil conf
| star_plus : forall conf tr conf'
(Hplus: plus conf tr conf'),
star conf tr conf'
.
Lemma star_trans: forall (conf: CONFIG) (tr1: _ EVENT) conf1
(Hstar1: star conf tr1 conf1)
conf2 tr2
(Hstar2: star conf1 tr2 conf2),
star conf (tr1 ++ tr2) conf2
.
Proof.
induction 1; simpl; trivial.
induction Hplus; simpl.
inversion 1; subst.
rewrite <- app_nil_end.
right; left; eauto.
right.
eright.
eleft.
eassumption.
assumption.
intros.
rewrite app_ass.
eauto.
Qed.
Corollary star_trans': forall (conf: CONFIG) (tr1: _ EVENT) conf1
(Hstar1: star conf tr1 conf1)
conf2 tr2
(Hstar2: star conf1 tr2 conf2)
tr (Htr: tr = tr1 ++ tr2),
star conf tr conf2
.
Proof.
intros; subst; eauto using star_trans.
Qed.
Lemma plus_star_left: forall conf tr conf'
(Hplus: plus conf tr conf'),
exists tr1, exists conf1, STEP conf tr1 conf1 /\
exists tr2, star conf1 tr2 conf' /\ tr = tr1 ++ tr2.
Proof.
induction 1; subst.
esplit.
esplit.
split.
eassumption.
esplit.
split.
eleft.
apply app_nil_end.
clear IHHplus2.
destruct IHHplus1.
destruct H.
destruct H.
destruct H0.
destruct H0.
subst.
esplit.
esplit.
split.
eassumption.
rewrite app_ass.
esplit.
split.
eapply star_trans.
eassumption.
eright.
eassumption.
reflexivity.
Qed.
Lemma plus_star_right: forall conf tr conf'
(Hplus: plus conf tr conf'),
exists tr1, exists conf1, star conf tr1 conf1 /\
exists tr2, STEP conf1 tr2 conf' /\ tr = tr1 ++ tr2.
Proof.
induction 1; subst.
esplit.
esplit.
split.
eleft.
esplit.
split.
eassumption.
reflexivity.
clear IHHplus1.
destruct IHHplus2.
destruct H.
destruct H.
destruct H0.
destruct H0.
subst.
esplit.
esplit.
split.
eapply star_trans.
eright.
eassumption.
eassumption.
esplit.
split.
eassumption.
rewrite app_ass.
reflexivity.
Qed.
Lemma star_ind2: forall
(P: _ -> _ EVENT -> CONFIG -> Prop)
(Hrefl: forall conf, P conf nil conf)
(Hstep: forall conf tr conf' (Hstep: STEP conf tr conf'), P conf tr conf')
(Htrans: forall conf tr1 conf1
(Hstar1: star conf tr1 conf1)
(IHstar1: P conf tr1 conf1)
tr2 conf2
(Hstar2: star conf1 tr2 conf2)
(IHstar2: P conf1 tr2 conf2),
P conf (tr1 ++ tr2) conf2
)
conf tr conf'
(Hstar: star conf tr conf'),
P conf tr conf'.
Proof.
intros.
induction Hstar; auto.
induction Hplus; auto.
eauto using star_plus.
Qed.
Inductive stepn_left : nat -> CONFIG -> list EVENT -> CONFIG -> Prop :=
| stepn_left_O : forall conf,
stepn_left O conf nil conf
| stepn_left_S : forall conf tr1 conf1
(Hstep: STEP conf tr1 conf1)
conf2 tr2 n
(Hstepn: stepn_left n conf1 tr2 conf2)
tr (Htr: tr = tr1 ++ tr2),
stepn_left (S n) conf tr conf2
.
Lemma stepn_left_app: forall n1 conf tr1 conf1
(Hstepn1: stepn_left n1 conf tr1 conf1)
n2 tr2 conf2
(Hstepn2: stepn_left n2 conf1 tr2 conf2),
stepn_left (n1 + n2) conf (tr1 ++ tr2) conf2.
Proof.
induction 1; simpl.
tauto.
subst.
intros.
rewrite app_ass.
econstructor.
eassumption.
eauto.
reflexivity.
Qed.
Lemma stepn_left_app_recip: forall n1 n2 conf tr conf'
(Hstepn: stepn_left (n1 + n2) conf tr conf'),
exists tr1, exists conf1,
stepn_left n1 conf tr1 conf1 /\
exists tr2,
stepn_left n2 conf1 tr2 conf' /\
tr = tr1 ++ tr2.
Proof.
induction n1; simpl.
intros; eauto 6 using stepn_left_O.
inversion 1; subst.
destruct (IHn1 _ _ _ _ Hstepn0).
destruct H.
destruct H.
destruct H0.
destruct H0.
subst.
eauto 7 using app_ass, eq_sym, stepn_left_S.
Qed.
Lemma plus_stepn_left: forall conf tr conf'
(Hplus: plus conf tr conf'),
exists n, stepn_left (S n) conf tr conf'.
Proof.
induction 1; subst.
esplit.
econstructor.
eassumption.
eleft.
apply app_nil_end.
destruct IHHplus1.
destruct IHHplus2.
exists (x + S x0).
replace (S (x + S x0)) with (S x + S x0) by omega.
eauto using stepn_left_app.
Qed.
Lemma stepn_left_plus: forall n conf tr conf'
(Hstepn: stepn_left (S n) conf tr conf'),
plus conf tr conf'.
Proof.
induction n; simpl; inversion 1; subst.
inversion Hstepn0; subst.
rewrite <- app_nil_end.
left.
assumption.
eright.
eleft.
eassumption.
apply IHn.
assumption.
Qed.
Lemma star_stepn_left: forall conf tr conf'
(Hstar: star conf tr conf'),
exists n, stepn_left n conf tr conf'.
Proof.
induction 1.
exists O; left.
destruct (plus_stepn_left Hplus); eauto.
Qed.
Lemma stepn_left_star: forall n conf tr conf'
(Hstepn: stepn_left n conf tr conf'),
star conf tr conf'.
Proof.
induction 1.
left.
subst.
eapply star_trans.
eright.
eleft.
eassumption.
assumption.
Qed.
Inductive stepn_right : nat -> CONFIG -> list EVENT -> CONFIG -> Prop :=
| stepn_right_O : forall conf,
stepn_right O conf nil conf
| stepn_right_S : forall conf tr1 conf1 n
(Hstepn: stepn_right n conf tr1 conf1)
conf2 tr2
(Hstep: STEP conf1 tr2 conf2)
tr (Htr: tr = tr1 ++ tr2),
stepn_right (S n) conf tr conf2
.
Lemma stepn_left_right: forall n conf tr conf'
(Hstep: stepn_left n conf tr conf'),
stepn_right n conf tr conf'.
Proof.
induction 1.
left.
subst.
clear Hstep0.
revert conf tr1 Hstep.
induction IHHstep.
intros.
rewrite <- app_nil_end.
eright.
eleft.
eassumption.
reflexivity.
subst.
intros.
econstructor.
eauto.
eassumption.
rewrite app_ass.
reflexivity.
Qed.
Lemma stepn_right_left: forall n conf tr conf'
(Hstep: stepn_right n conf tr conf'),
stepn_left n conf tr conf'.
Proof.
induction 1.
left.
subst.
clear Hstep.
revert conf2 tr2 Hstep0.
induction IHHstep.
intros.
eright.
eassumption.
eleft.
rewrite <- app_nil_end.
reflexivity.
subst.
intros.
econstructor.
eassumption.
eauto.
rewrite app_ass.
reflexivity.
Qed.
Lemma star_ind3_left: forall
(P: _ -> _ EVENT -> CONFIG -> Prop)
(Hrefl: forall conf, P conf nil conf)
(Hstep: forall conf tr1 conf1 (Hstep: STEP conf tr1 conf1)
tr2 conf2 (Hstar2: star conf1 tr2 conf2)
(IHstar2: P conf1 tr2 conf2),
P conf (tr1 ++ tr2) conf2)
conf tr conf'
(Hstar: star conf tr conf'),
P conf tr conf'.
Proof.
intros.
destruct (star_stepn_left Hstar).
clear Hstar.
induction H.
auto.
subst; eauto using stepn_left_star.
Qed.
Lemma star_ind3_right: forall
(P: _ -> _ EVENT -> CONFIG -> Prop)
(Hrefl: forall conf, P conf nil conf)
(Hstep: forall conf tr1 conf1 (Hstar1: star conf tr1 conf1)
(IHstar1: P conf tr1 conf1)
tr2 conf2 (Hstep: STEP conf1 tr2 conf2),
P conf (tr1 ++ tr2) conf2)
conf tr conf'
(Hstar: star conf tr conf'),
P conf tr conf'.
Proof.
intros.
destruct (star_stepn_left Hstar).
clear Hstar.
generalize (stepn_left_right H).
clear H.
induction 1.
auto.
subst; eauto using stepn_left_star, stepn_right_left.
Qed.
Lemma plus_ind3_left: forall
(P: _ -> _ EVENT -> CONFIG -> Prop)
(IHstep: forall conf tr1 conf1
(Hstep: STEP conf tr1 conf1),
P conf tr1 conf1)
(IHplus: forall conf tr1 conf1 (Hstep: STEP conf tr1 conf1)
tr2 conf2 (Hstar2: plus conf1 tr2 conf2)
(IHstar2: P conf1 tr2 conf2),
P conf (tr1 ++ tr2) conf2)
conf tr conf'
(Hplus: plus conf tr conf'),
P conf tr conf'.
Proof.
intros.
destruct (plus_stepn_left Hplus).
clear Hplus.
revert conf tr conf' H.
induction x; simpl.
inversion 1; subst.
inversion Hstepn; subst.
rewrite <- app_nil_end.
eauto.
inversion 1; subst.
eapply IHplus.
eassumption.
eapply stepn_left_plus.
eassumption.
eauto.
Qed.
Lemma plus_ind3_right: forall
(P: _ -> _ EVENT -> CONFIG -> Prop)
(IHstep: forall conf tr1 conf1
(Hstep: STEP conf tr1 conf1),
P conf tr1 conf1)
(IHplus: forall conf tr1 conf1 (Hplus1: plus conf tr1 conf1)
(IHplus1: P conf tr1 conf1)
tr2 conf2 (Hstep: STEP conf1 tr2 conf2),
P conf (tr1 ++ tr2) conf2)
conf tr conf'
(Hplus: plus conf tr conf'),
P conf tr conf'.
Proof.
intros.
destruct (plus_stepn_left Hplus).
clear Hplus.
generalize (stepn_left_right H).
clear H.
revert conf tr conf'.
induction x; simpl.
inversion 1; subst.
inversion Hstepn; subst.
simpl; eauto.
inversion 1; subst.
eapply IHplus.
eapply stepn_left_plus.
eapply stepn_right_left.
eassumption.
eauto.
assumption.
Qed.
Lemma star_not_nil: forall conf l conf'
(Hstar: star conf l conf')
e l'
(Hl: l = e :: l'),
exists conf1,
star conf nil conf1 /\
exists l1, exists conf2,
STEP conf1 (e :: l1) conf2 /\
exists l2,
star conf2 l2 conf' /\ l' = l1 ++ l2.
Proof.
induction 1 using star_ind3_left.
discriminate.
destruct tr1.
simpl.
intros; subst.
destruct (IHHstar _ _ (refl_equal _)).
destruct H.
destruct H0.
destruct H0.
destruct H0.
destruct H1.
destruct H1.
subst.
esplit.
split.
change nil with (nil ++ nil (A := EVENT)).
eapply star_trans.
eright.
eleft.
eassumption.
eassumption.
eauto 6.
simpl.
injection 1; intros; subst.
esplit.
split.
eleft.
eauto 6.
Qed.
Lemma star_plus_left: forall conf tr1 conf1
(Hplus: plus conf tr1 conf1)
tr2 conf2 (Hstar: star conf1 tr2 conf2)
tr (Htr: tr = tr1 ++ tr2),
plus conf tr conf2.
Proof.
intros.
destruct (plus_stepn_left Hplus).
destruct (star_stepn_left Hstar).
generalize (stepn_left_app H H0).
intros.
simpl in H1.
subst.
eapply stepn_left_plus.
eassumption.
Qed.
Lemma star_plus_right: forall conf tr1 conf1
(Hstar: star conf tr1 conf1)
tr2 conf2 (Hplus: plus conf1 tr2 conf2)
tr (Htr: tr = tr1 ++ tr2),
plus conf tr conf2.
Proof.
intros.
destruct (plus_stepn_left Hplus).
destruct (star_stepn_left Hstar).
generalize (stepn_left_app H0 H).
intros.
replace (x0 + S x) with (S (x0 + x)) in H1 by omega.
subst.
eapply stepn_left_plus.
eassumption.
Qed.
final is the set of results
Variable final : Type.
config_final is a relation associating final states with results.
Variable config_final : CONFIG -> final -> Prop.
Inductive config_terminates (conf: CONFIG) (evseq : list EVENT) (hfin : final) : Prop :=
| config_terminates_intro cfin
(Hterm: star conf evseq cfin)
(Hfin: config_final cfin hfin)
.
Inductive config_stuck (conf: CONFIG) (evseq: list EVENT) : Prop :=
| config_stuck_intro cstuck
(Hstar: star conf evseq cstuck)
(Hstuck_not_final: forall hfin (Habs: config_final cstuck hfin), False)
(Hconf_stuck: forall tr conf' (Hstep_abs: STEP cstuck tr conf'), False)
.
CoInductive config_silently_diverges : CONFIG -> Prop :=
| config_silently_diverges_intro
conf conf'
(Hstep: STEP conf nil conf')
(CIHsilently_diverges: config_silently_diverges conf') :
config_silently_diverges conf
.
Lemma config_silently_diverges_star: forall conf l conf1
(Hstar: star conf l conf1)
(l_nil: l = nil)
(Hdiverges: config_silently_diverges conf1),
config_silently_diverges conf.
Proof.
induction 1 using star_ind3_left.
intros. assumption.
intros.
destruct (app_eq_nil _ _ l_nil).
subst.
simpl in *.
econstructor.
eassumption.
apply IHHstar.
reflexivity.
assumption.
Qed.
CoInductive config_silently_diverges' : CONFIG -> Prop :=
| config_silently_diverges'_intro
conf conf'
(Hstep: plus conf nil conf')
(CIHsilently_diverges: config_silently_diverges' conf') :
config_silently_diverges' conf
.
Lemma config_silently_diverges'_diverges: forall conf
(Hdiv': config_silently_diverges' conf),
config_silently_diverges conf.
Proof.
cofix COINDHYP.
inversion 1; subst.
destruct (plus_stepn_left Hstep).
clear Hstep.
inversion H; subst.
symmetry in Htr.
destruct (app_eq_nil _ _ Htr).
clear Htr.
subst.
clear H.
assert (config_silently_diverges' conf1).
clear COINDHYP.
inversion Hstepn; subst.
assumption.
symmetry in Htr.
destruct (app_eq_nil _ _ Htr).
clear Htr.
subst.
econstructor.
eapply stepn_left_plus.
eassumption.
assumption.
clear conf' CIHsilently_diverges Hstepn.
econstructor.
eassumption.
apply COINDHYP.
assumption.
Qed.
Lemma config_silently_diverges_diverges': forall conf
(Hdiv': config_silently_diverges conf),
config_silently_diverges' conf.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor.
eleft.
eassumption.
apply COINDHYP.
assumption.
Qed.
Section Forever_silent_N.
Variable A: Type.
Variable order: A -> A -> Prop.
An alternate definition.
CoInductive forever_silent_N : A -> CONFIG -> Prop :=
| forever_silent_N_star: forall s1 s2 a1 a2,
star s1 nil s2 ->
order a2 a1 ->
forever_silent_N a2 s2 ->
forever_silent_N a1 s1
| forever_silent_N_plus: forall s1 s2 a1 a2,
plus s1 nil s2 ->
forever_silent_N a2 s2 ->
forever_silent_N a1 s1.
Hypothesis order_wf: well_founded order.
Lemma forever_silent_N_inv:
forall a s,
forever_silent_N a s ->
exists s', exists a',
STEP s nil s' /\ forever_silent_N a' s'.
Proof.
intros a0. pattern a0. apply (well_founded_ind order_wf).
intros. inversion H0; subst.
destruct (star_stepn_left H1); clear H1.
inversion H4; subst.
apply H with a2. auto. auto.
symmetry in Htr; destruct (app_eq_nil _ _ Htr); subst.
esplit.
esplit.
split.
eassumption.
eapply forever_silent_N_star; eauto using stepn_left_star.
destruct (plus_stepn_left H1); clear H1.
inversion H3; subst.
symmetry in Htr; destruct (app_eq_nil _ _ Htr); subst.
destruct x0.
inversion Hstepn; subst.
esplit.
esplit.
split.
eassumption.
eassumption.
esplit.
exists a2.
split.
eassumption.
eapply forever_silent_N_plus.
eapply stepn_left_plus.
eassumption.
eassumption.
Qed.
Lemma forever_silent_N_forever:
forall a s, forever_silent_N a s -> config_silently_diverges s.
Proof.
cofix COINDHYP; intros.
destruct (forever_silent_N_inv H) as [s' [a' [P Q]]].
econstructor. eassumption. eapply COINDHYP. eassumption.
Qed.
End Forever_silent_N.
Inductive config_diverges (conf: CONFIG) (evseq: list EVENT) : Prop :=
| config_diverges_intro conf'
(Hstar: star conf evseq conf')
(Hsilently_diverges: config_silently_diverges conf')
.
Lemma config_diverges_nil_silently_diverges: forall conf
(Hconf: config_diverges conf nil),
config_silently_diverges conf.
Proof.
inversion 1; subst.
eapply config_silently_diverges_star.
eassumption.
reflexivity.
assumption.
Qed.
Lemma config_silently_diverges_inv: forall conf (Hdiv: config_silently_diverges conf) n, exists conf', stepn_left n conf nil conf' /\ config_silently_diverges conf'.
Proof.
intros.
revert conf Hdiv.
induction n; intros.
esplit.
split.
econstructor.
assumption.
inversion Hdiv; subst.
destruct (IHn _ CIHsilently_diverges).
destruct H.
esplit.
split.
econstructor.
eassumption.
eassumption.
reflexivity.
assumption.
Qed.
CoInductive config_diverges_N_min : CONFIG -> nat -> list EVENT -> Prop :=
| config_diverges_N_min_nil conf conf'
(Hstep: STEP conf nil conf')
tr (Htr: tr <> nil) n (HCOIND: config_diverges_N_min conf' n tr):
config_diverges_N_min conf (S n) tr
| config_diverges_N_min_silent conf conf' (Hstep: STEP conf nil conf')
(HCOIND: config_diverges_N_min conf' O nil):
config_diverges_N_min conf O nil
| config_diverges_N_min_cons tr1 (Htr1: tr1 <> nil)
conf conf' (Hstep: STEP conf tr1 conf')
n' tr2 (HCOIND: config_diverges_N_min conf' n' tr2)
tr (Htr: tr = tr1 ++ tr2):
config_diverges_N_min conf O tr
.
Lemma config_silently_diverges_N_min: forall conf (Hdiv: config_silently_diverges conf),
config_diverges_N_min conf O nil.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor.
eassumption.
eapply COINDHYP.
assumption.
Qed.
Lemma config_diverges_config_diverges_N_min: forall conf tr (Hdiv: config_diverges conf tr),
exists n, config_diverges_N_min conf n tr.
Proof.
inversion 1; subst; clear Hdiv.
destruct (star_stepn_left Hstar); clear Hstar.
generalize (config_silently_diverges_N_min Hsilently_diverges); clear Hsilently_diverges.
induction H.
eauto.
intros.
destruct (IHstepn_left H0); clear IHstepn_left H0.
destruct tr1.
simpl in Htr; subst.
destruct tr2.
inversion H1; subst.
destruct Htr.
trivial.
esplit.
eapply config_diverges_N_min_silent.
eassumption.
assumption.
symmetry in Htr.
destruct (app_eq_nil _ _ Htr).
contradiction.
esplit.
econstructor.
eassumption.
congruence.
eassumption.
esplit.
eapply config_diverges_N_min_cons.
2: eassumption.
congruence.
eassumption.
assumption.
Qed.
CoInductive config_diverges_N_lt : CONFIG -> nat -> list EVENT -> Prop :=
| config_diverges_N_lt_intro conf tr1 conf'
(Hstep: STEP conf tr1 conf')
tr2 n'
(HCOIND: config_diverges_N_lt conf' n' tr2)
n
(Hn: tr1 = nil -> tr2 <> nil -> n' < n)
tr (Htr: tr = tr1 ++ tr2):
config_diverges_N_lt conf n tr.
Lemma config_diverges_N_min_lt: forall conf n tr (Hdiv: config_diverges_N_min conf n tr),
config_diverges_N_lt conf n tr.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
clear COINDHYP; intros; omega.
reflexivity.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
destruct 2; trivial.
reflexivity.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
intro; contradiction.
reflexivity.
Qed.
CoInductive config_diverges_N_plus : CONFIG -> nat -> list EVENT -> Prop :=
| config_diverges_N_plus_intro conf tr1 conf'
(Hstep: plus conf tr1 conf')
tr2 n'
(HCOIND: config_diverges_N_plus conf' n' tr2)
n
(Hn: tr1 = nil -> tr2 <> nil -> n' < n)
tr (Htr: tr = tr1 ++ tr2):
config_diverges_N_plus conf n tr.
Lemma config_diverges_N_lt_plus: forall conf n tr (Hdiv: config_diverges_N_lt conf n tr),
config_diverges_N_plus conf n tr.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor.
eleft.
eassumption.
eapply COINDHYP.
eassumption.
assumption.
reflexivity.
Qed.
Lemma config_diverges_N_plus_nil_silently_diverges': forall conf n
(Hdiv: config_diverges_N_plus conf n nil),
config_silently_diverges' conf.
Proof.
cofix COINDHYP.
inversion 1; subst.
symmetry in Htr.
destruct (app_eq_nil _ _ Htr).
subst.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
Qed.
Lemma config_diverges_N_plus_diverges: forall tr n conf
(Hdiv: config_diverges_N_plus conf n tr),
config_diverges conf tr.
Proof.
intro tr.
refine ((fun v (Hv: v = length tr) => _) _ (refl_equal _)).
revert tr Hv.
induction v using (well_founded_induction lt_wf).
intros until 1.
induction n using (well_founded_induction lt_wf).
inversion 1; subst.
destruct tr1.
destruct tr2.
simpl in *.
econstructor.
eleft.
eapply config_silently_diverges'_diverges.
eapply config_diverges_N_plus_nil_silently_diverges'.
eassumption.
simpl in *.
assert (e :: tr2 <> nil) by congruence.
generalize (H0 _ (Hn (refl_equal _) H1) _ HCOIND).
inversion 1; subst.
econstructor.
eapply star_trans'.
eright.
eassumption.
eassumption.
reflexivity.
assumption.
simpl in *.
assert (length tr2 < S (length (tr1 ++ tr2))).
rewrite app_length.
omega.
generalize (H _ H1 _ (refl_equal _) _ _ HCOIND).
inversion 1; subst.
econstructor.
eapply star_trans'.
eright.
eassumption.
eassumption.
reflexivity.
assumption.
Qed.
CoInductive config_reacts : CONFIG -> Stream EVENT -> Prop :=
| config_reacts_intro
conf conf' e l
(Hstep: star conf (e :: l) conf')
il
(CIHreacts: config_reacts conf' il)
il'
(Hil': il' = appinf (e :: l) il) :
config_reacts conf il'
.
Lemma config_reacts_bisim: forall conf etr1
(Hreacts: config_reacts conf etr1)
etr2 (Hbisim: bisim etr1 etr2),
config_reacts conf etr2.
Proof.
cofix COINDHYP.
inversion 1; subst.
intros.
edestruct (bisim_appinf_strong).
eassumption.
subst.
generalize (appinf_bisim Hbisim).
intro.
econstructor.
eassumption.
2: reflexivity.
eapply COINDHYP.
eassumption.
assumption.
Qed.
Lemma config_reacts_redecomp: forall conf etr
(Hreacts: config_reacts conf etr)
tr1 etr1 (Hbis1: bisim etr (appinf tr1 etr1)),
exists conf', exists tr2,
star conf (tr1 ++ tr2) conf' /\
exists etr2,
config_reacts conf' etr2 /\
bisim etr1 (appinf tr2 etr2).
Proof.
intros until tr1.
generalize (refl_equal (length tr1)).
generalize (length tr1) at 1.
intro.
revert conf etr Hreacts tr1.
induction n using (well_founded_induction lt_wf).
inversion 1; subst.
intros; subst.
destruct (le_lt_dec (length tr1) (length (e :: l))).
apply bisim_sym in Hbis1.
edestruct appinf_le_bisim.
eassumption.
eassumption.
destruct H0.
rewrite H0 in Hstep.
esplit.
esplit.
split.
eassumption.
eauto.
edestruct appinf_le_bisim.
2 : eassumption.
omega.
destruct H0.
assert (length x < length tr1).
rewrite H0.
rewrite app_length.
simpl.
omega.
edestruct H.
eassumption.
eexact CIHreacts.
reflexivity.
eassumption.
destruct H3.
destruct H3.
destruct H4.
destruct H4.
subst.
esplit.
esplit.
split.
rewrite app_ass.
eapply star_trans.
eassumption.
eassumption.
eauto.
Qed.
CoInductive config_reacts' : CONFIG -> Stream EVENT -> Prop :=
| config_reacts'_intro
conf conf1
(Hstar: star conf nil conf1)
conf2 e l
(Hstep: STEP conf1 (e :: l) conf2)
il
(CIHreacts: config_reacts' conf2 il)
il'
(Hil': il' = appinf (e :: l) il) :
config_reacts' conf il'
.
Lemma config_reacts'_reacts: forall conf st
(Hreacts: config_reacts' conf st),
config_reacts conf st.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor.
2 : eapply COINDHYP; eexact CIHreacts.
2 : reflexivity.
change (e :: l) with (nil ++ (e :: l)).
eapply star_trans.
eassumption.
eright.
eleft.
assumption.
Qed.
Lemma config_reacts_star: forall conf l conf1
(Hstar: star conf l conf1)
il
(Hreacts: config_reacts conf1 il),
config_reacts conf (appinf l il).
Proof.
destruct l.
simpl; inversion 2; subst.
econstructor.
2 : eexact CIHreacts.
2 : reflexivity.
change (e :: l) with (nil ++ (e :: l)).
eapply star_trans.
eassumption.
assumption.
intros; econstructor; eauto.
Qed.
Lemma config_reacts_reacts': forall conf st
(Hreacts: config_reacts conf st),
config_reacts' conf st.
Proof.
cofix COINDHYP.
inversion 1; subst.
destruct (star_not_nil Hstep (refl_equal _)).
destruct H.
destruct H0.
destruct H0.
destruct H0.
destruct H1.
destruct H1.
subst.
econstructor.
eassumption.
eassumption.
eapply COINDHYP.
eapply config_reacts_star.
eassumption.
eassumption.
change (e :: x0 ++ x2) with ((e :: x0) ++ x2).
symmetry.
apply appinf_app.
Qed.
CoInductive config_reacts_N_min : CONFIG -> nat -> Stream EVENT -> Prop :=
| config_reacts_N_min_nil
conf conf'
(Hstep_nil: STEP conf nil conf')
n tr
(CIHreacts: config_reacts_N_min conf' n tr) :
config_reacts_N_min conf (S n) tr
| config_reacts_N_min_cons
l (Hl_not_nil: l <> nil)
conf conf'
(Hstep_cons: STEP conf l conf')
n il
(CIHreacts: config_reacts_N_min conf' n il)
il'
(Hil': il' = appinf l il)
:
config_reacts_N_min conf O il'.
Lemma config_reacts'_config_reacts_N_min_aux: forall n conf conf'
(Hstepn_nil: stepn_left n conf nil conf')
l (Hl_not_nil: l <> nil)
conf''
(Hstep_ev: STEP conf' l conf'')
il
(CIHreacts: config_reacts' conf'' il)
il'
(Hil': il' = appinf l il),
config_reacts_N_min conf n il'.
Proof.
cofix COINDHYP.
inversion 1; subst.
inversion 3; subst.
clear CIHreacts.
destruct (star_stepn_left Hstar); clear Hstar.
intros; subst.
eright.
2: eassumption.
eassumption.
eapply COINDHYP.
eassumption.
2: eassumption.
clear COINDHYP. congruence.
eassumption.
reflexivity.
reflexivity.
symmetry in Htr.
destruct (app_eq_nil _ _ Htr).
subst.
intros.
eleft.
eassumption.
eapply COINDHYP.
eassumption.
eassumption.
eassumption.
eassumption.
assumption.
Qed.
Corollary config_reacts_config_reacts_N_min: forall conf tr
(Hreacts: config_reacts conf tr),
exists n, config_reacts_N_min conf n tr.
Proof.
intros.
generalize (config_reacts_reacts' Hreacts).
inversion 1; subst.
destruct (star_stepn_left Hstar).
esplit.
eapply config_reacts'_config_reacts_N_min_aux.
eassumption.
2: eassumption.
congruence.
eassumption.
reflexivity.
Qed.
CoInductive config_reacts_N : CONFIG -> nat -> Stream EVENT -> Prop :=
| config_reacts_N_nil
conf conf'
(Hstep_nil: STEP conf nil conf')
n tr
(CIHreacts: config_reacts_N conf' n tr) :
config_reacts_N conf (S n) tr
| config_reacts_N_cons
l (Hl_not_nil: l <> nil)
conf conf'
(Hstep_cons: STEP conf l conf')
n il
(CIHreacts: config_reacts_N conf' n il)
il'
(Hil': il' = appinf l il)
n' :
config_reacts_N conf n' il'.
Lemma config_reacts_N_min_config_reacts_N: forall conf n st
(Hreacts: config_reacts_N_min conf n st),
config_reacts_N conf n st.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor.
eassumption.
eapply COINDHYP.
assumption.
econstructor.
eassumption.
eassumption.
eapply COINDHYP.
eassumption.
reflexivity.
Qed.
CoInductive config_reacts_N_star : CONFIG -> nat -> Stream EVENT -> Prop :=
| config_reacts_N_star_nil
conf conf'
(Hstep_nil: star conf nil conf')
n tr
(CIHreacts: config_reacts_N_star conf' n tr) :
config_reacts_N_star conf (S n) tr
| config_reacts_N_star_cons
l (Hl: l <> nil)
conf conf'
(Hstep_cons: star conf l conf')
n tr
(CIHreacts: config_reacts_N_star conf' n tr)
il'
(Hil': il' = appinf l tr)
n' :
config_reacts_N_star conf n' il'.
Corollary config_reacts_N_config_reacts_N_star: forall conf tr n
(Hreacts: config_reacts_N conf n tr),
config_reacts_N_star conf n tr.
Proof.
cofix COINDHYP.
inversion 1; subst.
eleft.
eright.
eleft.
eassumption.
eapply COINDHYP.
assumption.
eright.
eassumption.
eright.
eleft.
eassumption.
eapply COINDHYP.
eassumption.
reflexivity.
Qed.
CoInductive config_reacts_N_star_lt' : CONFIG -> nat -> Stream EVENT -> Prop :=
| config_reacts_N_star_lt'_intro
conf l conf'
(Hstep_nil: star conf l conf')
n tr
(CIHreacts: config_reacts_N_star_lt' conf' n tr)
il'
(Hil': il' = appinf l tr)
n' (Hn: l = nil -> n < n') :
config_reacts_N_star_lt' conf n' il'.
Lemma config_reacts_N_star_config_reacts_N_star_lt': forall conf n st
(Hreacts: config_reacts_N_star conf n st),
config_reacts_N_star_lt' conf n st.
Proof.
cofix COINDHYP.
inversion 1; subst.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
reflexivity.
clear COINDHYP. omega.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
reflexivity.
intro; contradiction.
Qed.
CoInductive config_reacts_N_star_lt : CONFIG -> nat -> Stream EVENT -> Prop :=
| config_reacts_N_star_lt_nil
conf conf'
(Hstep_nil: star conf nil conf')
n n' (Hn: n < n') tr
(CIHreacts: config_reacts_N_star_lt conf' n tr) :
config_reacts_N_star_lt conf n' tr
| config_reacts_N_star_lt_cons
l (Hl: l <> nil)
conf conf'
(Hstep_cons: star conf l conf')
n tr
(CIHreacts: config_reacts_N_star_lt conf' n tr)
il'
(Hil': il' = appinf l tr)
n' :
config_reacts_N_star_lt conf n' il'.
Lemma config_reacts_N_star_lt'_lt: forall conf n st
(Hreacts: config_reacts_N_star_lt' conf n st),
config_reacts_N_star_lt conf n st.
Proof.
cofix COINDHYP.
inversion 1; subst.
destruct l.
eleft.
eassumption.
eapply Hn.
trivial.
eapply COINDHYP.
eassumption.
eright.
2: eassumption.
clear COINDHYP. congruence.
eapply COINDHYP.
eassumption.
reflexivity.
Qed.
Lemma config_reacts_N_star_lt_inv: forall n conf tr
(Hreacts: config_reacts_N_star_lt conf n tr),
exists conf',
star conf nil conf' /\
exists l, l <> nil /\ exists conf'',
STEP conf' l conf'' /\
exists n'', exists tr'',
config_reacts_N_star_lt conf'' n'' tr'' /\
tr = appinf l tr''.
Proof.
induction n using (well_founded_induction lt_wf).
inversion 1; subst.
edestruct H.
2: eexact CIHreacts.
assumption.
destruct H0.
esplit.
split.
eapply star_trans'.
eassumption.
eassumption.
reflexivity.
assumption.
clear n H Hreacts.
revert tr0 Hl n0 CIHreacts.
induction Hstep_cons using star_ind3_left.
destruct 1.
trivial.
destruct tr1.
simpl in *.
intros.
edestruct IHHstep_cons.
assumption.
eassumption.
destruct H.
esplit.
split.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
assumption.
clear IHHstep_cons.
intros.
clear Hl.
destruct tr2.
rewrite <- app_nil_end.
esplit.
split.
eleft.
eexists (_ :: _).
split.
intro.
discriminate.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eapply config_reacts_N_star_lt_nil with (n' := S n0).
eassumption.
2: eassumption.
omega.
reflexivity.
esplit.
split.
eleft.
eexists (_ :: _).
split.
intro.
discriminate.
esplit.
split.
eassumption.
esplit.
esplit.
split.
eapply config_reacts_N_star_lt_cons with (n' := O).
2: eassumption.
congruence.
eassumption.
reflexivity.
rewrite appinf_app.
reflexivity.
Qed.
Corollary config_reacts_N_star_lt_config_reacts: forall n conf tr
(Hreacts: config_reacts_N_star_lt conf n tr),
config_reacts conf tr.
Proof.
intros.
eapply config_reacts'_reacts.
revert n conf tr Hreacts.
cofix COINDHYP.
intros.
edestruct config_reacts_N_star_lt_inv.
eassumption.
destruct H.
destruct H0.
destruct H0.
destruct H1.
destruct H1.
destruct H2.
destruct H2.
destruct H2.
destruct x0.
destruct H0.
trivial.
econstructor.
eassumption.
eassumption.
eapply COINDHYP.
eassumption.
assumption.
Qed.
Lemma config_reacts_config_reacts_N_star: forall conf tr
(Hreacts: config_reacts conf tr)
n,
config_reacts_N_star conf n tr.
Proof.
cofix COINDHYP.
inversion 1; subst.
intros.
eright with (n := n).
2: eassumption.
clear COINDHYP. congruence.
eapply COINDHYP.
eassumption.
reflexivity.
Qed.
Inductive behavior : Type :=
| Terminates (evseq : list EVENT) (hfin: final)
| Stuck (evseq : list EVENT)
| Diverges (evseq : list EVENT)
| Reacts (evinfseq : nat -> EVENT)
.
Function cons_behavior (e: EVENT) (b: behavior) {struct b}: behavior :=
match b with
| Terminates l h => Terminates (e::l) h
| Stuck l => Stuck (e::l)
| Diverges l => Diverges (e::l)
| Reacts f => Reacts (func_of_stream (Cons e (stream_of_func f)))
end.
Lemma cons_behavior_inj: forall e1 e2 b1 b2
(Heq: cons_behavior e1 b1 = cons_behavior e2 b2),
(e1, b1) = (e2, b2).
Proof.
destruct b1; destruct b2; simpl; try congruence.
injection 1; intros; subst.
generalize (func_eq_bisim H).
inversion 1; subst.
f_equal.
f_equal.
rewrite <- (func_of_stream_of_func evinfseq).
rewrite <- (func_of_stream_of_func evinfseq0).
apply bisim_func_eq.
assumption.
Qed.
Function append1_behavior (el: list EVENT) (b: behavior) {struct el}: behavior :=
match el with
| nil => b
| e :: el' => cons_behavior e (append1_behavior el' b)
end.
Lemma append_behavior_app : forall tr1 tr2 b,
append1_behavior (tr1 ++ tr2) b = append1_behavior tr1 (append1_behavior tr2 b).
Proof.
induction tr1; simpl.
reflexivity.
intros.
rewrite IHtr1.
reflexivity.
Qed.
Function append2_behavior (el: list EVENT) (b: behavior) {struct b}: behavior :=
match b with
| Terminates l h => Terminates (el ++ l) h
| Stuck l => Stuck (el ++ l)
| Diverges l => Diverges (el ++ l)
| Reacts f => Reacts (func_of_stream (appinf el (stream_of_func f)))
end.
Lemma append_behavior: append1_behavior = append2_behavior.
Proof.
apply functional_extensionality.
intro l.
apply functional_extensionality.
induction l; intro b; destruct b; simpl; trivial.
rewrite func_of_stream_of_func. trivial.
rewrite IHl. reflexivity.
rewrite IHl. reflexivity.
rewrite IHl. reflexivity.
rewrite IHl. simpl.
f_equal.
apply bisim_func_eq.
constructor.
apply bisim_sym.
apply stream_of_func_of_stream.
Qed.
Lemma append_behavior_right_inj: forall e b1 b2
(Heq: append1_behavior e b1 = append1_behavior e b2),
b1 = b2.
Proof.
induction e; simpl. eauto.
intros.
generalize (cons_behavior_inj Heq).
injection 1; intros; subst.
eauto.
Qed.
Lemma append_behavior_elim_weak: forall a1 a2 b1 b2
(Heq: append1_behavior a1 b1 = append1_behavior a2 b2) ,
(exists a, a1 = a2 ++ a /\ b2 = append1_behavior a b1) \/
(exists a, a2 = a1 ++ a /\ b1 = append1_behavior a b2).
Proof.
cut (
forall a1 a2,
length a1 <= length a2 -> forall b1 b2
(Heq: append1_behavior a1 b1 = append1_behavior a2 b2) ,
(exists a, a1 = a2 ++ a /\ b2 = append1_behavior a b1) \/
(exists a, a2 = a1 ++ a /\ b1 = append1_behavior a b2)
).
intros.
destruct (le_lt_dec (length a1) (length a2)).
eauto.
assert (length a2 <= length a1) by auto with arith.
edestruct H; eauto.
intros.
right.
revert a2 H b1 b2 Heq.
induction a1.
simpl.
intros until 2; subst.
eauto.
destruct a2.
simpl; intros; exfalso; omega.
simpl.
intros.
generalize (cons_behavior_inj Heq).
injection 1; intros; subst.
assert (length a1 <= length a2) by omega.
destruct (IHa1 _ H2 _ _ H1).
destruct H3.
subst.
eauto.
Qed.
Lemma append_behavior_elim_strong: forall a1 a2 b1 b2
(Heq: append1_behavior a1 b1 = append1_behavior a2 b2) ,
(exists a, a1 = a2 ++ a /\ b2 = append1_behavior a b1) \/
(exists e, exists a, a2 = a1 ++ e :: a /\ b1 = append1_behavior (e :: a) b2).
Proof.
intros.
destruct (append_behavior_elim_weak Heq).
eauto.
destruct H.
destruct H.
destruct x.
rewrite <- app_nil_end in H.
simpl in *.
subst.
left.
exists nil.
split.
rewrite <- app_nil_end.
trivial.
reflexivity.
eauto.
Qed.
Lemma append_behavior_terminates: forall l r,
Terminates l r = append1_behavior l (Terminates nil r).
Proof.
rewrite append_behavior.
simpl.
intros.
rewrite <- app_nil_end.
reflexivity.
Qed.
Lemma append_behavior_stuck: forall l,
Stuck l = append1_behavior l (Stuck nil).
Proof.
rewrite append_behavior.
simpl.
intros.
rewrite <- app_nil_end.
reflexivity.
Qed.
Lemma append_behavior_diverges: forall l,
Diverges l = append1_behavior l (Diverges nil).
Proof.
rewrite append_behavior.
simpl.
intros.
rewrite <- app_nil_end.
reflexivity.
Qed.
Inductive behavior_elim (b: behavior): Type :=
| Nilterm r (Hr: b = Terminates nil r)
| Nilstuck (Hstuck: b = Stuck nil)
| Nildiv (Hdiv: b = Diverges nil)
| Consbeh e b' (Hb': b = cons_behavior e b')
.
Lemma behavior_elim_ex: forall b, behavior_elim b.
Proof.
destruct b.
destruct evseq.
eapply Nilterm. reflexivity.
eapply Consbeh with (b' := Terminates _ _). reflexivity.
destruct evseq.
eapply Nilstuck. reflexivity.
eapply Consbeh with (b' := Stuck _). reflexivity.
destruct evseq.
eapply Nildiv. reflexivity.
eapply Consbeh with (b' := Diverges _). reflexivity.
rewrite <- (func_of_stream_of_func evinfseq).
rewrite stream_of_func_unfold'.
eapply Consbeh with (b' := Reacts _). reflexivity.
Qed.
Function config_behaviors (conf: CONFIG) (tr : behavior) {struct tr} : Prop :=
match tr with
| Terminates tr h' => config_terminates conf tr h'
| Stuck tr => config_stuck conf tr
| Diverges tr => config_diverges conf tr
| Reacts trinf => config_reacts conf (stream_of_func trinf)
end.
CoInductive config_behaviors' : CONFIG -> behavior -> Prop :=
| config_behaviors'_terminates c hr (Hconf: config_terminates c nil hr):
config_behaviors' c (Terminates nil hr)
| config_behaviors'_stuck c (Hconf: config_stuck c nil):
config_behaviors' c (Stuck nil)
| config_behaviors'_diverges c (Hconf: config_diverges c nil):
config_behaviors' c (Diverges nil)
| config_behaviors'_cons c c1
(Hnil: star c nil c1)
e l c2 (Hcons: STEP c1 (e :: l) c2)
b
(Hcoind: config_behaviors' c2 b)
b' (Hb': b' = append1_behavior (e :: l) b):
config_behaviors' c b'
.
Lemma config_behaviors_behaviors': forall b c
(Hbeh: config_behaviors c b),
config_behaviors' c b.
Proof.
destruct b.
inversion 1; subst.
clear Hbeh.
destruct (star_stepn_left Hterm); clear Hterm.
revert x c H.
induction 1; simpl in *.
econstructor.
econstructor.
eleft.
assumption.
destruct tr1.
simpl in Htr; subst.
generalize (IHstepn_left Hfin).
inversion 1; subst.
inversion Hconf; subst.
econstructor.
econstructor.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
assumption.
econstructor.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
eassumption.
eassumption.
assumption.
econstructor.
eleft.
eassumption.
eapply IHstepn_left.
assumption.
rewrite append_behavior.
subst.
reflexivity.
inversion 1; subst.
clear Hbeh.
destruct (star_stepn_left Hstar); clear Hstar.
revert x c H.
induction 1; simpl in *.
econstructor.
econstructor.
eleft.
assumption.
assumption.
destruct tr1.
simpl in Htr; subst.
generalize (IHstepn_left Hstuck_not_final Hconf_stuck).
inversion 1; subst.
inversion Hconf; subst.
econstructor.
econstructor.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
assumption.
assumption.
econstructor.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
eassumption.
eassumption.
assumption.
econstructor.
eleft.
eassumption.
eapply IHstepn_left.
assumption.
assumption.
rewrite append_behavior.
subst.
reflexivity.
inversion 1; subst.
clear Hbeh.
destruct (star_stepn_left Hstar); clear Hstar.
revert x c H.
induction 1; simpl in *.
econstructor.
econstructor.
eleft.
assumption.
destruct tr1.
simpl in Htr; subst.
generalize (IHstepn_left Hsilently_diverges).
inversion 1; subst.
generalize (config_diverges_nil_silently_diverges Hconf).
intros.
econstructor.
econstructor.
eright.
eleft.
eassumption.
eassumption.
econstructor.
eapply star_trans'.
eright.
eleft.
eassumption.
eassumption.
reflexivity.
eassumption.
eassumption.
assumption.
econstructor.
eleft.
eassumption.
eapply IHstepn_left.
assumption.
rewrite append_behavior.
subst.
reflexivity.
intros.
generalize (config_reacts_reacts' Hbeh).
clear Hbeh.
rewrite <- (func_of_stream_of_func evinfseq) at 2.
generalize (stream_of_func evinfseq).
clear evinfseq.
revert c.
cofix COINDHYP.
inversion 1; subst.
econstructor.
eassumption.
eassumption.
eapply COINDHYP.
eassumption.
rewrite append_behavior.
simpl.
f_equal.
eapply bisim_func_eq.
constructor.
apply bisim_appinf.
apply stream_of_func_of_stream.
Qed.
CoInductive config_behaviors_star : CONFIG -> behavior -> Prop :=
| config_behaviors_star_terminates c hr (Hconf: config_terminates c nil hr):
config_behaviors_star c (Terminates nil hr)
| config_behaviors_star_stuck c (Hconf: config_stuck c nil):
config_behaviors_star c (Stuck nil)
| config_behaviors_star_diverges c (Hconf: config_diverges c nil):
config_behaviors_star c (Diverges nil)
| config_behaviors_star_cons c
e l c2 (Hcons: star c (e :: l) c2)
b
(Hcoind: config_behaviors_star c2 b)
b' (Hb': b' = append1_behavior (e :: l) b):
config_behaviors_star c b'
.
Lemma config_behaviors'_config_behaviors_star: forall c b
(Hbeh: config_behaviors' c b),
config_behaviors_star c b.
Proof.
cofix COINDHYP.
inversion 1; subst; try (clear COINDHYP; constructor; eauto; fail).
econstructor. eapply star_trans'. eassumption. eright. eleft. eassumption. simpl; reflexivity.
eapply COINDHYP. eassumption. reflexivity.
Qed.
Lemma config_behaviors_star_config_behaviors: forall c b
(Hbeh: config_behaviors_star c b),
config_behaviors c b.
Proof.
destruct b.
refine ((fun n (Hn: n = length evseq) => _) _ (refl_equal _)).
revert evseq Hn c.
induction n using (well_founded_induction lt_wf).
inversion 2; subst.
assumption.
rewrite append_behavior in Hb'.
symmetry in Hb'.
functional inversion Hb'; subst.
refine (_ (H _ _ _ _ _ Hcoind)).
3: reflexivity.
inversion 1; subst.
econstructor.
eapply star_trans'.
eassumption.
eassumption.
reflexivity.
assumption.
simpl.
rewrite app_length.
omega.
refine ((fun n (Hn: n = length evseq) => _) _ (refl_equal _)).
revert evseq Hn c.
induction n using (well_founded_induction lt_wf).
inversion 2; subst.
assumption.
rewrite append_behavior in Hb'.
symmetry in Hb'.
functional inversion Hb'; subst.
refine (_ (H _ _ _ _ _ Hcoind)).
3: reflexivity.
inversion 1; subst.
econstructor.
eapply star_trans'.
eassumption.
eassumption.
reflexivity.
assumption.
assumption.
simpl.
rewrite app_length.
omega.
refine ((fun n (Hn: n = length evseq) => _) _ (refl_equal _)).
revert evseq Hn c.
induction n using (well_founded_induction lt_wf).
inversion 2; subst.
assumption.
rewrite append_behavior in Hb'.
symmetry in Hb'.
functional inversion Hb'; subst.
refine (_ (H _ _ _ _ _ Hcoind)).
3: reflexivity.
inversion 1; subst.
econstructor.
eapply star_trans'.
eassumption.
eassumption.
reflexivity.
assumption.
simpl.
rewrite app_length.
omega.
simpl.
rewrite <- (func_of_stream_of_func evinfseq) at 1.
generalize (stream_of_func evinfseq).
intro.
generalize (bisim_refl s).
generalize s at 2 4.
revert s.
clear evinfseq.
revert c.
cofix COINDHYP.
inversion 2; subst.
symmetry in Hb'.
rewrite append_behavior in Hb'.
functional inversion Hb'; subst.
rewrite <- (func_of_stream_of_func f) in Hcoind.
generalize (func_eq_bisim H1).
inversion 1; subst.
destruct (bisim_appinf_strong CIHbisim).
subst.
generalize (appinf_bisim CIHbisim).
intros.
inversion H; subst.
destruct (bisim_appinf_strong CIHbisim0).
subst.
generalize (appinf_bisim CIHbisim0).
intros.
econstructor.
eassumption.
eapply COINDHYP.
2: eassumption.
2: reflexivity.
eapply bisim_trans.
eassumption.
assumption.
Qed.
Lemma star_config_behaviors: forall c1 tr c2,
star c1 tr c2 ->
forall b, config_behaviors c2 b ->
forall b', b' = append1_behavior tr b ->
config_behaviors c1 b'.
Proof.
intros.
subst.
generalize (config_behaviors_behaviors' H0).
intro.
generalize (config_behaviors'_config_behaviors_star H1).
intro.
eapply config_behaviors_star_config_behaviors.
destruct tr.
simpl.
inversion H2; subst.
inversion Hconf; subst. econstructor. econstructor. eapply star_trans'. eassumption.
eassumption.
reflexivity.
assumption.
inversion Hconf; subst. econstructor. econstructor. eapply star_trans'. eassumption. eassumption. reflexivity. assumption. assumption.
inversion Hconf; subst. econstructor. econstructor. eapply star_trans'. eassumption. eassumption. reflexivity. assumption.
econstructor. eapply star_trans'. eassumption. eassumption. simpl; reflexivity. eassumption. reflexivity.
econstructor. eassumption. eassumption. reflexivity.
Qed.
Lemma config_behavior_star_recip: forall a b s,
config_behaviors s (append1_behavior a b) ->
exists a', exists s', star s (a ++ a') s' /\
exists b', config_behaviors s' b' /\
b = append1_behavior a' b'.
Proof.
intro a.
generalize (refl_equal (length a)).
generalize (length a) at 1.
intro.
revert a.
induction n using (well_founded_induction lt_wf).
destruct a.
simpl.
intros.
esplit.
esplit.
split.
eleft.
simpl; eauto.
simpl.
intros; subst.
generalize (config_behaviors_behaviors' H1).
inversion 1; subst.
rewrite append_behavior in H4.
destruct b; discriminate.
rewrite append_behavior in H4.
destruct b; discriminate.
rewrite append_behavior in H4.
destruct b; discriminate.
simpl in Hb'.
generalize (cons_behavior_inj Hb').
injection 1; intros; subst.
symmetry in H3.
generalize (config_behaviors_star_config_behaviors (config_behaviors'_config_behaviors_star Hcoind)); clear Hcoind.
intro Hcoind.
destruct (append_behavior_elim_strong H3).
clear H.
destruct H4.
destruct H.
subst.
esplit.
esplit.
split.
eapply star_trans'.
eassumption.
eright.
eleft.
eassumption.
reflexivity.
eauto.
destruct H4.
destruct H4.
destruct H4.
subst.
assert (length (x :: x0) < S (length (l ++ x :: x0))).
rewrite app_length.
simpl.
omega.
destruct (H _ H4 _ (refl_equal _) _ _ Hcoind); clear H H4 Hcoind.
destruct H5.
destruct H.
destruct H4.
destruct H4.
subst.
esplit.
esplit.
split.
eapply star_trans'.
eassumption.
eapply star_trans.
eright.
eleft.
eassumption.
eassumption.
rewrite app_ass.
reflexivity.
eauto.
Qed.
Section Deterministic.
Hypothesis determ: forall conf tr1 conf1 (Hconf1: STEP conf tr1 conf1) tr2 conf2 (Hconf2: STEP conf tr2 conf2), (tr1,conf1) = (tr2,conf2).
Lemma determ_n: forall n conf tr1 conf1 (Hstepn: stepn_left n conf tr1 conf1) tr2 conf2 (Hstepn': stepn_left n conf tr2 conf2), (tr1,conf1) = (tr2,conf2).
Proof.
induction 1; inversion 1; subst.
auto.
generalize (determ Hstep Hstep0).
injection 1; intros; subst.
generalize (IHHstepn _ _ Hstepn0).
injection 1; intros; subst; eauto.
Qed.
Lemma triangle: forall n conf tr conf' (Hstepn: stepn_left n conf tr conf') n1 tr1 conf1 (Hstepn': stepn_left n1 conf tr1 conf1) n2, n = n1 + n2 -> exists tr2, stepn_left n2 conf1 tr2 conf' /\ tr = tr1 ++ tr2.
Proof.
intros; subst.
revert conf tr conf' n1 Hstepn tr1 conf1 Hstepn'.
induction n2.
intros.
replace (n1 + 0) with n1 in * by omega.
generalize (determ_n Hstepn Hstepn').
injection 1; intros; subst.
esplit.
split.
eleft.
rewrite <- app_nil_end.
trivial.
intros.
replace (n1 + S n2) with (S (n1 + n2)) in * by omega.
generalize (stepn_left_right Hstepn); clear Hstepn.
inversion 1; subst.
clear H.
destruct (IHn2 _ _ _ _ (stepn_right_left Hstepn) _ _ Hstepn').
destruct H.
esplit.
split.
eapply stepn_right_left.
econstructor.
eapply stepn_left_right.
eassumption.
eassumption.
reflexivity.
subst.
rewrite app_ass.
trivial.
Qed.
Unicity of diverging behaviors
Lemma reacts_min_stepn: forall n conf tr conf' (Hconf': stepn_left n conf tr conf') f n1 (Hreacts: config_reacts_N_min conf n1 f), exists f', exists n2, config_reacts_N_min conf' n2 f' /\ f = appinf tr f' /\ (tr = nil -> (n <= n1 /\ n2 = n1 - n)).
Proof.
induction n; inversion 1; subst; simpl.
intros.
replace (n1 - 0) with n1 by omega.
eauto 6 with arith.
inversion 1; subst.
generalize (determ Hstep Hstep_nil).
injection 1; intros; subst.
destruct (IHn _ _ _ Hstepn _ _ CIHreacts).
destruct H0.
destruct H0.
destruct H1.
subst.
simpl in *.
esplit.
esplit.
split.
eassumption.
split.
reflexivity.
intro.
destruct (H2 H1).
split; omega.
generalize (determ Hstep Hstep_cons).
injection 1; intros; subst.
destruct (IHn _ _ _ Hstepn _ _ CIHreacts).
destruct H0.
destruct H0.
destruct H1.
esplit.
esplit.
split.
eassumption.
split.
rewrite <- appinf_app.
congruence.
intros.
destruct (app_eq_nil _ _ H3).
contradiction.
Qed.
Lemma reacts_stepn: forall n conf tr conf' (Hconf': stepn_left n conf tr conf') f (Hreacts: config_reacts' conf f), exists f', config_reacts' conf' f' /\ f = (appinf tr f').
Proof.
induction n using (well_founded_induction lt_wf).
inversion 2; subst.
destruct (star_stepn_left Hstar); clear Hstar.
destruct (le_lt_dec n x).
assert (x = n + (x - n)) by omega.
destruct (triangle H0 Hconf' H1); clear H1.
destruct H2.
symmetry in H2.
destruct (app_eq_nil _ _ H2).
subst tr x0.
esplit.
split.
econstructor.
eapply stepn_left_star.
eassumption.
eassumption.
eassumption.
reflexivity.
reflexivity.
generalize (stepn_right_left (stepn_right_S (stepn_left_right H0) Hstep (refl_equal _))).
intros.
simpl in H1.
assert (n = (S x) + (n - S x)) by omega.
destruct (triangle Hconf' H1 H2); clear H1.
destruct H3.
assert (n - S x < n) by omega.
destruct (H _ H4 _ _ _ H1 _ CIHreacts).
destruct H5.
esplit.
split.
eassumption.
clear H2.
subst.
simpl.
rewrite appinf_app.
reflexivity.
Qed.
Lemma silently_diverges_step: forall conf (Hdiv: config_silently_diverges conf) tr conf1 (Hstep: STEP conf tr conf1), tr = nil /\ config_silently_diverges conf1.
Proof.
inversion 1; subst.
intros.
generalize (determ Hstep Hstep0).
injection 1; intros; subst; eauto.
Qed.
Corollary silently_diverges_stepn: forall n tr conf conf1 (Hstepn: stepn_left n conf tr conf1) (Hdiv: config_silently_diverges conf), tr = nil /\ config_silently_diverges conf1.
Proof.
induction 1.
eauto.
intros.
destruct (silently_diverges_step Hdiv Hstep).
subst.
simpl.
eauto.
Qed.
Lemma silently_diverges_not_reacts: forall conf (Hdiv: config_silently_diverges conf) tr' (Hreacts: config_reacts conf tr'), False.
Proof.
inversion 2; subst.
destruct (star_stepn_left Hstep).
destruct(silently_diverges_stepn H Hdiv).
discriminate.
Qed.
Lemma diverges_not_reacts: forall conf tr (Hdiv: config_diverges conf tr) tr' (Hreacts: config_reacts conf tr'), False.
Proof.
inversion 1; intros; subst.
generalize (config_reacts_reacts' Hreacts).
intros.
destruct (star_stepn_left Hstar).
destruct (reacts_stepn H0 H).
destruct H1.
eapply silently_diverges_not_reacts.
eassumption.
eapply config_reacts'_reacts.
eassumption.
Qed.
Lemma config_diverges_determ: forall conf tr1 (Hdiv1: config_diverges conf tr1) tr2 (Hdiv2: config_diverges conf tr2), tr1 = tr2.
Proof.
inversion 1; subst; clear Hdiv1.
destruct (star_stepn_left Hstar); clear Hstar.
inversion 1; subst; clear Hdiv2.
destruct (star_stepn_left Hstar); clear Hstar.
assert (x <= max x x0).
eapply le_max_l.
destruct (config_silently_diverges_inv Hsilently_diverges (max x x0 - x)).
destruct H2.
assert (x0 <= max x x0).
eapply le_max_r.
destruct (config_silently_diverges_inv Hsilently_diverges0 (max x x0 - x0)).
destruct H5.
generalize (stepn_left_app H0 H5).
replace (x0 + (max x x0 - x0)) with (max x x0) by omega.
rewrite <- app_nil_end.
intros.
generalize (stepn_left_app H H2).
replace (x + (max x x0 - x)) with (max x x0) by omega.
rewrite <- app_nil_end.
intros.
generalize (determ_n H8 H7).
injection 1; intros; subst; eauto.
Qed.
Lemma config_reacts_determ: forall conf tr1 (Hreacts: config_reacts conf tr1) tr2 (Hreacts': config_reacts conf tr2), bisim tr1 tr2.
Proof.
intros.
generalize (config_reacts_reacts' Hreacts); clear Hreacts; intro Hreacts.
generalize (config_reacts_reacts' Hreacts'); clear Hreacts'; intro Hreacts'.
eapply bisim'_bisim.
revert conf tr1 Hreacts tr2 Hreacts'.
cofix COINDHYP.
inversion 1; subst.
clear Hreacts.
destruct (star_stepn_left Hstar); clear Hstar.
inversion 1; subst.
clear Hreacts'.
destruct (star_stepn_left Hstar); clear Hstar.
destruct (le_lt_dec x x0).
destruct( le_lt_eq_dec _ _ l1).
exfalso.
clear COINDHYP.
assert (x0 = S x + (x0 - S x)) by omega.
destruct (triangle H0 (stepn_right_left (stepn_right_S (stepn_left_right H) Hstep (refl_equal _))) H1).
destruct H2.
discriminate.
subst.
generalize (determ_n H0 H).
injection 1; intros; subst.
generalize (determ Hstep Hstep0).
injection 1; intros; subst.
econstructor.
eapply COINDHYP.
eapply CIHreacts.
eapply CIHreacts0.
2: reflexivity.
clear COINDHYP; congruence.
reflexivity.
exfalso.
clear COINDHYP.
assert (x = S x0 + (x - S x0)) by omega.
destruct (triangle H (stepn_right_left (stepn_right_S (stepn_left_right H0) Hstep0 (refl_equal _))) H1).
destruct H2.
discriminate.
Qed.
Lemma diverges_min_stepn: forall n conf tr conf' (Hconf': stepn_left n conf tr conf') f n1 (Hdiv: config_diverges_N_min conf n1 f), exists f', exists n2, config_diverges_N_min conf' n2 f' /\ f = tr ++ f' /\ (tr = nil -> f <> nil -> (n <= n1 /\ n2 = n1 - n)).
Proof.
induction n; inversion 1; subst; simpl.
intros.
replace (n1 - 0) with n1 by omega.
eauto 6 with arith.
inversion 1; subst.
generalize (determ Hstep Hstep0).
injection 1; intros; subst.
simpl.
destruct (IHn _ _ _ Hstepn _ _ HCOIND).
destruct H0.
destruct H0.
destruct H1.
subst.
simpl in *.
esplit.
esplit.
split.
eassumption.
split.
reflexivity.
intros.
destruct (H2 H1 H3).
split; omega.
generalize (determ Hstep Hstep0).
injection 1; intros; subst.
simpl in *.
destruct (IHn _ _ _ Hstepn _ _ HCOIND).
destruct H0.
destruct H0.
destruct H1.
symmetry in H1.
destruct (app_eq_nil _ _ H1).
subst.
simpl in *.
esplit.
esplit.
split.
eassumption.
split.
reflexivity.
destruct 2.
reflexivity.
generalize (determ Hstep Hstep0).
injection 1; intros; subst.
destruct (IHn _ _ _ Hstepn _ _ HCOIND).
destruct H0.
destruct H0.
destruct H1.
subst.
esplit.
esplit.
split.
eassumption.
split.
rewrite app_ass.
reflexivity.
intro.
destruct (app_eq_nil _ _ H1).
contradiction.
Qed.
End Deterministic.
End Smallstep.
Section Simult.
Variable (CONFIG: Type) (STEP: CONFIG -> list EVENT -> CONFIG -> Prop).
Hypothesis determ: forall conf tr1 conf1 (Hconf1: STEP conf tr1 conf1) tr2 conf2 (Hconf2: STEP conf tr2 conf2), (tr1,conf1) = (tr2,conf2).
Variable (CONFIG': Type) (STEP': CONFIG' -> list EVENT -> CONFIG' -> Prop).
Definition simult (I: CONFIG -> CONFIG' -> Prop) : Prop :=
forall c1 c'1, I c1 c'1 -> exists tr, exists c2, plus STEP c1 tr c2 /\ exists c'2, plus STEP' c'1 tr c'2 /\ I c2 c'2.
Variable I: CONFIG -> CONFIG' -> Prop.
Hypothesis Hsimult: simult I.
Lemma config_diverges_simult: forall tr conf, config_diverges STEP conf tr -> forall conf', I conf conf' -> config_diverges STEP' conf' tr.
Proof.
intros.
destruct (config_diverges_config_diverges_N_min H); clear H.
eapply config_diverges_N_plus_diverges with (n := x).
revert tr conf conf' H0 x H1.
cofix COINDHYP.
intros.
destruct (Hsimult H0).
destruct H.
destruct H.
destruct H2.
destruct H2.
destruct (plus_stepn_left H).
destruct (diverges_min_stepn determ H4 H1).
destruct H5.
destruct H5.
destruct H6.
subst.
econstructor.
eassumption.
eapply COINDHYP.
eassumption.
eassumption.
intros; subst; simpl in *.
destruct (H7 (refl_equal _) H8).
omega.
reflexivity.
Qed.
Lemma config_reacts_simult: forall tr conf, config_reacts STEP conf tr -> forall conf', I conf conf' -> config_reacts STEP' conf' tr.
Proof.
intros.
destruct (config_reacts_config_reacts_N_min H); clear H.
eapply config_reacts_N_star_lt_config_reacts with (n := x).
eapply config_reacts_N_star_lt'_lt.
revert tr conf conf' H0 x H1.
cofix COINDHYP.
intros.
destruct (Hsimult H0).
destruct H.
destruct H.
destruct H2.
destruct H2.
destruct (plus_stepn_left H).
destruct (reacts_min_stepn determ H4 H1).
destruct H5.
destruct H5.
destruct H6.
subst.
econstructor.
eright.
eassumption.
eapply COINDHYP.
eassumption.
eassumption.
reflexivity.
intros; subst; simpl in *.
destruct (H7 (refl_equal _)).
omega.
Qed.
End Simult.
Section Simult_sym.
Variables (CONFIG: Type) (STEP: CONFIG -> list EVENT -> CONFIG -> Prop).
Variables (CONFIG': Type) (STEP': CONFIG' -> list EVENT -> CONFIG' -> Prop).
Variables (I: CONFIG -> CONFIG' -> Prop) (Hsimult: simult STEP STEP' I).
Let J c' c := I c c'.
Lemma simult_sym: simult STEP' STEP J.
Proof.
unfold simult.
unfold J.
intros.
destruct (Hsimult H).
destruct H0.
destruct H0.
destruct H1.
destruct H1.
eauto 6.
Qed.
End Simult_sym.
End Event.