Require Import Coqlib.
Require Import Maps.
Require Import AST.
Require Import Registers.
Require Import CSSA.
Require Import SSA.
Require Import SSAutils.
Require Import CSSAgen.
Require Import Utils.
Require Import Kildall.
Require Import KildallComp.
Require Import DLib.
Require Import CSSAgenspec.
Require Import CSSAutils.
Require Import CSSAproof.
Require Import Classical.
Unset Allow StrictProp.
Lemma handle_phi_block_spec_plt_arg_last :
forall maxreg preds pc block fs_init fs_last
phicode1 phicode2 parcode1 parcode2,
handle_phi_block_spec maxreg preds pc
block fs_init fs_last
phicode1 phicode2 parcode1 parcode2 ->
forall arg'
args'
dst'
phib',
phicode1 !
pc =
Some nil ->
phicode2 !
pc =
Some phib' ->
In (
Iphi args'
dst')
phib' ->
In arg'
args' ->
Plt arg'
fs_last.
Proof.
intros until parcode2.
intro H;
induction H;
intros.
+
assert (
RW:
phib' =
nil)
by congruence.
rewrite RW in *.
inv H1.
+
inv PBSPECnew_phi.
rewrite PTree.gss in H1.
case_eq (
list_eq_dec peq args'
args'0);
intros.
-
rewrite e in *.
apply Plt_Ple_trans with fs_next;
auto.
eapply gen_new_regs_spec_max_in;
eauto.
eauto.
-
assert (
RWphib':
phib'
=
Iphi args'
fs_init ::
phib'0)
by congruence.
rewrite RWphib'
in *.
simpl in H2.
destruct H2;
try congruence.
eauto.
Qed.
Lemma handle_phi_block_spec_plt_dst_last :
forall maxreg preds pc block fs_init fs_last
phicode1 phicode2 parcode1 parcode2,
handle_phi_block_spec maxreg preds pc
block fs_init fs_last
phicode1 phicode2 parcode1 parcode2 ->
forall args'
dst'
phib',
phicode1 !
pc =
Some nil ->
phicode2 !
pc =
Some phib' ->
In (
Iphi args'
dst')
phib' ->
Plt dst'
fs_last.
Proof.
intros until parcode2.
intro H;
induction H;
intros.
+
assert (
RW:
phib' =
nil)
by congruence.
rewrite RW in *.
inv H1.
+
inv PBSPECnew_phi.
rewrite PTree.gss in H1.
case_eq (
peq fs_init dst'0);
intros.
-
rewrite e in *.
apply Plt_Ple_trans with (
Pos.succ dst'0).
apply Plt_succ.
apply Ple_trans with fs_next;
eauto.
-
assert (
RWphib':
phib'
=
Iphi args'
fs_init ::
phib'0)
by congruence.
rewrite RWphib'
in *.
simpl in H2.
destruct H2;
try congruence.
eauto.
Qed.
Variant interval_freshness (
phicode :
phicode) (
pc :
node) (
r1 r2 :
positive) :
Prop :=
|
interval_freshness_intros:
forall phib,
phicode !
pc =
Some phib ->
Ple r1 r2 ->
(
forall args dst arg,
In (
Iphi args dst)
phib ->
In arg (
dst ::
args) ->
Ple r1 arg /\
Ple arg r2) ->
interval_freshness phicode pc r1 r2.
Variant disjoint_phis (
phicode :
phicode)
(
pc pc' :
node) :
Prop :=
|
disjoint_phis_intros:
forall phib phib',
phicode !
pc =
Some phib ->
phicode !
pc' =
Some phib' ->
(
forall args args'
dst dst'
arg arg',
In (
Iphi args dst)
phib ->
In arg (
dst ::
args) ->
In (
Iphi args'
dst')
phib' ->
In arg' (
dst' ::
args') ->
arg <>
arg') ->
disjoint_phis phicode pc pc'.
Lemma disjoint_interval_disjoint :
forall r1 r2 r3 r4 pc pc'
phicode,
Plt r2 r3 ->
interval_freshness phicode pc r1 r2 ->
interval_freshness phicode pc'
r3 r4 ->
disjoint_phis phicode pc pc'.
Proof.
intros.
inv H0.
inv H1.
econstructor;
eauto;
intros.
assert(
Ple arg r2).
exploit H4;
go.
tauto.
assert(
Ple r3 arg').
exploit H6;
go.
tauto.
assert(
Plt arg arg').
apply Ple_Plt_trans with (
q :=
r2).
auto.
apply Plt_Ple_trans with (
q :=
r3).
auto.
auto.
go.
Qed.
Lemma disjoint_interval_comm :
forall phicode pc pc',
disjoint_phis phicode pc pc' ->
disjoint_phis phicode pc'
pc.
Proof.
intros.
inv H.
econstructor; go.
intros.
assert(arg' <> arg).
eapply H2; go.
congruence.
Qed.
Inductive chained_intervals :
list node ->
phicode ->
positive ->
positive ->
Prop :=
|
chained_intervals_nil:
forall phicode init last,
Ple init last ->
chained_intervals nil phicode init last
|
chained_intervals_cons:
forall l pc phicode1 phicode2 init last new_init next,
interval_freshness phicode1 pc new_init next ->
chained_intervals l phicode2 init last ->
phicode1 !
pc =
phicode2 !
pc ->
Plt next init ->
chained_intervals (
pc ::
l)
phicode2 new_init last.
Lemma chained_intervals_ple :
forall l phicode init last,
chained_intervals l phicode init last ->
Ple init last.
Proof.
intros.
induction H.
auto.
apply Ple_trans with (
q :=
init).
apply Ple_trans with (
q :=
next).
inv H;
auto.
apply Plt_Ple;
auto.
auto.
Qed.
Lemma chained_intervals_interval :
forall l phicode init last,
chained_intervals l phicode init last ->
forall pc,
In pc l ->
exists r1 r2,
interval_freshness phicode pc r1 r2 /\
Ple init r1 /\
Ple r2 last.
Proof.
intros until last.
intro H.
induction H;
intros.
go.
inv H3.
+
exists new_init.
exists next.
split;
auto.
inv H;
go.
econstructor;
go.
congruence.
split;
auto.
inv H.
apply Ple_refl.
apply Ple_trans with (
q :=
init).
apply Plt_Ple;
auto.
eapply chained_intervals_ple;
eauto.
+
exploit IHchained_intervals;
eauto.
intros Hinterval.
destruct Hinterval as [
r1 [
r2 Hinterval]].
destruct Hinterval as [
Hinterval [
Hpleinit Hplelast]].
exists r1.
exists r2.
split;
auto.
split;
auto.
apply Ple_trans with (
q :=
init).
apply Ple_trans with (
q :=
next).
inv H;
auto.
apply Plt_Ple;
auto.
auto.
Qed.
Lemma chained_intervals_left_prolonged :
forall l phicode init last new_init,
Ple new_init init ->
chained_intervals l phicode init last ->
chained_intervals l phicode new_init last.
Proof.
intros.
induction H0.
constructor.
apply Ple_trans with (
q :=
init);
auto.
econstructor;
eauto.
inv H0.
econstructor;
eauto.
apply Ple_trans with (
q :=
new_init0);
auto.
intros.
exploit H6;
eauto;
intros HPles.
destruct HPles.
split.
apply Ple_trans with (
q :=
new_init0);
auto.
auto.
Qed.
Inductive chained_disjoint_phis (
phicode :
phicode) :
list node ->
Prop :=
|
chained_disjoint_phis_nil:
chained_disjoint_phis phicode nil
|
chained_disjoint_phis_cons:
forall l pc,
chained_disjoint_phis phicode l ->
(
forall pc',
In pc'
l ->
disjoint_phis phicode pc pc') ->
chained_disjoint_phis phicode (
pc ::
l).
Lemma chained_intervals_disjoint :
forall l phicode init last,
chained_intervals l phicode init last ->
chained_disjoint_phis phicode l.
Proof.
intros.
induction H.
+
go.
+
constructor.
auto.
intros.
exploit chained_intervals_interval;
eauto.
intros Hinterval.
destruct Hinterval as [
r1 [
r2 Hinterval]].
destruct Hinterval as [
Hinterval [
Hpleinit Hplelast]].
assert(
interval_freshness phicode2 pc new_init next).
inv H.
econstructor;
eauto.
congruence.
clear H.
apply disjoint_interval_disjoint
with (
r1 :=
new_init) (
r2 :=
next)
(
r3 :=
r1) (
r4 :=
r2);
auto.
apply Plt_Ple_trans with (
q :=
init);
auto.
Qed.
Lemma chained_disjoint_phis_correct :
forall l phicode,
chained_disjoint_phis phicode l ->
forall pc pc',
In pc l ->
In pc'
l ->
pc <>
pc' ->
disjoint_phis phicode pc pc'.
Proof.
intros until phicode.
intro H.
induction H;
intros.
+
inv H.
+
case_eq(
peq pc0 pc);
case_eq(
peq pc'
pc);
go;
intros.
rewrite e in *.
inv H2;
auto.
congruence.
rewrite e in *.
inv H1;
auto.
congruence.
apply disjoint_interval_comm.
eapply H0;
eauto.
inv H1;
go.
inv H2;
go.
Qed.
Definition has_really_phiblock f pc :=
match (
fn_phicode f) !
pc with
|
None =>
false
|
Some phib =>
match phib with
|
nil =>
false
|
_ =>
true
end
end.
Lemma Plt_Ple_succ :
forall x y,
Plt x (
Pos.succ y) ->
Ple x y.
Proof.
Lemma mfold_copy_node_other_preserves_phicode :
forall f pc l s u s'
INCR,
wf_ssa_function f ->
normalized_jp f ->
~
In pc l ->
NoDup l ->
Plt (
get_maxreg f) (
next_fresh_reg s) ->
mfold_unit
(
copy_node (
make_predecessors (
fn_code f)
successors_instr)
(
fn_code f) (
fn_phicode f))
l s =
OK u s'
INCR ->
(
st_phicode s) !
pc = (
st_phicode s') !
pc.
Proof.
Lemma copy_node_nophib_preserves_phicode :
forall f pc s u s'
INCR,
wf_ssa_function f ->
normalized_jp f ->
copy_node (
make_predecessors (
fn_code f)
successors_instr)
(
fn_code f) (
fn_phicode f)
pc s =
OK u s'
INCR ->
(
fn_phicode f) !
pc =
None ->
(
st_phicode s) !
pc = (
st_phicode s') !
pc.
Proof.
intros.
unfold copy_node in H1.
flatten H1.
unfold ret in H1.
go.
Qed.
Lemma mfold_copy_node_nophib_preserves_phicode :
forall f pc l s u s'
INCR,
wf_ssa_function f ->
normalized_jp f ->
NoDup l ->
Plt (
get_maxreg f) (
next_fresh_reg s) ->
mfold_unit
(
copy_node (
make_predecessors (
fn_code f)
successors_instr)
(
fn_code f) (
fn_phicode f))
l s =
OK u s'
INCR ->
(
fn_phicode f) !
pc =
None ->
(
st_phicode s) !
pc = (
st_phicode s') !
pc.
Proof.
Lemma handle_phi_real_block_does_incr :
forall l phi phib pc s x s'
INCR,
handle_phi_block l (
phi ::
phib)
pc s =
OK x s'
INCR ->
Plt (
next_fresh_reg s) (
next_fresh_reg s').
Proof.
Lemma mfold_copy_node_correct_more :
forall f,
wf_ssa_function f ->
normalized_jp f ->
forall l
u s s'
incr,
mfold_unit (
copy_node
(
make_predecessors (
fn_code f)
successors_instr)
(
fn_code f) (
fn_phicode f))
l s =
OK u s'
incr ->
Plt (
get_maxreg f) (
next_fresh_reg s) ->
NoDup l ->
chained_intervals (
filter (
has_really_phiblock f)
l) (
st_phicode s')
(
next_fresh_reg s) (
next_fresh_reg s').
Proof.
Lemma check_nodup_in_reglist_correct_aux :
forall l visited visited'
r,
check_nodup_in_reglist l visited = (
true,
visited') ->
SSARegSet.In r visited ->
~
In r l.
Proof.
induction l;
intros.
go.
simpl in *.
flatten H.
unfold not;
intros Hinornot.
destruct Hinornot.
rewrite H1 in *.
exploit SSARegSet.mem_1;
go.
go.
exploit IHl;
go.
unfold reg_use.
apply SSARegSet.MF.add_2;
auto.
Qed.
Lemma check_nodup_in_reglist_correct :
forall l visited visited',
check_nodup_in_reglist l visited = (
true,
visited') ->
NoDup l.
Proof.
Lemma check_nodup_in_reglist_correct_2aux :
forall l visited visited'
r,
check_nodup_in_reglist l visited = (
true,
visited') ->
SSARegSet.In r visited ->
SSARegSet.In r visited'.
Proof.
induction l;
intros;
simpl in *.
go.
flatten H.
eapply IHl;
eauto.
apply SSARegSet.MF.add_2;
auto.
Qed.
Lemma check_nodup_in_reglist_correct_2 :
forall l visited visited'
r,
check_nodup_in_reglist l visited = (
true,
visited') ->
In r l ->
SSARegSet.In r visited'.
Proof.
Lemma check_nodup_in_reglist_correct_3 :
forall l visited visited'
r,
check_nodup_in_reglist l visited = (
true,
visited') ->
SSARegSet.In r visited' ->
~
SSARegSet.In r visited ->
In r l.
Proof.
induction l;
intros;
simpl in *.
go.
flatten H.
case_eq (
peq a r);
intros.
+
left;
auto.
+
right.
eapply IHl;
eauto.
unfold not;
intros.
exploit SSARegSet.add_3;
eauto.
Qed.
Fixpoint check_nodup_in_phib phib visited :=
match phib with
|
nil =>
true
|
Iphi args dst ::
phib' =>
let '(
test,
visited') :=
check_nodup_in_reglist (
dst ::
args)
visited in
if test
then check_nodup_in_phib phib'
visited'
else false
end.
Lemma check_nodup_in_reglist_false :
forall l visited visited'
t,
(
forall r :
SSARegSet.elt,
SSARegSet.In r visited ->
SSARegSet.In r visited') ->
check_nodup_in_reglist l visited = (
false,
t) ->
exists t',
check_nodup_in_reglist l visited' = (
false,
t').
Proof.
Lemma check_nodup_in_phib_preserves_true :
forall phib visited visited',
(
forall r,
SSARegSet.In r visited ->
SSARegSet.In r visited') ->
check_nodup_in_phib phib visited' =
true ->
check_nodup_in_phib phib visited =
true.
Proof.
Lemma check_nodup_in_phib_correct_for_dst :
forall phib visited,
check_nodup_in_phib phib visited =
true ->
forall args dst,
In (
Iphi args dst)
phib ->
~
SSARegSet.In dst visited.
Proof.
Lemma check_nodup_in_phib_correct_for_arg :
forall phib visited,
check_nodup_in_phib phib visited =
true ->
forall args arg dst,
In (
Iphi args dst)
phib ->
In arg args ->
~
SSARegSet.In arg visited.
Proof.
Inductive nodups_in_phib_spec :
phiblock ->
Prop :=
|
nodups_in_phib_spec_nil :
nodups_in_phib_spec nil
|
nodups_in_phib_spec_cons :
forall args dst phib,
(
forall args'
dst'
x y,
In (
Iphi args'
dst')
phib ->
In x (
dst ::
args) ->
In y (
dst' ::
args') ->
x <>
y ) ->
NoDup (
dst ::
args) ->
nodups_in_phib_spec phib ->
nodups_in_phib_spec (
Iphi args dst ::
phib).
Lemma check_nodup_in_phib_correct :
forall phib visited,
check_nodup_in_phib phib visited =
true ->
nodups_in_phib_spec phib.
Proof.
Lemma not_in_dst_and_args :
forall phib,
nodups_in_phib_spec phib ->
forall args dst args'
dst',
In (
Iphi args dst)
phib ->
In (
Iphi args'
dst')
phib ->
~
In dst args'.
Proof.
intros phib H.
induction H;
intros.
inv H.
inv H2;
inv H3.
+
unfold not;
intros.
inv H0.
go.
+
inv H0.
assert(
EQ:
dst0 =
dst)
by go.
rewrite EQ in *.
unfold not;
intros.
exploit (
H args'
dst'
dst dst);
eauto.
+
unfold not;
intros.
assert(
EQ:
args' =
args)
by go.
rewrite EQ in *.
exploit (
H args0 dst0 dst0 dst0);
eauto.
+
go.
Qed.
Lemma no_dups_in_phib_dst :
forall phib,
nodups_in_phib_spec phib ->
NoDup (
map phib_dst phib).
Proof.
intros phib H.
induction H;
intros.
go.
simpl in *.
constructor.
unfold not;
intros.
exploit in_phib_dst_exists_args;
eauto.
intros Hin.
destruct Hin as [
args0 Hin].
exploit (
H args0 dst dst dst);
go.
auto.
Qed.
Lemma check_phicode_for_dups_in_phib_correct :
forall f pc phib,
check_phicode_for_dups_in_phib f =
true ->
(
CSSA.fn_phicode f) !
pc =
Some phib ->
nodups_in_phib_spec phib.
Proof.
Lemma nodups_in_phib_nodups_in_phi :
forall phib,
nodups_in_phib_spec phib ->
forall args dst,
In (
Iphi args dst)
phib ->
NoDup (
dst ::
args).
Proof.
intros phib H.
induction H; go.
intros.
inv H2; go.
Qed.
Lemma exists_phib_iff :
forall f tf pc,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
((
fn_phicode f) !
pc <>
None
<->
(
CSSA.fn_phicode tf) !
pc <>
None).
Proof.
Lemma exists_phib_iff_none :
forall f tf pc,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
((
fn_phicode f) !
pc =
None
<->
(
CSSA.fn_phicode tf) !
pc =
None).
Proof.
intros.
split;
intros;
exploit (
exists_phib_iff f tf pc);
eauto;
intros;
tauto.
Qed.
Definition get_tf s f :=
{|
CSSA.fn_sig :=
fn_sig f;
CSSA.fn_params :=
fn_params f;
CSSA.fn_stacksize :=
fn_stacksize f;
CSSA.fn_code :=
fn_code f;
CSSA.fn_phicode :=
st_phicode s;
fn_parcopycode :=
st_parcopycode s;
CSSA.fn_entrypoint :=
fn_entrypoint f |}.
Lemma cfgeq_to_star :
forall f tf pc pc',
(
forall x y,
cfg f x y ->
CSSA.cfg tf x y) ->
(
cfg f **)
pc pc' ->
(
CSSA.cfg tf **)
pc pc'.
Proof.
intros.
induction H0.
constructor. auto.
constructor 2.
constructor 3 with (y := y); auto.
Qed.
Lemma cfgeq_reached :
forall f tf pc pc',
(
forall x y,
cfg f x y ->
CSSA.cfg tf x y) ->
Dom.reached (
cfg f)
pc pc' ->
Dom.reached (
CSSA.cfg tf)
pc pc'.
Proof.
intros.
inv H0.
constructor.
auto.
constructor 2.
constructor 3
with (
y :=
y).
eapply cfgeq_to_star;
eauto.
eapply cfgeq_to_star;
eauto.
Qed.
Lemma check_nb_args_correct :
forall f,
check_nb_args f
(
CSSA.get_preds f)
=
true ->
CSSA.block_nb_args f.
Proof.
Lemma check_parcbSome_correct :
forall f phib pc pred,
check_parcbSome f
(
CSSA.get_preds f)
=
true ->
(
CSSA.fn_phicode f) !
pc =
Some phib ->
In pred (
CSSA.get_preds f) !!!
pc ->
(
CSSA.fn_parcopycode f) !
pred <>
None.
Proof.
Lemma check_parcb'
Some_correct :
forall f phib pc,
check_parcb'
Some f =
true ->
(
CSSA.fn_phicode f) !
pc =
Some phib ->
(
CSSA.fn_parcopycode f) !
pc <>
None.
Proof.
Lemma check_fn_parcb_inop_correct :
forall f pc parcb,
check_fn_parcb_inop f =
true ->
(
CSSA.fn_parcopycode f) !
pc =
Some parcb ->
exists s, (
CSSA.fn_code f) !
pc =
Some (
Inop s).
Proof.
Lemma wf_fn_phicode_inv :
forall f tf jp,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
(
CSSA.join_point jp tf <-> (
CSSA.fn_phicode tf) !
jp <>
None).
Proof.
Lemma check_fn_parcbjp_correct :
forall f tf pc parcb pc',
wf_ssa_function f ->
normalized_jp f ->
transl_function f =
Errors.OK tf ->
check_fn_parcbjp tf =
true ->
(
CSSA.fn_parcopycode tf) !
pc =
Some parcb ->
(
CSSA.fn_code tf) !
pc =
Some (
Inop pc') ->
~
CSSA.join_point pc'
tf ->
CSSA.join_point pc tf.
Proof.
Lemma cssa_fn_inop_in_jp:
forall f tf,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
forall pc,
(
CSSA.fn_phicode tf) !
pc <>
None ->
exists succ, (
CSSA.fn_code tf) !
pc =
Some (
Inop succ).
Proof.
Lemma equiv_phib_phibarg_parcdst :
forall phib maxreg parcb parcb'
phib'
k,
equiv_phib_spec maxreg k phib parcb phib'
parcb' ->
forall args dst arg,
In (
Iphi args dst)
phib' ->
nth_error args k =
Some arg ->
exists src,
In (
Iparcopy src arg)
parcb.
Proof.
intros until k.
intros H. induction H; intros.
go.
inv H9.
+ assert(EQ1: args0 = args') by congruence.
assert(EQ2: dst0 = dst') by congruence.
rewrite EQ1 in *.
rewrite EQ2 in *.
assert(EQ3: arg0 = arg') by go.
rewrite EQ3 in *.
exists arg; auto.
+ exploit IHequiv_phib_spec; eauto.
intros Hexists.
destruct Hexists as [src Hin].
exists src; go.
Qed.
Lemma equiv_phib_parcb'
src_phibdst' :
forall phib maxreg parcb parcb'
phib'
k,
equiv_phib_spec maxreg k phib parcb phib'
parcb' ->
forall src dst,
In (
Iparcopy src dst)
parcb' ->
exists args,
In (
Iphi args src)
phib'.
Proof.
intros until k.
intros H. induction H; intros.
go.
inv H9.
+ go.
+ exploit IHequiv_phib_spec; eauto.
intros Hexists.
destruct Hexists as [args0 Hin].
exists args0; go.
Qed.
Lemma equiv_phib_phibdst_parcb'
src :
forall phib maxreg parcb parcb'
phib'
k,
equiv_phib_spec maxreg k phib parcb phib'
parcb' ->
forall args dst,
In (
Iphi args dst)
phib ->
exists src,
In (
Iparcopy src dst)
parcb'.
Proof.
intros until k.
intros H. induction H; intros.
go.
inv H9.
+ go.
+ exploit IHequiv_phib_spec; eauto.
intros Hexists.
destruct Hexists as [src Hin].
exists src; go.
Qed.
Lemma phib_nil_phib'
_nil:
forall f tf pc,
normalized_jp f ->
wf_ssa_function f ->
transl_function f =
Errors.OK tf ->
(
fn_phicode f) !
pc =
Some nil ->
(
CSSA.fn_phicode tf) !
pc <>
Some nil ->
False.
Proof.
Lemma In_has_really_phiblock_filter:
forall f tf pc phib r args,
normalized_jp f ->
wf_ssa_function f ->
transl_function f =
Errors.OK tf ->
(
CSSA.fn_phicode tf) !
pc =
Some phib ->
In (
Iphi args r)
phib ->
In pc (
filter (
has_really_phiblock f)
(
map fst (
PTree.elements (
fn_code f)))).
Proof.
Lemma use_code_plemaxreg:
forall f r pc,
use_code f r pc ->
wf_ssa_function f ->
Ple r (
get_maxreg f).
Proof.
Lemma assigned_code_plemaxreg:
forall f r pc,
assigned_code_spec (
fn_code f)
pc r ->
wf_ssa_function f ->
Ple r (
get_maxreg f).
Proof.
Lemma assigned_phi_pltmaxreg_r:
forall f tf r pc,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
assigned_phi_spec (
CSSA.fn_phicode tf)
pc r ->
wf_ssa_function f ->
Plt (
get_maxreg f)
r.
Proof.
Lemma assigned_parc_pltmaxreg_notjp:
forall f tf r pc,
transl_function f =
Errors.OK tf ->
wf_ssa_function f ->
normalized_jp f ->
check_fn_parcb_inop tf =
true ->
check_fn_parcbjp tf =
true ->
assigned_parcopy_spec (
CSSA.fn_parcopycode tf)
pc r ->
~
join_point pc f ->
Plt (
get_maxreg f)
r.
Proof.
Lemma parcb'
dst_is_phibdst :
forall maxreg k phib parcb phib'
parcb',
equiv_phib_spec maxreg k phib parcb phib'
parcb' ->
forall src dst,
In (
Iparcopy src dst)
parcb' ->
exists args,
In (
Iphi args dst)
phib.
Proof.
intros until parcb'.
intro H.
induction H; go.
intros src dst0 Hin.
inv Hin.
assert(EQ: dst0 = dst) by go.
rewrite EQ in *.
exists args. go.
exploit IHequiv_phib_spec; eauto.
intros Hin.
destruct Hin as [args0 Hin].
exists args0. go.
Qed.
Lemma parcbsrc_is_phibarg :
forall maxreg k phib parcb phib'
parcb',
equiv_phib_spec maxreg k phib parcb phib'
parcb' ->
forall src dst,
In (
Iparcopy src dst)
parcb ->
exists args dst',
In (
Iphi args dst')
phib /\
nth_error args k =
Some src.
Proof.
intros until parcb'.
intro H.
induction H.
intros. inv H.
intros src dst0 Hin.
inv Hin.
go.
exploit (IHequiv_phib_spec src dst0); eauto.
intros Hexists.
destruct Hexists as [args0 Hexists].
destruct Hexists as [dst'0 Hexists].
exists args0. exists dst'0.
destruct Hexists.
split; go.
Qed.
Lemma assigned_parcjp_assigned_phi:
forall f tf r pc,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
assigned_parcopy_spec (
CSSA.fn_parcopycode tf)
pc r ->
join_point pc f ->
assigned_phi_spec (
fn_phicode f)
pc r.
Proof.
Lemma check_parcborparcb'
_correct:
forall f tf pc,
wf_ssa_function f ->
normalized_jp f ->
transl_function f =
Errors.OK tf ->
check_parcborparcb'
tf =
true ->
(
CSSA.fn_parcopycode tf) !
pc <>
None ->
exists pc',
(
fn_code f) !
pc =
Some (
Inop pc') /\
((
CSSA.fn_phicode tf) !
pc <>
None \/
(
CSSA.fn_phicode tf) !
pc' <>
None).
Proof.
intros until pc.
intros WF Hnorm Htransl Hcheck Hparcb.
unfold check_parcborparcb'
in Hcheck.
rewrite forallb_forall in Hcheck.
case_eq((
fn_parcopycode tf) !
pc);
try congruence;
intros.
specialize (
Hcheck (
pc,
p)).
exploit Hcheck;
eauto.
apply PTree.elements_correct;
auto.
intros Hcheck2.
erewrite instructions_preserved in Hcheck2;
eauto.
clear Hcheck.
flatten Hcheck2;
go.
Qed.
Lemma used_parcnotjp_use_phi:
forall f tf r pc,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
check_parcborparcb'
tf =
true ->
wf_ssa_function f ->
use_parcopycode tf r pc ->
~
join_point pc f ->
use_phicode f r pc.
Proof.
Lemma NoDupdstphib_NoDupphib:
forall phib,
NoDup (
map phib_dst phib) ->
NoDup phib.
Proof.
induction phib;
intros.
go.
simpl in *.
inv H.
constructor;
go.
unfold not;
intros.
contradict H2.
destruct a.
eapply in_phib_dst_in;
eauto.
Qed.
Lemma NoDupdstparcb_NoDupparcb:
forall parcb,
NoDup (
map parc_dst parcb) ->
NoDup parcb.
Proof.
induction parcb;
intros.
go.
simpl in *.
inv H.
constructor;
go.
unfold not;
intros.
contradict H2.
destruct a.
eapply parc_dst_in;
eauto.
Qed.
Lemma twoargs_not_NoDup:
forall phib args args'
r,
In (
Iphi args r)
phib ->
In (
Iphi args'
r)
phib ->
args <>
args' ->
~
NoDup (
map phib_dst phib).
Proof.
induction phib;
intros.
go.
simpl in *.
destruct H;
destruct H0;
go;
unfold not;
intros Hnodup.
+
inv Hnodup.
simpl in *.
contradict H4.
eapply in_phib_dst_in;
eauto.
+
inv Hnodup.
simpl in *.
contradict H4.
eapply in_phib_dst_in;
eauto.
+
inv Hnodup.
contradict H5;
go.
Qed.
Lemma twoparc_not_NoDup:
forall parcb src src'
r,
In (
Iparcopy src r)
parcb ->
In (
Iparcopy src'
r)
parcb ->
src <>
src' ->
~
NoDup (
map parc_dst parcb).
Proof.
induction parcb;
intros.
go.
simpl in *.
destruct H;
destruct H0;
go;
unfold not;
intros Hnodup.
+
inv Hnodup.
simpl in *.
contradict H4.
eapply parc_dst_in;
eauto.
+
inv Hnodup.
simpl in *.
contradict H4.
eapply parc_dst_in;
eauto.
+
inv Hnodup.
contradict H5;
go.
Qed.
Lemma phib_plemaxreg :
forall f tf r pc,
transl_function f =
Errors.OK tf ->
wf_ssa_function f ->
assigned_phi_spec (
fn_phicode f)
pc r ->
Ple r (
get_maxreg f).
Proof.
Lemma really_parcb_really_phib':
forall maxreg k phib parcb phib'
parcb',
equiv_phib maxreg k phib parcb phib'
parcb' ->
forall src dst,
In (
Iparcopy src dst)
parcb ->
exists args dst',
nth_error args k =
Some dst /\
In (
Iphi args dst')
phib'.
Proof.
intros until parcb'.
intro H.
induction H; intros.
inv H.
inv H15.
+ exists args'. exists dst'. split; go.
+ exploit IHequiv_phib; eauto. intros Hargs.
destruct Hargs as [args0 [dst1 Hin]].
destruct Hin.
exists args0. exists dst1. split; go.
Qed.
Lemma nodup_phib_dst_twoins :
forall phib args args'
r,
NoDup (
map phib_dst phib) ->
In (
Iphi args r)
phib ->
In (
Iphi args'
r)
phib ->
args =
args'.
Proof.
induction phib;
intros.
inv H0.
simpl in *.
destruct H0;
destruct H1.
+
go.
+
inv H.
simpl in *.
contradict H4.
eapply in_phib_dst_in;
eauto.
+
inv H.
simpl in *.
contradict H4.
eapply in_phib_dst_in;
eauto.
+
inv H.
go.
Qed.
Lemma nodups_in_phib_nodup_in_args :
forall phib,
nodups_in_phib_spec phib ->
forall r args args'
dst dst',
In r args ->
In r args' ->
In (
Iphi args dst)
phib ->
In (
Iphi args'
dst')
phib ->
args =
args' /\
dst =
dst'.
Proof.
intros phib H.
induction H; intros; eauto.
inv H1.
inv H4; inv H5.
+ go.
+ replace args with args0 in *; go.
replace dst with dst0 in *; go.
exploit (H args' dst' r r); eauto.
contradiction.
+ replace args with args' in *; go.
replace dst with dst' in *; go.
exploit (H args0 dst0 r r); eauto.
contradiction.
+ go.
Qed.
Lemma Ple_maxreg_and_maxreg_Plt :
forall (
r :
reg)
maxreg,
Ple r maxreg ->
Plt maxreg r ->
False.
Proof.
Lemma notassign_phi_and_parc:
forall f tf r pc pc',
transl_function f =
Errors.OK tf ->
normalized_jp f ->
CSSA.block_nb_args tf ->
check_parcbSome tf (
CSSA.get_preds tf) =
true ->
check_parcb'
Some tf =
true ->
check_fn_parcb_inop tf =
true ->
check_fn_parcbjp tf =
true ->
check_parcborparcb'
tf =
true ->
check_phicode_for_dups_in_phib tf =
true ->
wf_ssa_function f ->
assigned_phi_spec (
CSSA.fn_phicode tf)
pc r ->
assigned_parcopy_spec (
CSSA.fn_parcopycode tf)
pc'
r ->
False.
Proof.
Lemma no_successive_jp :
forall f pc pc',
wf_ssa_function f ->
normalized_jp f ->
join_point pc f ->
join_point pc'
f ->
(
fn_code f) !
pc =
Some (
Inop pc') ->
False.
Proof.
Lemma nodups_in_phib_nodups_in_args :
forall phib,
nodups_in_phib_spec phib ->
forall args dst,
In (
Iphi args dst)
phib ->
NoDup args.
Proof.
intros phib H.
induction H; intros; go.
simpl in *.
destruct H2; go.
inv H0; go.
Qed.
Lemma nodups_args_nth_same_k :
forall args k k' (
r :
reg),
NoDup args ->
nth_error args k =
Some r ->
nth_error args k' =
Some r ->
k =
k'.
Proof.
induction args;
intros.
destruct k;
destruct k';
simpl in *;
unfold Specif.error in *;
congruence.
destruct k;
destruct k';
simpl in *;
unfold Specif.error in *;
unfold value in *;
try congruence.
+
assert(
In r args).
eapply nth_In_reg;
eauto.
inv H;
go.
+
assert(
In r args).
eapply nth_In_reg;
eauto.
inv H;
go.
+
inv H.
assert(
k =
k')
by go.
go.
Qed.
Lemma wf_unique_def :
forall f tf,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
CSSA.block_nb_args tf ->
check_parcbSome tf (
CSSA.get_preds tf) =
true ->
check_parcb'
Some tf =
true ->
check_fn_parcb_inop tf =
true ->
check_fn_parcbjp tf =
true ->
check_parcborparcb'
tf =
true ->
check_phicode_for_dups_in_phib tf =
true ->
wf_ssa_function f ->
CSSA.unique_def_spec tf.
Proof.
Lemma wf_cssa_extrainv_1 :
forall f tf,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
forall pc r,
CSSA.join_point pc tf ->
use_parcopycode tf r pc ->
assigned_phi_spec (
CSSA.fn_phicode tf)
pc r.
Proof.
Lemma wf_cssa_extrainv_2 :
forall f tf,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
forall pc r,
CSSA.use_phicode tf r pc ->
assigned_parcopy_spec (
CSSA.fn_parcopycode tf)
pc r.
Proof.
Lemma assigned_phi_assigned_parcb':
forall f tf,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
forall pc r,
assigned_phi_spec (
fn_phicode f)
pc r ->
assigned_parcopy_spec (
CSSA.fn_parcopycode tf)
pc r.
Proof.
intros f tf Htrans
Hnorm WF.
intros.
exploit transl_function_spec_ok;
eauto.
eapply transl_function_charact;
eauto.
intros SPEC.
assert(
Hremembertrans:
transl_function f =
Errors.OK tf)
by auto.
unfold transl_function in Htrans.
flatten Htrans.
intros.
inv SPEC.
inv H.
simpl in *.
unfold inop_in_jp in H5.
exploit (
H5 pc);
eauto.
congruence.
intros.
destruct H as [
succ Hinop].
exploit (
H2 pc (
Inop succ) 0);
eauto;
intros.
inv H.
-
congruence.
-
intros.
assert(
Hjp:
join_point pc f).
apply fn_phicode_inv;
auto.
congruence.
induction Hjp.
unfold node in *.
assert(
EQ:
lpreds =
l0)
by go.
rewrite EQ in *.
destruct l0;
go.
-
econstructor;
eauto.
destruct H6.
assert(
EQ:
phiinstr =
phib)
by go.
rewrite EQ in *.
eapply equiv_phib_phibdst_parcb'
src;
eauto.
Qed.
Lemma param_ple_maxreg :
forall x f,
In x (
fn_params f) ->
Ple x (
get_maxreg f).
Proof.
Lemma use_phib_ple_maxreg :
forall f r pc,
wf_ssa_function f ->
use_phicode f r pc ->
Ple r (
get_maxreg f).
Proof.
Lemma parcb_src_ple_maxreg :
forall maxreg k phib parcb phib'
parcb',
equiv_phib_spec maxreg k phib parcb phib'
parcb' ->
forall src dst,
In (
Iparcopy src dst)
parcb ->
Ple src maxreg.
Proof.
intros until parcb'.
intros H. induction H; intros.
inv H.
simpl in *.
destruct H9; go.
Qed.
Lemma cssa_def_code_def_code:
forall f tf r pc,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
assigned_code_spec (
CSSA.fn_code tf)
r pc ->
assigned_code_spec (
fn_code f)
r pc.
Proof.
intros.
unfold transl_function in H.
flatten Htrans;
simpl in *.
inv H2;
go.
Qed.
Lemma cssa_use_code_use_code:
forall f tf r pc,
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
CSSA.use_code tf r pc ->
use_code f r pc.
Proof.
intros.
unfold transl_function in H.
flatten Htrans;
simpl in *.
inv H2;
simpl in * ;
go.
Qed.
Lemma cssa_ext_params_ext_params:
forall f tf r,
transl_function f =
Errors.OK tf ->
check_parcborparcb'
tf =
true ->
normalized_jp f ->
wf_ssa_function f ->
CSSA.ext_params tf r ->
ext_params f r.
Proof.
intros f tf r Htrans Hcheckparcborparcb'
Hnorm WF Hext.
inv Hext.
+
unfold transl_function in Htrans.
flatten Htrans.
constructor 1.
auto.
+
destruct H.
constructor 2.
-
inv H.
exploit cssa_use_code_use_code;
eauto.
intros.
exists x.
go.
exploit wf_cssa_extrainv_2;
eauto.
intros.
contradict H2.
eauto.
exploit (
check_parcborparcb'
_correct f tf x);
eauto.
inv H3;
go.
intros Hpc'.
destruct Hpc'
as [
pc' [
Hinop Hphicode]].
destruct Hphicode as [
Hphibpc |
Hphibpc'].
exploit wf_cssa_extrainv_1;
eauto.
eapply wf_fn_phicode_inv;
eauto.
intros.
contradict H.
eauto.
exploit used_parcnotjp_use_phi;
eauto.
unfold not;
intros.
apply no_successive_jp with (
f :=
f) (
pc :=
x)
(
pc' :=
pc');
eauto.
rewrite <-
exists_phib_iff in Hphibpc';
eauto.
apply fn_phicode_inv;
auto.
intros.
exists x.
go.
-
intros.
unfold not;
intros Hassignphi.
exploit assigned_phi_assigned_parcb';
eauto.
intros Hassign.
contradict Hassign.
eauto.
-
intros.
unfold not;
intros Hassigncode.
contradict Hassigncode.
unfold transl_function in Htrans.
flatten Htrans.
eauto.
Qed.
Lemma ext_params_ple_maxreg:
forall f r,
wf_ssa_function f ->
ext_params f r ->
Ple r (
get_maxreg f).
Proof.
Lemma cssa_cfg_cfg:
forall f tf pc pc',
transl_function f =
Errors.OK tf ->
cfg f pc pc' ->
CSSA.cfg tf pc pc'.
Proof.
intros.
unfold transl_function in H.
flatten H;
simpl in *.
inv H0;
go.
Qed.
Lemma cssa_cfg_cfg_2:
forall f tf pc pc',
transl_function f =
Errors.OK tf ->
CSSA.cfg tf pc pc' ->
cfg f pc pc'.
Proof.
intros.
unfold transl_function in H.
flatten H;
simpl in *.
inv H0;
go.
Qed.
Lemma cssa_cfg_star_cfg_star:
forall f tf pc pc',
transl_function f =
Errors.OK tf ->
(
cfg f **)
pc pc' ->
(
CSSA.cfg tf **)
pc pc'.
Proof.
intros.
induction H0.
+
constructor.
eapply cssa_cfg_cfg;
eauto.
+
constructor 2.
+
constructor 3
with (
y :=
y);
auto.
Qed.
Lemma cssa_cfg_star_cfg_star_2:
forall f tf pc pc',
transl_function f =
Errors.OK tf ->
(
CSSA.cfg tf **)
pc pc' ->
(
cfg f **)
pc pc'.
Proof.
intros.
induction H0.
+
constructor.
eapply cssa_cfg_cfg_2;
eauto.
+
constructor 2.
+
constructor 3
with (
y :=
y);
auto.
Qed.
Lemma cssa_reached_reached:
forall f tf pc,
transl_function f =
Errors.OK tf ->
reached f pc ->
CSSA.reached tf pc.
Proof.
Lemma cssa_reached_reached_2:
forall f tf pc,
transl_function f =
Errors.OK tf ->
CSSA.reached tf pc ->
reached f pc.
Proof.
Lemma cssa_path_path :
forall f tf l pc pc',
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
Dom.path (
CSSA.cfg tf) (
CSSA.exit tf) (
fn_entrypoint f)
pc l pc' ->
Dom.path (
cfg f) (
exit f) (
fn_entrypoint f)
pc l pc'.
Proof.
Lemma cssa_dom_dom:
forall f tf pc pc',
transl_function f =
Errors.OK tf ->
normalized_jp f ->
wf_ssa_function f ->
dom f pc pc' ->
CSSA.dom tf pc pc'.
Proof.
Section WF_CSSA.
Variable f :
function.
Hypothesis WF :
wf_ssa_function f.
Variable tf :
CSSA.function.
Hypothesis Htrans :
transl_function f =
Errors.OK tf.
Lemma Hnorm :
normalized_jp f.
Proof.
Hint Resolve Hnorm:
core.
Lemma Hblock_nb_args :
CSSA.block_nb_args tf.
Proof.
Hint Resolve Hblock_nb_args:
core.
Lemma HparcbSome :
forall (
phib :
phiblock) (
pc :
positive),
(
CSSA.fn_phicode tf) !
pc =
Some phib ->
forall pred :
positive,
In pred (
CSSA.get_preds tf) !!!
pc ->
(
fn_parcopycode tf) !
pred <>
None.
Proof.
Hint Resolve HparcbSome:
core.
Lemma Hparcb'
Some :
forall (
phib :
phiblock) (
pc :
positive),
(
CSSA.fn_phicode tf) !
pc =
Some phib ->
(
fn_parcopycode tf) !
pc <>
None.
Proof.
unfold transl_function in * ;
flatten Htrans ;
boolInv.
eapply check_parcb'
Some_correct;
eauto.
Qed.
Hint Resolve Hparcb'
Some:
core.
Lemma p_fn_entry :
exists s :
node,
(
CSSA.fn_code tf) ! (
CSSA.fn_entrypoint tf) =
Some (
Inop s).
Proof.
Lemma p_fn_entry_pred :
forall pc :
positive,
~
CSSA.cfg tf pc (
CSSA.fn_entrypoint tf).
Proof.
Lemma p_fn_no_parcb_at_entry : (
fn_parcopycode tf) ! (
CSSA.fn_entrypoint tf) =
None.
Proof.
Lemma p_fn_phicode_inv :
forall jp :
positive,
CSSA.join_point jp tf <->
(
CSSA.fn_phicode tf) !
jp <>
None.
Proof.
Lemma p_fn_normalized :
forall (
jp :
positive) (
pc :
positive),
CSSA.join_point jp tf ->
In jp (
CSSA.succs tf pc) ->
(
CSSA.fn_code tf) !
pc =
Some (
Inop jp).
Proof.
Lemma p_fn_inop_in_jp :
forall pc :
positive,
(
CSSA.fn_phicode tf) !
pc <>
None ->
exists succ :
node,
(
CSSA.fn_code tf) !
pc =
Some (
Inop succ).
Proof.
Hint Resolve normalisation_checker_correct:
core.
Lemma p_fn_normalized_jp :
forall (
pc :
positive) (
preds :
list positive),
(
CSSA.fn_phicode tf) !
pc <>
None ->
(
CSSA.get_preds tf) !
pc =
Some preds ->
forall pred :
positive,
In pred preds -> (
CSSA.fn_phicode tf) !
pred =
None.
Proof.
Lemma p_fn_parcbjp :
forall (
pc :
positive) (
pc' :
node) (
parcb :
parcopyblock),
(
fn_parcopycode tf) !
pc =
Some parcb ->
(
CSSA.fn_code tf) !
pc =
Some (
Inop pc') ->
~
CSSA.join_point pc'
tf ->
CSSA.join_point pc tf.
Proof.
Lemma p_parcb_inop :
forall pc :
positive,
(
fn_parcopycode tf) !
pc <>
None ->
exists s :
node, (
CSSA.fn_code tf) !
pc =
Some (
Inop s).
Proof.
Lemma p_fn_code_reached :
forall (
pc :
positive) (
ins :
instruction),
(
CSSA.fn_code tf) !
pc =
Some ins ->
CSSA.reached tf pc.
Proof.
Lemma p_fn_code_closed :
forall (
pc :
positive) (
pc' :
node)
(
instr :
instruction),
(
CSSA.fn_code tf) !
pc =
Some instr ->
In pc' (
successors_instr instr) ->
exists instr' :
instruction,
(
CSSA.fn_code tf) !
pc' =
Some instr'.
Proof.
Lemma p_fn_cssa :
CSSA.unique_def_spec tf.
Proof.
Lemma p_fn_cssa_params :
forall x :
reg,
In x (
CSSA.fn_params tf) ->
(
forall pc :
node,
~
assigned_code_spec (
CSSA.fn_code tf)
pc x) /\
(
forall pc :
node,
~
assigned_phi_spec (
CSSA.fn_phicode tf)
pc x) /\
(
forall pc :
positive,
~
assigned_parcopy_spec (
fn_parcopycode tf)
pc x).
Proof.
Hint Resolve ident_eq:
core.
Lemma p_fn_strict :
forall (
x :
reg) (
u d :
positive),
CSSA.use tf x u ->
CSSA.def tf x d ->
CSSA.dom tf d u.
Proof.
Lemma p_fn_use_def_code :
forall (
x :
reg) (
pc :
positive),
CSSA.use_code tf x pc ->
assigned_code_spec (
CSSA.fn_code tf)
pc x ->
False.
Proof.
Lemma p_fn_strict_use_parcb :
forall (
x :
reg) (
pc :
positive),
use_parcopycode tf x pc ->
~
assigned_parcopy_spec (
fn_parcopycode tf)
pc x.
Proof.
Lemma p_fn_jp_use_parcb' :
forall (
x :
reg) (
pc :
positive),
CSSA.join_point pc tf ->
use_parcopycode tf x pc ->
assigned_phi_spec (
CSSA.fn_phicode tf)
pc x.
Proof.
Lemma p_fn_jp_use_phib :
forall (
x :
reg) (
pc :
positive),
CSSA.use_phicode tf x pc ->
assigned_parcopy_spec (
fn_parcopycode tf)
pc x.
Proof.
Lemma wf_cssa_tf :
wf_cssa_function tf.
Proof.
End WF_CSSA.
Require Import Errors.
Lemma wf_cssa_tprog :
forall prog tprog,
wf_ssa_program prog ->
match_prog prog tprog ->
CSSA.wf_cssa_program tprog.
Proof.
red.
intros prog tprog WF_SSA MATCH.
intros.
red in WF_SSA.
inv MATCH.
revert H0 H WF_SSA.
clear.
generalize (
prog_defs tprog).
induction 1;
intros.
-
inv H.
-
inv H1.
+
inv H.
inv H2.
{
destruct f1 ;
simpl in * ;
try constructor;
auto.
*
monadInv H5.
constructor.
eapply wf_cssa_tf;
eauto.
destruct a1,
g.
exploit (
WF_SSA (
Internal f0)
i);
eauto.
simpl in * ;
eauto.
left.
inv H;
simpl in *;
auto.
intros.
inv H1;
auto.
simpl in *.
inv H.
*
monadInv H5.
constructor.
}
+
eapply IHlist_forall2;
eauto.
Qed.