aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend/SimplExprproof.v
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2015-04-02 16:25:25 +0200
committerBernhard Schommer <bernhardschommer@gmail.com>2015-04-02 16:25:25 +0200
commit170284b51766112e29d6bbfe519bad9f6da95269 (patch)
tree13a163ccee48cee0512a8af484b394623cdce030 /cfrontend/SimplExprproof.v
parent2e30ad9698a6f24a8a746f68b30c235913006392 (diff)
parent959432fa13a899290db5236f93575a8bfdc13bb5 (diff)
downloadcompcert-170284b51766112e29d6bbfe519bad9f6da95269.tar.gz
compcert-170284b51766112e29d6bbfe519bad9f6da95269.zip
Merge branch 'master' into dwarf
Diffstat (limited to 'cfrontend/SimplExprproof.v')
-rw-r--r--cfrontend/SimplExprproof.v24
1 files changed, 17 insertions, 7 deletions
diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v
index 74019061..7ef1cbe2 100644
--- a/cfrontend/SimplExprproof.v
+++ b/cfrontend/SimplExprproof.v
@@ -761,20 +761,30 @@ Proof.
inv H; simpl; auto.
Qed.
+Lemma static_bool_val_sound:
+ forall v t m b, bool_val v t Mem.empty = Some b -> bool_val v t m = Some b.
+Proof.
+ intros until b; unfold bool_val. destruct (classify_bool t); destruct v; auto.
+ intros E. unfold Mem.weak_valid_pointer, Mem.valid_pointer, proj_sumbool in E.
+ rewrite ! pred_dec_false in E by (apply Mem.perm_empty). discriminate.
+Qed.
+
Lemma step_makeif:
forall f a s1 s2 k e le m v1 b,
eval_expr tge e le m a v1 ->
- bool_val v1 (typeof a) = Some b ->
+ bool_val v1 (typeof a) m = Some b ->
star step1 tge (State f (makeif a s1 s2) k e le m)
E0 (State f (if b then s1 else s2) k e le m).
Proof.
intros. functional induction (makeif a s1 s2).
- exploit eval_simpl_expr_sound; eauto. rewrite e0. intro EQ; subst v.
- rewrite e1 in H0. inv H0. constructor.
- exploit eval_simpl_expr_sound; eauto. rewrite e0. intro EQ; subst v.
- rewrite e1 in H0. inv H0. constructor.
- apply star_one. eapply step_ifthenelse; eauto.
- apply star_one. eapply step_ifthenelse; eauto.
+- exploit eval_simpl_expr_sound; eauto. rewrite e0. intro EQ; subst v.
+ assert (bool_val v1 (typeof a) m = Some true) by (apply static_bool_val_sound; auto).
+ replace b with true by congruence. constructor.
+- exploit eval_simpl_expr_sound; eauto. rewrite e0. intro EQ; subst v.
+ assert (bool_val v1 (typeof a) m = Some false) by (apply static_bool_val_sound; auto).
+ replace b with false by congruence. constructor.
+- apply star_one. eapply step_ifthenelse; eauto.
+- apply star_one. eapply step_ifthenelse; eauto.
Qed.
Lemma step_make_set: