Library proc_comp_stuck

Relationship between procedural semantics and compositional semantics

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

Section Semworld.

Variable sw: semworld.

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

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

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

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

Definition be_stuck (_: args) (_: heap) : cbehavior -> Prop := eq (Stuck _ nil).

Definition psistuck (d: funcname -> bool) (f: funcname) : option (args -> heap -> cbehavior -> Prop) :=
  if d f then None else Some (be_stuck).

Section Language.

Variable lang: language sw.

Notation local := (local lang) (only parsing).
Notation frame := (frame lang) (only parsing).

Variable m: module lang.

Definition stuck_closure (f: funcname) : option (args -> heap -> cbehavior -> Prop) :=
  if in_dom_dec m f then None else Some be_stuck.

Lemma psistuck_stuck_closure: psistuck (dom m) = stuck_closure.
Proof.
  apply functional_extensionality.
  unfold psistuck, stuck_closure.
  intro.
  destruct (in_dom_dec m x).
   rewrite (in_dom_dom i).
   trivial.
  case_eq (dom m x).
   intro.
   destruct n.
   apply dom_in_dom.
   assumption.
  trivial.
Qed.

Lemma not_in_dom_stuck_closure: forall f,
  (~ in_dom m f) -> in_dom stuck_closure f.
Proof.
  intro.
  case_eq (stuck_closure f).
   intros until 1.
   case_eq (stuck_closure f); intros; rewrite H in *; try discriminate.
    esplit. eassumption.
   unfold stuck_closure. destruct (in_dom_dec m f); try discriminate.
   intros; contradiction.
Qed.

Lemma not_in_stuck_closure_dom: forall f,
  (~ in_dom stuck_closure f) -> in_dom m f.
Proof.
  intros.
  destruct (in_dom_dec m f).
   assumption.
  destruct H.
  eauto using not_in_dom_stuck_closure.
Qed.

Lemma stuck_closure_disjoint: disjoint (create_map (dom m) (csemantics m)) stuck_closure.
Proof.
  unfold disjoint.
  intros until 1.
  destruct 1.
  unfold lookup, lookup_strong in H0.
  unfold stuck_closure in H0.
  destruct (in_dom_dec m a).
   discriminate.
  destruct n.
  eapply dom_in_dom.
  rewrite <- (create_map_dom (dom m) (csemantics m)).
  apply in_dom_dom.
  assumption.
Qed.

First part of the proof: procedural semantics included in bsemantics (link (csemantics m) stuck_closure)


Lemma stepn_cstepn: forall n c tr c'
  (Hstepn: stepn_left (step m) n c tr c'),
  stepn_left (cstep m) n c (map (@Event _) tr) c'.
Proof.
  induction 1; subst; simpl.
   left.
  rewrite map_app.
  eright.
  eapply cstep_step.
  eassumption.
  reflexivity.
  eassumption.
  reflexivity.
Qed.

Lemma step_cstep_terminates: forall s tr hr,
  config_terminates (step m) (@config_final _ _) s tr hr ->
  config_terminates (cstep m) (@config_final _ _) s (map (@Event _) tr) hr.
Proof.
  inversion 1; subst.
  destruct (star_stepn_left Hterm).
  econstructor.
  eapply stepn_left_star.
  eapply stepn_cstepn.
  eassumption.
  assumption.
Qed.

Lemma step_cstep_stuck': forall s tr,
  config_stuck (step m) (@config_final _ _) s tr ->
   config_stuck (cstep m) (@config_final _ _) s (map (@Event _) tr) \/
   exists h', exists l', exists k',
     star (cstep m) s (map (@Event _) tr) (Config h' l' k') /\
     exists f, (~ in_dom m f) /\
       exists a,
         local_kind l' = Some (Call f a).
Proof.
  inversion 1; subst.
  destruct (star_stepn_left Hstar).
  generalize (stepn_cstepn H0).
  intros.
  destruct cstuck.
  case_eq (local_kind c_local).
   intros.
   destruct s0.
    case_eq (caller_call a c_heap c_local).
    intros.
    destruct (in_dom_dec m func).
     exfalso.
     destruct i.
     case_eq (callee_init a h x0).
     intros.
     eapply Hconf_stuck.
     eapply step_call.
     eassumption.
     eassumption.
     eassumption.
     eassumption.
    right.
    eauto 9 using stepn_left_star.
   exfalso.
   destruct c_stack.
    eapply Hstuck_not_final.
    econstructor.
    eassumption.
   case_eq (caller_return r c_heap f).
   intros.
   eapply Hconf_stuck.
   eapply step_return.
   eassumption.
   eassumption.
  intros.
  left.
  econstructor.
  eapply stepn_left_star.
  eassumption.
  intros.
  inversion Habs; subst.
  rewrite Hcmd in H2.
  discriminate.
 inversion 1; subst.
  eauto.
 rewrite Hcmd in H2.
 discriminate.
Qed.

Lemma step_cstep_stuck: forall s tr,
  config_stuck (step m) (@config_final _ _) s tr ->
   config_stuck (cstep m) (@config_final _ _) s (map (@Event _) tr) \/
   exists f, (~ in_dom m f) /\
     exists a, exists h1, exists r, exists h2, exists beh,
       config_behaviors (cstep m) (@config_final _ _) s (append1_behavior (map (@Event _) tr) (cons_behavior (External f a h1 r h2) beh)).
Proof.
  intros.
  destruct (step_cstep_stuck' H).
   eauto.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H1.
  destruct H1.
  destruct H2.
  destruct (ret_inh sw).
  case_eq (caller_call x3 x x0).
  intros.
  case_eq (caller_return X h f).
  intros.
  generalize (cstep_call_external H2 H1 H3 (heap_le_refl _) H4 x1).
  intro.
  destruct (exists_behavior (@config_final _ _) (cstep m) (Config h0 l x1)).
  generalize (star_trans H0 (star_plus (plus_step H5))).
  intro.
  generalize (star_config_behaviors H7 H6 (refl_equal _)).
  intro.
  rewrite append_behavior_app in H8.
  simpl in H8.
  eauto 9.
Qed.

Lemma step_cstep_silently_diverges: forall s,
  config_silently_diverges (step m) s ->
  config_silently_diverges (cstep m) s.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  econstructor.
  eapply cstep_step.
  eassumption.
  reflexivity.
  eapply COINDHYP.
  assumption.
Qed.

Lemma step_cstep_diverges: forall s tr,
  config_diverges (step m) s tr ->
  config_diverges (cstep m) s (map (@Event _) tr).
Proof.
  inversion 1; subst.
  destruct (star_stepn_left Hstar).
  econstructor.
  eapply stepn_left_star.
  eapply stepn_cstepn.
  eassumption.
  eapply step_cstep_silently_diverges.
  assumption.
Qed.

Lemma step_cstep_reacts: forall s tr,
  config_reacts (step m) s tr ->
  config_reacts (cstep m) s (mapinf (@Event _) tr).
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  destruct (star_stepn_left Hstep).
  generalize (stepn_cstepn H0).
  intro.
  simpl in H1.
  econstructor.
  eapply stepn_left_star.
  eassumption.
  eapply COINDHYP.
  eassumption.
  rewrite mapinf_appinf.
  reflexivity.
Qed.

Lemma semantics_in_bsemantics_stuck_closure: forall f a h b,
    semantics m f a h b ->
    bsemantics (merge (create_map (dom m) (csemantics m)) (stuck_closure )) f a h (cbehavior_of_behavior b).
Proof.
  inversion 1; subst.
  generalize (in_dom_dom (lookup_in_dom Hcmd)).
  generalize (create_map_dom (dom m) (csemantics m)).
  intros.
  rewrite <- H0 in H1.
  generalize (dom_in_dom H1).
  destruct 1.
  generalize (lookup_merge_lookup_1 (stuck_closure ) H2).
  intro.
  generalize (create_map_lookup_elim H2).
  intro; subst.
  destruct b; simpl in *.
   generalize (step_cstep_terminates Hbeh).
   intro.
   eapply bsemantics_intro with (cb := Terminates _ _).
   eassumption.
   econstructor.
   eassumption.
   eassumption.
   eassumption.
   reflexivity.
   simpl.
   destruct hfin.
   rewrite (append_behavior_terminates) at 1.
   econstructor.
   eapply bstar_append.
   intros ? _.
   clear.
   induction evseq; simpl.
    tauto.
   destruct 1.
    discriminate.
   eauto.
   constructor.
   destruct (step_cstep_stuck Hbeh).
   eapply bsemantics_intro with (cb := Stuck _ _).
   eassumption.
   econstructor.
   eassumption.
   eassumption.
   eassumption.
   reflexivity.
   simpl.
   rewrite (append_behavior_stuck) at 1.
   econstructor.
   eapply bstar_append.
   intros ? _.
   clear.
   induction evseq; simpl.
    tauto.
   destruct 1.
    discriminate.
   eauto.
   inversion 1.
   inversion 1; subst.
    destruct b; discriminate.
    destruct b; discriminate.
    destruct b; discriminate.
  destruct H4.
  destruct H4.
  destruct H5.
  destruct H5.
  destruct H5.
  destruct H5.
  destruct H5.
  assert (lookup (merge (B:=args -> heap -> cbehavior -> Prop)
            (create_map (dom (B:=code lang) m) (csemantics (lang:=lang) m))
            (stuck_closure )) x (be_stuck )).
   eapply lookup_disj_merge_lookup_2.
   unfold lookup, lookup_strong, stuck_closure.
   destruct (in_dom_dec m x).
    contradiction.
   trivial.
   apply stuck_closure_disjoint.
  econstructor.
  eassumption.
  econstructor.
  eassumption.
  eassumption.
  eassumption.
  reflexivity.
  econstructor.
  eapply star_trans'.
   eapply bstar_append.
   intros ? _.
   clear.
   induction evseq; simpl.
    tauto.
   destruct 1.
    discriminate.
   eauto.
  eright.
  eleft.
  eapply bstep_internal.
  eassumption.
  reflexivity.
  reflexivity.
  rewrite <- app_nil_end.
  reflexivity.
  inversion 1.
  inversion 1; subst; destruct b; discriminate.
   generalize (step_cstep_diverges Hbeh).
   intro.
   eapply bsemantics_intro with (cb := Diverges _ _).
   eassumption.
   econstructor.
   eassumption.
   eassumption.
   eassumption.
   reflexivity.
   simpl.
   rewrite (append_behavior_diverges) at 1.
   econstructor.
   eapply bstar_append.
   intros ? _.
   clear.
   induction evseq; simpl.
    tauto.
   destruct 1.
    discriminate.
   eauto.
   cofix COINDHYP.
   econstructor.
   eapply bstep_diverges.
   assumption.
   generalize (step_cstep_reacts Hbeh).
   intro.
   generalize (config_reacts_bisim H4 (mapinf_stream_of_func _ _)).
   intro.
   eapply bsemantics_intro with (cb := Reacts _ _).
   eassumption.
   econstructor.
   eassumption.
   eassumption.
   simpl.
   eassumption.
   reflexivity.
   simpl.
   clear.
   generalize (bisim_refl (stream_of_func (fmapinf (@Event _) evinfseq))).
   generalize ((stream_of_func (fmapinf (@Event _) evinfseq))) at 1 3.
   generalize (refl_equal (fmapinf (@Event _) evinfseq)).
   generalize (fmapinf (@Event _) evinfseq) at 1 4.
   revert evinfseq.
   cofix COINDHYP.
   intros.
   subst.
   rewrite <- (func_of_stream_of_func evinfseq).
   rewrite stream_of_func_unfold'.
   rewrite stream_of_func_unfold' in H0.
   inversion H0; subst.
   rewrite fmapinf_func_of_stream.
   rewrite mapinf_cons.
   change (
     fmapinf (@Event _) evinfseq 0
   ) with (Event (evinfseq 0)).
   replace (Reacts (heap * ret) (func_of_stream (Cons (Event (evinfseq 0)) (mapinf (@Event _) (stream_of_func (fun x => evinfseq (S x))))))) with (cons_behavior (Event (evinfseq 0)) (Reacts (heap * ret) (func_of_stream (mapinf (@Event _) (stream_of_func (fun x => evinfseq (S x))))))).
   econstructor.
   eright.
   eleft.
   eapply bstep_cons.
   reflexivity.
   eapply COINDHYP.
   rewrite <- func_of_stream_mapinf_of_func.
   reflexivity.
   2: simpl; reflexivity.
   assumption.
   simpl.
   f_equal.
   apply bisim_func_eq.
   constructor.
   apply bisim_sym.
   apply stream_of_func_of_stream.
Qed.

Second part of the proof: bsemantics (link (csemantics m) stuck_closure) included in procedural semantics


Lemma cstep_beh_trace_inv: forall s beh,
  config_behaviors (cstep m) (@config_final _ _) s beh ->
  forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False.
Proof.
  intros.
  subst.
  destruct (config_behavior_star_recip H).
  destruct H0.
  destruct H0.
  eapply cstarn_trace_inv.
  eassumption.
  eassumption.
  eapply in_or_app.
  left.
  eassumption.
Qed.

Lemma cstep_beh_trace_inv': forall s beh,
  config_behaviors (cstep m) (@config_final _ _) s beh ->
  forall tr f a h1 r h2 b, beh = append1_behavior tr (cons_behavior (External f a h1 r h2) b) ->
    in_dom m f -> False.
Proof.
  intros; subst.
  change (cons_behavior (External f a h1 r h2) b) with (append1_behavior (External f a h1 r h2 :: nil) b) in H.
  rewrite <- append_behavior_app in H.
  eapply cstep_beh_trace_inv.
  eassumption.
  reflexivity.
  eassumption.
  eapply in_or_app.
  right.
  left.
  reflexivity.
Qed.

Lemma cstepn_stepn: forall n c tr c',
  stepn_left (cstep m) n c tr c' ->
  exists n', exists etr, exists h, exists l, exists k,
    stepn_left (step m) n' c etr (Config h l k) /\
    ((tr = map (@Event _) etr /\ c' = Config h l k /\ n' = n) \/
      n' <= n /\
      exists f,
        ~ in_dom m f /\
        exists a,
          local_kind l = Some (Call f a) /\
          exists h1, exists fr1,
            caller_call a h l = (h1, fr1) /\
          exists r, exists h', exists tr',
            tr = map (@Event _) etr ++ External f a h1 r h' :: tr').
Proof.
  induction 1; subst.
  destruct conf.
   esplit.
   exists nil.
   esplit.
   esplit.
   esplit.
   split.
   left.
   left.
   auto.
  destruct IHstepn_left.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H0.
  inversion Hstep; subst.
  esplit.
  esplit.
  esplit.
  esplit.
  esplit.
  split.
  eright.
  eassumption.
  eassumption.
  reflexivity.
  destruct H1.
   left.
   destruct H1.
   destruct H2.
   subst.
   rewrite map_app.
   auto.
  right.
  rewrite map_app.
  destruct H1.
  destruct H2.
  destruct H2.
  destruct H3.
  destruct H3.
  destruct H4.
  destruct H4.
  destruct H4.
  destruct H5.
  destruct H5.
  destruct H5.
  subst.
  split.
  omega.
  esplit.
  esplit.
  eassumption.
  esplit.
  split.
  eassumption.
  esplit.
  esplit.
  split.
  eassumption.
  esplit.
  esplit.
  esplit.
  rewrite app_ass.
  reflexivity.
 esplit.
 esplit.
 esplit.
 esplit.
 esplit.
 split.
 left.
 right.
 split.
 omega.
 esplit.
 split.
 eassumption.
 esplit.
 split.
 eassumption.
 esplit.
 esplit.
 esplit.
 eassumption.
 esplit.
 esplit.
 esplit.
 simpl.
 reflexivity.
Qed.

Lemma cstepn_stepn_no_external: forall n c tr c',
  stepn_left (cstep m) n c (map (@Event _) tr) c' ->
  stepn_left (step m) n c tr c'.
Proof.
  intros.
  destruct (cstepn_stepn H).
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H1.
   destruct H1.
   destruct H2.
   subst.
   replace tr with x0.
   subst.
   assumption.
   revert H1.
   clear.
   revert x0.
   induction tr; destruct x0; try discriminate.
    auto.
   injection 1; intros; subst.
   f_equal.
   auto.
  destruct H1.
  destruct H2.
  destruct H2.
  destruct H2.
  destruct H3.
  destruct H2.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H4.
  destruct H4.
  destruct H4.
  exfalso.
  revert H4.
  clear.
  revert x0.
  induction tr; destruct x0; try discriminate.
  injection 1; intros; subst; eauto.
Qed.

Lemma cstep_step_no_external: forall c tr c',
  cstep m c (map (@Event _) tr) c' ->
  step m c tr c'.
Proof.
  intros.
  refine (_ (cstepn_stepn_no_external (stepn_left_S H (stepn_left_O _ _) _))).
  2: rewrite <- app_nil_end; reflexivity.
  inversion 1; subst.
  inversion Hstepn; subst.
  rewrite <- app_nil_end in *.
  assumption.
Qed.

Lemma cstep_step_silently_diverges: forall s,
  config_silently_diverges (cstep m) s ->
  config_silently_diverges (step m) s.
Proof.
  cofix COINDHYP.
  inversion 1; subst.
  generalize (cstep_step_no_external (tr := nil) Hstep).
  intro.
  econstructor.
  eassumption.
  eapply COINDHYP.
  assumption.
Qed.

Lemma cstepn_stepn_with_external: forall n c tr f a h1 r h' tr' c',
  stepn_left (cstep m) n c (map (@Event _) tr ++ External f a h1 r h' :: tr') c' ->
  exists h, exists l, exists k,
  star (step m) c tr (Config h l k) /\
  ~ in_dom m f /\
  local_kind l = Some (Call f a) /\
  exists fr1,
  caller_call a h l = (h1, fr1).
Proof.
  intros.
  destruct (cstepn_stepn H).
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H0.
  destruct H1.
   exfalso.
   destruct H1.
   revert H1.
   clear H0.
   revert x0.
   clear.
   induction tr; destruct x0; try discriminate.
   injection 1; intros; subst; eauto.
  destruct H1.
  destruct H2.
  destruct H2.
  destruct H3.
  destruct H3.
  destruct H4.
  destruct H4.
  destruct H4.
  destruct H5.
  destruct H5.
  destruct H5.
  assert (tr = x0).
   revert H5.
   clear H0.
   revert x0.
   clear.
   induction tr; destruct x0; try discriminate.
    trivial.
   injection 1; intros; subst.
   f_equal.
   eauto.
  subst.
  generalize (app_inv_head _ _ _ H5).
  injection 1; intros; subst.
  eauto 9 using stepn_left_star.
Qed.

Lemma cstep_step_no_external_reacts: forall c tr1 tr2,
  bisim tr1 (mapinf (@Event _) tr2) ->
  config_reacts (cstep m) c tr1 ->
  config_reacts (step m) c tr2.
Proof.
  cofix COINDHYP.
  inversion 2; subst.
  destruct (bisim_appinf_strong H).
  refine ((fun H2 : exists el, e :: l = map (@Event _) el => _) _).
  refine ((fun H3 : exists eil, x = mapinf (@Event _) eil => _) _).
  destruct H2.
  destruct H3.
  subst.
  rewrite H2 in Hstep.
  destruct (star_stepn_left Hstep).
  generalize (cstepn_stepn_no_external H3).
  destruct x0.
   discriminate.
  injection H2; intros; subst.
  rewrite H1 in H.
  simpl in H.
  inversion H; subst.
  generalize (appinf_bisim CIHbisim).
  intro.
  refine ((fun H5 : exists tr3, tr2 = appinf (e0 :: x0) tr3 => _) _).
  destruct H5.
  subst.
  simpl in H1.
  rewrite mapinf_cons in H1.
  injection H1; intros; subst.
  rewrite mapinf_appinf in H5.
  generalize (bisim_refl (appinf (map (@Event _) x0) (mapinf (@Event _) x2))).
  intro.
  rewrite H5 in H7 at 1.
  generalize (appinf_bisim H7).
  intro.
  econstructor.
  2: eapply COINDHYP.
  eapply stepn_left_star.
  eassumption.
  2: eassumption.
  eapply bisim_trans.
  eassumption.
  eassumption.
  reflexivity.

  clear COINDHYP.
   revert H1.
   clear.
   revert e0 x1 tr2.
   induction x0; simpl.
    destruct tr2.
    rewrite mapinf_cons.
    injection 1; intros; subst.
    eauto.
   intros until tr2.
   destruct tr2.
   rewrite mapinf_cons.
   injection 1; intros; subst.
   destruct (IHx0 _ _ _ H).
   subst.
   simpl.
   eauto.

  clear COINDHYP.
   revert H1.
   clear.
   revert tr2 e x.
   induction l; simpl.
    destruct tr2.
    rewrite mapinf_cons.
    injection 1; intros; subst.
    eauto.
   destruct tr2.
   rewrite mapinf_cons.
   injection 1; intros; subst.
   destruct (IHl _ _ _ H).
   rewrite H0.
   eauto.

  clear COINDHYP.
   revert H1.
   clear.
   revert tr2 e x.
   induction l; simpl.
    destruct tr2.
    intro.
    rewrite mapinf_cons.
    injection 1; intros; subst.
    eexists (_ :: nil); simpl; eauto.
   intros until x.
   destruct tr2.
   rewrite mapinf_cons.
   injection 1; intros; subst.
   destruct (IHl _ _ _ H).
   destruct x0.
    discriminate.
   injection H0; intros; subst.
   eexists (_ :: _ :: _).
   simpl.
   reflexivity.
Qed.

Lemma bstep_empty_stack_or_stuck: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr beh' k',
    bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )) (BBehavior beh nil) tr (BBehavior beh' k') ->
    (k' = nil /\ beh = append1_behavior tr beh' /\ exists etr, tr = map (@Event _) etr) \/ (exists f, exists a, exists h1, exists r, exists h2, exists beh'', beh = cons_behavior (External f a h1 r h2) beh'' /\ tr = nil /\ beh' = Stuck _ nil).
Proof.
  intros.
  inversion H0; subst; simpl; auto.
   left.
   split.
    trivial.
   split.
    trivial.
   eexists (_ :: nil).
   simpl. reflexivity.
  destruct Hf.
  destruct (in_dom_dec m f).
   apply indom_merge_indom_1.
   apply dom_in_dom.
   rewrite create_map_dom.
   apply in_dom_dom.
   assumption.
  apply indom_merge_indom_2.
  apply not_in_dom_stuck_closure.
  assumption.
  right.
  assert (~ in_dom m f).
   intro.
   exact (H (_ :: nil) _ (refl_equal _) _ H1 _ _ _ _ (or_introl _ (refl_equal _))).
  assert (~ in_dom (create_map (dom m) (csemantics m)) f).
  intro.
  generalize (in_dom_dom H2).
  rewrite create_map_dom.
  intro.
  apply H1.
  apply dom_in_dom.
  assumption.
  generalize (mergelookup_2 Hf H2).
  unfold lookup, lookup_strong, stuck_closure.
  destruct (in_dom_dec m f).
   contradiction.
  injection 1; intros; subst.
  unfold be_stuck in Hbf.
  subst.
  eauto 9.
 left.
 split.
 trivial.
 split.
 trivial.
 exists nil.
 reflexivity.
Qed.

Lemma bstepn_empty_stack_or_stuck: forall n beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr beh' k',
    stepn_left (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure ))) n (BBehavior beh nil) tr (BBehavior beh' k') ->
    (
      (k' = nil /\ beh = append1_behavior tr beh') \/ (exists f, exists a, exists h1, exists r, exists h2, exists beh'', beh = append1_behavior (tr ++ External f a h1 r h2 :: nil) beh'' /\ beh' = Stuck _ nil)
    ) /\
    exists etr, tr = map (@Event _) etr.
Proof.
  induction n; inversion 2; subst.
   split.
   left.
   auto.
   exists nil; auto.
  destruct conf1.
   inversion Hstep.
  destruct (bstep_empty_stack_or_stuck H Hstep).
   destruct H1.
   destruct H2.
   destruct H3.
   subst.
   refine (_ (IHn _ _ _ _ _ Hstepn)).
   destruct 1.
   destruct H1.
    destruct H1.
    subst.
    split.
    left.
    split.
    trivial.
    rewrite append_behavior_app.
    trivial.
    destruct H2.
    subst.
    eexists (_ ++ _).
    rewrite map_app.
    reflexivity.
   destruct H1.
   destruct H1.
   destruct H1.
   destruct H1.
   destruct H1.
   destruct H1.
   destruct H1.
   destruct H2.
   subst.
   split.
   right.
   rewrite <- append_behavior_app.
   rewrite <- app_ass.
   eauto 8.
   eexists (_ ++ _).
   rewrite map_app.
   reflexivity.
   intros; subst.
   eapply H.
   rewrite <- append_behavior_app.
   reflexivity.
   eassumption.
   eapply in_or_app.
   right.
   eassumption.
  subst.
  destruct H1.
  destruct H1.
  destruct H1.
  destruct H1.
  destruct H1.
  destruct H1.
  destruct H1.
  destruct H2.
  subst.
  inversion Hstepn; subst.
  rewrite <- app_nil_end.
  subst.
  simpl.
  split.
  eauto 9.
  exists nil; reflexivity.
  inversion Hstep0; subst; destruct b; discriminate.
Qed.

Corollary bstepn_stuck_not_spurious: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr,
    star (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
    (BBehavior beh nil) tr (@BSpurious _) -> False.
Proof.
  intros.
  destruct (star_stepn_left H0).
  generalize (stepn_left_right H1).
  inversion 1; subst.
  inversion Hstep; subst.
  destruct (bstepn_empty_stack_or_stuck H (stepn_right_left Hstepn)).
  destruct H3.
   destruct H3.
   discriminate.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  discriminate.
Qed.

Corollary bstepn_stuck_no_spurious_behaviors: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr,
    config_terminates (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
    (@bstate_final _) (BBehavior beh nil) tr (@Spurious _) -> False.
Proof.
  inversion 2; subst.
  inversion Hfin; subst.
  eauto using bstepn_stuck_not_spurious.
Qed.

Lemma bstepn_stuck_reacts_events: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr,
    config_reacts (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
    (BBehavior beh nil) tr ->
    forall l tr', bisim tr (appinf l tr') ->
      exists el, l = map (@Event _) el.
Proof.
  intros until l.
  generalize (refl_equal (length l)).
  generalize (length l) at 1.
  intro.
  revert beh H tr H0 l.
  induction n using (well_founded_induction lt_wf).
  inversion 2; subst.
  intros; subst.
  destruct conf'.
   exfalso; eauto using bstepn_stuck_not_spurious.
  destruct (star_stepn_left Hstep).
  destruct (bstepn_empty_stack_or_stuck H0 H2).
  destruct H5.
  destruct H4.
   destruct H4.
   subst.
   destruct (le_lt_dec (length l0) (length (e :: l))).
    apply bisim_sym in H3.
    destruct (appinf_le_bisim l1 H3).
    destruct H4.
    rewrite H4 in H5.
    revert H5.
    clear.
    revert l0 x1.
    induction x0.
     simpl. intros. destruct (app_eq_nil _ _ H5). subst. exists nil. reflexivity.
     destruct l0.
     intros.
      exists nil. reflexivity.
     injection 1; intros; subst.
     destruct (IHx0 _ _ H).
     subst.
     eexists (_ :: _).
     reflexivity.
   destruct x0.
   discriminate.
   injection H5; intros; subst.
   assert (length (Event e0 :: map (@Event _) x0) <= length l0) by omega.
   destruct (appinf_le_bisim H4 H3).
   destruct H6.
   assert (length x1 < length l0).
    subst.
    rewrite app_length.
    simpl.
    omega.
   subst.
   refine (_ (H _ H8 _ _ _ CIHreacts _ (refl_equal _) _ H7)).
   destruct 1.
   subst.
   eexists ((_ :: _) ++ _).
   rewrite map_app.
   simpl.
   reflexivity.
   intros until 1.
   subst.
   intros.
   eapply H0.
   rewrite <- append_behavior_app.
   reflexivity.
   eassumption.
   eapply in_or_app.
   right.
   eassumption.
exfalso.
destruct H4.
destruct H4.
destruct H4.
destruct H4.
destruct H4.
destruct H4.
destruct H4.
subst.
inversion CIHreacts; subst.
destruct (star_stepn_left Hstep0).
inversion H4; subst.
inversion Hstep1; subst; destruct b; discriminate.
Qed.

Lemma event_trace_pop: forall tr,
  (forall l tr', bisim tr (appinf l tr') ->
    exists el, l = map (@Event sw) el) ->
  {
    e : _ & { tr' : _ &
    (forall l tr'', bisim tr' (appinf l tr'') ->
      exists el, l = map (@Event _) el) /\
    tr = Cons (Event e) tr'
  }}.
Proof.
  destruct tr.
  destruct x.
   intros.
   exists e.
   exists tr.
   split.
   intros.
   destruct (H (Event e :: l) (tr'')).
   simpl.
   constructor.
   assumption.
   destruct x.
    discriminate.
   injection H1; intros; subst; eauto.
   trivial.
  intros.
  exfalso.
  destruct (H (_ :: nil) _ (bisim_refl _)).
  destruct x; discriminate.
Qed.

CoFixpoint event_trace_obj (tr: Stream xevent)
  (Htr: forall l tr', bisim tr (appinf l tr') ->
    exists el, l = map (@Event _) el) : Stream event :=
  match event_trace_pop Htr with
    | existT e (existT tr' (conj H _)) => Cons e (event_trace_obj H)
  end.

Lemma event_trace_obj_prop: forall tr Htr,
  bisim tr (mapinf (@Event _) (@event_trace_obj tr Htr)).
Proof.
  cofix COINDHYP.
  destruct tr.
  intros.
  rewrite (unfold_Stream (event_trace_obj Htr)).
  simpl.
  destruct (event_trace_pop Htr).
  destruct s.
  destruct a.
  rewrite mapinf_cons.
  injection e0; intros; subst.
  constructor.
  apply COINDHYP.
Qed.

Corollary event_trace_obj_ex: forall tr,
  (forall l tr', bisim tr (appinf l tr') ->
    exists el, l = map (@Event _) el) ->
  { etr | bisim tr (mapinf (@Event sw) etr) }.
Proof.
  intros.
  esplit.
  eapply (event_trace_obj_prop) with (Htr := H).
Qed.

Lemma bstepn_stuck_silently_diverges: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  config_silently_diverges (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
  (BBehavior beh nil) ->
    beh = Diverges _ nil.
Proof.
  inversion 2; subst.
  destruct conf'.
  exfalso.
  eapply bstepn_stuck_not_spurious.
  eassumption.
  eright.
  eleft.
  eassumption.
 destruct (bstep_empty_stack_or_stuck H Hstep).
  destruct H1.
  destruct H2.
  simpl in H2.
  subst.
  inversion Hstep; subst.
  auto.
 exfalso.
 destruct H1.
 destruct H1.
 destruct H1.
 destruct H1.
 destruct H1.
 destruct H1.
 destruct H1.
 destruct H2.
 subst.
 inversion CIHsilently_diverges; subst.
 inversion Hstep0; subst.
 destruct b; discriminate.
Qed.

Lemma bstepn_stuck_diverges: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr,
    config_diverges (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
  (BBehavior beh nil) tr ->
    beh = Diverges _ tr /\ exists etr, tr = map (@Event _) etr.
Proof.
  inversion 2; subst.
  destruct conf'.
   exfalso.
   eauto using bstepn_stuck_not_spurious.
  destruct (star_stepn_left Hstar).
  destruct (bstepn_empty_stack_or_stuck H H1).
  split; auto.
  destruct H3.
  subst.
  destruct H2.
   destruct H2.
   subst.
   refine (_ (bstepn_stuck_silently_diverges _ Hsilently_diverges)).
   intro; subst.
   rewrite append_behavior.
   simpl.
   rewrite <- app_nil_end.
   reflexivity.
  intros; subst.
  eapply H.
  rewrite <- append_behavior_app.
  reflexivity.
  eassumption.
  eapply in_or_app.
  right.
  eassumption.
 exfalso.
 destruct H2.
 destruct H2.
 destruct H2.
 destruct H2.
 destruct H2.
 destruct H2.
 destruct H2.
 subst.
 inversion Hsilently_diverges; subst.
 inversion Hstep; subst.
 destruct b; discriminate.
Qed.

Lemma bstepn_stuck_terminates: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr res,
    config_terminates (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure ))) (@bstate_final _)
    (BBehavior beh nil) tr res ->
    (exists etr, tr = map (@Event _) etr) /\ exists h, exists r, res = Normal h r /\
    beh = Terminates tr (h,r).
Proof.
    inversion 2; subst.
  inversion Hfin; subst.
  destruct (star_stepn_left Hterm).
  destruct (bstepn_empty_stack_or_stuck H H1).
  split; auto.
  destruct H3.
  subst.
  destruct H2.
   destruct H2.
   subst.
   esplit.
   esplit.
   split.
   reflexivity.
   rewrite <- append_behavior_terminates.
   reflexivity.
  exfalso.
  destruct H2.
  destruct H2.
  destruct H2.
  destruct H2.
  destruct H2.
  destruct H2.
  destruct H2.
  discriminate.
   exfalso.
   eauto using bstepn_stuck_not_spurious.
Qed.

Lemma bstepn_stuck_reacts_which: forall beh,
  (forall tr b, beh = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr,
    config_reacts (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
    (BBehavior beh nil) tr ->
    exists ftr, beh = Reacts _ ftr.
Proof.
  cut (
    forall l beh,
      (forall tr b, beh = append1_behavior tr b ->
        forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
      forall tr,
        config_reacts (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
        (BBehavior beh nil) tr ->
        forall beh', beh = append1_behavior l beh' ->
          ((exists hr, beh' = Terminates nil hr) \/
          beh' = Diverges _ nil \/ beh' = Stuck _ nil) -> False
  ).
   intros.
   destruct beh; simpl; eauto; intros; exfalso.
   eapply H.
   eassumption.
   eassumption.
   eapply append_behavior_terminates.
   eauto.
   eapply H.
   eassumption.
   eassumption.
   eapply append_behavior_stuck.
   eauto.
   eapply H.
   eassumption.
   eassumption.
   eapply append_behavior_diverges.
   eauto.
  intro l.
  generalize (refl_equal (length l)).
  generalize (length l) at 1.
  intro n.
  revert l.
  induction n using (well_founded_induction lt_wf).
  intros; subst.
  inversion H2; subst.
  destruct conf'.
   eauto using bstepn_stuck_not_spurious.
  destruct (star_stepn_left Hstep).
  destruct (bstepn_empty_stack_or_stuck H1 H0).
  destruct H5.
  destruct H3.
   destruct H3.
   subst.
   destruct (append_behavior_elim_strong H6).
    destruct H3.
    destruct H3.
    assert (length x1 < length l).
     rewrite H3.
     rewrite app_length.
     simpl.
     omega.
    subst.
    eapply H.
    eassumption.
    reflexivity.
    2: eexact CIHreacts.
    intros.
    eapply H1.
    4: reflexivity.
    2: eassumption.
    rewrite append_behavior_app.
    rewrite H3.
    rewrite <- append_behavior_app.
    reflexivity.
    eapply in_or_app.
    right.
    eassumption.
    assumption.
  destruct H3.
  destruct H3.
  destruct H3.
  subst.
  rewrite append_behavior in H4.
  destruct H4.
   destruct H4.
    destruct b; simpl in H4; discriminate.
   destruct H4; destruct b; discriminate.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  subst.
  inversion CIHreacts; subst.
  destruct (star_stepn_left Hstep0).
  inversion H6; subst.
  inversion Hstep1; subst; destruct b; discriminate.
Qed.

Lemma bstepn_stuck_reacts_trace: forall ftr,
  (forall tr b, Reacts (heap * ret) ftr = append1_behavior tr b ->
    forall f (Hf: in_dom m f) a h1 r h2, List.In (External f a h1 r h2) tr -> False) ->
  forall tr,
    config_reacts (bstep (merge (create_map (dom m) (csemantics m)) (stuck_closure )))
    (BBehavior (Reacts _ ftr) nil) tr ->
    forall s, bisim s (stream_of_func ftr) ->
      bisim s tr.
Proof.
  intros.
  eapply bisim'_bisim.
  revert ftr H tr H0 s H1.
  cofix COINDHYP.
  inversion 2; subst.
  destruct conf'.
   clear COINDHYP; exfalso; eauto using bstepn_stuck_not_spurious.
  destruct (star_stepn_left Hstep).
  destruct (bstepn_empty_stack_or_stuck H H1).
  destruct H2.
  destruct H3.
  destruct H2.
  rewrite append_behavior in H4.
  destruct b; try discriminate.
  injection H4; intros; subst.
  generalize (bisim_sym (bisim_trans H6 (bisim_sym (stream_of_func_of_stream (Cons e (appinf l (stream_of_func evinfseq))))))).
  intro.
  change (Cons e (appinf l (stream_of_func evinfseq))) with (appinf (e :: l) (stream_of_func evinfseq)) in H2.
  destruct (bisim_appinf_strong H2).
  subst.
  generalize (bisim_sym (appinf_bisim H2)).
  intro.
  econstructor.
  eapply COINDHYP.
  3: eassumption.
  2: eassumption.
  3: reflexivity.
  3: reflexivity.
  2: discriminate.
  clear COINDHYP.
  intros.
  change (Reacts (heap * ret) (func_of_stream (Cons e (appinf l (stream_of_func evinfseq))))) with (append2_behavior (e :: l) (Reacts (heap * ret) evinfseq)) in H.
  rewrite <- append_behavior in H.
  eapply H.
  rewrite H7.
  rewrite <- append_behavior_app.
  reflexivity.
  eassumption.
  eapply in_or_app.
  right.
  eassumption.
clear COINDHYP.
exfalso.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H2.
destruct H3.
subst.
inversion CIHreacts; subst.
destruct (star_stepn_left Hstep0).
inversion H4; subst.
inversion Hstep1; subst; destruct b; discriminate.
Qed.

Lemma bsemantics_in_semantics_stuck_closure: forall f,
  in_dom m f ->
 forall a h b',
   bsemantics (merge (create_map (dom m) (csemantics m)) (stuck_closure )) f a h b' ->
   exists b, b' = cbehavior_of_behavior b /\
     semantics m f a h b.
Proof.
  inversion 2; subst.
  refine (_ (mergelookup_1 Hcmd _)).
  intro.
  generalize (create_map_lookup_elim x).
  intro; subst.
  inversion Hcb; subst.
  generalize (cstep_beh_trace_inv Hbeh0).
  intro.
  destruct b'; simpl in *.
   destruct hfin.
   simpl in Hbeh.
   destruct (bstepn_stuck_terminates H1 Hbeh).
   destruct H2.
   destruct H3.
   destruct H3.
   destruct H3.
   injection H3; intros; subst.
   inversion Hbeh0; subst.
   destruct (star_stepn_left Hterm).
   generalize (cstepn_stepn_no_external H2).
   intro.
   eexists (Terminates _ _).
   split.
   reflexivity.
   econstructor.
   eassumption.
   eassumption.
   econstructor.
   eapply stepn_left_star.
   eassumption.
   assumption.
  inversion Hbeh; subst.
  destruct cstuck.
   edestruct Hstuck_not_final.
   econstructor.
  destruct (star_stepn_left Hstar).
  destruct (behavior_elim_ex b).
   subst.
   destruct r.
   destruct l.
    edestruct Hstuck_not_final.
    econstructor.
   destruct p.
   destruct p.
   destruct (classic ((h0,r) = (h1,r0))).
    injection H3; intros; subst.
    edestruct Hconf_stuck.
    eapply bstep_terminates.
    edestruct Hconf_stuck.
    eapply bstep_spurious.
    assumption.
   subst.
   destruct (bstepn_empty_stack_or_stuck H1 H2).
   destruct H4.
   subst.
   destruct H3.
   destruct H3.
   subst.
   rewrite append_behavior in Hbeh0.
   simpl in Hbeh0.
   inversion Hbeh0; subst.
   rewrite <- app_nil_end in *.
   destruct (star_stepn_left Hstar0).
   generalize (cstepn_stepn_no_external H3).
   intro.
   eexists (Stuck _ _).
   split.
   reflexivity.
   econstructor.
   eassumption.
   eassumption.
   econstructor.
   eapply stepn_left_star.
   eassumption.
   assumption.
   intros.
   eapply Hconf_stuck0.
   eapply cstep_step.
   eassumption.
   reflexivity.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  destruct H3.
  subst.
  destruct (config_behavior_star_recip Hbeh0).
  destruct H3.
  destruct H3.
  destruct H5.
  destruct H5.
  subst.
  rewrite app_ass in H3.
  simpl in H3.
  destruct (star_stepn_left H3).
  destruct (cstepn_stepn_with_external H6).
  destruct H7.
  destruct H7.
  destruct H7.
  destruct H8.
  destruct H9.
  destruct H10.
  eexists (Stuck _ _).
  split.
  reflexivity.
  econstructor.
  eassumption.
  eassumption.
  econstructor.
  eassumption.
  inversion 1; subst.
  rewrite H9 in Hcmd1.
  discriminate.
  inversion 1; subst.
   generalize (eval_internal_kind Hstep).
   rewrite H9; discriminate.
   rewrite H9 in Hl.
   injection Hl; intros; subst.
   refine (_ (lookup_in_dom _)).
   eassumption.
   eassumption.
   rewrite H9 in Hcmd1. discriminate.
  subst.
  edestruct Hconf_stuck.
  eapply bstep_diverges.
  subst.
  destruct e.
   edestruct Hconf_stuck.
   eapply bstep_cons.
   reflexivity.
  destruct (in_dom_dec m f0).
   generalize (in_dom_dom i).
   rewrite <- (create_map_dom (dom m) (csemantics m)).
   intro.
   generalize (dom_in_dom H3).
   destruct 1.
   generalize (create_map_lookup_elim H4).
   intro; subst.
   destruct i.
   case_eq (callee_init a0 h0 x1).
   intros.
   destruct (exists_behavior (@config_final _ _) (cstep m) (Config h1 l0 nil)).
   edestruct Hconf_stuck.
   eapply bstep_internal.
   eapply lookup_merge_lookup_1.
   eassumption.
   econstructor.
   eassumption.
   eassumption.
   eassumption.
   reflexivity.
  generalize (not_in_dom_stuck_closure n).
  intro.
  inversion H3; subst.
  generalize H4.
  unfold lookup, lookup_strong, stuck_closure.
  destruct (in_dom_dec m f0); try discriminate.
  injection 1; intros; subst.
  edestruct Hconf_stuck.
  eapply bstep_internal.
  eapply lookup_disj_merge_lookup_2.
  eassumption.
  apply stuck_closure_disjoint.
  reflexivity.
  reflexivity.
  simpl in Hbeh.
   destruct (bstepn_stuck_diverges H1 Hbeh).
   destruct H3.
   subst.
   inversion Hbeh0; subst.
   destruct (star_stepn_left Hstar).
   generalize (cstepn_stepn_no_external H2).
   intro.
   eexists (Diverges _ _).
   split.
   reflexivity.
   econstructor.
   eassumption.
   eassumption.
   econstructor.
   eapply stepn_left_star.
   eassumption.
   eapply cstep_step_silently_diverges.
   assumption.
 destruct (bstepn_stuck_reacts_which H1 Hbeh).
 subst.
 generalize (bstepn_stuck_reacts_trace H1 Hbeh (bisim_refl _)).
 intro.
 generalize (bisim_func_eq H2).
 rewrite func_of_stream_of_func.
 rewrite func_of_stream_of_func.
 intro; subst.
 clear H2.
 generalize (bstepn_stuck_reacts_events H1 Hbeh).
 intro.
 destruct (event_trace_obj_ex H2).
 generalize (cstep_step_no_external_reacts b Hbeh0).
 intro.
 exists (Reacts _ (func_of_stream x0)).
 split.
 simpl.
 rewrite (fmapinf_func_of_stream_bisim (@Event _) (bisim_refl _)).
 f_equal.
 rewrite <- (func_of_stream_of_func evinfseq).
 eapply bisim_func_eq.
 assumption.
 econstructor.
 eassumption.
 eassumption.
 simpl.
 eapply config_reacts_bisim.
 eassumption.
 apply stream_of_func_of_stream.
intro.
eapply stuck_closure_disjoint.
2: eassumption.
eapply dom_in_dom.
rewrite create_map_dom.
apply in_dom_dom.
assumption.
Qed.

CPP 2015 submission: Lemma 2.


Theorem semantics_bsemantics_stuck_closure :
  forall f, in_dom m f ->
    bsemantics (merge (create_map (dom m) (csemantics m)) (stuck_closure )) f =
    lift_sem (semantics m) f.
Proof.
  intros.
  apply functional_extensionality.
  intro a.
  apply functional_extensionality.
  intro h.
  apply Extensionality_Ensembles.
  split; unfold Included, Ensembles.In, lift_sem, lift_behset.
   intros; eauto using bsemantics_in_semantics_stuck_closure.
  destruct 1; destruct H0; subst; eauto using semantics_in_bsemantics_stuck_closure.
Qed.

End Language.

End Semworld.

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