Library instance

Require Import Coqlib Integers List Maps AST Events Values Memory Globalenvs Op Registers RTL.
Set Implicit Arguments.
Unset Strict Implicit.

An instantiation of CompCert's RTL language into our framework.
We define a language in our procedural small-step semantics framework (CPP 2015 submission, Section III), and we show that its compositional big-step semantics (op. cit., Section IV) is equivalent to the big-step semantics of RTL.
Then, we prove that common subexpression elimination is correct and compositional, i.e. if m2 = CSE(m1), then for any RTL module m, (m ++ m2) refines (m ++ m1).

Definition


Inductive local : Type :=
  | LState
    (f: function)
    (sp: val)
    (pc: node)
    (rs: regset)
  | LCallstateAux
to make the function call explicit
    (fid: ident)
    (fsig: signature)
    (args: list val)
    (sf: stackframe)
  | LCallstate
    (f: pfundef)
    (args: list val)
  | LReturnstate
    (v: val)
    .

Function local_kind (l: local) : option (proc.special_local_kind sw) :=
  match l with
    | LCallstateAux fid fsig args _ => Some (proc.Call (sw := sw) (fid, fsig) args)
    | LReturnstate v => Some (proc.Return (sw := sw) v)
    | _ => None
  end.

Function local_of_state (s: state): option local :=
  match s with
    | State _ f sp pc rs _ => Some (LState f sp pc rs)
    | Callstate _ (External (EF_true_external _ _)) _ _ => None
    | Callstate _ (External (EF_pseudo pef)) args _ => Some (LCallstate (External pef) args)
    | Callstate _ (Internal fd) args _ => Some (LCallstate (Internal fd) args)
    | Returnstate _ v _ => Some (LReturnstate v)
  end.

Definition stack_of_state (s: state) :=
  match s with
    | State stk _ _ _ _ _ => stk
    | Callstate stk _ _ _ => stk
    | Returnstate stk _ _ => stk
  end.

Definition mem_of_state (s: state) :=
  match s with
    | State _ _ _ _ _ m => m
    | Callstate _ _ _ m => m
    | Returnstate _ _ m => m
  end.

Function filter_local (l: local): bool :=
  match l with
    | LState f _ pc _ =>
      match (fn_code f) ! pc with
        | Some (Icall _ _ _ _ _) => false
        | _ => true
      end
    | LReturnstate _ => false
    | _ => true
  end.

Lemma local_of_state_step : forall s l
  (Hlocal: local_of_state s = Some l)
  (Hfilter: filter_local l = true)
  ge t s'
  (Hstep: step ge s t s'),
    exists l', local_of_state s' = Some l'.
Proof.
  inversion 3; subst; simpl in *; try discriminate; injection Hlocal; intros; subst; simpl in *; try rewrite H in *; simpl in *; try discriminate; eauto.
Qed.

Lemma local_of_state_stack: forall s l
  (Hlocal: local_of_state s = Some l)
  (Hfilter: filter_local l = true)
  ge t s'
  (Hstep: step ge s t s'),
  stack_of_state s' = stack_of_state s.
Proof.
  inversion 3; subst; simpl in *; try discriminate; injection Hlocal; intros; subst; simpl in *; try rewrite H in *; simpl in *; try discriminate; eauto.
Qed.

Lemma local_of_state_stack_indep: forall s1 l1
  (Hlocal1: local_of_state s1 = Some l1)
  (Hfilter: filter_local l1 = true)
  ge t s2
  (Hstep: step ge s1 t s2)
  l2
  (Hlocal2: local_of_state s2 = Some l2)
  s'1
  (Hlocal'1: local_of_state s'1 = Some l1)
  (Hmem: mem_of_state s1 = mem_of_state s'1),
  exists s'2, step ge s'1 t s'2 /\ local_of_state s'2 = Some l2 /\ mem_of_state s2 = mem_of_state s'2.
Proof.
  inversion 3; subst; simpl in *; try discriminate; injection Hlocal1; intros; subst; simpl in *; try rewrite H in *; simpl in *; try discriminate; injection Hlocal2; intros; subst; simpl in *; functional inversion Hlocal'1; subst; simpl in *; try discriminate.
   esplit. split. eapply exec_Inop. eassumption. split; reflexivity.
   esplit. split. eapply exec_Iop. eassumption. eassumption. split; reflexivity.
   esplit. split. eapply exec_Iload. eassumption. eassumption. eassumption. split; reflexivity.
   esplit. split. eapply exec_Istore. eassumption. eassumption. eassumption. split; reflexivity.
   esplit. split. eapply exec_Ibuiltin. eassumption. eassumption. split; reflexivity.
   esplit. split. eapply exec_Icond. eassumption. eassumption. reflexivity. split; reflexivity.
   esplit. split. eapply exec_Ijumptable. eassumption. eassumption. eassumption. split; reflexivity.
   esplit. split. eapply exec_Ireturn. eassumption. eassumption. split; reflexivity.
   esplit. split. eapply exec_function_internal. eassumption. split; reflexivity.
   destruct ef; discriminate.
   destruct ef; try discriminate.
   injection Hlocal1; intros; subst.
   esplit. split. eapply exec_function_external. eassumption. split; reflexivity.
   destruct ef; discriminate.
Qed.

Lemma pseudo_external_call_events : forall pef (FN F V: Type) (ge: _ FN F V) lv m tr v m',
  pseudo_external_call pef ge lv m tr v m' ->
  exists o, tr = map (@comp.Event _) (proc.trace_of_option_event o).
Proof.
  destruct pef; simpl; inversion 1; subst; simpl; try (exists None; simpl; reflexivity); try (eexists (Some _); simpl; reflexivity).
  inversion H0; subst; try (exists None; simpl; reflexivity); try (eexists (Some _); simpl; reflexivity).
  inversion H0; subst; try (exists None; simpl; reflexivity); try (eexists (Some _); simpl; reflexivity).
  inversion H1; subst; try (exists None; simpl; reflexivity); try (eexists (Some _); simpl; reflexivity).
  inversion H1; subst; try (exists None; simpl; reflexivity); try (eexists (Some _); simpl; reflexivity).
Qed.

Lemma local_of_state_events: forall s l
  (Hlocal: local_of_state s = Some l)
  (Hfilter: filter_local l = true)
  ge t s'
  (Hstep: step ge s t s'),
  exists o, t = map (@comp.Event _) (proc.trace_of_option_event o).
Proof.
  inversion 3; subst; simpl in *; try discriminate; injection Hlocal; intros; subst; simpl in *; try rewrite H in *; simpl in *; try discriminate; eauto using pseudo_external_call_events; try (exists None; reflexivity).
  destruct ef; try discriminate. simpl in H. eauto using pseudo_external_call_events.
Qed.

Lemma local_of_state_ge : forall s l
  (Hlocal: local_of_state s = Some l)
  (Hfilter: filter_local l = true)
  ge t s'
  (Hstep: step ge s t s') f,
  step (Genv.replace_fundefs f ge) s t s'.
Proof.
  inversion 3; subst; simpl in *; try discriminate; injection Hlocal; intros; subst; simpl in *; try rewrite H in *; simpl in *; try discriminate; eauto.
   eapply exec_Inop. assumption.
   eapply exec_Iop. eassumption. assumption.
   eapply exec_Iload. eassumption. eassumption. assumption.
   eapply exec_Istore. eassumption. eassumption. assumption.
   eapply exec_Ibuiltin. eassumption.
    eapply ec_symbols_preserved. eapply pseudo_external_call_spec. 3: eassumption. reflexivity. reflexivity.
    eapply exec_Icond. eassumption. eassumption. reflexivity.
    eapply exec_Ijumptable. eassumption. eassumption. assumption.
    eapply exec_Ireturn. assumption. assumption.
    eapply exec_function_internal. assumption.
   destruct ef; try discriminate.
   injection Hlocal; intros; subst.
   eapply exec_function_external.
    simpl in *.
    eapply ec_symbols_preserved. eapply pseudo_external_call_spec. 3: eassumption. reflexivity. reflexivity.
Qed.

Corollary local_of_state_ge_inv : forall s l
  (Hlocal: local_of_state s = Some l)
  (Hfilter: filter_local l = true)
  ge t s' f
  (Hstep: step (Genv.replace_fundefs f ge) s t s') ,
  step ge s t s'.
Proof.
  intros.
  rewrite <- (Genv.replace_fundefs_id ge).
  rewrite <- (Genv.replace_fundefs_idem _ f).
  eauto using local_of_state_ge.
Qed.

Section GENV.

Variable ge: genv.

Inductive eval_internal: mem -> local -> option aevent -> mem -> local -> Prop :=
| eval_local: forall s o s'
  (Hstep: RTL.step ge s (map (@comp.Event _) (proc.trace_of_option_event o)) s')
  m (Hm: mem_of_state s = m)
  l (Hl: local_of_state s = Some l)
  (Hfilter: filter_local l = true)
  m' (Hm' : mem_of_state s' = m')
  l' (Hl': local_of_state s' = Some l'),
  eval_internal m l o m' l'
| eval_Icall:
  forall f sp pc rs m sig ros args res pc' fd,
    (fn_code f)!pc = Some(Icall sig ros args res pc') ->
    find_function ge ros rs = Some (fd, sig) ->
    eval_internal m (LState f sp pc rs)
    None m (LCallstateAux fd sig rs##args (Stackframe res f sp pc' rs))
.

Lemma eval_internal_le: forall h l tr h' l',
  eval_internal h l tr h' l' ->
  mem_le h h'.
Proof.
  inversion 1; subst; eauto using mem_le_refl.
  inversion Hstep; subst; eauto using mem_le_refl, mem_le_alloc, mem_le_free.
  unfold Mem.storev in H4.
  destruct a; try discriminate.
  eauto using mem_le_store.
  generalize (pseudo_external_call_spec ef).
  inversion 1.
  split; eauto.
  generalize (external_call_spec ef).
  inversion 1.
  split; eauto.
Qed.

Lemma eval_internal_kind : forall h l tr h' l',
  eval_internal h l tr h' l' ->
  local_kind l = None.
Proof.
  inversion 1; subst; try reflexivity.
  functional inversion Hfilter; subst; try reflexivity.
  destruct l; try contradiction; try reflexivity.
  destruct s; simpl in *; try discriminate.
  destruct f; try discriminate.
  destruct e; discriminate.
Qed.

Definition caller_call (args: list val) (m: mem) (l: local) : mem * stackframe :=
  match l with
    | LCallstateAux _ _ _ sf =>
      (m, sf)
    | _ => (m, Stackframe xH (mkfunction (mksignature nil None) nil 0 (PTree.empty _) xH) Vundef xH (Vundef, PTree.empty _))
  end.

Lemma caller_call_le: forall a h tr h' fr,
  caller_call a h tr = (h', fr) ->
  mem_le h h'.
Proof.
  unfold caller_call.
  destruct tr; injection 1; intros; subst; eauto using mem_le_refl.
Qed.

Definition caller_return (vres: val) (m: mem) (fr: stackframe) : (mem * local) :=
  match fr with
    | Stackframe res f sp pc rs =>
      (m, LState f sp pc (rs#res <- vres))
  end.

Lemma caller_return_le : forall r h fr h' l',
  caller_return r h fr = (h', l') ->
  mem_le h h'.
Proof.
  unfold caller_return; destruct fr; injection 1; intros; subst; eauto using mem_le_refl.
Qed.

Definition callee_init (args: list val) (m: mem) (c: pfundef) : (mem * local) :=
  (m, LCallstate (c) args).

Lemma callee_init_le: forall a h c h' l',
  callee_init a h c = (h', l') ->
  mem_le h h'.
Proof.
  unfold callee_init; injection 1; intros; subst; eauto using mem_le_refl.
Qed.

Definition rtl : proc.language sw :=
  proc.Language (sw := sw) eval_internal_le eval_internal_kind caller_call_le caller_return_le callee_init_le.

Proof of equivalence


Section Module.

Variable m: RTL.module.
Notation mge := (module_genv m _ ge) (only parsing).
Notation cm := (module_pseudofuns m : proc.module rtl) (only parsing).

Inductive match_states : forall (allow_R: bool), nat -> proc.config rtl -> state -> Prop :=
  | match_states_L s h
    (Hm: h = mem_of_state s)
    l
    (Hl: local_of_state s = Some l)
    stk
    (Hstk: stk = stack_of_state s) b :
    match_states b O (proc.Config (lang := rtl) h l stk) s
  | match_states_C fid fsig f
    (Hf: f = module_funs m (fid, fsig))
    h args fr stk b :
    match_states b 1 (proc.Config (lang := rtl) h (LCallstateAux fid fsig args fr) stk)
    (Callstate (fr :: stk) f args h)
  | match_states_R rs rs' res v
    (Hrs': rs' = (rs#res <- v))
    h f sp pc stk :
    match_states true 1 (proc.Config (lang := rtl) h (LState f sp pc rs') stk)
    (Returnstate (Stackframe res f sp pc rs :: stk) v h)
.

Open Scope nat_scope.

Lemma rtlstep_cstep_O:
  forall s1 t s2, step mge s1 t s2 ->
    forall b c1 (MS: match_states b O c1 s1),
      exists i', exists c2,
        comp.cstep cm c1 t c2 /\ match_states b i' c2 s2.
Proof.
  inversion 2; subst.
  case_eq (filter_local l).
   intros.
   destruct (local_of_state_step Hl H0 H).
   destruct (local_of_state_events Hl H0 H).
   subst.
   esplit.
   esplit.
   split.
   eapply comp.cstep_step.
   eapply proc.step_internal.
   simpl.
   eleft.
   eapply local_of_state_ge_inv.
   3: eassumption.
   eassumption.
   assumption.
   reflexivity.
   assumption.
   assumption.
   reflexivity.
   eassumption.
   reflexivity.
   reflexivity.
   eapply match_states_L.
   reflexivity.
   assumption.
   symmetry; eauto using local_of_state_stack.
 intros.
 functional inversion H0; subst.
  functional inversion Hl; subst.
  inversion H; subst; try (exfalso; congruence; fail).
  destruct fd.
  change (find_function mge ros _x0) with (find_function ge ros _x0) in H11.
  esplit.
  esplit.
  split.
  eapply comp.cstep_step.
  eapply proc.step_internal.
  simpl.
  eright.
  eassumption.
  simpl.
  rewrite module_funs_sig.
  eassumption.
  reflexivity.
  reflexivity.
  eapply match_states_C.
  simpl.
  rewrite module_funs_sig.
  reflexivity.
 functional inversion Hl; subst.
 inversion H; subst; try (exfalso; congruence; fail).
 esplit.
 esplit.
 split.
 eapply comp.cstep_step.
 eapply proc.step_return.
 simpl. reflexivity.
 simpl. reflexivity.
 reflexivity.
 eapply match_states_L.
 reflexivity.
 reflexivity.
 reflexivity.
Qed.

Notation CSTEP := (fun _ : genv => comp.cstep cm) (only parsing).

Lemma rtlstep_cstep:
  forall s1 t s2, step mge s1 t s2 ->
    forall i c1 (MS: match_states true i c1 s1),
      exists i', exists c2,
        (Smallstep.plus CSTEP mge c1 t c2 \/ (t = nil /\ c1 = c2 /\ i' < i)) /\ match_states true i' c2 s2.
Proof.
  inversion 2; subst.
  destruct (rtlstep_cstep_O H MS).
  destruct H0.
  destruct H0.
  eauto 6 using Smallstep.plus_one.
 inversion H; subst; try (exfalso; congruence; fail).
  symmetry in H2.
  generalize H2.
  unfold module_funs.
  case_eq (module_pseudofuns m (fid, fsig)); try discriminate.
  intros until 1.
  destruct f0; try discriminate.
  injection 1; intro; subst.
  esplit.
  esplit.
  split.
  left.
  econstructor.
  eapply comp.cstep_step.
  eapply proc.step_call.
  simpl. reflexivity.
  2: simpl; reflexivity.
  2: simpl; reflexivity.
  2: simpl; reflexivity.
  eassumption.
  eright.
  eapply comp.cstep_step.
  eapply proc.step_internal.
  simpl.
  eleft with (o := None).
  simpl.
  eapply exec_function_internal with (s := nil).
  eassumption.
  reflexivity.
  reflexivity.
  reflexivity.
  simpl; reflexivity.
  simpl; reflexivity.
  simpl; reflexivity.
  simpl; reflexivity.
  eleft.
  simpl; reflexivity.
  reflexivity.
  eapply match_states_L.
  reflexivity.
  reflexivity.
  reflexivity.
  symmetry in H2.
  generalize H2.
  unfold module_funs.
  case_eq (module_pseudofuns m (fid, fsig)); try discriminate.
   destruct f; try discriminate.
   injection 2; intro; subst.
   simpl in H6.
   destruct (pseudo_external_call_events H6).
   subst.
   esplit.
   esplit.
   split.
   left.
   econstructor.
   eapply comp.cstep_step.
   eapply proc.step_call.
   simpl. reflexivity.
   eassumption.
   simpl; reflexivity.
   simpl; reflexivity.
   simpl; reflexivity.
   eright.
   eapply comp.cstep_step.
   eapply proc.step_internal.
   eleft.
   eapply exec_function_external with (ef := EF_pseudo p) (s := nil).
   simpl. eapply ec_symbols_preserved. eapply pseudo_external_call_spec. 3: simpl; eassumption.
   reflexivity. reflexivity. reflexivity. reflexivity. reflexivity.
   simpl; reflexivity. simpl; reflexivity. reflexivity.
   reflexivity. eleft.
   rewrite E0_right. reflexivity.
   reflexivity.
   eapply match_states_L.
   reflexivity.
   reflexivity.
   reflexivity.
  injection 2; intros; subst.
  simpl in H6.
  inversion H6; intros; subst.
  destruct fr.
  esplit.
  esplit.
  split.
  left.
  econstructor.
  eapply comp.cstep_call_external.
  simpl; reflexivity.
  intro.
  destruct H3.
  unfold map.lookup, map.lookup_strong in H3. congruence.
  simpl; reflexivity.
  eassumption.
  simpl. reflexivity.
  eleft.
  rewrite E0_right.
  reflexivity.
  eapply match_states_R.
  reflexivity.
 inversion H; subst; try (exfalso; congruence; fail).
 exists O.
 esplit.
 split.
 right.
 split.
 reflexivity.
 split.
 reflexivity.
 omega.
 eapply match_states_L.
 reflexivity.
 reflexivity.
 reflexivity.
Qed.

Lemma cstep_rtlstep_O:
  forall c1 t c2, (comp.cstep cm) c1 t c2 ->
    forall b s1 (MS: match_states b O c1 s1),
      exists i', exists s2,
        step mge s1 t s2 /\ match_states b i' c2 s2.
Proof.
  inversion 2; subst.
  inversion H; subst.
   inversion Hstep; subst.
    simpl in Hstep0.
    inversion Hstep0; subst.
     destruct (local_of_state_stack_indep Hl0 Hfilter Hstep1 Hl' Hl Hm).
     destruct H0.
     destruct H1.
     esplit.
     esplit.
     split.
     eapply local_of_state_ge.
     eassumption.
     assumption.
     eassumption.
     eapply match_states_L.
     assumption.
     assumption.
     symmetry.
     eapply local_of_state_stack.
     eassumption.
     assumption.
     eassumption.
     functional inversion Hl; subst; simpl in *.
    esplit.
    esplit.
    split.
    simpl.
    eapply exec_Icall.
    eassumption.
    eassumption.
    simpl. eapply module_funs_sig.
    eapply match_states_C.
    reflexivity.
  simpl in Hl0.
  functional inversion Hl0; subst.
  functional inversion Hl.
  simpl in Hcmd.
  functional inversion Hcmd; subst.
  functional inversion Hl; subst.
  simpl in Hreturn. unfold caller_return in Hreturn. destruct fr. simpl in H3. injection Hreturn; intros; subst.
  esplit.
  esplit.
  split.
  eapply exec_return.
  eapply match_states_L.
  reflexivity.
  reflexivity.
  reflexivity.
 simpl in Hcmd. functional inversion Hcmd; subst. functional inversion Hl.
Qed.

Lemma cstep_rtlstep:
  forall c1 t c2, (comp.cstep cm) c1 t c2 ->
    forall i s1 (MS: match_states false i c1 s1),
      exists i', exists s2,
        (Smallstep.plus step mge s1 t s2 \/ (t = nil /\ s1 = s2 /\ i' < i)) /\ match_states false i' c2 s2.
Proof.
  inversion 2; subst.
  destruct (cstep_rtlstep_O H MS).
  destruct H0.
  destruct H0.
  eauto 6 using Smallstep.plus_one.
 inversion H; subst.
  inversion Hstep; subst.
   simpl in Hstep0. inversion Hstep0; subst. functional inversion Hl.
   simpl in Hl. injection Hl; intros; subst.
   destruct fr. simpl in Hl''. unfold callee_init in Hl''. injection Hl''; intros; subst.
   unfold map.lookup, map.lookup_strong in Hfunc.
   simpl in Hframe.
   injection Hframe; intros; subst.
   exists O.
   esplit.
   split.
   right.
   split.
   reflexivity.
   split.
   reflexivity.
   omega.
   eapply match_states_L.
   reflexivity.
   unfold module_funs.
   rewrite Hfunc.
   destruct cmd; reflexivity.
   reflexivity.
  simpl in Hcmd. discriminate.
  simpl in Hcmd. injection Hcmd; intros; subst.
  simpl in Hfr. injection Hfr; intros; subst.
  simpl in Hl'. unfold caller_return in Hl'. destruct fr0. injection Hl'; intros; subst.
  unfold module_funs.
   case_eq (module_pseudofuns m (fid, fsig)).
    intros.
    destruct Hf.
    esplit.
    eassumption.
   intro.
  esplit.
  esplit.
  split.
  left.
  econstructor.
  eapply exec_function_external.
  simpl.
  econstructor.
  eassumption.
  eright.
  eapply exec_return.
  eleft.
  simpl; reflexivity.
  rewrite E0_right. reflexivity.
  eapply match_states_L.
  reflexivity.
  reflexivity.
  reflexivity.
Qed.

Inductive rtl_initial_state : ((ident * signature) * (list val * mem)) -> proc.config rtl -> Prop :=
| rtl_initial_state_intro: forall f s phi, module_pseudofuns m (f,s) = Some phi ->
      forall args h,
    rtl_initial_state ((f, s), (args, h)) (proc.Config (lang := rtl) h (LCallstate phi args) nil)
    .

Restrict the semantics of RTL to functions defined in the module

Inductive RTL_initial_state : ((ident * signature) * (list val * mem)) -> state -> Prop :=
| RTL_initial_state_intro:
  forall fs phi, module_pseudofuns m (fs) = Some phi ->
    forall ah st, RTL.initial_state mge (fs, ah) st -> RTL_initial_state (fs, ah) st.

Lemma rtl_RTL_initial_states: forall inp s1, rtl_initial_state inp s1 ->
  exists s2, RTL_initial_state inp s2 /\ forall b, match_states b O s1 s2.
Proof.
  inversion 1; subst.
  esplit.
  split.
  econstructor.
  eassumption.
  econstructor.
  simpl.
  apply module_funs_sig.
  intros.
  constructor.
  reflexivity.
  simpl. unfold module_funs. rewrite H0. destruct phi; reflexivity.
  reflexivity.
Qed.

Lemma RTL_rtl_initial_states: forall inp s1, RTL_initial_state inp s1 ->
  exists s2, rtl_initial_state inp s2 /\ forall b, match_states b O s2 s1.
Proof.
  inversion 1; subst.
  inversion H1.
  subst fs ah s1.
  simpl in H5.
  unfold module_funs in H5.
  rewrite H0 in H5.
   esplit.
   split.
   econstructor.
   eassumption.
   intro.
   constructor.
   reflexivity.
   simpl in *. unfold module_funs. rewrite H0. destruct phi; reflexivity.
   reflexivity.
Qed.

Lemma rtl_RTL_final_states:
  forall b i s1 s2,
    match_states b i s1 s2 -> forall out, proc.config_final (lang := rtl) s1 out -> final_state s2 out.
Proof.
  inversion 2; subst.
  inversion H; subst; try (exfalso; simpl in *; congruence).
  simpl in Hcmd.
  functional inversion Hcmd; subst.
  functional inversion Hl; subst; simpl in *.
  subst.
  constructor.
Qed.

Lemma RTL_rtl_final_states:
  forall b i s1 s2,
    match_states b i s1 s2 -> forall out, final_state s2 out -> proc.config_final (lang := rtl) s1 out.
Proof.
  inversion 2; subst.
  inversion H; subst.
  functional inversion Hl; subst.
  simpl in *.
  constructor.
  reflexivity.
Qed.

Definition rtl' : Smallstep.semantics ((ident * signature) * (list val * mem)) (mem * val) :=
  Smallstep.Semantics CSTEP
  rtl_initial_state
  (proc.config_final (lang := rtl))
  mge
  .

Definition RTL : Smallstep.semantics ((ident * signature) * (list val * mem)) (mem * val) :=
  Smallstep.Semantics step
  RTL_initial_state
  final_state
  mge
  .

Require Behaviors unify.

Lemma RTL'_eq_RTL : forall fs, cm fs <> None -> forall arh,
  Behaviors.program_behaves RTL (fs, arh) = Behaviors.program_behaves (RTL.module_sem m _ ge) (fs, arh).
Proof.
  intros.
  apply Ensembles.Extensionality_Ensembles.
  split; unfold Ensembles.Included, Ensembles.In, module_sem; inversion 1; subst; simpl in *.
   inversion H1; subst.
   econstructor.
   eassumption.
   inversion H2; subst; simpl in *; econstructor; eauto.
  case_eq (module_pseudofuns m fs).
   intros.
   destruct arh.
   destruct fs.
   edestruct H1.
   econstructor.
   eassumption.
   econstructor.
   simpl.
   apply RTL.module_funs_sig.
  intro; contradiction.
  case_eq (module_pseudofuns m fs).
  eleft.
  econstructor.
  eassumption.
  eassumption.
  inversion H2; subst; simpl in *; econstructor; eauto.
  intro; contradiction.
   intros.
   destruct fs.
   destruct arh.
   edestruct H1.
   econstructor.
   simpl.
   apply RTL.module_funs_sig.
Qed.

Lemma rtl_RTL_simul: Smallstep.forward_simulation rtl' RTL.
Proof.
  refine (Smallstep.Forward_simulation rtl' RTL _ (match_states false) _ _ _ _); simpl in *.
  eapply lt_wf.
  intros.
  destruct (rtl_RTL_initial_states H).
  destruct H0; eauto.
  intros; eauto using rtl_RTL_final_states.
  intros.
  destruct (cstep_rtlstep H H0).
  destruct H1.
  destruct H1.
  destruct H1.
   eauto.
  destruct H1.
  destruct H3.
  subst.
  change nil with E0.
  eauto 6 using Smallstep.star_refl.
  reflexivity.
Defined.

Lemma RTL_rtl_simul: Smallstep.forward_simulation RTL rtl'.
Proof.
  refine (Smallstep.Forward_simulation RTL rtl' _ (fun i s1 s2 => match_states true i s2 s1) _ _ _ _); simpl in *.
  eapply lt_wf.
  intros.
  destruct (RTL_rtl_initial_states H).
  destruct H0; eauto.
  intros; eauto using RTL_rtl_final_states.
  intros.
  destruct (rtlstep_cstep H H0).
  destruct H1.
  destruct H1.
  destruct H1.
   eauto.
  destruct H1.
  destruct H3.
  subst.
  eauto 6 using Smallstep.star_refl.
  reflexivity.
Defined.

Lemma rtlstep_cstep_stuck_O:
  forall b s1 c1 (MS: match_states b O c1 s1),
    Smallstep.nostep step mge s1 ->
    (forall r, ~ final_state s1 r) ->
    Smallstep.nostep CSTEP mge c1 /\ forall r, ~ proc.config_final c1 r.
Proof.
  inversion 1; subst.
   intros.
   split.
   intro.
   intros.
   intro.
   destruct (cstep_rtlstep_O H1 MS).
   destruct H2.
   destruct H2.
   eapply H; eauto.
   intros; intro.
   edestruct H0; eauto using rtl_RTL_final_states.
Qed.

Lemma rtlstep_cstep_stuck:
  forall b s1 i c1 (MS: match_states b i c1 s1),
    Smallstep.nostep step mge s1 ->
    (forall r, ~ final_state s1 r) ->
    exists c2,
      Smallstep.star CSTEP mge c1 E0 c2 /\ Smallstep.nostep CSTEP mge c2 /\ forall r, ~ proc.config_final c2 r.
Proof.
  inversion 1; subst.
   eauto using rtlstep_cstep_stuck_O, Smallstep.star_refl.
  intros.
  case_eq (module_pseudofuns m (fid, fsig)).
   intros.
   assert (exists c,
     comp.cstep cm (proc.Config (lang := rtl) h (LCallstateAux fid fsig args fr) stk) nil c /\
     match_states b O c (Callstate (fr :: stk) (module_funs m (fid, fsig)) args h)).
    esplit.
    split.
    eapply comp.cstep_step.
    eapply proc.step_call.
    simpl; reflexivity.
    eassumption.
    simpl; reflexivity.
    simpl; reflexivity.
    reflexivity.
    constructor.
    reflexivity.
    destruct f; unfold module_funs in *; rewrite H1 in *; reflexivity.
    reflexivity.
   destruct H2.
   destruct H2.
   eauto using Smallstep.star_one, rtlstep_cstep_stuck_O.
  intros.
  exfalso.
  revert H.
  unfold module_funs.
  rewrite H1.
  edestruct 1.
  eapply exec_function_external with (res := Vundef).
  simpl.
  econstructor.
  eapply mem_le_refl.
 edestruct 1.
 eapply exec_return.
Qed.

Lemma cstep_rtlstep_stuck:
  forall s1 i c1 (MS: match_states false i c1 s1),
    Smallstep.nostep CSTEP mge c1 ->
    (forall r, ~ proc.config_final c1 r) ->
    exists s2,
      Smallstep.star step mge s1 E0 s2 /\ Smallstep.nostep step mge s2 /\ forall r, ~ final_state s2 r.
Proof.
  inversion 1; subst.
  intros.
  esplit.
   split.
   eleft.
   split.
   intro.
   intros.
   intro.
   destruct (rtlstep_cstep_O H1 MS).
   destruct H2.
   destruct H2.
   eapply H; eauto.
   intros; intro.
   edestruct H0; eauto using RTL_rtl_final_states.
  intros.
  exfalso.
  case_eq (module_pseudofuns m (fid, fsig)).
   intros.
   edestruct H.
   eapply comp.cstep_step.
   eapply proc.step_call.
   simpl; reflexivity.
   eassumption.
   simpl; reflexivity.
   simpl; reflexivity.
   simpl; reflexivity.
  intros.
  destruct fr.
  edestruct H.
  eapply comp.cstep_call_external with (r := Vundef).
  simpl; reflexivity.
  intro.
  destruct H2.
  unfold map.lookup, map.lookup_strong in H2.
  congruence.
  simpl; reflexivity.
  eapply mem_le_refl.
  simpl. reflexivity.
Qed.

Lemma match_states_O_index_indep : forall b c s,
  match_states b O c s ->
  forall b', match_states b' O c s.
Proof.
  inversion 1; subst; intros; econstructor; eauto.
Qed.

Lemma rtl_eq_RTL_state_behaves : forall b s c,
  match_states b O c s ->
  Behaviors.state_behaves rtl' c = Behaviors.state_behaves RTL s.
Proof.
  intros.
  apply Ensembles.Extensionality_Ensembles.
  split; unfold Ensembles.Included, Ensembles.In; intros.
   eapply unify.forward_with_stuck_state_behaves with (fs := rtl_RTL_simul); simpl; eauto using cstep_rtlstep_stuck, match_states_O_index_indep.
   eapply unify.forward_with_stuck_state_behaves with (fs := RTL_rtl_simul); simpl; eauto using rtlstep_cstep_stuck, match_states_O_index_indep.
Qed.

Theorem rtl'_eq_RTL :
  Behaviors.program_behaves rtl' = Behaviors.program_behaves RTL.
Proof.
  apply FunctionalExtensionality.functional_extensionality.
  intro inp.
  apply Ensembles.Extensionality_Ensembles.
  split; unfold Ensembles.Included, Ensembles.In.
   inversion 1; subst; simpl in *.
   destruct (rtl_RTL_initial_states H0).
   destruct H2.
   rewrite (rtl_eq_RTL_state_behaves (H3 true)) in H1.
   econstructor; eauto.
  eright.
  intros.
  intro.
  destruct (RTL_rtl_initial_states H1).
  destruct H2.
  eapply H0.
  eassumption.
   inversion 1; subst; simpl in *.
   destruct (RTL_rtl_initial_states H0).
   destruct H2.
   rewrite <- (rtl_eq_RTL_state_behaves (H3 true)) in H1.
   econstructor; eauto.
  eright.
  intros.
  intro.
  destruct (rtl_RTL_initial_states H1).
  destruct H2.
  eapply H0.
  eassumption.
Qed.

Theorem RTL_eq_rtl : forall fs, cm fs <> None -> forall ar h beh,
  Behaviors.program_behaves (RTL.module_sem m _ ge) (fs, (ar, h)) beh <->
  comp.csemantics cm fs ar h (unify.beh_of_compcert_beh beh).
Proof.
  intros.
  rewrite <- (RTL'_eq_RTL H).
  rewrite <- rtl'_eq_RTL.
  split.
   inversion 1; subst; simpl in *.
   inversion H1; subst.
   econstructor.
   eassumption.
   simpl.
   reflexivity.
   exact (unify.config_behaviors_state_behaves H2).
   exfalso.
   case_eq (module_pseudofuns m fs).
    destruct fs.
    intros.
    edestruct H1.
    econstructor.
    eassumption.
   assumption.
   inversion 1; subst; simpl in *.
   destruct fs.
   eleft with (s := proc.Config (lang := rtl) _ _ _).
   econstructor.
   eassumption.
   injection Hinit.
   intros; subst.
   exact (let (_, K) := unify.config_behaviors_eq_state_behaves (S := rtl') _ _ in K Hbeh).
Qed.

End Module.

Fixpoint module_dom (m: module) (idsig: ident*signature) : bool :=
  match m with
    | nil => false
    | (i, f) :: m' => if idsig_eq_dec (i, pfunsig f) idsig then true else module_dom m' idsig
  end.

Lemma eq_dec_true : forall (A: Type) (f: forall a b: A, {a = b} + {a <> b}) (B: Type) (x y:B) a, (if f a a then x else y) = x.
Proof.
  intros.
  destruct (f a a).
   trivial.
  destruct n.
  trivial.
Qed.

Lemma eq_dec_false: forall (A: Type) a b (H: a <> b) (f: forall a b: A, {a = b} + {a <> b}) (B: Type) (x y:B), (if f a b then x else y) = y.
Proof.
  intros.
  destruct (f a b).
  contradiction.
  trivial.
Qed.

Lemma eq_dec_false': forall (A: Type) a b (H: b <> a) (f: forall a b: A, {a = b} + {a <> b}) (B: Type) (x y:B), (if f a b then x else y) = y.
Proof.
  intros.
  assert (a <> b) by congruence.
  eauto using eq_dec_false.
Qed.

Lemma module_dom_dom : forall m, module_dom m = map.dom (module_pseudofuns m).
Proof.
  intros.
  apply FunctionalExtensionality.functional_extensionality.
  unfold map.dom.
  induction m; simpl.
   trivial.
  intro.
  destruct a.
  destruct (idsig_eq_dec (i, pfunsig p) x).
   injection e; intros; subst.
   destruct p.
   simpl.
   rewrite eq_dec_true.
   trivial.
   simpl.
   rewrite eq_dec_true.
   trivial.
  destruct p.
   rewrite (eq_dec_false n).
   auto.
  simpl in *.
  rewrite (eq_dec_false n).
  auto.
Qed.

Lemma module_dom_join : forall m1 m2,
  map.disjoint (module_pseudofuns m1) (module_pseudofuns m2) ->
  module_pseudofuns (m1 ++ m2) = map.merge (module_pseudofuns m1) (module_pseudofuns m2).
Proof.
  intros.
  assert (forall x, module_dom m1 x = true -> forall z, module_pseudofuns m2 x = Some z -> False).
   intros.
   rewrite module_dom_dom in H0.
   generalize (map.dom_in_dom H0).
   generalize (map.lookup_in_dom H1).
   intros; eapply H; eauto.
  clear H.
  unfold map.merge.
  apply FunctionalExtensionality.functional_extensionality.
  revert m2 H0.
  induction m1; simpl.
   trivial.
  intros.
  destruct a.
  destruct p; simpl.
   destruct (idsig_eq_dec (i, fn_sig f)).
    trivial.
   eapply IHm1.
   intros.
   eapply H0.
   2: eassumption.
   destruct (idsig_eq_dec (i, pfunsig (Internal f)) x0); trivial.
  destruct (idsig_eq_dec (i, pef_sig p) x).
   trivial.
  eapply IHm1.
  intros.
  eapply H0.
  2: eassumption.
  destruct (idsig_eq_dec (i, pfunsig (External p)) x0); trivial.
Qed.

Section Cse.

Require CSEproof.

Common subexpression elimination is correct and compositional.


Theorem cse_correct_compositional : forall m m1 (Hdisjoint: map.disjoint (module_pseudofuns m) (module_pseudofuns m1))
  m2 (Hm2: RTL.transf_module CSE.transf_function m1 = Errors.OK m2)
  fs (Hfs: map.in_dom (module_pseudofuns (m ++ m1)) fs) ah,
  unify.compcert_refinement (Behaviors.program_behaves (module_sem (m ++ m2) _ ge) (fs, ah)) (Behaviors.program_behaves (module_sem (m ++ m1) _ ge) (fs, ah)).
Proof.
  intros.
  assert (map.dom (module_pseudofuns m1) = map.dom (module_pseudofuns m2)).
   unfold map.dom.
   apply FunctionalExtensionality.functional_extensionality.
   intros.
   refine (let K := _ in _ (RTL.transf_module_pseudofuns_some _ K _ _ Hm2) (RTL.transf_module_pseudofuns_none _ K _ _ Hm2)).
    unfold CSE.transf_function.
    intro.
    destruct (CSE.analyze f).
     injection 1; intros; subst; reflexivity.
     discriminate.
   intros.
   case_eq (module_pseudofuns m1 x).
    intros.
    destruct (x0 _ _ H).
    destruct H0.
    rewrite H1.
    trivial.
   intro.
   rewrite (x1 _ H).
   trivial.
  assert (map.disjoint (module_pseudofuns m) (module_pseudofuns m2)).
   eapply map.disjoint_same_dom.
   eassumption.
   assumption.
  assert (map.in_dom (module_pseudofuns (m ++ m2)) fs).
   rewrite (module_dom_join H0).
   apply map.dom_in_dom.
   rewrite map.dom_merge.
   rewrite <- H.
   rewrite <- map.dom_merge.
   rewrite <- (module_dom_join Hdisjoint).
   apply map.in_dom_dom.
   assumption.
  destruct ah.
  eapply unify.compcert_improves_of_improves_set.
  2: eapply RTL_eq_rtl.
  3: eapply RTL_eq_rtl.
  rewrite (module_dom_join H0).
  rewrite (module_dom_join Hdisjoint).
  rewrite (link.csemantics_bsemantics).
  2: assumption.
  rewrite (link.csemantics_bsemantics).
  2: assumption.
  refine (_ (refinement.refines_compositionality (ccimproves.compcert_refinement_is_refinement sw) (o := map.create_map (map.dom (module_pseudofuns m)) (comp.csemantics (lang := rtl) (module_pseudofuns m))) (o1 := map.create_map (map.dom (module_pseudofuns m2)) (comp.csemantics (lang := rtl) (module_pseudofuns m2))) _ _ (o2 := map.create_map (map.dom (module_pseudofuns m1)) (comp.csemantics (lang := rtl) (module_pseudofuns m1))) _ _)).
  destruct 1.
  clear H2.
  eapply H3.
  unfold resolution.link.
  eapply map.create_map_lookup_intro.
  rewrite map.dom_merge.
  rewrite map.create_map_dom.
  rewrite map.create_map_dom.
  rewrite <- map.dom_merge.
  apply map.in_dom_dom.
  rewrite <- (module_dom_join H0).
  assumption.
  eapply map.create_map_lookup_intro.
  rewrite map.dom_merge.
  rewrite map.create_map_dom.
  rewrite map.create_map_dom.
  rewrite <- map.dom_merge.
  apply map.in_dom_dom.
  rewrite <- (module_dom_join Hdisjoint).
  assumption.
  intro.
  intros.
  generalize (map.in_dom_dom H3).
  rewrite map.create_map_dom.
  generalize (map.in_dom_dom H2).
  rewrite map.create_map_dom.
  unfold map.disjoint in H1.
  intros; eauto using map.dom_in_dom.
 intro.
 intros.
 generalize (map.create_map_lookup_elim H2).
 intro; subst.
 generalize (map.in_dom_dom (map.lookup_in_dom H2)).
 rewrite map.create_map_dom.
 intro.
 destruct (map.dom_in_dom H3).
 case_eq (proc.callee_init (l:=rtl) a h x).
 intros.
 destruct (smallstep_exists_beh.exists_behavior (@proc.config_final _ _) (comp.cstep (lang := rtl) (module_pseudofuns m)) (proc.Config h0 l0 nil)).
 esplit.
 econstructor.
 eassumption.
 eassumption.
 eassumption.
 intro.
 intros.
 generalize (map.create_map_lookup_elim H2).
 intro; subst.
 generalize (map.in_dom_dom (map.lookup_in_dom H2)).
 rewrite map.create_map_dom.
 intro.
 destruct (map.dom_in_dom H3).
 case_eq (proc.callee_init (l:=rtl) a h x).
 intros.
 destruct (smallstep_exists_beh.exists_behavior (@proc.config_final _ _) (comp.cstep (lang := rtl) (module_pseudofuns m2)) (proc.Config h0 l0 nil)).
 esplit.
 econstructor.
 eassumption.
 eassumption.
 eassumption.
split.
 rewrite map.create_map_dom.
 rewrite map.create_map_dom.
 auto.
intros.
generalize (map.create_map_lookup_elim H2).
intro; subst.
generalize (map.create_map_lookup_elim H3).
intro; subst.
generalize (map.in_dom_dom (map.lookup_in_dom H3)).
rewrite map.create_map_dom.
intro.
destruct (map.dom_in_dom H4).
rewrite H in H4.
destruct (map.dom_in_dom H4).
eapply unify.improves_of_compcert_improves_set.
2: eapply RTL_eq_rtl.
3: eapply RTL_eq_rtl.
intro.
intros.
eapply CSEproof.transf_module_correct.
eassumption.
assumption.
unfold map.lookup, map.lookup_strong in *. congruence.
unfold map.lookup, map.lookup_strong in *. congruence.
destruct H1. unfold map.lookup, map.lookup_strong in *. congruence.
destruct Hfs. unfold map.lookup, map.lookup_strong in *. congruence.
Qed.

End Cse.

End GENV.