aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2023-05-09 10:10:05 +0100
committerYann Herklotz <git@yannherklotz.com>2023-05-09 10:10:05 +0100
commitc1778dc2f1a5de755b32f8c4655a718c109c6489 (patch)
tree826185424f8e081203b392824781f84a7bb58cfe
parentbad5c59b014a9baf18df0e2146edcb11fb931216 (diff)
downloadvericert-c1778dc2f1a5de755b32f8c4655a718c109c6489.tar.gz
vericert-c1778dc2f1a5de755b32f8c4655a718c109c6489.zip
Split proof up into more files
-rw-r--r--src/common/NonEmpty.v11
-rw-r--r--src/hls/Abstr.v1824
-rw-r--r--src/hls/GiblePargen.v27
-rw-r--r--src/hls/GiblePargenproof.v2085
-rw-r--r--src/hls/GiblePargenproofBackwards.v49
-rw-r--r--src/hls/GiblePargenproofEquiv.v1736
-rw-r--r--src/hls/GiblePargenproofEvaluable.v395
-rw-r--r--src/hls/GiblePargenproofForward.v1568
-rw-r--r--src/hls/Predicate.v6
9 files changed, 3935 insertions, 3766 deletions
diff --git a/src/common/NonEmpty.v b/src/common/NonEmpty.v
index cafa7ff..f15bb82 100644
--- a/src/common/NonEmpty.v
+++ b/src/common/NonEmpty.v
@@ -213,3 +213,14 @@ Inductive norepet {A : Type} : non_empty A -> Prop :=
| norepet_singleton a : norepet (singleton a)
| list_norepet_cons hd tl :
~ In hd tl -> norepet tl -> norepet (hd ::| tl).
+
+Lemma NEin_NEapp5 :
+ forall A (a: A) x y,
+ In a (app x y) ->
+ In a x \/ In a y.
+Proof.
+ induction x; cbn in *; intros.
+ - inv H. inv H1. left. constructor. tauto.
+ - inv H. inv H1. left. constructor; tauto.
+ apply IHx in H. inv H; intuition (constructor; auto).
+Qed.
diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v
index 2653755..e97e907 100644
--- a/src/hls/Abstr.v
+++ b/src/hls/Abstr.v
@@ -1,6 +1,6 @@
(*
* Vericert: Verified high-level synthesis.
- * Copyright (C) 2021-2023 Yann Herklotz <yann@yannherklotz.com>
+ * Copyright (C) 2021-2023 Yann Herklotz <git@yannherklotz.com>
*
* This program is free software: you can redistribute it and/or modify
* it under the terms of the GNU General Public License as published by
@@ -215,6 +215,69 @@ Definition get_m i := match i with mk_instr_state a b c => c end.
Definition eval_predf_opt pr p :=
match p with Some p' => eval_predf pr p' | None => true end.
+Lemma get_empty:
+ forall r, empty#r r = NE.singleton (Ptrue, Ebase r).
+Proof. unfold "#r"; intros. rewrite RTree.gempty. trivial. Qed.
+
+Lemma forest_reg_gso:
+ forall (f : forest) w dst dst',
+ dst <> dst' ->
+ (f #r dst <- w) #r dst' = f #r dst'.
+Proof.
+ unfold "#r"; intros.
+ unfold forest_regs. unfold set_forest.
+ rewrite RTree.gso; auto.
+Qed.
+
+Lemma forest_reg_pred:
+ forall (f : forest) w dst dst',
+ (f #r dst <- w) #p dst' = f #p dst'.
+Proof. auto. Qed.
+
+Lemma forest_reg_gss:
+ forall (f : forest) w dst,
+ (f #r dst <- w) #r dst = w.
+Proof.
+ unfold "#r"; intros.
+ unfold forest_regs. unfold set_forest.
+ rewrite RTree.gss; auto.
+Qed.
+
+Lemma forest_pred_gso:
+ forall (f : forest) w dst dst',
+ dst <> dst' ->
+ (f #p dst <- w) #p dst' = f #p dst'.
+Proof.
+ unfold "#p"; intros.
+ unfold forest_preds, set_forest_p, get_forest_p'.
+ rewrite PTree.gso; auto.
+Qed.
+
+Lemma forest_pred_reg:
+ forall (f : forest) w dst dst',
+ (f #p dst <- w) #r dst' = f #r dst'.
+Proof. auto. Qed.
+
+Lemma forest_pred_gss:
+ forall (f : forest) w dst,
+ (f #p dst <- w) #p dst = w.
+Proof.
+ unfold "#p"; intros.
+ unfold forest_preds, set_forest_p, get_forest_p'.
+ rewrite PTree.gss; auto.
+Qed.
+
+Lemma forest_pred_gss2 :
+ forall (f : forest) r p,
+ (forest_preds (f #p r <- p)) ! r = Some p.
+Proof. intros. unfold set_forest_p. simpl. now rewrite PTree.gss. Qed.
+
+Lemma forest_pred_gso2 :
+ forall (f : forest) r w p,
+ r <> w ->
+ (forest_preds (f #p w <- p)) ! r = (forest_preds f) ! r.
+Proof. intros. unfold set_forest_p. simpl. now rewrite PTree.gso by auto. Qed.
+
(*|
Finally we want to define the semantics of execution for the expressions with
symbolic values, so the result of executing the expressions will be an
@@ -415,1762 +478,3 @@ Inductive sem : ctx -> forest -> instr_state * option cf_instr -> Prop :=
sem ctx f (mk_instr_state rs' pr' m', cf).
End SEMANTICS.
-
-Fixpoint beq_expression (e1 e2: expression) {struct e1}: bool :=
- match e1, e2 with
- | Ebase r1, Ebase r2 => if resource_eq r1 r2 then true else false
- | Eop op1 el1, Eop op2 el2 =>
- if operation_eq op1 op2 then
- beq_expression_list el1 el2 else false
- | Eload chk1 addr1 el1 e1, Eload chk2 addr2 el2 e2 =>
- if memory_chunk_eq chk1 chk2
- then if addressing_eq addr1 addr2
- then if beq_expression_list el1 el2
- then beq_expression e1 e2 else false else false else false
- | Estore e1 chk1 addr1 el1 m1, Estore e2 chk2 addr2 el2 m2 =>
- if memory_chunk_eq chk1 chk2
- then if addressing_eq addr1 addr2
- then if beq_expression_list el1 el2
- then if beq_expression m1 m2
- then beq_expression e1 e2 else false else false else false else false
- | _, _ => false
- end
-with beq_expression_list (el1 el2: expression_list) {struct el1} : bool :=
- match el1, el2 with
- | Enil, Enil => true
- | Econs e1 t1, Econs e2 t2 => beq_expression e1 e2 && beq_expression_list t1 t2
- | _, _ => false
- end
-.
-
-Scheme expression_ind2 := Induction for expression Sort Prop
- with expression_list_ind2 := Induction for expression_list Sort Prop.
-
-Definition beq_pred_expression (e1 e2: pred_expression) : bool :=
- match e1, e2 with
- | PEbase p1, PEbase p2 => if peq p1 p2 then true else false
- | PEsetpred c1 el1, PEsetpred c2 el2 =>
- if condition_eq c1 c2
- then beq_expression_list el1 el2 else false
- | _, _ => false
- end.
-
-Definition beq_exit_expression (e1 e2: exit_expression) : bool :=
- match e1, e2 with
- | EEbase, EEbase => true
- | EEexit cf1, EEexit cf2 => if cf_instr_eq cf1 cf2 then true else false
- | _, _ => false
- end.
-
-Lemma beq_expression_correct:
- forall e1 e2, beq_expression e1 e2 = true -> e1 = e2.
-Proof.
- intro e1;
- apply expression_ind2 with
- (P := fun (e1 : expression) =>
- forall e2, beq_expression e1 e2 = true -> e1 = e2)
- (P0 := fun (e1 : expression_list) =>
- forall e2, beq_expression_list e1 e2 = true -> e1 = e2); simplify;
- try solve [repeat match goal with
- | [ H : context[match ?x with _ => _ end] |- _ ] => destruct x eqn:?
- | [ H : context[if ?x then _ else _] |- _ ] => destruct x eqn:?
- end; subst; f_equal; crush; eauto using Peqb_true_eq].
-Qed.
-
-Lemma beq_expression_refl: forall e, beq_expression e e = true.
-Proof.
- intros.
- induction e using expression_ind2 with (P0 := fun el => beq_expression_list el el = true);
- crush; repeat (destruct_match; crush); [].
- crush. rewrite IHe. rewrite IHe0. auto.
-Qed.
-
-Lemma beq_expression_list_refl: forall e, beq_expression_list e e = true.
-Proof. induction e; auto. simplify. rewrite beq_expression_refl. auto. Qed.
-
-Lemma beq_expression_correct2:
- forall e1 e2, beq_expression e1 e2 = false -> e1 <> e2.
-Proof.
- induction e1 using expression_ind2
- with (P0 := fun el1 => forall el2, beq_expression_list el1 el2 = false -> el1 <> el2).
- - intros. simplify. repeat (destruct_match; crush).
- - intros. simplify. repeat (destruct_match; crush). subst. apply IHe1 in H.
- unfold not in *. intros. apply H. inv H0. auto.
- - intros. simplify. repeat (destruct_match; crush); subst.
- unfold not in *; intros. inv H0. rewrite beq_expression_refl in H. discriminate.
- unfold not in *; intros. inv H. rewrite beq_expression_list_refl in Heqb. discriminate.
- - simplify. repeat (destruct_match; crush); subst;
- unfold not in *; intros.
- inv H0. rewrite beq_expression_refl in H; crush.
- inv H. rewrite beq_expression_refl in Heqb0; crush.
- inv H. rewrite beq_expression_list_refl in Heqb; crush.
- (* - simplify. repeat (destruct_match; crush); subst. *)
- (* unfold not in *; intros. inv H0. rewrite beq_expression_list_refl in H; crush. *)
- - simplify. repeat (destruct_match; crush); subst.
- - simplify. repeat (destruct_match; crush); subst.
- apply andb_false_iff in H. inv H. unfold not in *; intros.
- inv H. rewrite beq_expression_refl in H0; discriminate.
- unfold not in *; intros. inv H. rewrite beq_expression_list_refl in H0; discriminate.
-Qed.
-
-Definition expression_dec: forall e1 e2: expression, {e1 = e2} + {e1 <> e2}.
-Proof.
- intros.
- destruct (beq_expression e1 e2) eqn:?. apply beq_expression_correct in Heqb. auto.
- apply beq_expression_correct2 in Heqb. auto.
-Defined.
-
-Lemma beq_expression_list_correct:
- forall e1 e2, beq_expression_list e1 e2 = true -> e1 = e2.
-Proof.
- induction e1; crush.
- - destruct_match; crush.
- - destruct_match; crush.
- apply IHe1 in H1; subst.
- apply beq_expression_correct in H0; subst.
- trivial.
-Qed.
-
-Lemma beq_expression_list_correct2:
- forall e1 e2, beq_expression_list e1 e2 = false -> e1 <> e2.
-Proof.
- induction e1; crush.
- - destruct_match; crush.
- - destruct_match; crush.
- apply andb_false_iff in H. inv H. apply beq_expression_correct2 in H0.
- unfold not in *; intros. apply H0. inv H. auto.
- apply IHe1 in H0. unfold not in *; intros. apply H0. inv H.
- auto.
-Qed.
-
-Lemma beq_pred_expression_correct:
- forall e1 e2, beq_pred_expression e1 e2 = true -> e1 = e2.
-Proof.
- destruct e1, e2; crush.
- - destruct_match; crush.
- - destruct_match; subst; crush.
- apply beq_expression_list_correct in H; subst.
- trivial.
-Qed.
-
-Lemma beq_pred_expression_refl:
- forall e, beq_pred_expression e e = true.
-Proof.
- destruct e; crush; destruct_match; crush.
- apply beq_expression_list_refl.
-Qed.
-
-Lemma beq_pred_expression_correct2:
- forall e1 e2, beq_pred_expression e1 e2 = false -> e1 <> e2.
-Proof.
- destruct e1, e2; unfold not; crush.
- + destruct_match; crush.
- + destruct_match; crush. inv H0.
- now rewrite beq_expression_list_refl in H.
-Qed.
-
-Lemma beq_exit_expression_correct:
- forall e1 e2, beq_exit_expression e1 e2 = true <-> e1 = e2.
-Proof.
- destruct e1, e2; split; crush;
- destruct_match; subst; crush.
-Qed.
-
-Definition pred_expr_item_eq (pe1 pe2: pred_op * expression) : bool :=
- @equiv_dec _ SATSetoid _ (fst pe1) (fst pe2) && beq_expression (snd pe1) (snd pe2).
-
-Definition pred_eexpr_item_eq (pe1 pe2: pred_op * exit_expression) : bool :=
- @equiv_dec _ SATSetoid _ (fst pe1) (fst pe2) && beq_exit_expression (snd pe1) (snd pe2).
-
-Definition pred_expr_dec: forall (pe1 pe2: pred_op * expression),
- {pred_expr_item_eq pe1 pe2 = true} + {pred_expr_item_eq pe1 pe2 = false}.
-Proof.
- intros; destruct (pred_expr_item_eq pe1 pe2) eqn:?; unfold not; [tauto | now right].
-Defined.
-
-Definition pred_expr_dec2: forall (pe1 pe2: pred_op * expression),
- {pred_expr_item_eq pe1 pe2 = true} + {~ pred_expr_item_eq pe1 pe2 = true}.
-Proof.
- intros; destruct (pred_expr_item_eq pe1 pe2) eqn:?; unfold not; [tauto | now right].
-Defined.
-
-Definition pred_expression_dec:
- forall e1 e2: pred_expression, {e1 = e2} + {e1 <> e2}.
-Proof.
- intros. destruct (beq_pred_expression e1 e2) eqn:?.
- eauto using beq_pred_expression_correct.
- eauto using beq_pred_expression_correct2.
-Defined.
-
-Lemma exit_expression_refl:
- forall e, beq_exit_expression e e = true.
-Proof. destruct e; crush; destruct_match; crush. Qed.
-
-Definition exit_expression_dec:
- forall e1 e2: exit_expression, {e1 = e2} + {e1 <> e2}.
-Proof.
- intros. destruct (beq_exit_expression e1 e2) eqn:?.
- - left. eapply beq_exit_expression_correct; eauto.
- - right. unfold not; intros.
- assert (X: ~ (beq_exit_expression e1 e2 = true))
- by eauto with bool.
- subst. apply X. apply exit_expression_refl.
-Defined.
-
-Lemma pred_eexpression_dec:
- forall (e1 e2: exit_expression) (p1 p2: pred_op),
- {(p1, e1) = (p2, e2)} + {(p1, e1) <> (p2, e2)}.
-Proof.
- pose proof (Predicate.eq_dec peq).
- pose proof (exit_expression_dec).
- decide equality.
-Defined.
-
-(*Fixpoint encode_expression_ne (max: predicate) (pe: pred_expr_ne) (h: hash_tree)
- : (PTree.t pred_op) * hash_tree :=
- match pe with
- | NE.singleton (p, e) =>
- let (p', h') := hash_value max e h in
- (Por (Pnot p) (Pvar p'), h')
- | (p, e) ::| pr =>
- let (p', h') := hash_value max e h in
- let (p'', h'') := encode_expression_ne max pr h' in
- (Pand (Por (Pnot p) (Pvar p')) p'', h'')
- end.*)
-
-Fixpoint max_pred_expr (pe: pred_expr) : positive :=
- match pe with
- | NE.singleton (p, e) => max_predicate p
- | (p, e) ::| pe' => Pos.max (max_predicate p) (max_pred_expr pe')
- end.
-
-Definition ge_preserved {A B C D: Type} (ge: Genv.t A B) (tge: Genv.t C D) : Prop :=
- (forall sp op vl m, Op.eval_operation ge sp op vl m =
- Op.eval_operation tge sp op vl m)
- /\ (forall sp addr vl, Op.eval_addressing ge sp addr vl =
- Op.eval_addressing tge sp addr vl).
-
-Lemma ge_preserved_same:
- forall A B ge, @ge_preserved A B A B ge ge.
-Proof. unfold ge_preserved; auto. Qed.
-#[local] Hint Resolve ge_preserved_same : core.
-
-Inductive match_states : instr_state -> instr_state -> Prop :=
-| match_states_intro:
- forall ps ps' rs rs' m m',
- (forall x, rs !! x = rs' !! x) ->
- (forall x, ps !! x = ps' !! x) ->
- m = m' ->
- match_states (mk_instr_state rs ps m) (mk_instr_state rs' ps' m').
-
-Lemma match_states_refl x : match_states x x.
-Proof. destruct x; constructor; crush. Qed.
-
-Lemma match_states_commut x y : match_states x y -> match_states y x.
-Proof. inversion 1; constructor; crush. Qed.
-
-Lemma match_states_trans x y z :
- match_states x y -> match_states y z -> match_states x z.
-Proof. repeat inversion 1; constructor; crush. Qed.
-
-#[global] Instance match_states_Equivalence : Equivalence match_states :=
- { Equivalence_Reflexive := match_states_refl ;
- Equivalence_Symmetric := match_states_commut ;
- Equivalence_Transitive := match_states_trans ; }.
-
-Inductive similar {A B} : @ctx A -> @ctx B -> Prop :=
-| similar_intro :
- forall ist ist' sp ge tge,
- ge_preserved ge tge ->
- match_states ist ist' ->
- similar (mk_ctx ist sp ge) (mk_ctx ist' sp tge).
-
-Lemma ge_preserved_refl:
- forall A B (a: Genv.t A B), ge_preserved a a.
-Proof. auto. Qed.
-
-Lemma similar_refl:
- forall A (a: @Abstr.ctx A), similar a a.
-Proof. intros; destruct a; constructor; auto. reflexivity. Qed.
-
-Lemma similar_commut:
- forall A B (a: @Abstr.ctx A) (b: @Abstr.ctx B), similar a b -> similar b a.
-Proof.
- inversion 1; constructor; auto.
- - unfold ge_preserved in *; inv H0; split; intros.
- rewrite H4; auto. rewrite H5; auto.
- - symmetry; auto.
-Qed.
-
-Lemma similar_trans:
- forall A B C (a: @Abstr.ctx A) (b: @Abstr.ctx B) (c: @Abstr.ctx C),
- similar a b -> similar b c -> similar a c.
-Proof.
- repeat inversion 1; constructor.
- - unfold ge_preserved in *; inv H0; inv H9; split; intros.
- rewrite H11. rewrite H0; auto.
- rewrite H12. rewrite H2. auto.
- - transitivity ist'; auto.
-Qed.
-
-Module HashExpr' <: MiniDecidableType.
- Definition t := expression.
- Definition eq_dec := expression_dec.
-End HashExpr'.
-
-Module HashExpr := Make_UDT(HashExpr').
-Module Import HT := HashTree(HashExpr).
-
-Module PHashExpr' <: MiniDecidableType.
- Definition t := pred_expression.
- Definition eq_dec := pred_expression_dec.
-End PHashExpr'.
-
-Module PHashExpr := Make_UDT(PHashExpr').
-Module PHT := HashTree(PHashExpr).
-
-Module EHashExpr' <: MiniDecidableType.
- Definition t := exit_expression.
- Definition eq_dec := exit_expression_dec.
-End EHashExpr'.
-
-Module EHashExpr := Make_UDT(EHashExpr').
-Module EHT := HashTree(EHashExpr).
-
-Fixpoint hash_predicate (max: predicate) (ap: pred_pexpr) (h: PHT.hash_tree)
- : pred_op * PHT.hash_tree :=
- match ap with
- | Ptrue => (Ptrue, h)
- | Pfalse => (Pfalse, h)
- | Plit (b, ap') =>
- let (p', h') := PHT.hash_value max ap' h in
- (Plit (b, p'), h')
- | Pand p1 p2 =>
- let (p1', h') := hash_predicate max p1 h in
- let (p2', h'') := hash_predicate max p2 h' in
- (Pand p1' p2', h'')
- | Por p1 p2 =>
- let (p1', h') := hash_predicate max p1 h in
- let (p2', h'') := hash_predicate max p2 h' in
- (Por p1' p2', h'')
- end.
-
-Definition predicated_mutexcl {A: Type} (pe: predicated A): Prop :=
- (forall x y, NE.In x pe -> NE.In y pe -> x <> y -> fst x ⇒ ¬ fst y)
- /\ NE.norepet pe.
-
-Lemma predicated_cons :
- forall A (a: pred_op * A) x,
- predicated_mutexcl (a ::| x) ->
- predicated_mutexcl x.
-Proof.
- unfold predicated_mutexcl; intros. inv H. inv H1. split; auto.
- intros. apply H0; auto; constructor; tauto.
-Qed.
-
-Lemma predicated_singleton :
- forall A (a: (pred_op * A)), predicated_mutexcl (NE.singleton a).
-Proof.
- unfold predicated_mutexcl; intros; split; intros.
- { inv H. now inv H0. }
- constructor.
-Qed.
-
-(*
-
-Lemma norm_expr_constant :
- forall x m h t h' e p,
- HN.norm_expression m x h = (t, h') ->
- h ! e = Some p ->
- h' ! e = Some p.
-Proof. Abort.
-
-Definition sat_aequiv ap1 ap2 :=
- exists h p1 p2,
- sat_equiv p1 p2
- /\ hash_predicate 1 ap1 h = (p1, h)
- /\ hash_predicate 1 ap2 h = (p2, h).
-
-Lemma aequiv_symm : forall a b, sat_aequiv a b -> sat_aequiv b a.
-Proof.
- unfold sat_aequiv; simplify; do 3 eexists; simplify; [symmetry | |]; eauto.
-Qed.
-
-Lemma existsh :
- forall ap1,
- exists h p1,
- hash_predicate 1 ap1 h = (p1, h).
-Proof. Admitted.
-
-Lemma aequiv_refl : forall a, sat_aequiv a a.
-Proof.
- unfold sat_aequiv; intros.
- pose proof (existsh a); simplify.
- do 3 eexists; simplify; eauto. reflexivity.
-Qed.
-
-Lemma aequiv_trans :
- forall a b c,
- sat_aequiv a b ->
- sat_aequiv b c ->
- sat_aequiv a c.
-Proof.
- unfold sat_aequiv; intros.
- simplify.
-Abort.
-
-Lemma norm_expr_mutexcl :
- forall m pe h t h' e e' p p',
- HN.norm_expression m pe h = (t, h') ->
- predicated_mutexcl pe ->
- t ! e = Some p ->
- t ! e' = Some p' ->
- e <> e' ->
- p ⇒ ¬ p'.
-Proof. Abort.*)
-
-Definition pred_expr_eqb: forall pe1 pe2: pred_expr,
- {pe1 = pe2} + {pe1 <> pe2}.
-Proof.
- pose proof expression_dec.
- pose proof NE.eq_dec.
- pose proof (Predicate.eq_dec peq).
- assert (forall a b: pred_op * expression, {a = b} + {a <> b})
- by decide equality.
- decide equality.
-Defined.
-
-Definition pred_pexpr_eqb: forall pe1 pe2: pred_pexpr,
- {pe1 = pe2} + {pe1 <> pe2}.
-Proof.
- pose proof pred_expression_dec.
- pose proof (Predicate.eq_dec pred_expression_dec).
- apply X.
-Defined.
-
-Definition beq_pred_pexpr (pe1 pe2: pred_pexpr): bool :=
- if pred_pexpr_eqb pe1 pe2 then true else
- let (np1, h) := hash_predicate 1 pe1 (PTree.empty _) in
- let (np2, h') := hash_predicate 1 pe2 h in
- equiv_check np1 np2.
-
-Lemma inj_asgn_eg : forall a b, (a =? b)%positive = (a =? a)%positive -> a = b.
-Proof.
- intros. destruct (peq a b); subst.
- auto. rewrite OrdersEx.Positive_as_OT.eqb_refl in H.
- now apply Peqb_true_eq.
-Qed.
-
-Lemma inj_asgn :
- forall a b, (forall (f: positive -> bool), f a = f b) -> a = b.
-Proof. intros. apply inj_asgn_eg. eauto. Qed.
-
-Lemma inj_asgn_false:
- forall n1 n2 , ~ (forall c: positive -> bool, c n1 = negb (c n2)).
-Proof.
- unfold not; intros. specialize (H (fun x => true)).
- simplify. discriminate.
-Qed.
-
-Lemma negb_inj:
- forall a b,
- negb a = negb b -> a = b.
-Proof. destruct a, b; crush. Qed.
-
-Lemma sat_predicate_Plit_inj :
- forall p1 p2,
- Plit p1 == Plit p2 -> p1 = p2.
-Proof.
- simplify. destruct p1, p2.
- destruct b, b0. f_equal. unfold sat_equiv in H. cbn in H. now apply inj_asgn.
- solve [exfalso; eapply inj_asgn_false; eauto].
- solve [exfalso; eapply inj_asgn_false; eauto].
- assert (p = p0). eapply inj_asgn. intros. specialize (H f).
- apply negb_inj in H. auto. rewrite H0; auto.
-Qed.
-
-Definition ind_preds t :=
- forall e1 e2 p1 p2 c,
- e1 <> e2 ->
- t ! e1 = Some p1 ->
- t ! e2 = Some p2 ->
- sat_predicate p1 c = true ->
- sat_predicate p2 c = false.
-
-Definition ind_preds_l t :=
- forall (e1: predicate) e2 p1 p2 c,
- e1 <> e2 ->
- In (e1, p1) t ->
- In (e2, p2) t ->
- sat_predicate p1 c = true ->
- sat_predicate p2 c = false.
-
-(*Lemma pred_equivalence_Some' :
- forall ta tb e pe pe',
- list_norepet (map fst ta) ->
- list_norepet (map fst tb) ->
- In (e, pe) ta ->
- In (e, pe') tb ->
- fold_right (fun p a => mk_pred_stmnt' (fst p) (snd p) ∧ a) T ta ==
- fold_right (fun p a => mk_pred_stmnt' (fst p) (snd p) ∧ a) T tb ->
- pe == pe'.
-Proof.
- induction ta as [|hd tl Hta]; try solve [crush].
- - intros * NRP1 NRP2 IN1 IN2 FOLD. inv NRP1. inv IN1.
- simpl in FOLD.
-
-Lemma pred_equivalence_Some :
- forall (ta tb: PTree.t pred_op) e pe pe',
- ta ! e = Some pe ->
- tb ! e = Some pe' ->
- mk_pred_stmnt ta == mk_pred_stmnt tb ->
- pe == pe'.
-Proof.
- intros * SMEA SMEB EQ. unfold mk_pred_stmnt in *.
- repeat rewrite PTree.fold_spec in EQ.
-
-Lemma pred_equivalence_None :
- forall (ta tb: PTree.t pred_op) e pe,
- (mk_pred_stmnt ta) == (mk_pred_stmnt tb) ->
- ta ! e = Some pe ->
- tb ! e = None ->
- equiv pe ⟂.
-Abort.
-
-Lemma pred_equivalence :
- forall (ta tb: PTree.t pred_op) e pe,
- equiv (mk_pred_stmnt ta) (mk_pred_stmnt tb) ->
- ta ! e = Some pe ->
- equiv pe ⟂ \/ (exists pe', tb ! e = Some pe' /\ equiv pe pe').
-Proof.
- intros * EQ SME. destruct (tb ! e) eqn:HTB.
- { right. econstructor. split; eauto. eapply pred_equivalence_Some; eauto. }
- { left. eapply pred_equivalence_None; eauto. }
-Qed.*)
-
-Section CORRECT.
-
- Context {FUN TFUN: Type}.
-
- Context (ictx: @ctx FUN) (octx: @ctx TFUN) (HSIM: similar ictx octx).
-
- Lemma sem_value_mem_det:
- forall e v v' m m',
- (sem_value ictx e v -> sem_value octx e v' -> v = v')
- /\ (sem_mem ictx e m -> sem_mem octx e m' -> m = m').
- Proof using HSIM.
- induction e using expression_ind2
- with (P0 := fun p => forall v v',
- sem_val_list ictx p v -> sem_val_list octx p v' -> v = v');
- inv HSIM; match goal with H: context [match_states] |- _ => inv H end; repeat progress simplify;
- try solve [match goal with
- | H: sem_value _ _ _, H2: sem_value _ _ _ |- _ => inv H; inv H2; auto
- | H: sem_mem _ _ _, H2: sem_mem _ _ _ |- _ => inv H; inv H2; auto
- | H: sem_val_list _ _ _, H2: sem_val_list _ _ _ |- _ => inv H; inv H2; auto
- end].
- - repeat match goal with
- | H: sem_value _ _ _ |- _ => inv H
- | H: sem_val_list {| ctx_ge := ge; |} ?e ?l1,
- H2: sem_val_list {| ctx_ge := tge |} ?e ?l2,
- IH: forall _ _, sem_val_list _ _ _ -> sem_val_list _ _ _ -> _ = _ |- _ =>
- assert (X: l1 = l2) by (apply IH; auto)
- | H: ge_preserved _ _ |- _ => inv H
- | |- context [ctx_rs] => unfold ctx_rs; cbn
- | H: context [ctx_mem] |- _ => unfold ctx_mem in H; cbn
- end; crush.
- - repeat match goal with H: sem_value _ _ _ |- _ => inv H end; simplify;
- assert (lv0 = lv) by (apply IHe; eauto); subst;
- match goal with H: ge_preserved _ _ |- _ => inv H end;
- match goal with H: context [Op.eval_addressing _ _ _ _ = Op.eval_addressing _ _ _ _] |- _
- => rewrite H in * end;
- assert (a0 = a1) by crush;
- assert (m'2 = m'1) by (apply IHe0; eauto); crush.
- - inv H0; inv H3. simplify.
- assert (lv = lv0) by ( apply IHe2; eauto). subst.
- assert (a1 = a0). { inv H. rewrite H3 in *. crush. }
- assert (v0 = v1). { apply IHe1; auto. }
- assert (m'1 = m'2). { apply IHe3; auto. } crush.
- - inv H0. inv H3. f_equal. apply IHe; auto.
- apply IHe0; auto.
- Qed.
-
- Lemma sem_value_mem_corr:
- forall e v m,
- (sem_value ictx e v -> sem_value octx e v)
- /\ (sem_mem ictx e m -> sem_mem octx e m).
- Proof using HSIM.
- induction e using expression_ind2
- with (P0 := fun p => forall v,
- sem_val_list ictx p v -> sem_val_list octx p v);
- inv HSIM; match goal with H: context [match_states] |- _ => inv H end; repeat progress simplify.
- - inv H0. unfold ctx_rs, ctx_ps, ctx_mem in *; cbn. rewrite H1. constructor.
- - inv H0. unfold ctx_rs, ctx_ps, ctx_mem in *; cbn. constructor.
- - inv H0. apply IHe in H6. econstructor; try eassumption.
- unfold ctx_rs, ctx_ps, ctx_mem in *; cbn in *. inv H. crush.
- - inv H0.
- - inv H0. eapply IHe in H10. eapply IHe0 in H8; auto.
- econstructor; try eassumption.
- unfold ctx_rs, ctx_ps, ctx_mem in *; cbn in *. inv H; crush.
- - inv H0.
- - inv H0.
- - inv H0. eapply IHe1 in H11; auto. eapply IHe2 in H12; auto. eapply IHe3 in H9; auto.
- econstructor; try eassumption.
- unfold ctx_rs, ctx_ps, ctx_mem in *; cbn in *. inv H; crush.
- - inv H0. econstructor.
- - inv H0. eapply IHe in H6; auto. eapply IHe0 in H8.
- econstructor; eassumption.
- Qed.
-
- Lemma sem_value_det:
- forall e v v', sem_value ictx e v -> sem_value octx e v' -> v = v'.
- Proof using HSIM.
- intros. eapply sem_value_mem_det; eauto; apply Mem.empty.
- Qed.
-
- Lemma sem_value_corr:
- forall e v, sem_value ictx e v -> sem_value octx e v.
- Proof using HSIM.
- intros. eapply sem_value_mem_corr; eauto; apply Mem.empty.
- Qed.
-
- Lemma sem_mem_det:
- forall e v v', sem_mem ictx e v -> sem_mem octx e v' -> v = v'.
- Proof using HSIM.
- intros. eapply sem_value_mem_det; eauto; apply (Vint (Int.repr 0%Z)).
- Qed.
-
- Lemma sem_mem_corr:
- forall e v, sem_mem ictx e v -> sem_mem octx e v.
- Proof using HSIM.
- intros. eapply sem_value_mem_corr; eauto; apply (Vint (Int.repr 0%Z)).
- Qed.
-
- Lemma sem_val_list_det:
- forall e l l', sem_val_list ictx e l -> sem_val_list octx e l' -> l = l'.
- Proof using HSIM.
- induction e; simplify.
- - inv H; inv H0; auto.
- - inv H; inv H0. f_equal. eapply sem_value_det; eauto; try apply Mem.empty.
- apply IHe; eauto.
- Qed.
-
- Lemma sem_val_list_corr:
- forall e l, sem_val_list ictx e l -> sem_val_list octx e l.
- Proof using HSIM.
- induction e; simplify.
- - inv H; constructor.
- - inv H. apply sem_value_corr in H3; auto; try apply Mem.empty;
- apply IHe in H5; constructor; assumption.
- Qed.
-
- Lemma sem_pred_det:
- forall e v v', sem_pred ictx e v -> sem_pred octx e v' -> v = v'.
- Proof using HSIM.
- try solve [inversion 1]; pose proof sem_value_det; pose proof sem_val_list_det; inv HSIM;
- match goal with H: match_states _ _ |- _ => inv H end; simplify.
- inv H2; inv H5; auto. assert (lv = lv0) by (eapply H0; eauto). subst. unfold ctx_mem in *.
- crush.
- Qed.
-
- Lemma sem_pred_corr:
- forall e v, sem_pred ictx e v -> sem_pred octx e v.
- Proof using HSIM.
- try solve [inversion 1]; pose proof sem_value_corr; pose proof sem_val_list_corr; inv HSIM;
- match goal with H: match_states _ _ |- _ => inv H end; simplify.
- inv H2; auto. apply H0 in H5. econstructor; eauto.
- unfold ctx_ps; cbn. rewrite H4. constructor.
- Qed.
-
- Lemma sem_exit_det:
- forall e v v', sem_exit ictx e v -> sem_exit octx e v' -> v = v'.
- Proof using HSIM.
- try solve [inversion 1]; pose proof sem_value_det; pose proof sem_val_list_det; inv HSIM;
- match goal with H: match_states _ _ |- _ => inv H end; simplify.
- inv H2; inv H5; auto.
- Qed.
-
- Lemma sem_exit_corr:
- forall e v, sem_exit ictx e v -> sem_exit octx e v.
- Proof using HSIM.
- try solve [inversion 1]; pose proof sem_value_corr; pose proof sem_val_list_corr; inv HSIM;
- match goal with H: match_states _ _ |- _ => inv H end; simplify.
- inv H2; auto; constructor.
- Qed.
-
- Lemma sem_pexpr_det :
- forall p b1 b2, sem_pexpr ictx p b1 -> sem_pexpr octx p b2 -> b1 = b2.
- Proof.
- induction p; crush; inv H; inv H0; firstorder.
- destruct b.
- - apply sem_pred_det with (e:=p0); auto.
- - apply negb_inj. apply sem_pred_det with (e:=p0); auto.
- Qed.
-
- Lemma sem_pexpr_corr :
- forall p b, sem_pexpr ictx p b -> sem_pexpr octx p b.
- Proof.
- induction p; crush; inv H; constructor;
- try solve [try inv H3; firstorder].
- now apply sem_pred_corr.
- Qed.
-
- Lemma sem_pred_exec_beq_correct2 :
- forall A B (sem: forall G, @Abstr.ctx G -> A -> B -> Prop) a p r R,
- (forall x y,
- sem _ ictx x y ->
- exists y', sem _ octx x y' /\ R y y') ->
- sem_pred_expr a (sem _) ictx p r ->
- exists r', sem_pred_expr a (sem _) octx p r' /\ R r r'.
- Proof.
- induction p; crush.
- - inv H0. apply H in H4. simplify.
- exists x; split; auto.
- constructor; auto.
- now apply sem_pexpr_corr.
- - inv H0.
- + apply H in H6; simplify.
- exists x; split; auto.
- constructor; auto.
- now apply sem_pexpr_corr.
- + exploit IHp; auto. exact H6. intros. simplify.
- exists x; split; auto.
- apply sem_pred_expr_cons_false; auto.
- now apply sem_pexpr_corr.
- Qed.
-
- Lemma sem_pred_expr_corr :
- forall A B (sem: forall G, @Abstr.ctx G -> A -> B -> Prop) a p r,
- (forall x y, sem _ ictx x y -> sem _ octx x y) ->
- sem_pred_expr a (sem _) ictx p r ->
- sem_pred_expr a (sem _) octx p r.
- Proof.
- intros.
- assert
- (forall x y,
- sem0 _ ictx x y ->
- exists y', sem0 _ octx x y' /\ eq y y') by firstorder.
- pose proof (sem_pred_exec_beq_correct2 _ _ sem0 a p r _ H1 H0).
- crush.
- Qed.
-
- Lemma sem_correct:
- forall f i cf, sem ictx f (i, cf) -> sem octx f (i, cf).
- Proof.
- intros. inv H. constructor.
- - inv H2. constructor; intros. specialize (H x).
- apply sem_pred_expr_corr; auto. exact sem_value_corr.
- - inv H3; constructor; intros. specialize (H x).
- now apply sem_pexpr_corr.
- - apply sem_pred_expr_corr; auto. exact sem_mem_corr.
- - apply sem_pred_expr_corr; auto. exact sem_exit_corr.
- Qed.
-
-End CORRECT.
-
-Section SEM_PRED_EXEC.
-
- Context (A: Type).
- Context (ctx: @Abstr.ctx A).
-
- Lemma sem_pexpr_negate :
- forall p b,
- sem_pexpr ctx p b ->
- sem_pexpr ctx (¬ p) (negb b).
- Proof.
- induction p; crush.
- - destruct_match. destruct b0; crush. inv Heqp0.
- constructor. inv H. rewrite negb_involutive. auto.
- constructor. inv H. auto.
- - inv H. constructor.
- - inv H. constructor.
- - inv H. inv H3.
- + apply IHp1 in H. solve [constructor; auto].
- + apply IHp2 in H. solve [constructor; auto].
- + apply IHp1 in H2. apply IHp2 in H4. solve [constructor; auto].
- - inv H. inv H3.
- + apply IHp1 in H. solve [constructor; auto].
- + apply IHp2 in H. solve [constructor; auto].
- + apply IHp1 in H2. apply IHp2 in H4. solve [constructor; auto].
- Qed.
-
- Lemma sem_pexpr_negate2 :
- forall p b,
- sem_pexpr ctx (¬ p) (negb b) ->
- sem_pexpr ctx p b.
- Proof.
- induction p; crush.
- - destruct_match. destruct b0; crush. inv Heqp0.
- constructor. inv H. rewrite negb_involutive in *. auto.
- constructor. inv H. auto.
- - inv H. destruct b; try discriminate. constructor.
- - inv H. destruct b; try discriminate. constructor.
- - inv H. destruct b; try discriminate.
- + constructor. inv H1; eauto.
- + destruct b; try discriminate. constructor; eauto.
- - inv H. destruct b; try discriminate.
- + constructor. inv H1; eauto.
- + destruct b; try discriminate. constructor; eauto.
- Qed.
-
- Lemma sem_pexpr_evaluable :
- forall f_p ps,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p, exists b, sem_pexpr ctx (from_pred_op f_p p) b.
- Proof.
- induction p; crush.
- - destruct_match. inv Heqp0. destruct b. econstructor. eauto.
- econstructor. eapply sem_pexpr_negate. eauto.
- - econstructor. constructor.
- - econstructor. constructor.
- - destruct x0, x; solve [eexists; constructor; auto].
- - destruct x0, x; solve [eexists; constructor; auto].
- Qed.
-
- Lemma sem_pexpr_eval1 :
- forall f_p ps,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p,
- eval_predf ps p = false ->
- sem_pexpr ctx (from_pred_op f_p p) false.
- Proof.
- induction p; crush.
- - destruct_match. inv Heqp0.
- destruct b.
- + cbn in H0. rewrite <- H0. eauto.
- + replace false with (negb true) by auto.
- apply sem_pexpr_negate. cbn in H0.
- apply negb_true_iff in H0. rewrite negb_involutive in H0.
- rewrite <- H0. eauto.
- - constructor.
- - rewrite eval_predf_Pand in H0.
- apply andb_false_iff in H0. inv H0. eapply IHp1 in H1.
- pose proof (sem_pexpr_evaluable _ _ H p2) as EVAL.
- inversion_clear EVAL as [x EVAL2].
- replace false with (false && x) by auto.
- constructor; auto.
- eapply IHp2 in H1.
- pose proof (sem_pexpr_evaluable _ _ H p1) as EVAL.
- inversion_clear EVAL as [x EVAL2].
- replace false with (x && false) by eauto with bool.
- apply sem_pexpr_Pand; auto.
- - rewrite eval_predf_Por in H0.
- apply orb_false_iff in H0. inv H0.
- replace false with (false || false) by auto.
- apply sem_pexpr_Por; auto.
- Qed.
-
- Lemma sem_pexpr_eval2 :
- forall f_p ps,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p,
- eval_predf ps p = true ->
- sem_pexpr ctx (from_pred_op f_p p) true.
- Proof.
- induction p; crush.
- - destruct_match. inv Heqp0.
- destruct b.
- + cbn in H0. rewrite <- H0. eauto.
- + replace true with (negb false) by auto.
- apply sem_pexpr_negate. cbn in H0.
- apply negb_true_iff in H0.
- rewrite <- H0. eauto.
- - constructor.
- - rewrite eval_predf_Pand in H0.
- apply andb_true_iff in H0. inv H0.
- replace true with (true && true) by auto.
- constructor; auto.
- - rewrite eval_predf_Por in H0.
- apply orb_true_iff in H0. inv H0. eapply IHp1 in H1.
- pose proof (sem_pexpr_evaluable _ _ H p2) as EVAL.
- inversion_clear EVAL as [x EVAL2].
- replace true with (true || x) by auto.
- apply sem_pexpr_Por; auto.
- eapply IHp2 in H1.
- pose proof (sem_pexpr_evaluable _ _ H p1) as EVAL.
- inversion_clear EVAL as [x EVAL2].
- replace true with (x || true) by eauto with bool.
- apply sem_pexpr_Por; auto.
- Qed.
-
- Lemma sem_pexpr_eval :
- forall f_p ps b,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p,
- eval_predf ps p = b ->
- sem_pexpr ctx (from_pred_op f_p p) b.
- Proof.
- intros; destruct b; eauto using sem_pexpr_eval1, sem_pexpr_eval2.
- Qed.
-
- Lemma sem_pexpr_eval_inv :
- forall f_p ps b,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p,
- sem_pexpr ctx (from_pred_op f_p p) b ->
- eval_predf ps p = b.
- Proof.
- induction p; intros.
- - cbn in H0. destruct_match. destruct b0; cbn in *.
- + specialize (H p0). eapply sem_pexpr_det; eauto. apply similar_refl.
- + rewrite <- negb_involutive in H0. apply sem_pexpr_negate2 in H0.
- symmetry; apply negb_sym. eapply sem_pexpr_det; eauto.
- apply similar_refl.
- - now inv H0.
- - now inv H0.
- - inv H0; try inv H4; rewrite eval_predf_Pand.
- + apply IHp1 in H0. rewrite H0. auto.
- + apply IHp2 in H0. rewrite H0. auto with bool.
- + apply IHp2 in H5. apply IHp1 in H3. rewrite H3. rewrite H5. auto.
- - inv H0; try inv H4; rewrite eval_predf_Por.
- + apply IHp1 in H0. rewrite H0. auto.
- + apply IHp2 in H0. rewrite H0. auto with bool.
- + apply IHp2 in H5. apply IHp1 in H3. rewrite H3. rewrite H5. auto.
- Qed.
-
- Context {C B: Type}.
- Context (f: PTree.t pred_pexpr).
- Context (ps: PMap.t bool).
- Context (a_sem: @Abstr.ctx A -> C -> B -> Prop).
-
- Context (F_EVALULABLE: forall x, sem_pexpr ctx (get_forest_p' x f) ps !! x).
-
- Lemma sem_pexpr_equiv :
- forall p1 p2 b,
- p1 == p2 ->
- sem_pexpr ctx (from_pred_op f p1) b ->
- sem_pexpr ctx (from_pred_op f p2) b.
- Proof.
- intros.
- eapply sem_pexpr_eval_inv in H0; eauto.
- eapply sem_pexpr_eval; eauto.
- Qed.
-
-End SEM_PRED_EXEC.
-
-Module HashExprNorm(HS: Hashable).
- Module H := HashTree(HS).
-
- Definition norm_tree: Type := PTree.t pred_op * H.hash_tree.
-
- Fixpoint norm_expression (max: predicate) (pe: predicated HS.t) (h: H.hash_tree)
- : norm_tree :=
- match pe with
- | NE.singleton (p, e) =>
- let (hashed_e, h') := H.hash_value max e h in
- (PTree.set hashed_e p (PTree.empty _), h')
- | (p, e) ::| pr =>
- let (hashed_e, h') := H.hash_value max e h in
- let (norm_pr, h'') := norm_expression max pr h' in
- match norm_pr ! hashed_e with
- | Some pr_op =>
- (PTree.set hashed_e (pr_op ∨ p) norm_pr, h'')
- | None =>
- (PTree.set hashed_e p norm_pr, h'')
- end
- end.
-
- Definition mk_pred_stmnt' (e: predicate) p_e := ¬ p_e ∨ Plit (true, e).
-
- Definition mk_pred_stmnt t := PTree.fold (fun x a b => mk_pred_stmnt' a b ∧ x) t T.
-
- Definition mk_pred_stmnt_l (t: list (predicate * pred_op)) :=
- fold_left (fun x a => uncurry mk_pred_stmnt' a ∧ x) t T.
-
- Definition encode_expression max pe h :=
- let (tree, h) := norm_expression max pe h in
- (mk_pred_stmnt tree, h).
-
- Definition pred_expr_dec: forall pe1 pe2: predicated HS.t,
- {pe1 = pe2} + {pe1 <> pe2}.
- Proof.
- pose proof HS.eq_dec as X.
- pose proof (Predicate.eq_dec peq).
- pose proof (NE.eq_dec _ X).
- assert (forall a b: pred_op * HS.t, {a = b} + {a <> b})
- by decide equality.
- decide equality.
- Defined.
-
- Definition beq_pred_expr' (pe1 pe2: predicated HS.t) : bool :=
- if pred_expr_dec pe1 pe2 then true else
- let (p1, h) := encode_expression 1%positive pe1 (PTree.empty _) in
- let (p2, h') := encode_expression 1%positive pe2 h in
- equiv_check p1 p2.
-
- Lemma mk_pred_stmnt_equiv' :
- forall l l' e p,
- mk_pred_stmnt_l l == mk_pred_stmnt_l l' ->
- In (e, p) l ->
- list_norepet (map fst l) ->
- (exists p', In (e, p') l' /\ p == p')
- \/ ~ In e (map fst l') /\ p == ⟂.
- Proof. Abort.
-
- Lemma mk_pred_stmnt_equiv :
- forall tree tree',
- mk_pred_stmnt tree == mk_pred_stmnt tree'.
- Proof. Abort.
-
- Definition tree_equiv_check_el (np2: PTree.t pred_op) (n: positive) (p: pred_op): bool :=
- match np2 ! n with
- | Some p' => equiv_check p p'
- | None => equiv_check p ⟂
- end.
-
- Definition tree_equiv_check_None_el (np2: PTree.t pred_op) (n: positive) (p: pred_op): bool :=
- match np2 ! n with
- | Some p' => true
- | None => equiv_check p ⟂
- end.
-
- Definition beq_pred_expr (pe1 pe2: predicated HS.t) : bool :=
- if pred_expr_dec pe1 pe2 then true else
- let (np1, h) := norm_expression 1 pe1 (PTree.empty _) in
- let (np2, h') := norm_expression 1 pe2 h in
- forall_ptree (tree_equiv_check_el np2) np1
- && forall_ptree (tree_equiv_check_None_el np1) np2.
-
- Lemma beq_pred_expr_correct:
- forall np1 np2 e p p',
- forall_ptree (tree_equiv_check_el np2) np1 = true ->
- np1 ! e = Some p ->
- np2 ! e = Some p' ->
- p == p'.
- Proof.
- intros.
- eapply forall_ptree_true in H; try eassumption.
- unfold tree_equiv_check_el in H.
- rewrite H1 in H. now apply equiv_check_correct.
- Qed.
-
- Lemma beq_pred_expr_correct2:
- forall np1 np2 e p,
- forall_ptree (tree_equiv_check_el np2) np1 = true ->
- np1 ! e = Some p ->
- np2 ! e = None ->
- p == ⟂.
- Proof.
- intros.
- eapply forall_ptree_true in H; try eassumption.
- unfold tree_equiv_check_el in H.
- rewrite H1 in H. now apply equiv_check_correct.
- Qed.
-
- Lemma beq_pred_expr_correct3:
- forall np1 np2 e p,
- forall_ptree (tree_equiv_check_None_el np1) np2 = true ->
- np1 ! e = None ->
- np2 ! e = Some p ->
- p == ⟂.
- Proof.
- intros.
- eapply forall_ptree_true in H; try eassumption.
- unfold tree_equiv_check_None_el in H.
- rewrite H0 in H. now apply equiv_check_correct.
- Qed.
-
- Section PRED_PROOFS.
-
- Context {G B: Type}.
- Context (f: PTree.t pred_pexpr).
- Context (ps: PMap.t bool).
- Context (a_sem: @Abstr.ctx G -> HS.t -> B -> Prop).
- Context (ctx: @Abstr.ctx G).
-
- Context (F_EVALULABLE: forall x, sem_pexpr ctx (get_forest_p' x f) ps !! x).
-
- Variant sem_pred_tree: PTree.t HS.t -> PTree.t pred_op -> B -> Prop :=
- | sem_pred_tree_intro :
- forall expr e v et pt pr,
- sem_pexpr ctx (from_pred_op f pr) true ->
- a_sem ctx expr v ->
- pt ! e = Some pr ->
- et ! e = Some expr ->
- sem_pred_tree et pt v.
-
- Lemma norm_expression_in :
- forall pe et pt h x y,
- H.wf_hash_table h ->
- norm_expression 1 pe h = (pt, et) ->
- h ! x = Some y ->
- et ! x = Some y.
- Proof.
- induction pe; crush; repeat (destruct_match; try discriminate; []).
- - inv H0. eauto using H.hash_constant.
- - destruct_match.
- + inv H0. eapply IHpe.
- eapply H.wf_hash_table_distr; eauto. eauto.
- eauto using H.hash_constant.
- + inv H0. eapply IHpe.
- eapply H.wf_hash_table_distr; eauto. eauto.
- eauto using H.hash_constant.
- Qed.
-
- Lemma norm_expression_exists :
- forall pe et pt h x y,
- H.wf_hash_table h ->
- norm_expression 1 pe h = (pt, et) ->
- pt ! x = Some y ->
- exists z, et ! x = Some z.
- Proof.
- induction pe; crush; repeat (destruct_match; try discriminate; []).
- - inv H0. destruct (peq x h0); subst; inv H1.
- + eexists. eauto using H.hash_value_in.
- + rewrite PTree.gso in H2 by auto. now rewrite PTree.gempty in H2.
- - assert (H.wf_hash_table h1) by eauto using H.wf_hash_table_distr.
- destruct_match; inv H0.
- + destruct (peq h0 x); subst; eauto.
- rewrite PTree.gso in H1 by auto. eauto.
- + destruct (peq h0 x); subst; eauto.
- * pose proof Heqp0 as X.
- eapply H.hash_value_in in Heqp0.
- eapply norm_expression_in in Heqn; eauto.
- * rewrite PTree.gso in H1 by auto. eauto.
- Qed.
-
- Lemma norm_expression_wf :
- forall pe et pt h,
- H.wf_hash_table h ->
- norm_expression 1 pe h = (pt, et) ->
- H.wf_hash_table et.
- Proof.
- induction pe; crush; repeat (destruct_match; try discriminate; []).
- - inv H0. eauto using H.wf_hash_table_distr.
- - destruct_match.
- + inv H0. eapply IHpe.
- eapply H.wf_hash_table_distr; eauto. eauto.
- + inv H0. eapply IHpe.
- eapply H.wf_hash_table_distr; eauto. eauto.
- Qed.
-
- Lemma sem_pred_expr_in_true :
- forall pe v,
- sem_pred_expr f a_sem ctx pe v ->
- exists p e, NE.In (p, e) pe
- /\ sem_pexpr ctx (from_pred_op f p) true
- /\ a_sem ctx e v.
- Proof.
- induction pe; crush.
- - inv H. do 2 eexists; split; try split; eauto. constructor.
- - inv H.
- + do 2 eexists; split; try split; eauto. constructor; tauto.
- + exploit IHpe; eauto. simplify.
- do 2 eexists; split; try split; eauto. constructor; tauto.
- Qed.
-
- Definition pred_Ht_dec :
- forall x y: pred_op * HS.t, {x = y} + {x <> y}.
- Proof.
- pose proof HS.eq_dec.
- pose proof (@Predicate.eq_dec positive peq).
- decide equality.
- Defined.
-
- Lemma sem_pred_mutexcl :
- forall pe p t v,
- predicated_mutexcl ((p, t) ::| pe) ->
- sem_pred_expr f a_sem ctx pe v ->
- sem_pexpr ctx (from_pred_op f p) false.
- Proof.
- intros. unfold predicated_mutexcl in H.
- exploit sem_pred_expr_in_true; eauto; simplify.
- unfold "⇒" in *. inv H5.
- destruct (pred_Ht_dec (x, x0) (p, t)); subst.
- { inv e; exfalso; apply H7; auto. }
- assert (NE.In (x, x0) ((p, t) ::| pe)) by (constructor; tauto).
- assert (NE.In (p, t) ((p, t) ::| pe)) by (constructor; tauto).
- pose proof (H3 _ _ H H5 n).
- assert (forall c, eval_predf c x = true -> eval_predf c (¬ p) = true)
- by eauto.
- eapply sem_pexpr_eval_inv in H1; eauto.
- eapply sem_pexpr_eval; eauto. apply H9 in H1.
- unfold eval_predf in *. rewrite negate_correct in H1.
- symmetry in H1. apply negb_sym in H1. auto.
- Qed.
-
- Lemma exec_tree_exec_pe :
- forall pe et pt v h
- (MUTEXCL: predicated_mutexcl pe),
- H.wf_hash_table h ->
- norm_expression 1 pe h = (pt, et) ->
- sem_pred_tree et pt v ->
- sem_pred_expr f a_sem ctx pe v.
- Proof.
- induction pe; simplify; repeat (destruct_match; try discriminate; []).
- - inv Heqp. inv H0. inv H1.
- destruct (peq e h0); subst.
- 2: { rewrite PTree.gso in H3 by auto.
- rewrite PTree.gempty in H3. discriminate. }
- assert (expr = t).
- { apply H.hash_value_in in Heqp0. rewrite H4 in Heqp0. now inv Heqp0. }
- subst. constructor; auto. rewrite PTree.gss in H3. inv H3; auto.
- - inv Heqp. inv H1. destruct_match; inv H0; destruct (peq h0 e); subst.
- + rewrite PTree.gss in H4. inv H4. inv H2. inv H1.
- * exploit IHpe. eauto using predicated_cons.
- eapply H.wf_hash_table_distr; eauto. eauto.
- econstructor. eauto. eauto. eauto. eauto. intros.
- assert (sem_pexpr ctx (from_pred_op f p) false)
- by (eapply sem_pred_mutexcl; eauto).
- eapply sem_pred_expr_cons_false; auto.
- * assert (et ! e = Some t).
- { eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
- eauto. apply H.hash_value_in in Heqp0. auto. }
- rewrite H1 in H5. inv H5.
- constructor; auto.
- + exploit IHpe. eauto using predicated_cons.
- eapply H.wf_hash_table_distr; eauto. eauto.
- econstructor. eauto. eauto. rewrite PTree.gso in H4; eauto. auto.
- intros.
- assert (sem_pexpr ctx (from_pred_op f p) false)
- by (eapply sem_pred_mutexcl; eauto).
- eapply sem_pred_expr_cons_false; auto.
- + rewrite PTree.gss in H4. inv H4.
- econstructor; auto.
- assert (et ! e = Some t).
- { eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
- eauto. apply H.hash_value_in in Heqp0. auto. }
- rewrite H0 in H5; inv H5. auto.
- + rewrite PTree.gso in H4 by auto.
- exploit IHpe. eauto using predicated_cons.
- eapply H.wf_hash_table_distr; eauto. eauto.
- econstructor. eauto. eauto. eauto. eauto. intros.
- assert (sem_pexpr ctx (from_pred_op f p) false)
- by (eapply sem_pred_mutexcl; eauto).
- eapply sem_pred_expr_cons_false; auto.
- Qed.
-
- Lemma exec_pe_exec_tree :
- forall pe et pt v h
- (MUTEXCL: predicated_mutexcl pe),
- H.wf_hash_table h ->
- norm_expression 1 pe h = (pt, et) ->
- sem_pred_expr f a_sem ctx pe v ->
- sem_pred_tree et pt v.
- Proof.
- induction pe; simplify; repeat (destruct_match; try discriminate; []).
- - inv H0. inv H1. econstructor; eauto. apply PTree.gss.
- eapply H.hash_value_in; eauto.
- - inv H1.
- + destruct_match.
- * inv H0. econstructor.
- 2: { eauto. }
- 2: { apply PTree.gss. }
- constructor; tauto.
- eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
- eauto. eapply H.hash_value_in; eauto.
- * inv H0. econstructor. eauto. eauto. apply PTree.gss.
- eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
- eauto. eapply H.hash_value_in; eauto.
- + destruct_match.
- * inv H0. exploit IHpe.
- eauto using predicated_cons.
- eapply H.wf_hash_table_distr; eauto.
- eauto. eauto. intros. inv H0.
- destruct (peq e h0); subst.
- -- rewrite H3 in Heqo. inv Heqo.
- econstructor.
- 3: { apply PTree.gss. }
- constructor; tauto. eauto. auto.
- -- econstructor. eauto. eauto. rewrite PTree.gso by eauto. auto.
- auto.
- * inv H0. exploit IHpe.
- eauto using predicated_cons.
- eapply H.wf_hash_table_distr; eauto.
- eauto. eauto. intros. inv H0.
- destruct (peq e h0); subst.
- -- rewrite H3 in Heqo; discriminate.
- -- econstructor; eauto. rewrite PTree.gso by auto. auto.
- Qed.
-
- Lemma beq_pred_expr_correct_top:
- forall p1 p2 v
- (MUTEXCL1: predicated_mutexcl p1)
- (MUTEXCL2: predicated_mutexcl p2),
- beq_pred_expr p1 p2 = true ->
- sem_pred_expr f a_sem ctx p1 v ->
- sem_pred_expr f a_sem ctx p2 v.
- Proof.
- unfold beq_pred_expr; intros.
- destruct_match; subst; auto.
- repeat (destruct_match; []).
- symmetry in H. apply andb_true_eq in H. inv H.
- symmetry in H1. symmetry in H2.
- pose proof Heqn0. eapply norm_expression_wf in H.
- 2: { unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3. }
- eapply exec_tree_exec_pe; eauto.
- eapply exec_pe_exec_tree in H0; auto.
- 3: { eauto. }
- 2: { unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3. }
- inv H0. destruct (t0 ! e) eqn:?.
- - assert (pr == p) by eauto using beq_pred_expr_correct.
- assert (sem_pexpr ctx (from_pred_op f p) true).
- { eapply sem_pexpr_eval; eauto. eapply sem_pexpr_eval_inv in H3; eauto. }
- econstructor. apply H7. eauto. eauto.
- eapply norm_expression_in; eauto.
- - assert (pr == ⟂) by eauto using beq_pred_expr_correct2.
- eapply sem_pexpr_eval_inv in H3; eauto. now rewrite H0 in H3.
- Qed.
-
- Lemma beq_pred_expr_correct_top2:
- forall p1 p2 v
- (MUTEXCL1: predicated_mutexcl p1)
- (MUTEXCL2: predicated_mutexcl p2),
- beq_pred_expr p1 p2 = true ->
- sem_pred_expr f a_sem ctx p2 v ->
- sem_pred_expr f a_sem ctx p1 v.
- Proof.
- unfold beq_pred_expr; intros.
- destruct_match; subst; auto.
- repeat (destruct_match; []).
- symmetry in H. apply andb_true_eq in H. inv H.
- symmetry in H1. symmetry in H2.
- pose proof Heqn0. eapply norm_expression_wf in H.
- 2: { unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3. }
- eapply exec_tree_exec_pe; auto.
- 2: { eauto. }
- unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3.
- eapply exec_pe_exec_tree in H0; auto.
- 3: { eauto. }
- 2: { auto. }
- inv H0. destruct (t ! e) eqn:?.
- - assert (p == pr) by eauto using beq_pred_expr_correct.
- assert (sem_pexpr ctx (from_pred_op f p) true).
- { eapply sem_pexpr_eval; eauto. eapply sem_pexpr_eval_inv in H3; eauto. }
- econstructor. apply H7. eauto. eauto.
- exploit norm_expression_exists.
- 2: { eapply Heqn0. } unfold H.wf_hash_table; intros * YH.
- now rewrite PTree.gempty in YH. eauto. simplify.
- exploit norm_expression_in. 2: { eauto. } auto. eauto. intros.
- crush.
- - assert (pr == ⟂) by eauto using beq_pred_expr_correct3.
- eapply sem_pexpr_eval_inv in H3; eauto. now rewrite H0 in H3.
- Qed.
-
- End PRED_PROOFS.
-
-End HashExprNorm.
-
-Module HN := HashExprNorm(HashExpr).
-Module EHN := HashExprNorm(EHashExpr).
-
-Definition and_list {A} (p: list (@Predicate.pred_op A)): @Predicate.pred_op A :=
- fold_left Pand p T.
-
-Definition or_list {A} (p: list (@Predicate.pred_op A)): @Predicate.pred_op A :=
- fold_left Por p ⟂.
-
-Definition check_mutexcl {A} (pe: predicated A) :=
- let preds := map fst (NE.to_list pe) in
- let pairs := map (fun x => x → or_list (remove (Predicate.eq_dec peq) x preds)) preds in
- match sat_pred_simple (simplify (negate (and_list pairs))) with
- | None => true
- | _ => false
- end.
-
-Lemma check_mutexcl_correct:
- forall A (pe: predicated A),
- check_mutexcl pe = true ->
- predicated_mutexcl pe.
-Proof. Admitted.
-
-Definition check_mutexcl_tree {A} (f: PTree.t (predicated A)) :=
- forall_ptree (fun _ => check_mutexcl) f.
-
-Lemma check_mutexcl_tree_correct:
- forall A (f: PTree.t (predicated A)) i pe,
- check_mutexcl_tree f = true ->
- f ! i = Some pe ->
- predicated_mutexcl pe.
-Proof.
- unfold check_mutexcl_tree; intros.
- eapply forall_ptree_true in H; eauto using check_mutexcl_correct.
-Qed.
-
-Definition check f1 f2 :=
- RTree.beq HN.beq_pred_expr f1.(forest_regs) f2.(forest_regs)
- && PTree.beq beq_pred_pexpr f1.(forest_preds) f2.(forest_preds)
- && EHN.beq_pred_expr f1.(forest_exit) f2.(forest_exit)
- && check_mutexcl_tree f1.(forest_regs)
- && check_mutexcl_tree f2.(forest_regs)
- && check_mutexcl f1.(forest_exit)
- && check_mutexcl f2.(forest_exit).
-
-Lemma sem_pexpr_forward_eval1 :
- forall A ctx f_p ps,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p,
- @sem_pexpr A ctx (from_pred_op f_p p) false ->
- eval_predf ps p = false.
-Proof.
- induction p; crush.
- - destruct_match. inv Heqp0. destruct b.
- cbn. specialize (H p0).
- eapply sem_pexpr_det; eauto. apply similar_refl.
- specialize (H p0).
- replace false with (negb true) in H0 by auto.
- eapply sem_pexpr_negate2 in H0. cbn.
- symmetry; apply negb_sym. cbn.
- eapply sem_pexpr_det; eauto. apply similar_refl.
- - inv H0.
- - inv H0. inv H2. rewrite eval_predf_Pand. rewrite IHp1; eauto.
- rewrite eval_predf_Pand. rewrite IHp2; eauto with bool.
- - inv H0. rewrite eval_predf_Por. rewrite IHp1; eauto.
-Qed.
-
-Lemma sem_pexpr_forward_eval2 :
- forall A ctx f_p ps,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p,
- @sem_pexpr A ctx (from_pred_op f_p p) true ->
- eval_predf ps p = true.
-Proof.
- induction p; crush.
- - destruct_match. inv Heqp0. destruct b.
- cbn. specialize (H p0).
- eapply sem_pexpr_det; eauto. apply similar_refl.
- cbn. symmetry. apply negb_sym; cbn.
- replace true with (negb false) in H0 by auto.
- eapply sem_pexpr_negate2 in H0.
- eapply sem_pexpr_det; eauto. apply similar_refl.
- - inv H0.
- - inv H0. rewrite eval_predf_Pand. rewrite IHp1; eauto.
- - inv H0. inv H2. rewrite eval_predf_Por. rewrite IHp1; eauto.
- rewrite eval_predf_Por. rewrite IHp2; eauto with bool.
-Qed.
-
-Lemma sem_pexpr_forward_eval :
- forall A ctx f_p ps,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- forall p b,
- @sem_pexpr A ctx (from_pred_op f_p p) b ->
- eval_predf ps p = b.
-Proof.
- intros; destruct b; eauto using sem_pexpr_forward_eval1, sem_pexpr_forward_eval2.
-Qed.
-
-Lemma get_empty:
- forall r, empty#r r = NE.singleton (Ptrue, Ebase r).
-Proof. unfold "#r"; intros. rewrite RTree.gempty. trivial. Qed.
-
-Section BOOLEAN_EQUALITY.
-
- Context {A B: Type}.
- Context (beqA: A -> B -> bool).
-
- Fixpoint beq2' (m1: PTree.tree' A) (m2: PTree.tree' B) {struct m1} : bool :=
- match m1, m2 with
- | PTree.Node001 r1, PTree.Node001 r2 => beq2' r1 r2
- | PTree.Node010 x1, PTree.Node010 x2 => beqA x1 x2
- | PTree.Node011 x1 r1, PTree.Node011 x2 r2 => beqA x1 x2 && beq2' r1 r2
- | PTree.Node100 l1, PTree.Node100 l2 => beq2' l1 l2
- | PTree.Node101 l1 r1, PTree.Node101 l2 r2 => beq2' l1 l2 && beq2' r1 r2
- | PTree.Node110 l1 x1, PTree.Node110 l2 x2 => beqA x1 x2 && beq2' l1 l2
- | PTree.Node111 l1 x1 r1, PTree.Node111 l2 x2 r2 => beqA x1 x2 && beq2' l1 l2 && beq2' r1 r2
- | _, _ => false
- end.
-
- Definition beq2 (m1: PTree.t A) (m2 : PTree.t B) : bool :=
- match m1, m2 with
- | PTree.Empty, PTree.Empty => true
- | PTree.Nodes m1', PTree.Nodes m2' => beq2' m1' m2'
- | _, _ => false
- end.
-
- Let beq2_optA (o1: option A) (o2: option B) : bool :=
- match o1, o2 with
- | None, None => true
- | Some a1, Some a2 => beqA a1 a2
- | _, _ => false
- end.
-
- Lemma beq_correct_bool:
- forall m1 m2,
- beq2 m1 m2 = true <-> (forall x, beq2_optA (m1 ! x) (m2 ! x) = true).
- Proof.
- Local Transparent PTree.Node.
- assert (beq_NN: forall l1 o1 r1 l2 o2 r2,
- PTree.not_trivially_empty l1 o1 r1 ->
- PTree.not_trivially_empty l2 o2 r2 ->
- beq2 (PTree.Node l1 o1 r1) (PTree.Node l2 o2 r2) =
- beq2 l1 l2 && beq2_optA o1 o2 && beq2 r1 r2).
- { intros.
- destruct l1, o1, r1; try contradiction; destruct l2, o2, r2; try contradiction;
- simpl; rewrite ? andb_true_r, ? andb_false_r; auto.
- rewrite andb_comm; auto.
- f_equal; rewrite andb_comm; auto. }
- induction m1 using PTree.tree_ind; [|induction m2 using PTree.tree_ind].
- - intros. simpl; split; intros.
- + destruct m2; try discriminate. simpl; auto.
- + replace m2 with (@PTree.Empty B); auto. apply PTree.extensionality; intros x.
- specialize (H x). destruct (m2 ! x); simpl; easy.
- - split; intros.
- + destruct (PTree.Node l o r); try discriminate. simpl; auto.
- + replace (PTree.Node l o r) with (@PTree.Empty A); auto. apply PTree.extensionality; intros x.
- specialize (H0 x). destruct ((PTree.Node l o r) ! x); simpl in *; easy.
- - rewrite beq_NN by auto. split; intros.
- + InvBooleans. rewrite ! PTree.gNode. destruct x.
- * apply IHm0; auto.
- * apply IHm1; auto.
- * auto.
- + apply andb_true_intro; split; [apply andb_true_intro; split|].
- * apply IHm1. intros. specialize (H1 (xO x)); rewrite ! PTree.gNode in H1; auto.
- * specialize (H1 xH); rewrite ! PTree.gNode in H1; auto.
- * apply IHm0. intros. specialize (H1 (xI x)); rewrite ! PTree.gNode in H1; auto.
- Qed.
-
- Theorem beq2_correct:
- forall m1 m2,
- beq2 m1 m2 = true <->
- (forall (x: PTree.elt),
- match m1 ! x, m2 ! x with
- | None, None => True
- | Some y1, Some y2 => beqA y1 y2 = true
- | _, _ => False
- end).
- Proof.
- intros. rewrite beq_correct_bool. unfold beq2_optA. split; intros.
- - specialize (H x). destruct (m1 ! x), (m2 ! x); intuition congruence.
- - specialize (H x). destruct (m1 ! x), (m2 ! x); intuition auto.
- Qed.
-
-End BOOLEAN_EQUALITY.
-
-Lemma forest_reg_gso:
- forall (f : forest) w dst dst',
- dst <> dst' ->
- (f #r dst <- w) #r dst' = f #r dst'.
-Proof.
- unfold "#r"; intros.
- unfold forest_regs. unfold set_forest.
- rewrite RTree.gso; auto.
-Qed.
-
-Lemma forest_reg_pred:
- forall (f : forest) w dst dst',
- (f #r dst <- w) #p dst' = f #p dst'.
-Proof. auto. Qed.
-
-Lemma forest_reg_gss:
- forall (f : forest) w dst,
- (f #r dst <- w) #r dst = w.
-Proof.
- unfold "#r"; intros.
- unfold forest_regs. unfold set_forest.
- rewrite RTree.gss; auto.
-Qed.
-
-Lemma forest_pred_gso:
- forall (f : forest) w dst dst',
- dst <> dst' ->
- (f #p dst <- w) #p dst' = f #p dst'.
-Proof.
- unfold "#p"; intros.
- unfold forest_preds, set_forest_p, get_forest_p'.
- rewrite PTree.gso; auto.
-Qed.
-
-Lemma forest_pred_reg:
- forall (f : forest) w dst dst',
- (f #p dst <- w) #r dst' = f #r dst'.
-Proof. auto. Qed.
-
-Lemma forest_pred_gss:
- forall (f : forest) w dst,
- (f #p dst <- w) #p dst = w.
-Proof.
- unfold "#p"; intros.
- unfold forest_preds, set_forest_p, get_forest_p'.
- rewrite PTree.gss; auto.
-Qed.
-
-Lemma forest_pred_gss2 :
- forall (f : forest) r p,
- (forest_preds (f #p r <- p)) ! r = Some p.
-Proof. intros. unfold set_forest_p. simpl. now rewrite PTree.gss. Qed.
-
-Lemma forest_pred_gso2 :
- forall (f : forest) r w p,
- r <> w ->
- (forest_preds (f #p w <- p)) ! r = (forest_preds f) ! r.
-Proof. intros. unfold set_forest_p. simpl. now rewrite PTree.gso by auto. Qed.
-
-Section GENERIC_CONTEXT.
-
-Context {A: Type}.
-Context (ctx: @ctx A).
-
-(*|
-Suitably restrict the predicate set so that one can evaluate a hashed predicate
-using that predicate set. However, one issue might be that we do not know that
-all the atoms of the original formula are actually evaluable.
-|*)
-
-Definition match_pred_states
- (ht: PHT.hash_tree) (p_out: pred_op) (pred_set: predset) :=
- forall (p: positive) (br: bool) (p_in: pred_expression),
- PredIn p p_out ->
- ht ! p = Some p_in ->
- sem_pred ctx p_in (pred_set !! p).
-
-Lemma eval_hash_predicate :
- forall max p_in ht p_out ht' br pred_set,
- hash_predicate max p_in ht = (p_out, ht') ->
- sem_pexpr ctx p_in br ->
- match_pred_states ht' p_out pred_set ->
- eval_predf pred_set p_out = br.
-Proof.
- induction p_in; simplify.
- + repeat destruct_match. inv H.
- unfold eval_predf. cbn.
- inv H0. inv H4. unfold match_pred_states in H1.
- specialize (H1 h br).
-Abort.
-
-Lemma sem_pexpr_beq_correct :
- forall p1 p2 b,
- beq_pred_pexpr p1 p2 = true ->
- sem_pexpr ctx p1 b ->
- sem_pexpr ctx p2 b.
-Proof.
- unfold beq_pred_pexpr.
- induction p1; intros; destruct_match; crush.
- Admitted.
-
-(*|
-This should only require a proof of sem_pexpr_beq_correct, the rest is
-straightforward.
-|*)
-
-Lemma pred_pexpr_beq_pred_pexpr :
- forall pr a b br,
- PTree.beq beq_pred_pexpr a b = true ->
- sem_pexpr ctx (from_pred_op a pr) br ->
- sem_pexpr ctx (from_pred_op b pr) br.
-Proof.
- induction pr; crush.
- - destruct_match. inv Heqp0. destruct b0.
- + unfold get_forest_p' in *.
- apply PTree.beq_correct with (x := p0) in H.
- destruct a ! p0; destruct b ! p0; try (exfalso; assumption); auto.
- eapply sem_pexpr_beq_correct; eauto.
- + replace br with (negb (negb br)) by (now apply negb_involutive).
- replace br with (negb (negb br)) in H0 by (now apply negb_involutive).
- apply sem_pexpr_negate. apply sem_pexpr_negate2 in H0.
- unfold get_forest_p' in *.
- apply PTree.beq_correct with (x := p0) in H.
- destruct a ! p0; destruct b ! p0; try (exfalso; assumption); auto.
- eapply sem_pexpr_beq_correct; eauto.
- - inv H0; try inv H4.
- + eapply IHpr1 in H0; eauto. apply sem_pexpr_Pand_false; tauto.
- + eapply IHpr2 in H0; eauto. apply sem_pexpr_Pand_false; tauto.
- + eapply IHpr1 in H3; eauto. eapply IHpr2 in H5; eauto.
- apply sem_pexpr_Pand_true; auto.
- - inv H0; try inv H4.
- + eapply IHpr1 in H0; eauto. apply sem_pexpr_Por_true; tauto.
- + eapply IHpr2 in H0; eauto. apply sem_pexpr_Por_true; tauto.
- + eapply IHpr1 in H3; eauto. eapply IHpr2 in H5; eauto.
- apply sem_pexpr_Por_false; auto.
-Qed.
-
-(*|
-This lemma requires a theorem that equivalence of symbolic predicates means they
-will be. This further needs three-valued logic to be able to compare arbitrary
-predicates with each other, that will also show that the predicate will also be
-computable.
-|*)
-
-Lemma sem_pred_exec_beq_correct :
- forall A B (sem: Abstr.ctx -> A -> B -> Prop) p a b r,
- PTree.beq beq_pred_pexpr a b = true ->
- sem_pred_expr a sem ctx p r ->
- sem_pred_expr b sem ctx p r.
-Proof.
- induction p; intros; inv H0.
- - constructor; auto. eapply pred_pexpr_beq_pred_pexpr; eauto.
- - constructor; auto. eapply pred_pexpr_beq_pred_pexpr; eauto.
- - apply sem_pred_expr_cons_false; eauto.
- eapply pred_pexpr_beq_pred_pexpr; eauto.
-Qed.
-
-End GENERIC_CONTEXT.
-
-Lemma tree_beq_pred_pexpr :
- forall a b x,
- RTree.beq beq_pred_pexpr (forest_preds a) (forest_preds b) = true ->
- beq_pred_pexpr a #p x b #p x = true.
-Proof.
- intros. unfold "#p". unfold get_forest_p'.
- apply PTree.beq_correct with (x := x) in H.
- destruct_match; destruct_match; auto.
- unfold beq_pred_pexpr. destruct_match; auto.
-Qed.
-
-Lemma tree_beq_pred_expr :
- forall a b x,
- RTree.beq HN.beq_pred_expr (forest_regs a) (forest_regs b) = true ->
- HN.beq_pred_expr a #r x b #r x = true.
-Proof.
- intros. unfold "#r" in *.
- apply PTree.beq_correct with (x := (R_indexed.index x)) in H.
- unfold RTree.get in *.
- unfold pred_expr in *.
- destruct_match; destruct_match; auto.
- unfold HN.beq_pred_expr. destruct_match; auto.
-Qed.
-
-Section CORRECT.
-
-Context {FUN TFUN: Type}.
-Context (ictx: @ctx FUN) (octx: @ctx TFUN) (HSIM: similar ictx octx).
-
-Lemma abstr_check_correct :
- forall i' a b cf,
- (exists ps, forall x, sem_pexpr ictx (get_forest_p' x (forest_preds a)) ps !! x) ->
- check a b = true ->
- sem ictx a (i', cf) ->
- exists ti', sem octx b (ti', cf) /\ match_states i' ti'.
-Proof.
- intros * EVALUABLE **. unfold check in H. simplify.
- inv H0. inv H10. inv H11.
- eexists; split; constructor; auto.
- - constructor. intros.
- eapply sem_pred_exec_beq_correct; eauto.
- eapply sem_pred_expr_corr; eauto. now apply sem_value_corr.
- eapply HN.beq_pred_expr_correct_top; eauto.
- { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
- { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
- eapply tree_beq_pred_expr; eauto.
- - (* This is where three-valued logic would be needed. *)
- constructor. intros. apply sem_pexpr_beq_correct with (p1 := a #p x0).
- apply tree_beq_pred_pexpr; auto.
- apply sem_pexpr_corr with (ictx:=ictx); auto.
- - eapply sem_pred_exec_beq_correct; eauto.
- eapply sem_pred_expr_corr; eauto. now apply sem_mem_corr.
- eapply HN.beq_pred_expr_correct_top; eauto.
- { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
- { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
- eapply tree_beq_pred_expr; eauto.
- - eapply sem_pred_exec_beq_correct; eauto.
- eapply sem_pred_expr_corr; eauto. now apply sem_exit_corr.
- eapply EHN.beq_pred_expr_correct_top; eauto using check_mutexcl_correct.
-Qed.
-
-(*|
-Proof Sketch:
-
-Two abstract states can be equivalent, without it being obvious that they can
-actually both be executed assuming one can be executed. One will therefore have
-to add a few more assumptions to makes it possible to execute the other.
-
-It currently assumes that all the predicates in the predicate tree are
-evaluable, which is actually something that can either be checked, or something
-that can be proven constructively. I believe that it should already be possible
-using the latter, so here it will only be assumed.
-
-Similarly, the current assumption is that mutual exclusivity of predicates is
-being checked within the ``check`` function, which could possibly also be proven
-constructively about the update function. This is a simpler short-term fix
-though.
-|*)
-
-End CORRECT.
diff --git a/src/hls/GiblePargen.v b/src/hls/GiblePargen.v
index 3000211..2c7bfc8 100644
--- a/src/hls/GiblePargen.v
+++ b/src/hls/GiblePargen.v
@@ -36,6 +36,8 @@ Require Import vericert.hls.Gible.
Require Import vericert.hls.Predicate.
Require Import vericert.hls.Abstr.
Require Import vericert.common.DecEq.
+Require Import vericert.hls.GiblePargenproofEquiv.
+
Import NE.NonEmptyNotation.
#[local] Open Scope positive.
@@ -80,6 +82,31 @@ Definition predicated_prod {A B: Type} (p1: predicated A) (p2: predicated B) :=
Definition predicated_map {A B: Type} (f: A -> B) (p: predicated A)
: predicated B := NE.map (fun x => (fst x, f (snd x))) p.
+Lemma NEin_map :
+ forall A B p y (f: A -> B) a,
+ NE.In (p, y) (predicated_map f a) ->
+ exists x, NE.In (p, x) a /\ y = f x.
+Proof.
+ induction a; intros.
+ - inv H. destruct a. econstructor. split; eauto. constructor.
+ - inv H. inv H1. inv H. destruct a. cbn in *. econstructor; econstructor; eauto.
+ constructor; tauto.
+ specialize (IHa H). inv IHa. inv H0.
+ econstructor; econstructor; eauto. constructor; tauto.
+Qed.
+
+Lemma NEin_map2 :
+ forall A B (f: A -> B) a p y,
+ NE.In (p, y) a ->
+ NE.In (p, f y) (predicated_map f a).
+Proof.
+ induction a; crush.
+ inv H. constructor.
+ inv H. inv H1.
+ - constructor; auto.
+ - constructor; eauto.
+Qed.
+
Definition cons_pred_expr (pel: pred_expr) (tpel: predicated expression_list) :=
predicated_map (uncurry Econs) (predicated_prod pel tpel).
diff --git a/src/hls/GiblePargenproof.v b/src/hls/GiblePargenproof.v
index 0b9bd78..9e58e03 100644
--- a/src/hls/GiblePargenproof.v
+++ b/src/hls/GiblePargenproof.v
@@ -1,6 +1,6 @@
(*
* Vericert: Verified high-level synthesis.
- * Copyright (C) 2020-2023 Yann Herklotz <yann@yannherklotz.com>
+ * Copyright (C) 2020-2023 Yann Herklotz <git@yannherklotz.com>
*
* This program is free software: you can redistribute it and/or modify
* it under the terms of the GNU General Public License as published by
@@ -29,10 +29,12 @@ Require Import vericert.common.Vericertlib.
Require Import vericert.hls.GibleSeq.
Require Import vericert.hls.GiblePar.
Require Import vericert.hls.Gible.
+Require Import vericert.hls.GiblePargenproofEquiv.
Require Import vericert.hls.GiblePargen.
Require Import vericert.hls.Predicate.
Require Import vericert.hls.Abstr.
Require Import vericert.common.Monad.
+Require Import vericert.hls.GiblePargenproofForward.
Module Import OptionExtra := MonadExtra(Option).
@@ -45,168 +47,7 @@ Module Import OptionExtra := MonadExtra(Option).
Ltac destr := destruct_match; try discriminate; [].
-(*|
-==============
-RTLPargenproof
-==============
-
-RTLBlock to abstract translation
-================================
-
-Correctness of translation from RTLBlock to the abstract interpretation
-language.
-|*)
-
-Definition is_regs i := match i with mk_instr_state rs _ _ => rs end.
-Definition is_mem i := match i with mk_instr_state _ _ m => m end.
-Definition is_ps i := match i with mk_instr_state _ p _ => p end.
-
-Definition evaluable {A B C} (sem: ctx -> B -> C -> Prop) (ctx: @ctx A) p := exists b, sem ctx p b.
-
-Definition p_evaluable {A} := @evaluable A _ _ sem_pexpr.
-
-Definition evaluable_expr {A} := @evaluable A _ _ sem_pred.
-
-Definition all_evaluable {A B} (ctx: @ctx A) f_p (l: predicated B) :=
- forall p y, NE.In (p, y) l -> p_evaluable ctx (from_pred_op f_p p).
-
-Definition all_evaluable2 {A B C} (ctx: @ctx A) (sem: Abstr.ctx -> B -> C -> Prop) (l: predicated B) :=
- forall p y, NE.In (p, y) l -> evaluable sem ctx y.
-
-Definition pred_forest_evaluable {A} (ctx: @ctx A) (ps: PTree.t pred_pexpr) :=
- forall i p, ps ! i = Some p -> p_evaluable ctx p.
-
-Definition forest_evaluable {A} (ctx: @ctx A) (f: forest) :=
- pred_forest_evaluable ctx f.(forest_preds).
-
-(* Lemma all_evaluable2_NEmap : *)
-(* forall G A (ctx: @ctx G) (f: (pred_op * A) -> (pred_op * pred_expression)) (x: predicated A), *)
-(* all_evaluable2 ctx sem_pred (NE.map f x). *)
-(* Proof. *)
-(* induction x. *)
-
-Lemma all_evaluable_cons :
- forall A B pr ctx b a,
- all_evaluable ctx pr (NE.cons a b) ->
- @all_evaluable A B ctx pr b.
-Proof.
- unfold all_evaluable; intros.
- enough (NE.In (p, y) (NE.cons a b)); eauto.
- constructor; tauto.
-Qed.
-
-Lemma all_evaluable2_cons :
- forall A B C sem ctx b a,
- all_evaluable2 ctx sem (NE.cons a b) ->
- @all_evaluable2 A B C ctx sem b.
-Proof.
- unfold all_evaluable2; intros.
- enough (NE.In (p, y) (NE.cons a b)); eauto.
- constructor; tauto.
-Qed.
-
-Lemma all_evaluable_cons2 :
- forall A B pr ctx b a p,
- @all_evaluable A B ctx pr (NE.cons (p, a) b) ->
- p_evaluable ctx (from_pred_op pr p).
-Proof.
- unfold all_evaluable; intros.
- eapply H. constructor; eauto.
-Qed.
-
-Lemma all_evaluable2_cons2 :
- forall A B C sem ctx b a p,
- @all_evaluable2 A B C ctx sem (NE.cons (p, a) b) ->
- evaluable sem ctx a.
-Proof.
- unfold all_evaluable; intros.
- eapply H. constructor; eauto.
-Qed.
-
-Lemma all_evaluable_cons3 :
- forall A B pr ctx b p a,
- all_evaluable ctx pr b ->
- p_evaluable ctx (from_pred_op pr p) ->
- @all_evaluable A B ctx pr (NE.cons (p, a) b).
-Proof.
- unfold all_evaluable; intros. inv H1. inv H3. inv H1. auto.
- eauto.
-Qed.
-
-Lemma all_evaluable_singleton :
- forall A B pr ctx b p,
- @all_evaluable A B ctx pr (NE.singleton (p, b)) ->
- p_evaluable ctx (from_pred_op pr p).
-Proof.
- unfold all_evaluable; intros. eapply H. constructor.
-Qed.
-
-Lemma get_forest_p_evaluable :
- forall A (ctx: @ctx A) f r,
- forest_evaluable ctx f ->
- p_evaluable ctx (f #p r).
-Proof.
- intros. unfold "#p", get_forest_p'.
- destruct_match. unfold forest_evaluable in *.
- unfold pred_forest_evaluable in *. eauto.
- unfold p_evaluable, evaluable. eexists.
- constructor. constructor.
-Qed.
-
-Lemma set_forest_p_evaluable :
- forall A (ctx: @ctx A) f r p,
- forest_evaluable ctx f ->
- p_evaluable ctx p ->
- forest_evaluable ctx (f #p r <- p).
-Proof.
- unfold forest_evaluable, pred_forest_evaluable; intros.
- destruct (peq i r); subst.
- - rewrite forest_pred_gss2 in H1. now inv H1.
- - rewrite forest_pred_gso2 in H1 by auto; eauto.
-Qed.
-
-Definition check_dest i r' :=
- match i with
- | RBop p op rl r => (r =? r')%positive
- | RBload p chunk addr rl r => (r =? r')%positive
- | _ => false
- end.
-
-Lemma check_dest_dec i r :
- {check_dest i r = true} + {check_dest i r = false}.
-Proof. destruct (check_dest i r); tauto. Qed.
-
-Fixpoint check_dest_l l r :=
- match l with
- | nil => false
- | a :: b => check_dest a r || check_dest_l b r
- end.
-
-Lemma check_dest_l_forall :
- forall l r,
- check_dest_l l r = false ->
- Forall (fun x => check_dest x r = false) l.
-Proof. induction l; crush. Qed.
-
-Lemma check_dest_l_dec i r :
- {check_dest_l i r = true} + {check_dest_l i r = false}.
-Proof. destruct (check_dest_l i r); tauto. Qed.
-
-Lemma match_states_list :
- forall A (rs: Regmap.t A) rs',
- (forall r, rs !! r = rs' !! r) ->
- forall l, rs ## l = rs' ## l.
-Proof. induction l; crush. Qed.
-
-Lemma PTree_matches :
- forall A (v: A) res rs rs',
- (forall r, rs !! r = rs' !! r) ->
- forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x.
-Proof.
- intros; destruct (Pos.eq_dec x res); subst;
- [ repeat rewrite Regmap.gss by auto
- | repeat rewrite Regmap.gso by auto ]; auto.
-Qed.
+Definition state_lessdef := GiblePargenproofEquiv.match_states.
Definition match_prog (prog : GibleSeq.program) (tprog : GiblePar.program) :=
match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog.
@@ -355,13 +196,80 @@ Section CORRECTNESS.
Qed.
Hint Resolve eval_addressing_eq : rtlgp.
- Lemma ge_preserved_lem:
+(*|
+==============
+RTLPargenproof
+==============
+
+RTLBlock to abstract translation
+================================
+
+Correctness of translation from RTLBlock to the abstract interpretation
+language.
+|*)
+
+ Lemma abstr_seq_reverse_correct :
+ forall sp x i i' ti cf x',
+ abstract_sequence x = Some x' ->
+ sem (mk_ctx i sp ge) x' (i', (Some cf)) ->
+ state_lessdef i ti ->
+ exists ti', SeqBB.step ge sp (Iexec ti) x (Iterm ti' cf)
+ /\ state_lessdef i' ti'.
+ Proof.
+
+(*|
+Proof Sketch:
+
+We do an induction over the list of instructions ``x``. This is trivial for the
+empty case and then for the inductive case we know that there exists an
+execution that matches the abstract execution, so we need to know that adding
+another instructions to it will still mean that the execution will result in the
+same value.
+
+Arithmetic operations will be a problem because we will have to show that these
+can be executed. However, this should mostly be a problem in the abstract state
+comparison, because there two abstract states can be equal without one being
+evaluable.
+|*)
+
+ Admitted.
+
+(*|
+This is the top-level lemma which uses the following proofs to complete the
+square:
+
+- ``abstr_sequence_correct``: This is the lemma that states the forward
+ translation form ``GibleSeq`` to ``Abstr`` was correct.
+- ``abstr_check_correct``: This is the lemma that states that if a check between
+ two ``Abstr`` programs succeeds, that they will also behave the same. This
+ depends on the SAT solver correctness, as the predicates might be
+ syntactically different to each other.
+- ``abstr_seq_reverse_correct``: This is the lemma that shows that the backwards
+ simulation between the abstract translation and the concrete execution also
+ holds. We only have a translation from the concrete into the abstract, but
+ then prove that if we have an execution in the abstract, we can observe that
+ same execution in the concrete.
+- ``seqbb_step_parbb_step``: Finally, this lemma states that the parallel
+ execution of the basic block is equivalent to the sequential execution of the
+ concatenation of that parallel block. This is because even in the translation
+ to HTL, the Verilog semantics are sequential within a clock cycle, but will
+ then be parallelised by the synthesis tool. The argument for why this is
+ still useful is because we are identifying and scheduling instructions into
+ clock cycles.
+|*)
+
+ Definition local_abstr_check_correct :=
+ @abstr_check_correct GibleSeq.fundef GiblePar.fundef.
+
+ Definition local_abstr_check_correct2 :=
+ @abstr_check_correct GibleSeq.fundef GibleSeq.fundef.
+
+ Lemma ge_preserved_local :
ge_preserved ge tge.
- Proof using TRANSL.
- unfold ge_preserved.
- eauto with rtlgp.
+ Proof.
+ unfold ge_preserved;
+ eauto using eval_op_eq, eval_addressing_eq.
Qed.
- Hint Resolve ge_preserved_lem : rtlgp.
Lemma lessdef_regmap_optget:
forall or rs rs',
@@ -533,1848 +441,13 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed.
eapply IHx; eauto.
Qed.
- Lemma eval_predf_negate :
- forall ps p,
- eval_predf ps (negate p) = negb (eval_predf ps p).
- Proof.
- unfold eval_predf; intros. rewrite negate_correct. auto.
- Qed.
-
- Lemma is_truthy_negate :
- forall ps p pred,
- truthy ps p ->
- falsy ps (combine_pred (Some (negate (Option.default T p))) pred).
- Proof.
- inversion 1; subst; simplify.
- - destruct pred; constructor; auto.
- - destruct pred; constructor.
- rewrite eval_predf_Pand. rewrite eval_predf_negate. rewrite H0. auto.
- rewrite eval_predf_negate. rewrite H0. auto.
- Qed.
-
- Lemma sem_pred_expr_NEapp :
- forall A B C sem f_p ctx a b v,
- sem_pred_expr f_p sem ctx a v ->
- @sem_pred_expr A B C f_p sem ctx (NE.app a b) v.
- Proof.
- induction a; crush.
- - inv H. constructor; auto.
- - inv H. constructor; eauto.
- eapply sem_pred_expr_cons_false; eauto.
- Qed.
-
- Lemma sem_pred_expr_NEapp2 :
- forall A B C sem f_p ctx a b v ps,
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) (ps !! x)) ->
- (forall x, NE.In x a -> eval_predf ps (fst x) = false) ->
- sem_pred_expr f_p sem ctx b v ->
- @sem_pred_expr A B C f_p sem ctx (NE.app a b) v.
- Proof.
- induction a; crush.
- - assert (IN: NE.In a (NE.singleton a)) by constructor.
- specialize (H0 _ IN). destruct a.
- eapply sem_pred_expr_cons_false; eauto. cbn [fst] in *.
- eapply sem_pexpr_eval; eauto.
- - assert (NE.In a (NE.cons a a0)) by (constructor; tauto).
- apply H0 in H2.
- destruct a. cbn [fst] in *.
- eapply sem_pred_expr_cons_false.
- eapply sem_pexpr_eval; eauto. eapply IHa; eauto.
- intros. eapply H0. constructor; tauto.
- Qed.
-
- Lemma sem_pred_expr_NEapp3 :
- forall A B C sem f_p ctx (a b: predicated B) v,
- (forall x, NE.In x a -> sem_pexpr ctx (from_pred_op f_p (fst x)) false) ->
- sem_pred_expr f_p sem ctx b v ->
- @sem_pred_expr A B C f_p sem ctx (NE.app a b) v.
- Proof.
- induction a; crush.
- - assert (IN: NE.In a (NE.singleton a)) by constructor.
- specialize (H _ IN). destruct a.
- eapply sem_pred_expr_cons_false; eauto.
- - assert (NE.In a (NE.cons a a0)) by (constructor; tauto).
- apply H in H1.
- destruct a. cbn [fst] in *.
- eapply sem_pred_expr_cons_false; auto.
- eapply IHa; eauto.
- intros. eapply H. constructor; tauto.
- Qed.
-
- Lemma sem_pred_expr_map_not :
- forall A p ps y x0,
- eval_predf ps p = false ->
- NE.In x0 (NE.map (fun x => (p ∧ fst x, snd x)) y) ->
- eval_predf ps (@fst _ A x0) = false.
- Proof.
- induction y; crush.
- - inv H0. simplify. rewrite eval_predf_Pand. rewrite H. auto.
- - inv H0. inv H2. cbn -[eval_predf]. rewrite eval_predf_Pand.
- rewrite H. auto. eauto.
- Qed.
-
- Lemma sem_pred_expr_map_Pand :
- forall A B C sem ctx f_p ps x v p,
- (forall x : positive, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
- eval_predf ps p = true ->
- sem_pred_expr f_p sem ctx x v ->
- @sem_pred_expr A B C f_p sem ctx
- (NE.map (fun x0 => (p ∧ fst x0, snd x0)) x) v.
- Proof.
- induction x; crush.
- inv H1. simplify. constructor; eauto.
- simplify. replace true with (true && true) by auto.
- constructor; auto.
- eapply sem_pexpr_eval; eauto.
- inv H1. simplify. constructor; eauto.
- simplify. replace true with (true && true) by auto.
- constructor; auto.
- eapply sem_pexpr_eval; eauto.
- eapply sem_pred_expr_cons_false. cbn.
- replace false with (true && false) by auto. apply sem_pexpr_Pand; auto.
- eapply sem_pexpr_eval; eauto. eauto.
- Qed.
-
- Lemma sem_pred_expr_app_predicated :
- forall A B C sem ctx f_p y x v p ps,
- sem_pred_expr f_p sem ctx x v ->
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) (ps !! x)) ->
- eval_predf ps p = true ->
- @sem_pred_expr A B C f_p sem ctx (app_predicated p y x) v.
- Proof.
- intros * SEM PS EVAL. unfold app_predicated.
- eapply sem_pred_expr_NEapp2; eauto.
- intros. eapply sem_pred_expr_map_not; eauto.
- rewrite eval_predf_negate. rewrite EVAL. auto.
- eapply sem_pred_expr_map_Pand; eauto.
- Qed.
-
- Lemma sem_pred_expr_app_predicated_false :
- forall A B C sem ctx f_p y x v p ps,
- sem_pred_expr f_p sem ctx y v ->
- (forall x, sem_pexpr ctx (get_forest_p' x f_p) (ps !! x)) ->
- eval_predf ps p = false ->
- @sem_pred_expr A B C f_p sem ctx (app_predicated p y x) v.
- Admitted.
-
- Lemma sem_pred_expr_prune_predicated :
- forall A B C sem ctx f_p y x v,
- sem_pred_expr f_p sem ctx x v ->
- prune_predicated x = Some y ->
- @sem_pred_expr A B C f_p sem ctx y v.
- Proof.
- intros * SEM PRUNE. unfold prune_predicated in *.
- Admitted.
-
- Inductive sem_ident {B A: Type} (c: @ctx B) (a: A): A -> Prop :=
- | sem_ident_intro : sem_ident c a a.
-
- Lemma sem_pred_expr_pred_ret :
- forall G A (ctx: @Abstr.ctx G) (i: A) ps,
- sem_pred_expr ps sem_ident ctx (pred_ret i) i.
- Proof. intros; constructor; constructor. Qed.
-
- Lemma sem_pred_expr_ident :
- forall G A B ps (ctx: @Abstr.ctx G) (l: predicated A) (s: @Abstr.ctx G -> A -> B -> Prop) l_,
- sem_pred_expr ps sem_ident ctx l l_ ->
- forall (v: B),
- s ctx l_ v ->
- sem_pred_expr ps s ctx l v.
- Proof.
- induction 1.
- - intros. constructor; auto. inv H0. auto.
- - intros. apply sem_pred_expr_cons_false; auto.
- - intros. inv H0. constructor; auto.
- Qed.
-
- Lemma sem_pred_expr_ident2 :
- forall G A B ps (ctx: @Abstr.ctx G) (l: predicated A) (s: @Abstr.ctx G -> A -> B -> Prop) (v: B),
- sem_pred_expr ps s ctx l v ->
- exists l_, sem_pred_expr ps sem_ident ctx l l_ /\ s ctx l_ v.
- Proof.
- induction 1.
- - exists e; split; auto. constructor; auto. constructor.
- - inversion_clear IHsem_pred_expr as [x IH].
- inversion_clear IH as [SEM EVALS].
- exists x; split; auto. apply sem_pred_expr_cons_false; auto.
- - exists e; split; auto; constructor; auto; constructor.
- Qed.
-
- Fixpoint of_elist (e: expression_list): list expression :=
- match e with
- | Econs a b => a :: of_elist b
- | Enil => nil
- end.
-
- Fixpoint to_elist (e: list expression): expression_list :=
- match e with
- | a :: b => Econs a (to_elist b)
- | nil => Enil
- end.
-
- Lemma sem_val_list_elist :
- forall G (ctx: @Abstr.ctx G) l j,
- sem_val_list ctx (to_elist l) j ->
- Forall2 (sem_value ctx) l j.
- Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
-
- Lemma sem_val_list_elist2 :
- forall G (ctx: @Abstr.ctx G) l j,
- Forall2 (sem_value ctx) l j ->
- sem_val_list ctx (to_elist l) j.
- Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
-
- Lemma sem_val_list_elist3 :
- forall G (ctx: @Abstr.ctx G) l j,
- sem_val_list ctx l j ->
- Forall2 (sem_value ctx) (of_elist l) j.
- Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
-
- Lemma sem_val_list_elist4 :
- forall G (ctx: @Abstr.ctx G) l j,
- Forall2 (sem_value ctx) (of_elist l) j ->
- sem_val_list ctx l j.
- Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
-
- Lemma sem_pred_expr_predicated_map :
- forall A B C pr (f: C -> B) ctx (pred: predicated C) (pred': C),
- sem_pred_expr pr sem_ident ctx pred pred' ->
- @sem_pred_expr A _ _ pr sem_ident ctx (predicated_map f pred) (f pred').
- Proof.
- induction pred; unfold predicated_map; intros.
- - inv H. inv H3. constructor; eauto. constructor.
- - inv H. inv H5. constructor; eauto. constructor.
- eapply sem_pred_expr_cons_false; eauto.
- Qed.
-
- Lemma NEapp_NEmap :
- forall A B (f: A -> B) a b,
- NE.map f (NE.app a b) = NE.app (NE.map f a) (NE.map f b).
- Proof. induction a; crush. Qed.
-
- Lemma sem_pred_expr_predicated_prod :
- forall A B C pr ctx (a: predicated C) (b: predicated B) a' b',
- sem_pred_expr pr sem_ident ctx a a' ->
- sem_pred_expr pr sem_ident ctx b b' ->
- @sem_pred_expr A _ _ pr sem_ident ctx (predicated_prod a b) (a', b').
- Proof.
- induction a; intros.
- - inv H. inv H4. unfold predicated_prod.
- generalize dependent b. induction b; intros.
- + cbn. destruct_match. inv Heqp. inv H0. inv H6.
- constructor. simplify. replace true with (true && true) by auto.
- eapply sem_pexpr_Pand; eauto. constructor.
- + inv H0. inv H6. cbn. constructor; cbn.
- replace true with (true && true) by auto.
- constructor; auto. constructor.
- eapply sem_pred_expr_cons_false; eauto. cbn.
- replace false with (true && false) by auto.
- apply sem_pexpr_Pand; auto.
- - unfold predicated_prod. simplify.
- rewrite NEapp_NEmap.
- inv H. eapply sem_pred_expr_NEapp.
- { induction b; crush.
- destruct a. inv H0. constructor.
- replace true with (true && true) by auto.
- eapply sem_pexpr_Pand; auto. inv H6. inv H7.
- constructor.
-
- destruct a. inv H0. constructor.
- replace true with (true && true) by auto.
- constructor; auto. inv H8. inv H6. constructor.
-
- specialize (IHb H8). eapply sem_pred_expr_cons_false; auto.
- replace false with (true && false) by auto.
- apply sem_pexpr_Pand; auto.
- }
- { exploit IHa. eauto. eauto. intros.
- unfold predicated_prod in *.
- eapply sem_pred_expr_NEapp3; eauto; [].
- clear H. clear H0.
- induction b.
- { intros. inv H. destruct a. simpl.
- constructor; tauto. }
- { intros. inv H. inv H1. destruct a. simpl.
- constructor; tauto. eauto. } }
- Qed.
-
- Lemma sem_pred_expr_seq_app :
- forall G A B (f: predicated (A -> B))
- ps (ctx: @ctx G) l f_ l_,
- sem_pred_expr ps sem_ident ctx l l_ ->
- sem_pred_expr ps sem_ident ctx f f_ ->
- sem_pred_expr ps sem_ident ctx (seq_app f l) (f_ l_).
- Proof.
- unfold seq_app; intros.
- remember (fun x : (A -> B) * A => fst x (snd x)) as app.
- replace (f_ l_) with (app (f_, l_)) by (rewrite Heqapp; auto).
- eapply sem_pred_expr_predicated_map. eapply sem_pred_expr_predicated_prod; auto.
- Qed.
-
- Lemma sem_pred_expr_map :
- forall A B C (ctx: @ctx A) ps (f: B -> C) y v,
- sem_pred_expr ps sem_ident ctx y v ->
- sem_pred_expr ps sem_ident ctx (NE.map (fun x => (fst x, f (snd x))) y) (f v).
- Proof.
- induction y; crush. inv H. constructor; crush. inv H3. constructor.
- inv H. inv H5. constructor; eauto. constructor.
- eapply sem_pred_expr_cons_false; eauto.
- Qed.
-
- Lemma sem_pred_expr_flap :
- forall G A B C (f: predicated (A -> B -> C))
- ps (ctx: @ctx G) l f_,
- sem_pred_expr ps sem_ident ctx f f_ ->
- sem_pred_expr ps sem_ident ctx (flap f l) (f_ l).
- Proof.
- induction f. unfold flap2; intros. inv H. inv H3.
- constructor; eauto. constructor.
- intros. inv H. cbn.
- constructor; eauto. inv H5. constructor.
- eapply sem_pred_expr_cons_false; eauto.
- Qed.
-
- Lemma sem_pred_expr_flap2 :
- forall G A B C (f: predicated (A -> B -> C))
- ps (ctx: @ctx G) l1 l2 f_,
- sem_pred_expr ps sem_ident ctx f f_ ->
- sem_pred_expr ps sem_ident ctx (flap2 f l1 l2) (f_ l1 l2).
- Proof.
- induction f. unfold flap2; intros. inv H. inv H3.
- constructor; eauto. constructor.
- intros. inv H. cbn.
- constructor; eauto. inv H5. constructor.
- eapply sem_pred_expr_cons_false; eauto.
- Qed.
-
- Lemma sem_pred_expr_seq_app_val :
- forall G A B V (s: @Abstr.ctx G -> B -> V -> Prop)
- (f: predicated (A -> B))
- ps ctx l v f_ l_,
- sem_pred_expr ps sem_ident ctx l l_ ->
- sem_pred_expr ps sem_ident ctx f f_ ->
- s ctx (f_ l_) v ->
- sem_pred_expr ps s ctx (seq_app f l) v.
- Proof.
- intros. eapply sem_pred_expr_ident; [|eassumption].
- eapply sem_pred_expr_seq_app; eauto.
- Qed.
-
- Fixpoint Eapp a b :=
- match a with
- | Enil => b
- | Econs ax axs => Econs ax (Eapp axs b)
- end.
-
- Lemma Eapp_right_nil :
- forall a, Eapp a Enil = a.
- Proof. induction a; cbn; try rewrite IHa; auto. Qed.
-
- Lemma Eapp_left_nil :
- forall a, Eapp Enil a = a.
- Proof. auto. Qed.
-
- Lemma sem_pred_expr_cons_pred_expr :
- forall A (ctx: @ctx A) pr s s' a e,
- sem_pred_expr pr sem_ident ctx s s' ->
- sem_pred_expr pr sem_ident ctx a e ->
- sem_pred_expr pr sem_ident ctx (cons_pred_expr a s) (Econs e s').
- Proof.
- intros. unfold cons_pred_expr.
- replace (Econs e s') with ((uncurry Econs) (e, s')) by auto.
- eapply sem_pred_expr_predicated_map.
- eapply sem_pred_expr_predicated_prod; eauto.
- Qed.
-
- Lemma evaluable_singleton :
- forall A B ctx fp r,
- @all_evaluable A B ctx fp (NE.singleton (T, r)).
- Proof.
- unfold all_evaluable, evaluable; intros.
- inv H. simpl. exists true. constructor.
- Qed.
-
-Lemma evaluable_negate :
- forall G (ctx: @ctx G) p,
- p_evaluable ctx p ->
- p_evaluable ctx (¬ p).
-Proof.
- unfold p_evaluable, evaluable in *; intros.
- inv H. eapply sem_pexpr_negate in H0. econstructor; eauto.
-Qed.
-
-Lemma predicated_evaluable :
- forall G (ctx: @ctx G) ps (p: pred_op),
- pred_forest_evaluable ctx ps ->
- p_evaluable ctx (from_pred_op ps p).
-Proof.
- unfold pred_forest_evaluable; intros. induction p; intros; cbn.
- - repeat destruct_match; subst; unfold get_forest_p'.
- destruct_match. eapply H; eauto. econstructor. constructor. constructor.
- eapply evaluable_negate.
- destruct_match. eapply H; eauto. econstructor. constructor. constructor.
- - repeat econstructor.
- - repeat econstructor.
- - inv IHp1. inv IHp2. econstructor. apply sem_pexpr_Pand; eauto.
- - inv IHp1. inv IHp2. econstructor. apply sem_pexpr_Por; eauto.
-Qed.
-
-Lemma predicated_evaluable_all :
- forall A G (ctx: @ctx G) ps (p: predicated A),
- pred_forest_evaluable ctx ps ->
- all_evaluable ctx ps p.
-Proof.
- unfold all_evaluable; intros.
- eapply predicated_evaluable. eauto.
-Qed.
-
-Lemma predicated_evaluable_all_list :
- forall A G (ctx: @ctx G) ps (p: list (predicated A)),
- pred_forest_evaluable ctx ps ->
- Forall (all_evaluable ctx ps) p.
-Proof.
- induction p.
- - intros. constructor.
- - intros. constructor; eauto. apply predicated_evaluable_all; auto.
-Qed.
-
-Hint Resolve evaluable_singleton : core.
-Hint Resolve predicated_evaluable : core.
-Hint Resolve predicated_evaluable_all : core.
-Hint Resolve predicated_evaluable_all_list : core.
-
- Lemma sem_pred_expr_fold_right :
- forall A pr ctx s a a' s',
- sem_pred_expr pr sem_ident ctx s s' ->
- Forall2 (sem_pred_expr pr sem_ident ctx) (NE.to_list a) (of_elist a') ->
- @sem_pred_expr A _ _ pr sem_ident ctx (NE.fold_right cons_pred_expr s a) (Eapp a' s').
- Proof.
- induction a; crush. inv H0. inv H5. destruct a'; crush. destruct a'; crush.
- inv H3. eapply sem_pred_expr_cons_pred_expr; eauto.
- inv H0. destruct a'; crush. inv H3.
- eapply sem_pred_expr_cons_pred_expr; eauto.
- Qed.
-
- Lemma sem_pred_expr_fold_right2 :
- forall A pr ctx s a a' s',
- sem_pred_expr pr sem_ident ctx s s' ->
- @sem_pred_expr A _ _ pr sem_ident ctx (NE.fold_right cons_pred_expr s a) (Eapp a' s') ->
- Forall2 (sem_pred_expr pr sem_ident ctx) (NE.to_list a) (of_elist a').
- Proof.
- induction a. Admitted.
-
- Lemma NEof_list_some :
- forall A a a' (e: A),
- NE.of_list a = Some a' ->
- NE.of_list (e :: a) = Some (NE.cons e a').
- Proof.
- induction a; [crush|].
- intros.
- cbn in H. destruct a0. inv H. auto.
- destruct_match; [|discriminate].
- inv H. specialize (IHa n a ltac:(trivial)).
- cbn. destruct_match. unfold NE.of_list in IHa.
- fold (@NE.of_list A) in IHa. rewrite IHa in Heqo0. inv Heqo0. auto.
- unfold NE.of_list in IHa. fold (@NE.of_list A) in IHa. rewrite IHa in Heqo0. inv Heqo0.
- Qed.
-
- Lemma NEof_list_contra :
- forall A b (a: A),
- ~ NE.of_list (a :: b) = None.
- Proof.
- induction b; try solve [crush].
- intros.
- specialize (IHb a).
- enough (X: exists x, NE.of_list (a :: b) = Some x).
- inversion_clear X as [x X'].
- erewrite NEof_list_some; eauto; discriminate.
- destruct (NE.of_list (a :: b)) eqn:?; [eauto|contradiction].
- Qed.
-
- Lemma NEof_list_exists :
- forall A b (a: A),
- exists x, NE.of_list (a :: b) = Some x.
- Proof.
- intros. pose proof (NEof_list_contra _ b a).
- destruct (NE.of_list (a :: b)); try contradiction.
- eauto.
- Qed.
-
- Lemma NEof_list_exists2 :
- forall A b (a c: A),
- exists x,
- NE.of_list (c :: a :: b) = Some (NE.cons c x)
- /\ NE.of_list (a :: b) = Some x.
- Proof.
- intros. pose proof (NEof_list_exists _ b a).
- inversion_clear H as [x B].
- econstructor; split; eauto.
- eapply NEof_list_some; eauto.
- Qed.
-
- Lemma NEof_list_to_list :
- forall A (x: list A) y,
- NE.of_list x = Some y ->
- NE.to_list y = x.
- Proof.
- induction x; [crush|].
- intros. destruct x; [crush|].
- pose proof (NEof_list_exists2 _ x a0 a).
- inversion_clear H0 as [x0 [HN1 HN2]]. rewrite HN1 in H. inv H.
- cbn. f_equal. eauto.
- Qed.
-
- Lemma sem_pred_expr_merge :
- forall G (ctx: @Abstr.ctx G) ps l args,
- Forall2 (sem_pred_expr ps sem_ident ctx) args l ->
- sem_pred_expr ps sem_ident ctx (merge args) (to_elist l).
- Proof.
- induction l; intros.
- - inv H. cbn; repeat constructor.
- - inv H. cbn. unfold merge. Admitted.
-
- Lemma sem_pred_expr_merge2 :
- forall A (ctx: @ctx A) pr l l',
- sem_pred_expr pr sem_ident ctx (merge l) l' ->
- Forall2 (sem_pred_expr pr sem_ident ctx) l (of_elist l').
- Proof.
- induction l; crush.
- - unfold merge in *; cbn in *.
- inv H. inv H5. constructor.
- - unfold merge in H.
- pose proof (NEof_list_exists _ l a) as Y.
- inversion_clear Y as [x HNE]. rewrite HNE in H.
- erewrite <- (NEof_list_to_list _ (a :: l)) by eassumption.
- apply sem_pred_expr_fold_right2 with (s := (NE.singleton (T, Enil))) (s' := Enil).
- repeat constructor.
- rewrite Eapp_right_nil. auto.
- Qed.
-
- Lemma sem_merge_list :
- forall A ctx f rs ps m args,
- sem ctx f ((mk_instr_state rs ps m), None) ->
- @sem_pred_expr A _ _ (forest_preds f) sem_val_list ctx
- (merge (list_translation args f)) (rs ## args).
- Proof.
- induction args; crush. cbn. constructor; constructor.
- unfold merge. specialize (IHargs H).
- eapply sem_pred_expr_ident2 in IHargs.
- inversion_clear IHargs as [l_ [HSEM HVAL]].
- destruct_match; [|exfalso; eapply NEof_list_contra; eauto].
- destruct args; cbn -[NE.of_list] in *.
- - unfold merge in *. simplify.
- inv H. inv H6. specialize (H a).
- eapply sem_pred_expr_ident2 in H.
- inversion_clear H as [l_2 [HSEM2 HVAL2]].
- unfold cons_pred_expr.
- eapply sem_pred_expr_ident.
- eapply sem_pred_expr_predicated_map.
- eapply sem_pred_expr_predicated_prod; [eassumption|repeat constructor].
- repeat constructor; auto.
- - pose proof (NEof_list_exists2 _ (list_translation args f) (f #r (Reg r)) (f #r (Reg a))) as Y.
- inversion_clear Y as [x [Y1 Y2]]. rewrite Heqo in Y1. inv Y1.
- inversion_clear H as [? ? ? ? ? ? REG PRED MEM EXIT].
- inversion_clear REG as [? ? ? REG'].
- inversion_clear PRED as [? ? ? PRED'].
- pose proof (REG' a) as REGa. pose proof (REG' r) as REGr.
- exploit sem_pred_expr_ident2; [exact REGa|].
- intro REGI'; inversion_clear REGI' as [a' [SEMa SEMa']].
- exploit sem_pred_expr_ident2; [exact REGr|].
- intro REGI'; inversion_clear REGI' as [r' [SEMr SEMr']].
- apply sem_pred_expr_ident with (l_ := Econs a' l_); [|constructor; auto].
- replace (Econs a' l_) with (Eapp (Econs a' l_) Enil) by (now apply Eapp_right_nil).
- apply sem_pred_expr_fold_right; eauto.
- repeat constructor.
- constructor; eauto.
- erewrite NEof_list_to_list; eauto.
- eapply sem_pred_expr_merge2; auto.
- Qed.
-
- Lemma sem_pred_expr_list_translation :
- forall G ctx args f i,
- @sem G ctx f (i, None) ->
- exists l,
- Forall2 (sem_pred_expr (forest_preds f) sem_ident ctx) (list_translation args f) l
- /\ sem_val_list ctx (to_elist l) ((is_rs i)##args).
- Proof.
- induction args; intros.
- - exists nil; cbn; split; auto; constructor.
- - exploit IHargs; try eassumption; intro EX.
- inversion_clear EX as [l [FOR SEM]].
- destruct i as [rs' m' ps'].
- inversion_clear H as [? ? ? ? ? ? REGSET PREDSET MEM EXIT].
- inversion_clear REGSET as [? ? ? REG].
- pose proof (REG a).
- exploit sem_pred_expr_ident2; eauto; intro IDENT.
- inversion_clear IDENT as [l_ [SEMP SV]].
- exists (l_ :: l). split. constructor; auto.
- cbn; constructor; auto.
- Qed.
-
-Lemma evaluable_and_true :
- forall G (ctx: @ctx G) ps p,
- p_evaluable ctx (from_pred_op ps p) ->
- p_evaluable ctx (from_pred_op ps (p ∧ T)).
-Proof.
- intros. unfold evaluable in *. inv H. exists (x && true). cbn.
- apply sem_pexpr_Pand; auto. constructor.
-Qed.
-
-Lemma NEin_map :
- forall A B p y (f: A -> B) a,
- NE.In (p, y) (predicated_map f a) ->
- exists x, NE.In (p, x) a /\ y = f x.
-Proof.
- induction a; intros.
- - inv H. destruct a. econstructor. split; eauto. constructor.
- - inv H. inv H1. inv H. destruct a. cbn in *. econstructor; econstructor; eauto.
- constructor; tauto.
- specialize (IHa H). inv IHa. inv H0.
- econstructor; econstructor; eauto. constructor; tauto.
-Qed.
-
-Lemma NEin_map2 :
- forall A B (f: A -> B) a p y,
- NE.In (p, y) a ->
- NE.In (p, f y) (predicated_map f a).
-Proof.
- induction a; crush.
- inv H. constructor.
- inv H. inv H1.
- - constructor; auto.
- - constructor; eauto.
-Qed.
-
-Lemma all_evaluable_predicated_map :
- forall A B G (ctx: @ctx G) ps (f: A -> B) p,
- all_evaluable ctx ps p ->
- all_evaluable ctx ps (predicated_map f p).
-Proof.
- induction p.
- - unfold all_evaluable; intros.
- exploit NEin_map; eauto; intros. inv H1. inv H2.
- eapply H; eauto.
- - intros. cbn.
- eapply all_evaluable_cons3. eapply IHp. eapply all_evaluable_cons; eauto.
- cbn. destruct a. cbn in *. eapply all_evaluable_cons2; eauto.
-Qed.
-
-Lemma all_evaluable_predicated_map2 :
- forall A B G (ctx: @ctx G) ps (f: A -> B) p,
- all_evaluable ctx ps (predicated_map f p) ->
- all_evaluable ctx ps p.
-Proof.
- induction p.
- - unfold all_evaluable in *; intros.
- eapply H. eapply NEin_map2; eauto.
- - intros. cbn. destruct a.
- cbn in H. pose proof H. eapply all_evaluable_cons in H; eauto.
- eapply all_evaluable_cons2 in H0; eauto.
- unfold all_evaluable. specialize (IHp H).
- unfold all_evaluable in IHp.
- intros. inv H1. inv H3. inv H1; eauto.
- specialize (IHp _ _ H1). eauto.
-Qed.
-
-Lemma all_evaluable_map_and :
- forall A B G (ctx: @ctx G) ps (a: NE.non_empty ((pred_op * A) * (pred_op * B))),
- (forall p1 x p2 y,
- NE.In ((p1, x), (p2, y)) a ->
- p_evaluable ctx (from_pred_op ps p1)
- /\ p_evaluable ctx (from_pred_op ps p2)) ->
- all_evaluable ctx ps (NE.map (fun x => match x with ((a, b), (c, d)) => (Pand a c, (b, d)) end) a).
-Proof.
- induction a.
- - intros. cbn. repeat destruct_match. inv Heqp.
- unfold all_evaluable; intros. inv H0.
- unfold evaluable.
- exploit H. constructor.
- intros [Ha Hb]. inv Ha. inv Hb.
- exists (x && x0). apply sem_pexpr_Pand; auto.
- - intros. unfold all_evaluable; cbn; intros. inv H0. inv H2.
- + repeat destruct_match. inv Heqp0. inv H0.
- specialize (H p2 a1 p3 b ltac:(constructor; eauto)).
- inv H. inv H0. inv H1. exists (x && x0).
- apply sem_pexpr_Pand; eauto.
- + eapply IHa; intros. eapply H. econstructor. right. eauto.
- eauto.
-Qed.
-
-Lemma all_evaluable_map_add :
- forall A B G (ctx: @ctx G) ps (a: pred_op * A) (b: predicated B) p1 x p2 y,
- p_evaluable ctx (from_pred_op ps (fst a)) ->
- all_evaluable ctx ps b ->
- NE.In (p1, x, (p2, y)) (NE.map (fun x : pred_op * B => (a, x)) b) ->
- p_evaluable ctx (from_pred_op ps p1) /\ p_evaluable ctx (from_pred_op ps p2).
-Proof.
- induction b; intros.
- - cbn in *. inv H1. unfold all_evaluable in *. specialize (H0 _ _ ltac:(constructor)).
- auto.
- - cbn in *. inv H1. inv H3.
- + inv H1. unfold all_evaluable in H0. specialize (H0 _ _ ltac:(constructor; eauto)); auto.
- + eapply all_evaluable_cons in H0. specialize (IHb _ _ _ _ H H0 H1). auto.
-Qed.
-
-Lemma NEin_NEapp5 :
- forall A (a: A) x y,
- NE.In a (NE.app x y) ->
- NE.In a x \/ NE.In a y.
-Proof.
- induction x; crush.
- - inv H. inv H1. left. constructor. tauto.
- - inv H. inv H1. left. constructor; tauto.
- exploit IHx; eauto; intros. inv H0.
- left. constructor; tauto. tauto.
-Qed.
-
-Lemma all_evaluable_non_empty_prod :
- forall A B G (ctx: @ctx G) ps p1 x p2 y (a: predicated A) (b: predicated B),
- all_evaluable ctx ps a ->
- all_evaluable ctx ps b ->
- NE.In ((p1, x), (p2, y)) (NE.non_empty_prod a b) ->
- p_evaluable ctx (from_pred_op ps p1)
- /\ p_evaluable ctx (from_pred_op ps p2).
-Proof.
- induction a; intros.
- - cbn in *. eapply all_evaluable_map_add; eauto. destruct a; cbn in *. eapply H; constructor.
- - cbn in *. eapply NEin_NEapp5 in H1. inv H1. eapply all_evaluable_map_add; eauto.
- destruct a; cbn in *. eapply all_evaluable_cons2; eauto.
- eapply all_evaluable_cons in H. eauto.
-Qed.
-
-Lemma all_evaluable_predicated_prod :
- forall A B G (ctx: @ctx G) ps (a: predicated A) (b: predicated B),
- all_evaluable ctx ps a ->
- all_evaluable ctx ps b ->
- all_evaluable ctx ps (predicated_prod a b).
-Proof.
- intros. unfold all_evaluable; intros.
- unfold predicated_prod in *. exploit all_evaluable_map_and.
- 2: { eauto. }
- intros. 2: { eauto. }
-Admitted. (* Requires simple lemma to prove, but this lemma is not needed. *)
-
-Lemma cons_pred_expr_evaluable :
- forall G (ctx: @ctx G) ps a b,
- all_evaluable ctx ps a ->
- all_evaluable ctx ps b ->
- all_evaluable ctx ps (cons_pred_expr a b).
-Proof.
- unfold cons_pred_expr; intros.
- apply all_evaluable_predicated_map.
- now apply all_evaluable_predicated_prod.
-Qed.
-
-Lemma fold_right_all_evaluable :
- forall G (ctx: @ctx G) ps n,
- Forall (all_evaluable ctx ps) (NE.to_list n) ->
- all_evaluable ctx ps (NE.fold_right cons_pred_expr (NE.singleton (T, Enil)) n).
-Proof.
- induction n; cbn in *; intros.
- - unfold cons_pred_expr. eapply all_evaluable_predicated_map. eapply all_evaluable_predicated_prod.
- inv H. auto. unfold all_evaluable; intros. inv H0. exists true. constructor.
- - inv H. specialize (IHn H3). now eapply cons_pred_expr_evaluable.
-Qed.
-
-Lemma merge_all_evaluable :
- forall G (ctx: @ctx G) ps,
- pred_forest_evaluable ctx ps ->
- forall f args,
- all_evaluable ctx ps (merge (list_translation args f)).
-Proof.
- intros. eapply predicated_evaluable_all. eauto.
-Qed.
-
-(*|
-Here we can finally assume that everything in the forest is evaluable, which
-will allow us to prove that translating the list of register accesses will also
-all be evaluable.
-|*)
-
- Lemma sem_update_Iop :
- forall A op rs args m v f ps ctx,
- Op.eval_operation (ctx_ge ctx) (ctx_sp ctx) op rs ## args (is_mem (ctx_is ctx)) = Some v ->
- sem ctx f ((mk_instr_state rs ps m), None) ->
- @sem_pred_expr A _ _ (forest_preds f) sem_value ctx
- (seq_app (pred_ret (Eop op)) (merge (list_translation args f))) v.
- Proof.
- intros * EVAL SEM.
- exploit sem_pred_expr_list_translation; try eassumption.
- intro H; inversion_clear H as [x [HS HV]].
- eapply sem_pred_expr_seq_app_val.
- - cbn in *; eapply sem_pred_expr_merge; eauto.
- - apply sem_pred_expr_pred_ret.
- - econstructor; [eauto|]; auto.
- Qed.
-
- Lemma sem_update_Iload :
- forall A rs args m v f ps ctx addr a0 chunk,
- Op.eval_addressing (ctx_ge ctx) (ctx_sp ctx) addr rs ## args = Some a0 ->
- Mem.loadv chunk m a0 = Some v ->
- sem ctx f ((mk_instr_state rs ps m), None) ->
- @sem_pred_expr A _ _ (forest_preds f) sem_value ctx
- (seq_app (seq_app (pred_ret (Eload chunk addr)) (merge (list_translation args f))) (f #r Mem)) v.
- Proof.
- intros * EVAL MEM SEM.
- exploit sem_pred_expr_list_translation; try eassumption.
- intro H; inversion_clear H as [x [HS HV]].
- inversion SEM as [? ? ? ? ? ? REG PRED HMEM EXIT]; subst.
- exploit sem_pred_expr_ident2; [eapply HMEM|].
- intros H; inversion_clear H as [x' [HS' HV']].
- eapply sem_pred_expr_seq_app_val; eauto.
- { eapply sem_pred_expr_seq_app; eauto.
- - eapply sem_pred_expr_merge; eauto.
- - apply sem_pred_expr_pred_ret.
- }
- econstructor; eauto.
- Qed.
-
- Lemma storev_valid_pointer1 :
- forall chunk m m' s d b z,
- Mem.storev chunk m s d = Some m' ->
- Mem.valid_pointer m b z = true ->
- Mem.valid_pointer m' b z = true.
- Proof using.
- intros. unfold Mem.storev in *. destruct_match; try discriminate; subst.
- apply Mem.valid_pointer_nonempty_perm. apply Mem.valid_pointer_nonempty_perm in H0.
- eapply Mem.perm_store_1; eauto.
- Qed.
-
- Lemma storev_valid_pointer2 :
- forall chunk m m' s d b z,
- Mem.storev chunk m s d = Some m' ->
- Mem.valid_pointer m' b z = true ->
- Mem.valid_pointer m b z = true.
- Proof using.
- intros. unfold Mem.storev in *. destruct_match; try discriminate; subst.
- apply Mem.valid_pointer_nonempty_perm. apply Mem.valid_pointer_nonempty_perm in H0.
- eapply Mem.perm_store_2; eauto.
- Qed.
-
- Definition valid_mem m m' :=
- forall b z, Mem.valid_pointer m b z = Mem.valid_pointer m' b z.
-
- #[global] Program Instance valid_mem_Equivalence : Equivalence valid_mem.
- Next Obligation. unfold valid_mem; auto. Qed. (* Reflexivity *)
- Next Obligation. unfold valid_mem; auto. Qed. (* Symmetry *)
- Next Obligation. unfold valid_mem; eauto. Qed. (* Transitity *)
-
- Lemma storev_valid_pointer :
- forall chunk m m' s d,
- Mem.storev chunk m s d = Some m' ->
- valid_mem m m'.
- Proof using.
- unfold valid_mem. symmetry.
- intros. destruct (Mem.valid_pointer m b z) eqn:?;
- eauto using storev_valid_pointer1.
- apply not_true_iff_false.
- apply not_true_iff_false in Heqb0.
- eauto using storev_valid_pointer2.
- Qed.
-
- Lemma storev_cmpu_bool :
- forall m m' c v v0,
- valid_mem m m' ->
- Val.cmpu_bool (Mem.valid_pointer m) c v v0 = Val.cmpu_bool (Mem.valid_pointer m') c v v0.
- Proof using.
- unfold valid_mem.
- intros. destruct v, v0; crush.
- { destruct_match; crush.
- destruct_match; crush.
- destruct_match; crush.
- apply orb_true_iff in H1.
- inv H1. rewrite H in H2. rewrite H2 in Heqb1.
- simplify. rewrite H0 in Heqb1. crush.
- rewrite H in H2. rewrite H2 in Heqb1.
- rewrite H0 in Heqb1. crush.
- destruct_match; auto. simplify.
- apply orb_true_iff in H1.
- inv H1. rewrite <- H in H2. rewrite H2 in Heqb1.
- simplify. rewrite H0 in Heqb1. crush.
- rewrite <- H in H2. rewrite H2 in Heqb1.
- rewrite H0 in Heqb1. crush. }
-
- { destruct_match; crush.
- destruct_match; crush.
- destruct_match; crush.
- apply orb_true_iff in H1.
- inv H1. rewrite H in H2. rewrite H2 in Heqb1.
- simplify. rewrite H0 in Heqb1. crush.
- rewrite H in H2. rewrite H2 in Heqb1.
- rewrite H0 in Heqb1. crush.
- destruct_match; auto. simplify.
- apply orb_true_iff in H1.
- inv H1. rewrite <- H in H2. rewrite H2 in Heqb1.
- simplify. rewrite H0 in Heqb1. crush.
- rewrite <- H in H2. rewrite H2 in Heqb1.
- rewrite H0 in Heqb1. crush. }
-
- { destruct_match; auto. destruct_match; auto; crush.
- { destruct_match; crush.
- { destruct_match; crush.
- setoid_rewrite H in H0; eauto.
- setoid_rewrite H in H1; eauto.
- rewrite H0 in Heqb. rewrite H1 in Heqb; crush.
- }
- { destruct_match; crush.
- setoid_rewrite H in Heqb0; eauto.
- rewrite H0 in Heqb0. rewrite H1 in Heqb0; crush. } }
- { destruct_match; crush.
- { destruct_match; crush.
- setoid_rewrite H in H0; eauto.
- setoid_rewrite H in H1; eauto.
- rewrite H0 in Heqb0. rewrite H1 in Heqb0; crush.
- }
- { destruct_match; crush.
- setoid_rewrite H in Heqb0; eauto.
- rewrite H0 in Heqb0. rewrite H1 in Heqb0; crush. } } }
- Qed.
-
- Lemma storev_eval_condition :
- forall m m' c rs,
- valid_mem m m' ->
- Op.eval_condition c rs m = Op.eval_condition c rs m'.
- Proof using.
- intros. destruct c; crush.
- destruct rs; crush.
- destruct rs; crush.
- destruct rs; crush.
- eapply storev_cmpu_bool; eauto.
- destruct rs; crush.
- destruct rs; crush.
- eapply storev_cmpu_bool; eauto.
- Qed.
-
- Lemma eval_operation_valid_pointer :
- forall A B (ge: Genv.t A B) sp op args m m' v,
- valid_mem m m' ->
- Op.eval_operation ge sp op args m' = Some v ->
- Op.eval_operation ge sp op args m = Some v.
- Proof.
- intros * VALID OP. destruct op; auto.
- - destruct cond; auto; cbn in *.
- + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
- + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
- - destruct c; auto; cbn in *.
- + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
- + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
- Qed.
-
- Lemma bb_memory_consistency_eval_operation :
- forall A B (ge: Genv.t A B) sp state i state' args op v,
- step_instr ge sp (Iexec state) i (Iexec state') ->
- Op.eval_operation ge sp op args (is_mem state) = Some v ->
- Op.eval_operation ge sp op args (is_mem state') = Some v.
- Proof.
- inversion_clear 1; intro EVAL; auto.
- cbn in *.
- eapply eval_operation_valid_pointer. unfold valid_mem. symmetry.
- eapply storev_valid_pointer; eauto.
- auto.
- Qed.
-
- Lemma truthy_dflt :
- forall ps p,
- truthy ps p -> eval_predf ps (dfltp p) = true.
- Proof. intros. destruct p; cbn; inv H; auto. Qed.
-
- Lemma sem_update_Istore :
- forall A rs args m v f ps ctx addr a0 chunk m' v'
- (EVALF: forest_evaluable ctx f),
- Op.eval_addressing (ctx_ge ctx) (ctx_sp ctx) addr rs ## args = Some a0 ->
- Mem.storev chunk m a0 v' = Some m' ->
- sem_value ctx v v' ->
- sem ctx f ((mk_instr_state rs ps m), None) ->
- @sem_pred_expr A _ _ (forest_preds f) sem_mem ctx
- (seq_app (seq_app (pred_ret (Estore v chunk addr))
- (merge (list_translation args f))) (f #r Mem)) m'.
- Proof.
- intros * EVALF EVAL STOREV SEMVAL SEM.
- exploit sem_merge_list; try eassumption.
- intro MERGE. exploit sem_pred_expr_ident2; eauto.
- intro TMP; inversion_clear TMP as [x [HS HV]].
- inversion_clear SEM as [? ? ? ? ? ? REG PRED HMEM EXIT].
- exploit sem_pred_expr_ident2; [eapply HMEM|].
- intros TMP; inversion_clear TMP as [x' [HS' HV']].
- eapply sem_pred_expr_ident.
- eapply sem_pred_expr_seq_app; eauto.
- eapply sem_pred_expr_seq_app; eauto.
- eapply sem_pred_expr_pred_ret.
- econstructor; eauto.
- Qed.
-
- Lemma sem_update_Iop_top:
- forall A f p p' f' rs m pr op res args p0 v state,
- Op.eval_operation (ctx_ge state) (ctx_sp state) op rs ## args m = Some v ->
- truthy pr p0 ->
- valid_mem (is_mem (ctx_is state)) m ->
- sem state f ((mk_instr_state rs pr m), None) ->
- update (p, f) (RBop p0 op args res) = Some (p', f') ->
- eval_predf pr p = true ->
- @sem A state f' (mk_instr_state (rs # res <- v) pr m, None).
- Proof.
- intros * EVAL_OP TRUTHY VALID SEM UPD EVAL_PRED.
- pose proof SEM as SEM2.
- inversion UPD as [PRUNE]. unfold Option.bind in PRUNE.
- destr. inversion_clear PRUNE.
- rename Heqo into PRUNE.
- inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
- cbn [is_ps] in *. constructor.
- + constructor; intro x. inversion_clear REG as [? ? ? REG']. specialize (REG' x).
- destruct f as [fr fp fe]. cbn [forest_preds set_forest] in *.
- destruct (peq x res); subst.
- * rewrite forest_reg_gss in *. rewrite Regmap.gss in *.
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated; [| |eauto].
- replace fp with (forest_preds {| forest_regs := fr; forest_preds := fp; forest_exit := fe |}) by auto.
- eapply sem_update_Iop; eauto. cbn in *.
- eapply eval_operation_valid_pointer; eauto.
- inversion_clear SEM2 as [? ? ? ? ? ? REG2 PRED2 MEM2 EXIT2].
- inversion_clear PRED2; eauto.
- cbn -[eval_predf]. rewrite eval_predf_Pand.
- rewrite EVAL_PRED. rewrite truthy_dflt; auto.
- * rewrite forest_reg_gso. rewrite Regmap.gso; auto.
- unfold not in *; intros. apply n0. inv H; auto.
- + constructor; intros. inv PRED. rewrite forest_reg_pred. auto.
- + rewrite forest_reg_gso; auto; discriminate.
- + auto.
- Qed.
-
- Lemma sem_update_Iload_top:
- forall A f p p' f' rs m pr res args p0 v state addr a chunk,
- Op.eval_addressing (ctx_ge state) (ctx_sp state) addr rs ## args = Some a ->
- Mem.loadv chunk m a = Some v ->
- truthy pr p0 ->
- valid_mem (is_mem (ctx_is state)) m ->
- sem state f ((mk_instr_state rs pr m), None) ->
- update (p, f) (RBload p0 chunk addr args res) = Some (p', f') ->
- eval_predf pr p = true ->
- @sem A state f' (mk_instr_state (rs # res <- v) pr m, None).
- Proof.
- intros * EVAL_OP LOAD TRUTHY VALID SEM UPD EVAL_PRED.
- pose proof SEM as SEM2.
- inversion UPD as [PRUNE]. unfold Option.bind in PRUNE. destr.
- inversion_clear PRUNE.
- rename Heqo into PRUNE.
- inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
- cbn [is_ps] in *. constructor.
- + constructor; intro x. inversion_clear REG as [? ? ? REG']. specialize (REG' x).
- destruct f as [fr fp fe]. cbn [forest_preds set_forest] in *.
- destruct (peq x res); subst.
- * rewrite forest_reg_gss in *. rewrite Regmap.gss in *.
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated; [| |eauto].
- replace fp with (forest_preds {| forest_regs := fr; forest_preds := fp; forest_exit := fe |}) by auto.
- eapply sem_update_Iload; eauto.
- inversion_clear PRED; eauto.
- cbn -[eval_predf]. rewrite eval_predf_Pand.
- rewrite EVAL_PRED. rewrite truthy_dflt; auto.
- * rewrite forest_reg_gso. rewrite Regmap.gso; auto.
- unfold not in *; intros. apply n0. inv H; auto.
- + constructor; intros. inv PRED. rewrite forest_reg_pred. auto.
- + rewrite forest_reg_gso; auto; discriminate.
- + auto.
- Qed.
-
- Lemma exists_sem_pred :
- forall A B C (ctx: @ctx A) s pr r v,
- @sem_pred_expr A B C pr s ctx r v ->
- exists r',
- NE.In r' r /\ s ctx (snd r') v.
- Proof.
- induction r; crush.
- - inv H. econstructor. split. constructor. auto.
- - inv H.
- + econstructor. split. constructor. left; auto. auto.
- + exploit IHr; eauto. intros HH. inversion_clear HH as [x HH']. inv HH'.
- econstructor. split. constructor. right. eauto. auto.
- Qed.
-
- Lemma sem_update_Istore_top:
- forall A f p p' f' rs m pr res args p0 state addr a chunk m',
- Op.eval_addressing (ctx_ge state) (ctx_sp state) addr rs ## args = Some a ->
- Mem.storev chunk m a rs !! res = Some m' ->
- truthy pr p0 ->
- valid_mem (is_mem (ctx_is state)) m ->
- sem state f ((mk_instr_state rs pr m), None) ->
- update (p, f) (RBstore p0 chunk addr args res) = Some (p', f') ->
- eval_predf pr p = true ->
- @sem A state f' (mk_instr_state rs pr m', None).
- Proof.
- intros * EVAL_OP STORE TRUTHY VALID SEM UPD EVAL_PRED.
- pose proof SEM as SEM2.
- inversion UPD as [PRUNE]. unfold Option.bind in PRUNE. destr.
- inversion_clear PRUNE.
- rename Heqo into PRUNE.
- inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
- cbn [is_ps] in *. constructor.
- + constructor; intros. inv REG. rewrite forest_reg_gso; eauto. discriminate.
- + constructor; intros. inv PRED. rewrite forest_reg_pred. auto.
- + destruct f as [fr fp fm]; cbn -[seq_app] in *.
- rewrite forest_reg_gss.
- exploit sem_pred_expr_ident2; [exact MEM|]; intro HSEM_;
- inversion_clear HSEM_ as [x [HSEM1 HSEM2]].
- inv REG. specialize (H res).
- pose proof H as HRES.
- eapply sem_pred_expr_ident2 in HRES.
- inversion_clear HRES as [r2 [HRES1 HRES2]].
- apply exists_sem_pred in H. inversion_clear H as [r' [HNE HVAL]].
- exploit sem_merge_list. eapply SEM2. instantiate (2 := args). intro HSPE. eapply sem_pred_expr_ident2 in HSPE.
- inversion_clear HSPE as [l_ [HSPE1 HSPE2]].
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated.
- eapply sem_pred_expr_seq_app_val; [solve [eauto]| |].
- eapply sem_pred_expr_seq_app; [solve [eauto]|].
- eapply sem_pred_expr_flap2.
- eapply sem_pred_expr_seq_app; [solve [eauto]|].
- eapply sem_pred_expr_pred_ret. econstructor. eauto. 3: { eauto. }
- eauto. auto. eauto. inv PRED. eauto.
- rewrite eval_predf_Pand. rewrite EVAL_PRED.
- rewrite truthy_dflt. auto. auto.
- + auto.
- Qed.
-
- Definition predicated_not_inP {A} (p: Gible.predicate) (l: predicated A) :=
- forall op e, NE.In (op, e) l -> ~ PredIn p op.
-
- Lemma predicated_not_inP_cons :
- forall A p (a: (pred_op * A)) l,
- predicated_not_inP p (NE.cons a l) ->
- predicated_not_inP p l.
- Proof.
- unfold predicated_not_inP; intros. eapply H. econstructor. right; eauto.
- Qed.
-
- Lemma sem_pexpr_not_in :
- forall G (ctx: @ctx G) p0 ps p e b,
- ~ PredIn p p0 ->
- sem_pexpr ctx (from_pred_op ps p0) b ->
- sem_pexpr ctx (from_pred_op (PTree.set p e ps) p0) b.
- Proof.
- induction p0; auto; intros.
- - cbn. destruct p. unfold get_forest_p'.
- assert (p0 <> p) by
- (unfold not; intros; apply H; subst; constructor).
- rewrite PTree.gso; auto.
- - cbn in *.
- assert (X: ~ PredIn p p0_1 /\ ~ PredIn p p0_2) by
- (split; unfold not; intros; apply H; constructor; tauto).
- inversion_clear X as [X1 X2].
- inv H0. inv H4.
- specialize (IHp0_1 _ p e _ X1 H0). constructor. tauto.
- specialize (IHp0_2 _ p e _ X2 H0). constructor. tauto.
- constructor; auto.
- - cbn in *.
- assert (X: ~ PredIn p p0_1 /\ ~ PredIn p p0_2) by
- (split; unfold not; intros; apply H; constructor; tauto).
- inversion_clear X as [X1 X2].
- inv H0. inv H4.
- specialize (IHp0_1 _ p e _ X1 H0). constructor. tauto.
- specialize (IHp0_2 _ p e _ X2 H0). constructor. tauto.
- constructor; auto.
- Qed.
-
- Lemma sem_pred_not_in :
- forall A B G (ctx: @ctx G) (s: @Abstr.ctx G -> A -> B -> Prop) l v p e ps,
- sem_pred_expr ps s ctx l v ->
- predicated_not_inP p l ->
- sem_pred_expr (PTree.set p e ps) s ctx l v.
- Proof.
- induction l.
- - intros. unfold predicated_not_inP in H0.
- destruct a. constructor. apply sem_pexpr_not_in.
- eapply H0. econstructor. inv H. auto. inv H. auto.
- - intros. inv H. constructor. unfold predicated_not_inP in H0.
- eapply sem_pexpr_not_in. eapply H0. constructor. left. eauto.
- auto. auto.
- apply sem_pred_expr_cons_false. apply sem_pexpr_not_in. eapply H0.
- constructor. tauto. auto. auto.
- eapply IHl. eauto. eapply predicated_not_inP_cons; eauto.
- Qed.
-
- Lemma pred_not_in_forestP :
- forall pred f,
- predicated_not_in_forest pred f = true ->
- forall x, predicated_not_inP pred (f #r x).
- Proof. Admitted.
-
- Lemma pred_not_in_forest_exitP :
- forall pred f,
- predicated_not_in_forest pred f = true ->
- predicated_not_inP pred (forest_exit f).
- Proof. Admitted.
-
- Lemma from_predicated_sem_pred_expr :
- forall A (ctx: @ctx A) preds pe b,
- sem_pred_expr preds sem_pred ctx pe b ->
- sem_pexpr ctx (from_predicated true preds pe) b.
- Proof. Admitted.
-
- Lemma sem_update_Isetpred:
- forall A (ctx: @ctx A) f pr p0 c args b rs m,
- valid_mem (ctx_mem ctx) m ->
- sem ctx f (mk_instr_state rs pr m, None) ->
- Op.eval_condition c rs ## args m = Some b ->
- truthy pr p0 ->
- sem_pexpr ctx
- (from_predicated true (forest_preds f) (seq_app (pred_ret (PEsetpred c)) (merge (list_translation args f)))) b.
- Proof.
- intros. eapply from_predicated_sem_pred_expr.
- pose proof (sem_merge_list _ ctx f rs pr m args H0).
- apply sem_pred_expr_ident2 in H3; simplify.
- eapply sem_pred_expr_seq_app_val; [eauto| |].
- constructor. constructor. constructor.
- econstructor; eauto.
- erewrite storev_eval_condition; eauto.
- Qed.
-
- Lemma sem_update_Isetpred_top:
- forall A f p p' f' rs m pr args p0 state c pred b,
- Op.eval_condition c rs ## args m = Some b ->
- truthy pr p0 ->
- valid_mem (is_mem (ctx_is state)) m ->
- sem state f ((mk_instr_state rs pr m), None) ->
- update (p, f) (RBsetpred p0 c args pred) = Some (p', f') ->
- eval_predf pr p = true ->
- @sem A state f' (mk_instr_state rs (pr # pred <- b) m, None).
- Proof.
- intros * EVAL_OP TRUTHY VALID SEM UPD EVAL_PRED.
- pose proof SEM as SEM2.
- inversion UPD as [PRUNE]. unfold Option.bind in PRUNE. destr. destr.
- inversion_clear PRUNE.
- rename Heqo into PRUNE.
- inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
- cbn [is_ps] in *. constructor.
- + constructor. intros. apply sem_pred_not_in. rewrite forest_pred_reg.
- inv REG. auto. rewrite forest_pred_reg. apply pred_not_in_forestP.
- unfold assert_ in *. repeat (destruct_match; try discriminate); auto.
- + constructor; intros. destruct (peq x pred); subst.
- * rewrite Regmap.gss.
- rewrite forest_pred_gss.
- cbn [update] in *. unfold Option.bind in *. destr. destr. inv UPD.
- replace b with (b && true) by eauto with bool.
- apply sem_pexpr_Pand.
- destruct b. constructor. right.
- eapply sem_update_Isetpred; eauto.
- constructor. constructor. replace false with (negb true) by auto.
- apply sem_pexpr_negate. inv TRUTHY. constructor.
- cbn. eapply sem_pexpr_eval. inv PRED. eauto. auto.
- replace false with (negb true) by auto.
- apply sem_pexpr_negate.
- eapply sem_pexpr_eval. inv PRED. eauto. auto.
- eapply sem_update_Isetpred; eauto.
- constructor. constructor. constructor.
- replace true with (negb false) by auto. apply sem_pexpr_negate.
- eapply sem_pexpr_eval. inv PRED. eauto. inv TRUTHY. auto. cbn -[eval_predf].
- rewrite eval_predf_negate. rewrite H; auto.
- replace true with (negb false) by auto. apply sem_pexpr_negate.
- eapply sem_pexpr_eval. inv PRED. eauto. rewrite eval_predf_negate.
- rewrite EVAL_PRED. auto.
- * rewrite Regmap.gso by auto. inv PRED. specialize (H x).
- rewrite forest_pred_gso by auto; auto.
- + rewrite forest_pred_reg. apply sem_pred_not_in. auto. apply pred_not_in_forestP.
- unfold assert_ in *. now repeat (destruct_match; try discriminate).
- + cbn -[from_predicated from_pred_op seq_app]. apply sem_pred_not_in; auto.
- apply pred_not_in_forest_exitP.
- unfold assert_ in *. now repeat (destruct_match; try discriminate).
- Qed.
-
- Lemma sem_pexpr_impl :
- forall A (state: @ctx A) a b res,
- sem_pexpr state b res ->
- sem_pexpr state a true ->
- sem_pexpr state (a → b) res.
- Proof.
- intros. destruct res.
- constructor; tauto.
- constructor; auto. replace false with (negb true) by auto.
- now apply sem_pexpr_negate.
- Qed.
-
- Lemma eval_predf_simplify :
- forall ps x,
- eval_predf ps (simplify x) = eval_predf ps x.
- Proof.
- unfold eval_predf; intros.
- rewrite simplify_correct. auto.
- Qed.
-
- Lemma sem_update_falsy:
- forall A state f f' rs ps m p a p',
- instr_falsy ps a ->
- update (p, f) a = Some (p', f') ->
- sem state f (mk_instr_state rs ps m, None) ->
- @sem A state f' (mk_instr_state rs ps m, None).
- Proof.
- inversion 1; cbn [update] in *; intros.
- - unfold Option.bind in *. destr. inv H2.
- constructor.
- * constructor. intros. destruct (peq x res); subst.
- rewrite forest_reg_gss. cbn.
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated_false. inv H3. inv H8. auto.
- inv H3. inv H9. eauto. rewrite eval_predf_Pand. cbn -[eval_predf].
- rewrite H0. auto.
- rewrite forest_reg_gso. inv H3. inv H8. auto.
- unfold not; intros; apply n0. now inv H1.
- * constructor; intros. rewrite forest_reg_pred. inv H3. inv H9. auto.
- * rewrite forest_reg_gso. inv H3. auto. unfold not; intros. inversion H1.
- * inv H3. auto.
- - unfold Option.bind in *. destr. inv H2.
- constructor.
- * constructor. intros. destruct (peq x dst); subst.
- rewrite forest_reg_gss. cbn.
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated_false. inv H3. inv H8. auto.
- inv H3. inv H9. eauto. rewrite eval_predf_Pand. cbn -[eval_predf].
- rewrite H0. auto.
- rewrite forest_reg_gso. inv H3. inv H8. auto.
- unfold not; intros; apply n0. now inv H1.
- * constructor; intros. rewrite forest_reg_pred. inv H3. inv H9. auto.
- * rewrite forest_reg_gso. inv H3. auto. unfold not; intros. inversion H1.
- * inv H3. auto.
- - unfold Option.bind in *. destr. inv H2.
- constructor.
- * constructor. intros. rewrite forest_reg_gso by discriminate. inv H3. inv H8. auto.
- * constructor; intros. rewrite forest_reg_pred. inv H3. inv H9. auto.
- * rewrite forest_reg_gss. cbn. eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated_false. inv H3. auto. inv H3. inv H9. eauto.
- rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H0. auto.
- * inv H3. auto.
- - unfold Option.bind in *. destr. destr. inv H2.
- constructor.
- * constructor; intros. rewrite forest_pred_reg. apply sem_pred_not_in.
- inv H3. inv H8. auto. apply pred_not_in_forestP. unfold assert_ in Heqo. now destr.
- * constructor. intros. destruct (peq x pred); subst.
- rewrite forest_pred_gss. replace (ps !! pred) with (true && ps !! pred) by auto.
- assert (sem_pexpr state0 (¬ (from_pred_op (forest_preds f) p0 ∧ from_pred_op (forest_preds f) p')) true).
- { replace true with (negb false) by auto. apply sem_pexpr_negate.
- constructor. left. eapply sem_pexpr_eval. inv H3. inv H9. eauto.
- auto.
- }
- apply sem_pexpr_Pand. constructor; tauto.
- apply sem_pexpr_impl. inv H3. inv H10. eauto.
- { constructor. left. eapply sem_pexpr_eval. inv H3. inv H10. eauto.
- rewrite eval_predf_negate. rewrite H0. auto.
- }
- rewrite forest_pred_gso by auto. inv H3. inv H9. auto.
- * rewrite forest_pred_reg. apply sem_pred_not_in. inv H3. auto.
- apply pred_not_in_forestP. unfold assert_ in Heqo. now destr.
- * apply sem_pred_not_in. inv H3; auto. cbn.
- apply pred_not_in_forest_exitP. unfold assert_ in Heqo. now destr.
- - unfold Option.bind in *. destr. inv H2. inv H3. constructor.
- * constructor. inv H8. auto.
- * constructor. intros. inv H9. eauto.
- * auto.
- * cbn. eapply sem_pred_expr_prune_predicated; [|eauto].
- eapply sem_pred_expr_app_predicated_false; auto.
- inv H9. eauto.
- rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H0. auto.
- Qed.
-
- Lemma sem_update_falsy_input:
- forall A state f f' rs ps m p a p' exitcf,
- eval_predf ps p = false ->
- update (p, f) a = Some (p', f') ->
- sem state f (mk_instr_state rs ps m, exitcf) ->
- @sem A state f' (mk_instr_state rs ps m, exitcf)
- /\ eval_predf ps p' = false.
- Proof.
- intros; destruct a; cbn [update] in *; intros.
- - inv H0. auto.
- - unfold Option.bind in *. destr. inv H0. split; [|solve [auto]].
- constructor.
- * constructor. intros. destruct (peq x r); subst.
- rewrite forest_reg_gss. cbn.
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated_false. inv H1. inv H7. auto.
- inv H1. inv H8. eauto. rewrite eval_predf_Pand.
- rewrite H. eauto with bool.
- rewrite forest_reg_gso. inv H1. inv H7. auto.
- unfold not; intros; apply n0. now inv H0.
- * constructor; intros. rewrite forest_reg_pred. inv H1. inv H8. auto.
- * rewrite forest_reg_gso. inv H1. auto. unfold not; intros. inversion H0.
- * inv H1. auto.
- - unfold Option.bind in *. destr. inv H0. split; [|solve [auto]].
- constructor.
- * constructor. intros. destruct (peq x r); subst.
- rewrite forest_reg_gss. cbn.
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated_false. inv H1. inv H7. auto.
- inv H1. inv H8. eauto. rewrite eval_predf_Pand. cbn -[eval_predf].
- rewrite H. eauto with bool.
- rewrite forest_reg_gso. inv H1. inv H7. auto.
- unfold not; intros; apply n0. now inv H0.
- * constructor; intros. rewrite forest_reg_pred. inv H1. inv H8. auto.
- * rewrite forest_reg_gso. inv H1. auto. unfold not; intros. inversion H0.
- * inv H1. auto.
- - unfold Option.bind in *. destr. inv H0. split; [|solve [auto]].
- constructor.
- * constructor. intros. rewrite forest_reg_gso by discriminate. inv H1. inv H7. auto.
- * constructor; intros. rewrite forest_reg_pred. inv H1. inv H8. auto.
- * rewrite forest_reg_gss. cbn. eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated_false. inv H1. auto. inv H1. inv H8. eauto.
- rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H. eauto with bool.
- * inv H1. auto.
- - unfold Option.bind in *. destr. destr. inv H0. split; [|solve [auto]].
- constructor.
- * constructor; intros. rewrite forest_pred_reg. apply sem_pred_not_in.
- inv H1. inv H7. auto. apply pred_not_in_forestP. unfold assert_ in Heqo0. now destr.
- * constructor. intros. destruct (peq x p0); subst.
- rewrite forest_pred_gss. replace (ps !! p0) with (true && ps !! p0) by auto.
- assert (sem_pexpr state0 (¬ (from_pred_op (forest_preds f) (dfltp o) ∧ from_pred_op (forest_preds f) p')) true).
- { replace true with (negb false) by auto. apply sem_pexpr_negate.
- constructor. right. eapply sem_pexpr_eval. inv H1. inv H8. eauto.
- auto.
- }
- apply sem_pexpr_Pand. constructor; tauto.
- apply sem_pexpr_impl. inv H1. inv H9. eauto.
- { constructor. right. eapply sem_pexpr_eval. inv H1. inv H9. eauto.
- rewrite eval_predf_negate. rewrite H. auto.
- }
- rewrite forest_pred_gso by auto. inv H1. inv H8. auto.
- * rewrite forest_pred_reg. apply sem_pred_not_in. inv H1. auto.
- apply pred_not_in_forestP. unfold assert_ in Heqo0. now destr.
- * apply sem_pred_not_in. inv H1; auto. cbn.
- apply pred_not_in_forest_exitP. unfold assert_ in Heqo0. now destr.
- - unfold Option.bind in *. destr. inv H0. inv H1. split.
- -- constructor.
- * constructor. inv H7. auto.
- * constructor. intros. inv H8. eauto.
- * auto.
- * cbn. eapply sem_pred_expr_prune_predicated; [|eauto].
- eapply sem_pred_expr_app_predicated_false; auto.
- inv H8. eauto.
- rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H. eauto with bool.
- -- rewrite eval_predf_simplify. rewrite eval_predf_Pand. rewrite H. eauto with bool.
- Qed.
-
- Definition setpred_wf (i: instr): bool :=
- match i with
- | RBsetpred (Some op) _ _ p => negb (predin peq p op)
- | _ => true
- end.
-
- Lemma sem_update_instr :
- forall f i' i'' a sp p i p' f',
- step_instr ge sp (Iexec i') a (Iexec i'') ->
- valid_mem (is_mem i) (is_mem i') ->
- sem (mk_ctx i sp ge) f (i', None) ->
- update (p, f) a = Some (p', f') ->
- eval_predf (is_ps i') p = true ->
- sem (mk_ctx i sp ge) f' (i'', None).
- Proof.
- inversion 1; subst; clear H; intros VALID SEM UPD EVAL_PRED; pose proof SEM as SEM2.
- - inv UPD; auto.
- - eauto using sem_update_Iop_top.
- - eauto using sem_update_Iload_top.
- - eauto using sem_update_Istore_top.
- - eauto using sem_update_Isetpred_top.
- - destruct i''. eauto using sem_update_falsy.
- Qed.
-
- Lemma Pand_true_left :
- forall ps a b,
- eval_predf ps a = false ->
- eval_predf ps (a ∧ b) = false.
- Proof.
- intros.
- rewrite eval_predf_Pand. now rewrite H.
- Qed.
-
- Lemma Pand_true_right :
- forall ps a b,
- eval_predf ps b = false ->
- eval_predf ps (a ∧ b) = false.
- Proof.
- intros.
- rewrite eval_predf_Pand. rewrite H.
- eauto with bool.
- Qed.
-
- Lemma sem_update_instr_term :
- forall f i' i'' sp i cf p p' p'' f',
- sem (mk_ctx i sp ge) f (i', None) ->
- step_instr ge sp (Iexec i') (RBexit p cf) (Iterm i'' cf) ->
- update (p', f) (RBexit p cf) = Some (p'', f') ->
- eval_predf (is_ps i') p' = true ->
- sem (mk_ctx i sp ge) f' (i'', Some cf)
- /\ eval_predf (is_ps i') p'' = false.
- Proof.
- intros. inv H0. simpl in *.
- unfold Option.bind in *. destruct_match; try discriminate.
- apply truthy_dflt in H6. inv H1.
- assert (eval_predf (Gible.is_ps i'') (¬ dfltp p) = false).
- { rewrite eval_predf_negate. now rewrite negb_false_iff. }
- apply Pand_true_left with (b := p') in H0.
- rewrite <- eval_predf_simplify in H0. split; auto.
- unfold "<-e". destruct i''.
- inv H. constructor; auto.
- constructor. inv H9. intros. cbn. eauto.
- inv H10. constructor; intros. eauto.
- cbn.
- eapply sem_pred_expr_prune_predicated; eauto.
- eapply sem_pred_expr_app_predicated.
- constructor. constructor. constructor.
- inv H10. eauto. cbn -[eval_predf] in *.
- rewrite eval_predf_Pand. rewrite H2. now rewrite H6.
- Qed.
-
- Lemma step_instr_lessdef_term :
- forall sp a i i' ti cf,
- step_instr ge sp (Iexec i) a (Iterm i' cf) ->
- Abstr.match_states i ti ->
- exists ti', step_instr ge sp (Iexec ti) a (Iterm ti' cf) /\ Abstr.match_states i' ti'.
- Proof.
- inversion 1; intros; subst.
- econstructor. split; eauto. constructor.
- destruct p. constructor. erewrite eval_predf_pr_equiv. inv H4.
- eauto. inv H6. eauto. constructor.
- Qed.
-
- Lemma combined_falsy :
- forall i o1 o,
- falsy i o1 ->
- falsy i (combine_pred o o1).
- Proof.
- inversion 1; subst; crush. destruct o; simplify.
- constructor. rewrite eval_predf_Pand. rewrite H0. crush.
- constructor. auto.
- Qed.
-
- Lemma Abstr_match_states_sem :
- forall i sp f i' ti cf,
- sem (mk_ctx i sp ge) f (i', cf) ->
- Abstr.match_states i ti ->
- exists ti', sem (mk_ctx ti sp ge) f (ti', cf) /\ Abstr.match_states i' ti'.
- Proof. Admitted. (* This needs a bit more in Abstr.v *)
-
- Lemma mfold_left_update_Some :
- forall xs x v,
- mfold_left update xs x = Some v ->
- exists y, x = Some y.
- Proof.
- induction xs; intros.
- { cbn in *. inv H. eauto. }
- cbn in *. unfold Option.bind in *. exploit IHxs; eauto.
- intros. simplify. destruct x; crush.
- eauto.
- Qed.
-
- Lemma step_instr_term_exists :
- forall A B ge sp v x v2 cf,
- @step_instr A B ge sp (Iexec v) x (Iterm v2 cf) ->
- exists p, x = RBexit p cf.
- Proof using. inversion 1; eauto. Qed.
-
- Lemma eval_predf_update_true :
- forall i i' curr_p next_p f f_next instr sp,
- update (curr_p, f) instr = Some (next_p, f_next) ->
- step_instr ge sp (Iexec i) instr (Iexec i') ->
- eval_predf (is_ps i) curr_p = true ->
- eval_predf (is_ps i') next_p = true.
- Proof.
- intros * UPD STEP EVAL. destruct instr; cbn [update] in UPD;
- try solve [unfold Option.bind in *; try destr; inv UPD; inv STEP; auto].
- - unfold Option.bind in *. destr. destr. inv UPD. inv STEP; auto. cbn [is_ps] in *.
- unfold is_initial_pred_and_notin in Heqo1. unfold assert_ in Heqo1. destr. destr.
- destr. destr. destr. destr. subst. assert (~ PredIn p2 next_p).
- unfold not; intros. apply negb_true_iff in Heqb0. apply not_true_iff_false in Heqb0.
- apply Heqb0. now apply predin_PredIn. rewrite eval_predf_not_PredIn; auto.
- - unfold Option.bind in *. destr. inv UPD. inv STEP. inv H3. cbn.
- rewrite eval_predf_simplify. rewrite eval_predf_Pand. rewrite eval_predf_negate.
- destruct i'; cbn in *. rewrite H0. auto.
- Qed.
-
- Lemma forest_evaluable_regset :
- forall A f (ctx: @ctx A) n x,
- forest_evaluable ctx f ->
- forest_evaluable ctx f #r x <- n.
- Proof.
- unfold forest_evaluable, pred_forest_evaluable; intros.
- eapply H. eauto.
- Qed.
-
- Lemma evaluable_impl :
- forall A (ctx: @ctx A) a b,
- p_evaluable ctx a ->
- p_evaluable ctx b ->
- p_evaluable ctx (a → b).
- Proof.
- unfold evaluable.
- inversion_clear 1 as [b1 SEM1].
- inversion_clear 1 as [b2 SEM2].
- unfold Pimplies.
- econstructor. apply sem_pexpr_Por;
- eauto using sem_pexpr_negate.
- Qed.
-
- Lemma parts_evaluable :
- forall A (ctx: @ctx A) b p0,
- evaluable sem_pred ctx p0 ->
- evaluable sem_pexpr ctx (Plit (b, p0)).
- Proof.
- intros. unfold evaluable in *. inv H.
- destruct b. do 2 econstructor. eauto.
- exists (negb x). constructor. rewrite negb_involutive. auto.
- Qed.
-
- Lemma p_evaluable_Pand :
- forall A (ctx: @ctx A) a b,
- p_evaluable ctx a ->
- p_evaluable ctx b ->
- p_evaluable ctx (a ∧ b).
- Proof.
- intros. inv H. inv H0.
- econstructor. apply sem_pexpr_Pand; eauto.
- Qed.
-
- Lemma from_predicated_evaluable :
- forall A (ctx: @ctx A) f b a,
- pred_forest_evaluable ctx f ->
- all_evaluable2 ctx sem_pred a ->
- p_evaluable ctx (from_predicated b f a).
- Proof.
- induction a.
- cbn. destruct_match; intros.
- eapply evaluable_impl.
- eauto using predicated_evaluable.
- unfold all_evaluable2 in H0. unfold p_evaluable.
- eapply parts_evaluable. eapply H0. econstructor.
-
- intros. cbn. destruct_match.
- eapply p_evaluable_Pand.
- eapply all_evaluable2_cons2 in H0.
- eapply evaluable_impl.
- eauto using predicated_evaluable.
- unfold all_evaluable2 in H0. unfold p_evaluable.
- eapply parts_evaluable. eapply H0.
- eapply all_evaluable2_cons in H0. eauto.
- Qed.
-
- Lemma seq_app_cons :
- forall A B f a l p0 p1,
- @seq_app A B (pred_ret f) (NE.cons a l) = NE.cons p0 p1 ->
- @seq_app A B (pred_ret f) l = p1.
- Proof. intros. cbn in *. inv H. eauto. Qed.
-
- Lemma p_evaluable_imp :
- forall A (ctx: @ctx A) a b,
- sem_pexpr ctx a false ->
- p_evaluable ctx (a → b).
- Proof.
- intros. unfold "→".
- unfold p_evaluable, evaluable. exists true.
- constructor. replace true with (negb false) by trivial. left.
- now apply sem_pexpr_negate.
- Qed.
-
- Lemma sem_update_valid_mem :
- forall i i' i'' curr_p next_p f f_next instr sp,
- step_instr ge sp (Iexec i') instr (Iexec i'') ->
- update (curr_p, f) instr = Some (next_p, f_next) ->
- sem (mk_ctx i sp ge) f (i', None) ->
- valid_mem (is_mem i') (is_mem i'').
- Proof.
- inversion 1; crush.
- unfold Option.bind in *. destr. inv H7.
- eapply storev_valid_pointer; eauto.
- Qed.
-
- Lemma eval_predf_lessdef :
- forall p a b,
- Abstr.match_states a b ->
- eval_predf (is_ps a) p = eval_predf (is_ps b) p.
- Proof using.
- induction p; crush.
- - inv H. simpl. unfold eval_predf. simpl.
- repeat destr. inv Heqp0. rewrite H1. auto.
- - rewrite !eval_predf_Pand.
- erewrite IHp1 by eassumption.
- now erewrite IHp2 by eassumption.
- - rewrite !eval_predf_Por.
- erewrite IHp1 by eassumption.
- now erewrite IHp2 by eassumption.
- Qed.
-
-(*|
-``abstr_fold_falsy``: This lemma states that when we are at the end of an
-execution, the values in the register map do not continue to change.
-|*)
-
- Lemma abstr_fold_falsy :
- forall A ilist i sp ge f res p f' p',
- @sem A (mk_ctx i sp ge) f res ->
- mfold_left update ilist (Some (p, f)) = Some (p', f') ->
- eval_predf (is_ps (fst res)) p = false ->
- sem (mk_ctx i sp ge) f' res.
- Proof.
- induction ilist.
- - intros. cbn in *. inv H0. auto.
- - intros. cbn -[update] in H0.
- exploit mfold_left_update_Some. eauto. intros. inv H2.
- rewrite H3 in H0. destruct x.
- destruct res. destruct i0.
- exploit sem_update_falsy_input; try eassumption; intros.
- inversion_clear H2.
- eapply IHilist; eassumption.
- Qed.
-
- Lemma abstr_fold_correct :
- forall sp x i i' i'' cf f p f' curr_p
- (VALID: valid_mem (is_mem i) (is_mem i')),
- SeqBB.step ge sp (Iexec i') x (Iterm i'' cf) ->
- sem (mk_ctx i sp ge) f (i', None) ->
- eval_predf (is_ps i') curr_p = true ->
- mfold_left update x (Some (curr_p, f)) = Some (p, f') ->
- forall ti,
- Abstr.match_states i ti ->
- exists ti', sem (mk_ctx ti sp ge) f' (ti', Some cf)
- /\ Abstr.match_states i'' ti'
- /\ valid_mem (is_mem i) (is_mem i'').
- Proof.
- induction x as [| x xs IHx ]; intros; cbn -[update] in *; inv H; cbn [fst snd] in *.
- - (* inductive case *)
- exploit mfold_left_update_Some; eauto; intros Hexists;
- inversion Hexists as [[curr_p_inter f_inter] EXEQ]; clear Hexists.
- exploit eval_predf_update_true;
- eauto; intros EVALTRUE.
- rewrite EXEQ in H2. eapply IHx in H2; cbn [fst snd] in *.
- eauto.
- transitivity (is_mem i'); auto.
- eapply sem_update_valid_mem; eauto.
- eauto.
- eapply sem_update_instr; eauto. eauto. eauto.
- - (* terminal case *)
- exploit mfold_left_update_Some; eauto; intros Hexists;
- inversion Hexists as [[curr_p_inter f_inter] EXEQ]; clear Hexists. rewrite EXEQ in H2.
- exploit Abstr_match_states_sem; (* TODO *)
- eauto; intros H; inversion H as [v [Hi LESSDEF]]; clear H.
- exploit step_instr_lessdef_term;
- eauto; intros H; inversion H as [v2 [Hi2 LESSDEF2]]; clear H.
- exploit step_instr_term_exists; eauto; inversion 1 as [? ?]; subst; clear H.
- erewrite eval_predf_lessdef in H1 by eassumption.
- exploit sem_update_instr_term;
- eauto; intros [A B].
- exists v2.
- exploit abstr_fold_falsy.
- apply A.
- eassumption. auto. cbn. inversion Hi2; subst. auto. intros.
- split; auto. split; auto.
- inversion H7; subst; auto.
- Qed.
-
- Lemma sem_regset_empty :
- forall A ctx, @sem_regset A ctx empty (ctx_rs ctx).
- Proof using.
- intros; constructor; intros.
- constructor; auto. constructor.
- constructor.
- Qed.
-
- Lemma sem_predset_empty :
- forall A ctx, @sem_predset A ctx empty (ctx_ps ctx).
- Proof using.
- intros; constructor; intros.
- constructor; auto. constructor.
- Qed.
-
- Lemma sem_empty :
- forall A ctx, @sem A ctx empty (ctx_is ctx, None).
- Proof using.
- intros. destruct ctx. destruct ctx_is.
- constructor; try solve [constructor; constructor; crush].
- eapply sem_regset_empty.
- eapply sem_predset_empty.
- Qed.
-
- Lemma abstr_sequence_correct :
- forall sp x i i'' cf x',
- SeqBB.step ge sp (Iexec i) x (Iterm i'' cf) ->
- abstract_sequence x = Some x' ->
- forall ti,
- Abstr.match_states i ti ->
- exists ti', sem (mk_ctx ti sp ge) x' (ti', Some cf)
- /\ Abstr.match_states i'' ti'.
- Proof.
- unfold abstract_sequence. intros. unfold Option.map in H0.
- destruct_match; try easy.
- destruct p as [p TMP]; simplify.
- exploit abstr_fold_correct; eauto; crush.
- { apply sem_empty. }
- exists x0. auto.
- Qed.
-
- Lemma abstr_seq_reverse_correct :
- forall sp x i i' ti cf x',
- abstract_sequence x = Some x' ->
- sem (mk_ctx i sp ge) x' (i', (Some cf)) ->
- Abstr.match_states i ti ->
- exists ti', SeqBB.step ge sp (Iexec ti) x (Iterm ti' cf)
- /\ Abstr.match_states i' ti'.
- Proof.
-
-(*|
-Proof Sketch:
-
-We do an induction over the list of instructions ``x``. This is trivial for the
-empty case and then for the inductive case we know that there exists an
-execution that matches the abstract execution, so we need to know that adding
-another instructions to it will still mean that the execution will result in the
-same value.
-
-Arithmetic operations will be a problem because we will have to show that these
-can be executed. However, this should mostly be a problem in the abstract state
-comparison, because there two abstract states can be equal without one being
-evaluable.
-|*)
-
- Admitted.
-
-(*|
-This is the top-level lemma which uses the following proofs to complete the
-square:
-
-- ``abstr_sequence_correct``: This is the lemma that states the forward
- translation form ``GibleSeq`` to ``Abstr`` was correct.
-- ``abstr_check_correct``: This is the lemma that states that if a check between
- two ``Abstr`` programs succeeds, that they will also behave the same. This
- depends on the SAT solver correctness, as the predicates might be
- syntactically different to each other.
-- ``abstr_seq_reverse_correct``: This is the lemma that shows that the backwards
- simulation between the abstract translation and the concrete execution also
- holds. We only have a translation from the concrete into the abstract, but
- then prove that if we have an execution in the abstract, we can observe that
- same execution in the concrete.
-- ``seqbb_step_parbb_step``: Finally, this lemma states that the parallel
- execution of the basic block is equivalent to the sequential execution of the
- concatenation of that parallel block. This is because even in the translation
- to HTL, the Verilog semantics are sequential within a clock cycle, but will
- then be parallelised by the synthesis tool. The argument for why this is
- still useful is because we are identifying and scheduling instructions into
- clock cycles.
-|*)
-
- Definition local_abstr_check_correct :=
- @abstr_check_correct GibleSeq.fundef GiblePar.fundef.
-
- Definition local_abstr_check_correct2 :=
- @abstr_check_correct GibleSeq.fundef GibleSeq.fundef.
-
- Lemma ge_preserved_local :
- ge_preserved ge tge.
- Proof.
- unfold ge_preserved;
- eauto using Op.eval_operation_preserved, Op.eval_addressing_preserved.
- Qed.
-
Lemma schedule_oracle_correct :
forall x y sp i i' ti cf,
schedule_oracle x y = true ->
SeqBB.step ge sp (Iexec i) x (Iterm i' cf) ->
- Abstr.match_states i ti ->
+ state_lessdef i ti ->
exists ti', ParBB.step tge sp (Iexec ti) y (Iterm ti' cf)
- /\ Abstr.match_states i' ti'.
+ /\ state_lessdef i' ti'.
Proof.
unfold schedule_oracle; intros. repeat (destruct_match; try discriminate). simplify.
exploit abstr_sequence_correct; eauto; simplify.
@@ -2406,7 +479,7 @@ Proof Sketch: Trivial because of structural equality.
Lemma match_states_stepBB :
forall s f sp pc rs pr m sf' f' trs tps tm rs' pr' m' trs' tpr' tm',
match_states (GibleSeq.State s f sp pc rs pr m) (State sf' f' sp pc trs tps tm) ->
- Abstr.match_states (mk_instr_state rs' pr' m') (mk_instr_state trs' tpr' tm') ->
+ state_lessdef (mk_instr_state rs' pr' m') (mk_instr_state trs' tpr' tm') ->
match_states (GibleSeq.State s f sp pc rs' pr' m') (State sf' f' sp pc trs' tpr' tm').
Proof.
inversion 1; subst; simplify.
diff --git a/src/hls/GiblePargenproofBackwards.v b/src/hls/GiblePargenproofBackwards.v
new file mode 100644
index 0000000..50b733a
--- /dev/null
+++ b/src/hls/GiblePargenproofBackwards.v
@@ -0,0 +1,49 @@
+(*
+ * Vericert: Verified high-level synthesis.
+ * Copyright (C) 2023 Yann Herklotz <git@yannherklotz.com>
+ *
+ * This program is free software: you can redistribute it and/or modify
+ * it under the terms of the GNU General Public License as published by
+ * the Free Software Foundation, either version 3 of the License, or
+ * (at your option) any later version.
+ *
+ * This program is distributed in the hope that it will be useful,
+ * but WITHOUT ANY WARRANTY; without even the implied warranty of
+ * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+ * GNU General Public License for more details.
+ *
+ * You should have received a copy of the GNU General Public License
+ * along with this program. If not, see <https://www.gnu.org/licenses/>.
+ *)
+
+Require Import compcert.backend.Registers.
+Require Import compcert.common.AST.
+Require Import compcert.common.Errors.
+Require Import compcert.common.Linking.
+Require Import compcert.common.Globalenvs.
+Require Import compcert.common.Memory.
+Require Import compcert.common.Values.
+Require Import compcert.lib.Maps.
+
+Require Import vericert.common.Vericertlib.
+Require Import vericert.hls.GibleSeq.
+Require Import vericert.hls.GiblePar.
+Require Import vericert.hls.Gible.
+Require Import vericert.hls.GiblePargenproofEquiv.
+Require Import vericert.hls.GiblePargen.
+Require Import vericert.hls.Predicate.
+Require Import vericert.hls.Abstr.
+Require Import vericert.common.Monad.
+
+Module Import OptionExtra := MonadExtra(Option).
+
+#[local] Open Scope positive.
+#[local] Open Scope forest.
+#[local] Open Scope pred_op.
+
+#[local] Opaque simplify.
+#[local] Opaque deep_simplify.
+
+#[local] Ltac destr := destruct_match; try discriminate; [].
+
+Definition state_lessdef := GiblePargenproofEquiv.match_states.
diff --git a/src/hls/GiblePargenproofEquiv.v b/src/hls/GiblePargenproofEquiv.v
new file mode 100644
index 0000000..5f28d53
--- /dev/null
+++ b/src/hls/GiblePargenproofEquiv.v
@@ -0,0 +1,1736 @@
+(*
+ * Vericert: Verified high-level synthesis.
+ * Copyright (C) 2023 Yann Herklotz <git@yannherklotz.com>
+ *
+ * This program is free software: you can redistribute it and/or modify
+ * it under the terms of the GNU General Public License as published by
+ * the Free Software Foundation, either version 3 of the License, or
+ * (at your option) any later version.
+ *
+ * This program is distributed in the hope that it will be useful,
+ * but WITHOUT ANY WARRANTY; without even the implied warranty of
+ * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+ * GNU General Public License for more details.
+ *
+ * You should have received a copy of the GNU General Public License
+ * along with this program. If not, see <https://www.gnu.org/licenses/>.
+ *)
+
+Require Import Coq.Logic.Decidable.
+Require Import Coq.Structures.Equalities.
+
+Require Import compcert.backend.Registers.
+Require Import compcert.common.AST.
+Require Import compcert.common.Globalenvs.
+Require Import compcert.common.Memory.
+Require Import compcert.common.Values.
+Require Import compcert.lib.Floats.
+Require Import compcert.lib.Integers.
+Require Import compcert.lib.Maps.
+Require compcert.verilog.Op.
+
+Require Import vericert.common.Vericertlib.
+Require Import vericert.hls.GibleSeq.
+Require Import vericert.hls.GiblePar.
+Require Import vericert.hls.Gible.
+Require Import vericert.hls.HashTree.
+Require Import vericert.hls.Predicate.
+Require Import vericert.common.DecEq.
+Require Import vericert.hls.Abstr.
+Require vericert.common.NonEmpty.
+Module NE := NonEmpty.
+Import NE.NonEmptyNotation.
+
+#[local] Open Scope non_empty_scope.
+#[local] Open Scope positive.
+#[local] Open Scope pred_op.
+#[local] Open Scope forest.
+
+Fixpoint beq_expression (e1 e2: expression) {struct e1}: bool :=
+ match e1, e2 with
+ | Ebase r1, Ebase r2 => if resource_eq r1 r2 then true else false
+ | Eop op1 el1, Eop op2 el2 =>
+ if operation_eq op1 op2 then
+ beq_expression_list el1 el2 else false
+ | Eload chk1 addr1 el1 e1, Eload chk2 addr2 el2 e2 =>
+ if memory_chunk_eq chk1 chk2
+ then if addressing_eq addr1 addr2
+ then if beq_expression_list el1 el2
+ then beq_expression e1 e2 else false else false else false
+ | Estore e1 chk1 addr1 el1 m1, Estore e2 chk2 addr2 el2 m2 =>
+ if memory_chunk_eq chk1 chk2
+ then if addressing_eq addr1 addr2
+ then if beq_expression_list el1 el2
+ then if beq_expression m1 m2
+ then beq_expression e1 e2 else false else false else false else false
+ | _, _ => false
+ end
+with beq_expression_list (el1 el2: expression_list) {struct el1} : bool :=
+ match el1, el2 with
+ | Enil, Enil => true
+ | Econs e1 t1, Econs e2 t2 => beq_expression e1 e2 && beq_expression_list t1 t2
+ | _, _ => false
+ end
+.
+
+Scheme expression_ind2 := Induction for expression Sort Prop
+ with expression_list_ind2 := Induction for expression_list Sort Prop.
+Definition beq_pred_expression (e1 e2: pred_expression) : bool :=
+ match e1, e2 with
+ | PEbase p1, PEbase p2 => if peq p1 p2 then true else false
+ | PEsetpred c1 el1, PEsetpred c2 el2 =>
+ if condition_eq c1 c2
+ then beq_expression_list el1 el2 else false
+ | _, _ => false
+ end.
+
+Definition beq_exit_expression (e1 e2: exit_expression) : bool :=
+ match e1, e2 with
+ | EEbase, EEbase => true
+ | EEexit cf1, EEexit cf2 => if cf_instr_eq cf1 cf2 then true else false
+ | _, _ => false
+ end.
+
+Lemma beq_expression_correct:
+ forall e1 e2, beq_expression e1 e2 = true -> e1 = e2.
+Proof.
+ intro e1;
+ apply expression_ind2 with
+ (P := fun (e1 : expression) =>
+ forall e2, beq_expression e1 e2 = true -> e1 = e2)
+ (P0 := fun (e1 : expression_list) =>
+ forall e2, beq_expression_list e1 e2 = true -> e1 = e2); simplify;
+ try solve [repeat match goal with
+ | [ H : context[match ?x with _ => _ end] |- _ ] => destruct x eqn:?
+ | [ H : context[if ?x then _ else _] |- _ ] => destruct x eqn:?
+ end; subst; f_equal; crush; eauto using Peqb_true_eq].
+Qed.
+
+Lemma beq_expression_refl: forall e, beq_expression e e = true.
+Proof.
+ intros.
+ induction e using expression_ind2 with (P0 := fun el => beq_expression_list el el = true);
+ crush; repeat (destruct_match; crush); [].
+ crush. rewrite IHe. rewrite IHe0. auto.
+Qed.
+
+Lemma beq_expression_list_refl: forall e, beq_expression_list e e = true.
+Proof. induction e; auto. simplify. rewrite beq_expression_refl. auto. Qed.
+
+Lemma beq_expression_correct2:
+ forall e1 e2, beq_expression e1 e2 = false -> e1 <> e2.
+Proof.
+ induction e1 using expression_ind2
+ with (P0 := fun el1 => forall el2, beq_expression_list el1 el2 = false -> el1 <> el2).
+ - intros. simplify. repeat (destruct_match; crush).
+ - intros. simplify. repeat (destruct_match; crush). subst. apply IHe1 in H.
+ unfold not in *. intros. apply H. inv H0. auto.
+ - intros. simplify. repeat (destruct_match; crush); subst.
+ unfold not in *; intros. inv H0. rewrite beq_expression_refl in H. discriminate.
+ unfold not in *; intros. inv H. rewrite beq_expression_list_refl in Heqb. discriminate.
+ - simplify. repeat (destruct_match; crush); subst;
+ unfold not in *; intros.
+ inv H0. rewrite beq_expression_refl in H; crush.
+ inv H. rewrite beq_expression_refl in Heqb0; crush.
+ inv H. rewrite beq_expression_list_refl in Heqb; crush.
+ (* - simplify. repeat (destruct_match; crush); subst. *)
+ (* unfold not in *; intros. inv H0. rewrite beq_expression_list_refl in H; crush. *)
+ - simplify. repeat (destruct_match; crush); subst.
+ - simplify. repeat (destruct_match; crush); subst.
+ apply andb_false_iff in H. inv H. unfold not in *; intros.
+ inv H. rewrite beq_expression_refl in H0; discriminate.
+ unfold not in *; intros. inv H. rewrite beq_expression_list_refl in H0; discriminate.
+Qed.
+
+Definition expression_dec: forall e1 e2: expression, {e1 = e2} + {e1 <> e2}.
+Proof.
+ intros.
+ destruct (beq_expression e1 e2) eqn:?. apply beq_expression_correct in Heqb. auto.
+ apply beq_expression_correct2 in Heqb. auto.
+Defined.
+
+Lemma beq_expression_list_correct:
+ forall e1 e2, beq_expression_list e1 e2 = true -> e1 = e2.
+Proof.
+ induction e1; crush.
+ - destruct_match; crush.
+ - destruct_match; crush.
+ apply IHe1 in H1; subst.
+ apply beq_expression_correct in H0; subst.
+ trivial.
+Qed.
+
+Lemma beq_expression_list_correct2:
+ forall e1 e2, beq_expression_list e1 e2 = false -> e1 <> e2.
+Proof.
+ induction e1; crush.
+ - destruct_match; crush.
+ - destruct_match; crush.
+ apply andb_false_iff in H. inv H. apply beq_expression_correct2 in H0.
+ unfold not in *; intros. apply H0. inv H. auto.
+ apply IHe1 in H0. unfold not in *; intros. apply H0. inv H.
+ auto.
+Qed.
+
+Lemma beq_pred_expression_correct:
+ forall e1 e2, beq_pred_expression e1 e2 = true -> e1 = e2.
+Proof.
+ destruct e1, e2; crush.
+ - destruct_match; crush.
+ - destruct_match; subst; crush.
+ apply beq_expression_list_correct in H; subst.
+ trivial.
+Qed.
+
+Lemma beq_pred_expression_refl:
+ forall e, beq_pred_expression e e = true.
+Proof.
+ destruct e; crush; destruct_match; crush.
+ apply beq_expression_list_refl.
+Qed.
+
+Lemma beq_pred_expression_correct2:
+ forall e1 e2, beq_pred_expression e1 e2 = false -> e1 <> e2.
+Proof.
+ destruct e1, e2; unfold not; crush.
+ + destruct_match; crush.
+ + destruct_match; crush. inv H0.
+ now rewrite beq_expression_list_refl in H.
+Qed.
+
+Lemma beq_exit_expression_correct:
+ forall e1 e2, beq_exit_expression e1 e2 = true <-> e1 = e2.
+Proof.
+ destruct e1, e2; split; crush;
+ destruct_match; subst; crush.
+Qed.
+
+Definition pred_expr_item_eq (pe1 pe2: pred_op * expression) : bool :=
+ @equiv_dec _ SATSetoid _ (fst pe1) (fst pe2) && beq_expression (snd pe1) (snd pe2).
+
+Definition pred_eexpr_item_eq (pe1 pe2: pred_op * exit_expression) : bool :=
+ @equiv_dec _ SATSetoid _ (fst pe1) (fst pe2) && beq_exit_expression (snd pe1) (snd pe2).
+
+Definition pred_expr_dec: forall (pe1 pe2: pred_op * expression),
+ {pred_expr_item_eq pe1 pe2 = true} + {pred_expr_item_eq pe1 pe2 = false}.
+Proof.
+ intros; destruct (pred_expr_item_eq pe1 pe2) eqn:?; unfold not; [tauto | now right].
+Defined.
+
+Definition pred_expr_dec2: forall (pe1 pe2: pred_op * expression),
+ {pred_expr_item_eq pe1 pe2 = true} + {~ pred_expr_item_eq pe1 pe2 = true}.
+Proof.
+ intros; destruct (pred_expr_item_eq pe1 pe2) eqn:?; unfold not; [tauto | now right].
+Defined.
+
+Definition pred_expression_dec:
+ forall e1 e2: pred_expression, {e1 = e2} + {e1 <> e2}.
+Proof.
+ intros. destruct (beq_pred_expression e1 e2) eqn:?.
+ eauto using beq_pred_expression_correct.
+ eauto using beq_pred_expression_correct2.
+Defined.
+
+Lemma exit_expression_refl:
+ forall e, beq_exit_expression e e = true.
+Proof. destruct e; crush; destruct_match; crush. Qed.
+
+Definition exit_expression_dec:
+ forall e1 e2: exit_expression, {e1 = e2} + {e1 <> e2}.
+Proof.
+ intros. destruct (beq_exit_expression e1 e2) eqn:?.
+ - left. eapply beq_exit_expression_correct; eauto.
+ - right. unfold not; intros.
+ assert (X: ~ (beq_exit_expression e1 e2 = true))
+ by eauto with bool.
+ subst. apply X. apply exit_expression_refl.
+Defined.
+
+Lemma pred_eexpression_dec:
+ forall (e1 e2: exit_expression) (p1 p2: pred_op),
+ {(p1, e1) = (p2, e2)} + {(p1, e1) <> (p2, e2)}.
+Proof.
+ pose proof (Predicate.eq_dec peq).
+ pose proof (exit_expression_dec).
+ decide equality.
+Defined.
+
+(*Fixpoint encode_expression_ne (max: predicate) (pe: pred_expr_ne) (h: hash_tree)
+ : (PTree.t pred_op) * hash_tree :=
+ match pe with
+ | NE.singleton (p, e) =>
+ let (p', h') := hash_value max e h in
+ (Por (Pnot p) (Pvar p'), h')
+ | (p, e) ::| pr =>
+ let (p', h') := hash_value max e h in
+ let (p'', h'') := encode_expression_ne max pr h' in
+ (Pand (Por (Pnot p) (Pvar p')) p'', h'')
+ end.*)
+
+Fixpoint max_pred_expr (pe: pred_expr) : positive :=
+ match pe with
+ | NE.singleton (p, e) => max_predicate p
+ | (p, e) ::| pe' => Pos.max (max_predicate p) (max_pred_expr pe')
+ end.
+
+Definition ge_preserved {A B C D: Type} (ge: Genv.t A B) (tge: Genv.t C D) : Prop :=
+ (forall sp op vl m, Op.eval_operation ge sp op vl m =
+ Op.eval_operation tge sp op vl m)
+ /\ (forall sp addr vl, Op.eval_addressing ge sp addr vl =
+ Op.eval_addressing tge sp addr vl).
+
+Lemma ge_preserved_same:
+ forall A B ge, @ge_preserved A B A B ge ge.
+Proof. unfold ge_preserved; auto. Qed.
+#[local] Hint Resolve ge_preserved_same : core.
+
+Inductive match_states : instr_state -> instr_state -> Prop :=
+| match_states_intro:
+ forall ps ps' rs rs' m m',
+ (forall x, rs !! x = rs' !! x) ->
+ (forall x, ps !! x = ps' !! x) ->
+ m = m' ->
+ match_states (mk_instr_state rs ps m) (mk_instr_state rs' ps' m').
+
+Lemma match_states_refl x : match_states x x.
+Proof. destruct x; constructor; crush. Qed.
+
+Lemma match_states_commut x y : match_states x y -> match_states y x.
+Proof. inversion 1; constructor; crush. Qed.
+
+Lemma match_states_trans x y z :
+ match_states x y -> match_states y z -> match_states x z.
+Proof. repeat inversion 1; constructor; crush. Qed.
+
+#[global] Instance match_states_Equivalence : Equivalence match_states :=
+ { Equivalence_Reflexive := match_states_refl ;
+ Equivalence_Symmetric := match_states_commut ;
+ Equivalence_Transitive := match_states_trans ; }.
+
+Inductive similar {A B} : @ctx A -> @ctx B -> Prop :=
+| similar_intro :
+ forall ist ist' sp ge tge,
+ ge_preserved ge tge ->
+ match_states ist ist' ->
+ similar (mk_ctx ist sp ge) (mk_ctx ist' sp tge).
+
+Lemma ge_preserved_refl:
+ forall A B (a: Genv.t A B), ge_preserved a a.
+Proof. auto. Qed.
+
+Lemma similar_refl:
+ forall A (a: @Abstr.ctx A), similar a a.
+Proof. intros; destruct a; constructor; auto. reflexivity. Qed.
+
+Lemma similar_commut:
+ forall A B (a: @Abstr.ctx A) (b: @Abstr.ctx B), similar a b -> similar b a.
+Proof.
+ inversion 1; constructor; auto.
+ - unfold ge_preserved in *; inv H0; split; intros.
+ rewrite H4; auto. rewrite H5; auto.
+ - symmetry; auto.
+Qed.
+
+Lemma similar_trans:
+ forall A B C (a: @Abstr.ctx A) (b: @Abstr.ctx B) (c: @Abstr.ctx C),
+ similar a b -> similar b c -> similar a c.
+Proof.
+ repeat inversion 1; constructor.
+ - unfold ge_preserved in *; inv H0; inv H9; split; intros.
+ rewrite H11. rewrite H0; auto.
+ rewrite H12. rewrite H2. auto.
+ - transitivity ist'; auto.
+Qed.
+
+Module HashExpr' <: MiniDecidableType.
+ Definition t := expression.
+ Definition eq_dec := expression_dec.
+End HashExpr'.
+
+Module HashExpr := Make_UDT(HashExpr').
+Module Import HT := HashTree(HashExpr).
+
+Module PHashExpr' <: MiniDecidableType.
+ Definition t := pred_expression.
+ Definition eq_dec := pred_expression_dec.
+End PHashExpr'.
+
+Module PHashExpr := Make_UDT(PHashExpr').
+Module PHT := HashTree(PHashExpr).
+
+Module EHashExpr' <: MiniDecidableType.
+ Definition t := exit_expression.
+ Definition eq_dec := exit_expression_dec.
+End EHashExpr'.
+
+Module EHashExpr := Make_UDT(EHashExpr').
+Module EHT := HashTree(EHashExpr).
+
+Fixpoint hash_predicate (max: predicate) (ap: pred_pexpr) (h: PHT.hash_tree)
+ : pred_op * PHT.hash_tree :=
+ match ap with
+ | Ptrue => (Ptrue, h)
+ | Pfalse => (Pfalse, h)
+ | Plit (b, ap') =>
+ let (p', h') := PHT.hash_value max ap' h in
+ (Plit (b, p'), h')
+ | Pand p1 p2 =>
+ let (p1', h') := hash_predicate max p1 h in
+ let (p2', h'') := hash_predicate max p2 h' in
+ (Pand p1' p2', h'')
+ | Por p1 p2 =>
+ let (p1', h') := hash_predicate max p1 h in
+ let (p2', h'') := hash_predicate max p2 h' in
+ (Por p1' p2', h'')
+ end.
+
+Definition predicated_mutexcl {A: Type} (pe: predicated A): Prop :=
+ (forall x y, NE.In x pe -> NE.In y pe -> x <> y -> fst x ⇒ ¬ fst y)
+ /\ NE.norepet pe.
+
+Lemma predicated_cons :
+ forall A (a: pred_op * A) x,
+ predicated_mutexcl (a ::| x) ->
+ predicated_mutexcl x.
+Proof.
+ unfold predicated_mutexcl; intros. inv H. inv H1. split; auto.
+ intros. apply H0; auto; constructor; tauto.
+Qed.
+
+Lemma predicated_singleton :
+ forall A (a: (pred_op * A)), predicated_mutexcl (NE.singleton a).
+Proof.
+ unfold predicated_mutexcl; intros; split; intros.
+ { inv H. now inv H0. }
+ constructor.
+Qed.
+
+(*
+
+Lemma norm_expr_constant :
+ forall x m h t h' e p,
+ HN.norm_expression m x h = (t, h') ->
+ h ! e = Some p ->
+ h' ! e = Some p.
+Proof. Abort.
+
+Definition sat_aequiv ap1 ap2 :=
+ exists h p1 p2,
+ sat_equiv p1 p2
+ /\ hash_predicate 1 ap1 h = (p1, h)
+ /\ hash_predicate 1 ap2 h = (p2, h).
+
+Lemma aequiv_symm : forall a b, sat_aequiv a b -> sat_aequiv b a.
+Proof.
+ unfold sat_aequiv; simplify; do 3 eexists; simplify; [symmetry | |]; eauto.
+Qed.
+
+Lemma existsh :
+ forall ap1,
+ exists h p1,
+ hash_predicate 1 ap1 h = (p1, h).
+Proof. Admitted.
+
+Lemma aequiv_refl : forall a, sat_aequiv a a.
+Proof.
+ unfold sat_aequiv; intros.
+ pose proof (existsh a); simplify.
+ do 3 eexists; simplify; eauto. reflexivity.
+Qed.
+
+Lemma aequiv_trans :
+ forall a b c,
+ sat_aequiv a b ->
+ sat_aequiv b c ->
+ sat_aequiv a c.
+Proof.
+ unfold sat_aequiv; intros.
+ simplify.
+Abort.
+
+Lemma norm_expr_mutexcl :
+ forall m pe h t h' e e' p p',
+ HN.norm_expression m pe h = (t, h') ->
+ predicated_mutexcl pe ->
+ t ! e = Some p ->
+ t ! e' = Some p' ->
+ e <> e' ->
+ p ⇒ ¬ p'.
+Proof. Abort.*)
+
+Definition pred_expr_eqb: forall pe1 pe2: pred_expr,
+ {pe1 = pe2} + {pe1 <> pe2}.
+Proof.
+ pose proof expression_dec.
+ pose proof NE.eq_dec.
+ pose proof (Predicate.eq_dec peq).
+ assert (forall a b: pred_op * expression, {a = b} + {a <> b})
+ by decide equality.
+ decide equality.
+Defined.
+
+Definition pred_pexpr_eqb: forall pe1 pe2: pred_pexpr,
+ {pe1 = pe2} + {pe1 <> pe2}.
+Proof.
+ pose proof pred_expression_dec.
+ pose proof (Predicate.eq_dec pred_expression_dec).
+ apply X.
+Defined.
+
+Definition beq_pred_pexpr (pe1 pe2: pred_pexpr): bool :=
+ if pred_pexpr_eqb pe1 pe2 then true else
+ let (np1, h) := hash_predicate 1 pe1 (PTree.empty _) in
+ let (np2, h') := hash_predicate 1 pe2 h in
+ equiv_check np1 np2.
+
+Lemma inj_asgn_eg : forall a b, (a =? b)%positive = (a =? a)%positive -> a = b.
+Proof.
+ intros. destruct (peq a b); subst.
+ auto. rewrite OrdersEx.Positive_as_OT.eqb_refl in H.
+ now apply Peqb_true_eq.
+Qed.
+
+Lemma inj_asgn :
+ forall a b, (forall (f: positive -> bool), f a = f b) -> a = b.
+Proof. intros. apply inj_asgn_eg. eauto. Qed.
+
+Lemma inj_asgn_false:
+ forall n1 n2 , ~ (forall c: positive -> bool, c n1 = negb (c n2)).
+Proof.
+ unfold not; intros. specialize (H (fun x => true)).
+ simplify. discriminate.
+Qed.
+
+Lemma negb_inj:
+ forall a b,
+ negb a = negb b -> a = b.
+Proof. destruct a, b; crush. Qed.
+
+Lemma sat_predicate_Plit_inj :
+ forall p1 p2,
+ Plit p1 == Plit p2 -> p1 = p2.
+Proof.
+ simplify. destruct p1, p2.
+ destruct b, b0. f_equal. unfold sat_equiv in H. cbn in H. now apply inj_asgn.
+ solve [exfalso; eapply inj_asgn_false; eauto].
+ solve [exfalso; eapply inj_asgn_false; eauto].
+ assert (p = p0). eapply inj_asgn. intros. specialize (H f).
+ apply negb_inj in H. auto. rewrite H0; auto.
+Qed.
+
+Definition ind_preds t :=
+ forall e1 e2 p1 p2 c,
+ e1 <> e2 ->
+ t ! e1 = Some p1 ->
+ t ! e2 = Some p2 ->
+ sat_predicate p1 c = true ->
+ sat_predicate p2 c = false.
+
+Definition ind_preds_l t :=
+ forall (e1: predicate) e2 p1 p2 c,
+ e1 <> e2 ->
+ In (e1, p1) t ->
+ In (e2, p2) t ->
+ sat_predicate p1 c = true ->
+ sat_predicate p2 c = false.
+
+(*Lemma pred_equivalence_Some' :
+ forall ta tb e pe pe',
+ list_norepet (map fst ta) ->
+ list_norepet (map fst tb) ->
+ In (e, pe) ta ->
+ In (e, pe') tb ->
+ fold_right (fun p a => mk_pred_stmnt' (fst p) (snd p) ∧ a) T ta ==
+ fold_right (fun p a => mk_pred_stmnt' (fst p) (snd p) ∧ a) T tb ->
+ pe == pe'.
+Proof.
+ induction ta as [|hd tl Hta]; try solve [crush].
+ - intros * NRP1 NRP2 IN1 IN2 FOLD. inv NRP1. inv IN1.
+ simpl in FOLD.
+
+Lemma pred_equivalence_Some :
+ forall (ta tb: PTree.t pred_op) e pe pe',
+ ta ! e = Some pe ->
+ tb ! e = Some pe' ->
+ mk_pred_stmnt ta == mk_pred_stmnt tb ->
+ pe == pe'.
+Proof.
+ intros * SMEA SMEB EQ. unfold mk_pred_stmnt in *.
+ repeat rewrite PTree.fold_spec in EQ.
+
+Lemma pred_equivalence_None :
+ forall (ta tb: PTree.t pred_op) e pe,
+ (mk_pred_stmnt ta) == (mk_pred_stmnt tb) ->
+ ta ! e = Some pe ->
+ tb ! e = None ->
+ equiv pe ⟂.
+Abort.
+
+Lemma pred_equivalence :
+ forall (ta tb: PTree.t pred_op) e pe,
+ equiv (mk_pred_stmnt ta) (mk_pred_stmnt tb) ->
+ ta ! e = Some pe ->
+ equiv pe ⟂ \/ (exists pe', tb ! e = Some pe' /\ equiv pe pe').
+Proof.
+ intros * EQ SME. destruct (tb ! e) eqn:HTB.
+ { right. econstructor. split; eauto. eapply pred_equivalence_Some; eauto. }
+ { left. eapply pred_equivalence_None; eauto. }
+Qed.*)
+
+Section CORRECT.
+
+ Context {FUN TFUN: Type}.
+
+ Context (ictx: @ctx FUN) (octx: @ctx TFUN) (HSIM: similar ictx octx).
+
+ Lemma sem_value_mem_det:
+ forall e v v' m m',
+ (sem_value ictx e v -> sem_value octx e v' -> v = v')
+ /\ (sem_mem ictx e m -> sem_mem octx e m' -> m = m').
+ Proof using HSIM.
+ induction e using expression_ind2
+ with (P0 := fun p => forall v v',
+ sem_val_list ictx p v -> sem_val_list octx p v' -> v = v');
+ inv HSIM; match goal with H: context [match_states] |- _ => inv H end; repeat progress simplify;
+ try solve [match goal with
+ | H: sem_value _ _ _, H2: sem_value _ _ _ |- _ => inv H; inv H2; auto
+ | H: sem_mem _ _ _, H2: sem_mem _ _ _ |- _ => inv H; inv H2; auto
+ | H: sem_val_list _ _ _, H2: sem_val_list _ _ _ |- _ => inv H; inv H2; auto
+ end].
+ - repeat match goal with
+ | H: sem_value _ _ _ |- _ => inv H
+ | H: sem_val_list {| ctx_ge := ge; |} ?e ?l1,
+ H2: sem_val_list {| ctx_ge := tge |} ?e ?l2,
+ IH: forall _ _, sem_val_list _ _ _ -> sem_val_list _ _ _ -> _ = _ |- _ =>
+ assert (X: l1 = l2) by (apply IH; auto)
+ | H: ge_preserved _ _ |- _ => inv H
+ | |- context [ctx_rs] => unfold ctx_rs; cbn
+ | H: context [ctx_mem] |- _ => unfold ctx_mem in H; cbn
+ end; crush.
+ - repeat match goal with H: sem_value _ _ _ |- _ => inv H end; simplify;
+ assert (lv0 = lv) by (apply IHe; eauto); subst;
+ match goal with H: ge_preserved _ _ |- _ => inv H end;
+ match goal with H: context [Op.eval_addressing _ _ _ _ = Op.eval_addressing _ _ _ _] |- _
+ => rewrite H in * end;
+ assert (a0 = a1) by crush;
+ assert (m'2 = m'1) by (apply IHe0; eauto); crush.
+ - inv H0; inv H3. simplify.
+ assert (lv = lv0) by ( apply IHe2; eauto). subst.
+ assert (a1 = a0). { inv H. rewrite H3 in *. crush. }
+ assert (v0 = v1). { apply IHe1; auto. }
+ assert (m'1 = m'2). { apply IHe3; auto. } crush.
+ - inv H0. inv H3. f_equal. apply IHe; auto.
+ apply IHe0; auto.
+ Qed.
+
+ Lemma sem_value_mem_corr:
+ forall e v m,
+ (sem_value ictx e v -> sem_value octx e v)
+ /\ (sem_mem ictx e m -> sem_mem octx e m).
+ Proof using HSIM.
+ induction e using expression_ind2
+ with (P0 := fun p => forall v,
+ sem_val_list ictx p v -> sem_val_list octx p v);
+ inv HSIM; match goal with H: context [match_states] |- _ => inv H end; repeat progress simplify.
+ - inv H0. unfold ctx_rs, ctx_ps, ctx_mem in *; cbn. rewrite H1. constructor.
+ - inv H0. unfold ctx_rs, ctx_ps, ctx_mem in *; cbn. constructor.
+ - inv H0. apply IHe in H6. econstructor; try eassumption.
+ unfold ctx_rs, ctx_ps, ctx_mem in *; cbn in *. inv H. crush.
+ - inv H0.
+ - inv H0. eapply IHe in H10. eapply IHe0 in H8; auto.
+ econstructor; try eassumption.
+ unfold ctx_rs, ctx_ps, ctx_mem in *; cbn in *. inv H; crush.
+ - inv H0.
+ - inv H0.
+ - inv H0. eapply IHe1 in H11; auto. eapply IHe2 in H12; auto. eapply IHe3 in H9; auto.
+ econstructor; try eassumption.
+ unfold ctx_rs, ctx_ps, ctx_mem in *; cbn in *. inv H; crush.
+ - inv H0. econstructor.
+ - inv H0. eapply IHe in H6; auto. eapply IHe0 in H8.
+ econstructor; eassumption.
+ Qed.
+
+ Lemma sem_value_det:
+ forall e v v', sem_value ictx e v -> sem_value octx e v' -> v = v'.
+ Proof using HSIM.
+ intros. eapply sem_value_mem_det; eauto; apply Mem.empty.
+ Qed.
+
+ Lemma sem_value_corr:
+ forall e v, sem_value ictx e v -> sem_value octx e v.
+ Proof using HSIM.
+ intros. eapply sem_value_mem_corr; eauto; apply Mem.empty.
+ Qed.
+
+ Lemma sem_mem_det:
+ forall e v v', sem_mem ictx e v -> sem_mem octx e v' -> v = v'.
+ Proof using HSIM.
+ intros. eapply sem_value_mem_det; eauto; apply (Vint (Int.repr 0%Z)).
+ Qed.
+
+ Lemma sem_mem_corr:
+ forall e v, sem_mem ictx e v -> sem_mem octx e v.
+ Proof using HSIM.
+ intros. eapply sem_value_mem_corr; eauto; apply (Vint (Int.repr 0%Z)).
+ Qed.
+
+ Lemma sem_val_list_det:
+ forall e l l', sem_val_list ictx e l -> sem_val_list octx e l' -> l = l'.
+ Proof using HSIM.
+ induction e; simplify.
+ - inv H; inv H0; auto.
+ - inv H; inv H0. f_equal. eapply sem_value_det; eauto; try apply Mem.empty.
+ apply IHe; eauto.
+ Qed.
+
+ Lemma sem_val_list_corr:
+ forall e l, sem_val_list ictx e l -> sem_val_list octx e l.
+ Proof using HSIM.
+ induction e; simplify.
+ - inv H; constructor.
+ - inv H. apply sem_value_corr in H3; auto; try apply Mem.empty;
+ apply IHe in H5; constructor; assumption.
+ Qed.
+
+ Lemma sem_pred_det:
+ forall e v v', sem_pred ictx e v -> sem_pred octx e v' -> v = v'.
+ Proof using HSIM.
+ try solve [inversion 1]; pose proof sem_value_det; pose proof sem_val_list_det; inv HSIM;
+ match goal with H: match_states _ _ |- _ => inv H end; simplify.
+ inv H2; inv H5; auto. assert (lv = lv0) by (eapply H0; eauto). subst. unfold ctx_mem in *.
+ crush.
+ Qed.
+
+ Lemma sem_pred_corr:
+ forall e v, sem_pred ictx e v -> sem_pred octx e v.
+ Proof using HSIM.
+ try solve [inversion 1]; pose proof sem_value_corr; pose proof sem_val_list_corr; inv HSIM;
+ match goal with H: match_states _ _ |- _ => inv H end; simplify.
+ inv H2; auto. apply H0 in H5. econstructor; eauto.
+ unfold ctx_ps; cbn. rewrite H4. constructor.
+ Qed.
+
+ Lemma sem_exit_det:
+ forall e v v', sem_exit ictx e v -> sem_exit octx e v' -> v = v'.
+ Proof using HSIM.
+ try solve [inversion 1]; pose proof sem_value_det; pose proof sem_val_list_det; inv HSIM;
+ match goal with H: match_states _ _ |- _ => inv H end; simplify.
+ inv H2; inv H5; auto.
+ Qed.
+
+ Lemma sem_exit_corr:
+ forall e v, sem_exit ictx e v -> sem_exit octx e v.
+ Proof using HSIM.
+ try solve [inversion 1]; pose proof sem_value_corr; pose proof sem_val_list_corr; inv HSIM;
+ match goal with H: match_states _ _ |- _ => inv H end; simplify.
+ inv H2; auto; constructor.
+ Qed.
+
+ Lemma sem_pexpr_det :
+ forall p b1 b2, sem_pexpr ictx p b1 -> sem_pexpr octx p b2 -> b1 = b2.
+ Proof.
+ induction p; crush; inv H; inv H0; firstorder.
+ destruct b.
+ - apply sem_pred_det with (e:=p0); auto.
+ - apply negb_inj. apply sem_pred_det with (e:=p0); auto.
+ Qed.
+
+ Lemma sem_pexpr_corr :
+ forall p b, sem_pexpr ictx p b -> sem_pexpr octx p b.
+ Proof.
+ induction p; crush; inv H; constructor;
+ try solve [try inv H3; firstorder].
+ now apply sem_pred_corr.
+ Qed.
+
+ Lemma sem_pred_exec_beq_correct2 :
+ forall A B (sem: forall G, @Abstr.ctx G -> A -> B -> Prop) a p r R,
+ (forall x y,
+ sem _ ictx x y ->
+ exists y', sem _ octx x y' /\ R y y') ->
+ sem_pred_expr a (sem _) ictx p r ->
+ exists r', sem_pred_expr a (sem _) octx p r' /\ R r r'.
+ Proof.
+ induction p; crush.
+ - inv H0. apply H in H4. simplify.
+ exists x; split; auto.
+ constructor; auto.
+ now apply sem_pexpr_corr.
+ - inv H0.
+ + apply H in H6; simplify.
+ exists x; split; auto.
+ constructor; auto.
+ now apply sem_pexpr_corr.
+ + exploit IHp; auto. exact H6. intros. simplify.
+ exists x; split; auto.
+ apply sem_pred_expr_cons_false; auto.
+ now apply sem_pexpr_corr.
+ Qed.
+
+ Lemma sem_pred_expr_corr :
+ forall A B (sem: forall G, @Abstr.ctx G -> A -> B -> Prop) a p r,
+ (forall x y, sem _ ictx x y -> sem _ octx x y) ->
+ sem_pred_expr a (sem _) ictx p r ->
+ sem_pred_expr a (sem _) octx p r.
+ Proof.
+ intros.
+ assert
+ (forall x y,
+ sem _ ictx x y ->
+ exists y', sem _ octx x y' /\ eq y y') by firstorder.
+ pose proof (sem_pred_exec_beq_correct2 _ _ sem a p r _ H1 H0).
+ crush.
+ Qed.
+
+ Lemma sem_correct:
+ forall f i cf, sem ictx f (i, cf) -> sem octx f (i, cf).
+ Proof.
+ intros. inv H. constructor.
+ - inv H2. constructor; intros. specialize (H x).
+ apply sem_pred_expr_corr; auto. exact sem_value_corr.
+ - inv H3; constructor; intros. specialize (H x).
+ now apply sem_pexpr_corr.
+ - apply sem_pred_expr_corr; auto. exact sem_mem_corr.
+ - apply sem_pred_expr_corr; auto. exact sem_exit_corr.
+ Qed.
+
+End CORRECT.
+
+Section SEM_PRED_EXEC.
+
+ Context (A: Type).
+ Context (ctx: @Abstr.ctx A).
+
+ Lemma sem_pexpr_negate :
+ forall p b,
+ sem_pexpr ctx p b ->
+ sem_pexpr ctx (¬ p) (negb b).
+ Proof.
+ induction p; crush.
+ - destruct_match. destruct b0; crush. inv Heqp0.
+ constructor. inv H. rewrite negb_involutive. auto.
+ constructor. inv H. auto.
+ - inv H. constructor.
+ - inv H. constructor.
+ - inv H. inv H3.
+ + apply IHp1 in H. solve [constructor; auto].
+ + apply IHp2 in H. solve [constructor; auto].
+ + apply IHp1 in H2. apply IHp2 in H4. solve [constructor; auto].
+ - inv H. inv H3.
+ + apply IHp1 in H. solve [constructor; auto].
+ + apply IHp2 in H. solve [constructor; auto].
+ + apply IHp1 in H2. apply IHp2 in H4. solve [constructor; auto].
+ Qed.
+
+ Lemma sem_pexpr_negate2 :
+ forall p b,
+ sem_pexpr ctx (¬ p) (negb b) ->
+ sem_pexpr ctx p b.
+ Proof.
+ induction p; crush.
+ - destruct_match. destruct b0; crush. inv Heqp0.
+ constructor. inv H. rewrite negb_involutive in *. auto.
+ constructor. inv H. auto.
+ - inv H. destruct b; try discriminate. constructor.
+ - inv H. destruct b; try discriminate. constructor.
+ - inv H. destruct b; try discriminate.
+ + constructor. inv H1; eauto.
+ + destruct b; try discriminate. constructor; eauto.
+ - inv H. destruct b; try discriminate.
+ + constructor. inv H1; eauto.
+ + destruct b; try discriminate. constructor; eauto.
+ Qed.
+
+ Lemma sem_pexpr_evaluable :
+ forall f_p ps,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p, exists b, sem_pexpr ctx (from_pred_op f_p p) b.
+ Proof.
+ induction p; crush.
+ - destruct_match. inv Heqp0. destruct b. econstructor. eauto.
+ econstructor. eapply sem_pexpr_negate. eauto.
+ - econstructor. constructor.
+ - econstructor. constructor.
+ - destruct x0, x; solve [eexists; constructor; auto].
+ - destruct x0, x; solve [eexists; constructor; auto].
+ Qed.
+
+ Lemma sem_pexpr_eval1 :
+ forall f_p ps,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p,
+ eval_predf ps p = false ->
+ sem_pexpr ctx (from_pred_op f_p p) false.
+ Proof.
+ induction p; crush.
+ - destruct_match. inv Heqp0.
+ destruct b.
+ + cbn in H0. rewrite <- H0. eauto.
+ + replace false with (negb true) by auto.
+ apply sem_pexpr_negate. cbn in H0.
+ apply negb_true_iff in H0. rewrite negb_involutive in H0.
+ rewrite <- H0. eauto.
+ - constructor.
+ - rewrite eval_predf_Pand in H0.
+ apply andb_false_iff in H0. inv H0. eapply IHp1 in H1.
+ pose proof (sem_pexpr_evaluable _ _ H p2) as EVAL.
+ inversion_clear EVAL as [x EVAL2].
+ replace false with (false && x) by auto.
+ constructor; auto.
+ eapply IHp2 in H1.
+ pose proof (sem_pexpr_evaluable _ _ H p1) as EVAL.
+ inversion_clear EVAL as [x EVAL2].
+ replace false with (x && false) by eauto with bool.
+ apply sem_pexpr_Pand; auto.
+ - rewrite eval_predf_Por in H0.
+ apply orb_false_iff in H0. inv H0.
+ replace false with (false || false) by auto.
+ apply sem_pexpr_Por; auto.
+ Qed.
+
+ Lemma sem_pexpr_eval2 :
+ forall f_p ps,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p,
+ eval_predf ps p = true ->
+ sem_pexpr ctx (from_pred_op f_p p) true.
+ Proof.
+ induction p; crush.
+ - destruct_match. inv Heqp0.
+ destruct b.
+ + cbn in H0. rewrite <- H0. eauto.
+ + replace true with (negb false) by auto.
+ apply sem_pexpr_negate. cbn in H0.
+ apply negb_true_iff in H0.
+ rewrite <- H0. eauto.
+ - constructor.
+ - rewrite eval_predf_Pand in H0.
+ apply andb_true_iff in H0. inv H0.
+ replace true with (true && true) by auto.
+ constructor; auto.
+ - rewrite eval_predf_Por in H0.
+ apply orb_true_iff in H0. inv H0. eapply IHp1 in H1.
+ pose proof (sem_pexpr_evaluable _ _ H p2) as EVAL.
+ inversion_clear EVAL as [x EVAL2].
+ replace true with (true || x) by auto.
+ apply sem_pexpr_Por; auto.
+ eapply IHp2 in H1.
+ pose proof (sem_pexpr_evaluable _ _ H p1) as EVAL.
+ inversion_clear EVAL as [x EVAL2].
+ replace true with (x || true) by eauto with bool.
+ apply sem_pexpr_Por; auto.
+ Qed.
+
+ Lemma sem_pexpr_eval :
+ forall f_p ps b,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p,
+ eval_predf ps p = b ->
+ sem_pexpr ctx (from_pred_op f_p p) b.
+ Proof.
+ intros; destruct b; eauto using sem_pexpr_eval1, sem_pexpr_eval2.
+ Qed.
+
+ Lemma sem_pexpr_eval_inv :
+ forall f_p ps b,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p,
+ sem_pexpr ctx (from_pred_op f_p p) b ->
+ eval_predf ps p = b.
+ Proof.
+ induction p; intros.
+ - cbn in H0. destruct_match. destruct b0; cbn in *.
+ + specialize (H p0). eapply sem_pexpr_det; eauto. apply similar_refl.
+ + rewrite <- negb_involutive in H0. apply sem_pexpr_negate2 in H0.
+ symmetry; apply negb_sym. eapply sem_pexpr_det; eauto.
+ apply similar_refl.
+ - now inv H0.
+ - now inv H0.
+ - inv H0; try inv H4; rewrite eval_predf_Pand.
+ + apply IHp1 in H0. rewrite H0. auto.
+ + apply IHp2 in H0. rewrite H0. auto with bool.
+ + apply IHp2 in H5. apply IHp1 in H3. rewrite H3. rewrite H5. auto.
+ - inv H0; try inv H4; rewrite eval_predf_Por.
+ + apply IHp1 in H0. rewrite H0. auto.
+ + apply IHp2 in H0. rewrite H0. auto with bool.
+ + apply IHp2 in H5. apply IHp1 in H3. rewrite H3. rewrite H5. auto.
+ Qed.
+
+ Context {C B: Type}.
+ Context (f: PTree.t pred_pexpr).
+ Context (ps: PMap.t bool).
+ Context (a_sem: @Abstr.ctx A -> C -> B -> Prop).
+
+ Context (F_EVALULABLE: forall x, sem_pexpr ctx (get_forest_p' x f) ps !! x).
+
+ Lemma sem_pexpr_equiv :
+ forall p1 p2 b,
+ p1 == p2 ->
+ sem_pexpr ctx (from_pred_op f p1) b ->
+ sem_pexpr ctx (from_pred_op f p2) b.
+ Proof.
+ intros.
+ eapply sem_pexpr_eval_inv in H0; eauto.
+ eapply sem_pexpr_eval; eauto.
+ Qed.
+
+End SEM_PRED_EXEC.
+
+Module HashExprNorm(HS: Hashable).
+ Module H := HashTree(HS).
+
+ Definition norm_tree: Type := PTree.t pred_op * H.hash_tree.
+
+ Fixpoint norm_expression (max: predicate) (pe: predicated HS.t) (h: H.hash_tree)
+ : norm_tree :=
+ match pe with
+ | NE.singleton (p, e) =>
+ let (hashed_e, h') := H.hash_value max e h in
+ (PTree.set hashed_e p (PTree.empty _), h')
+ | (p, e) ::| pr =>
+ let (hashed_e, h') := H.hash_value max e h in
+ let (norm_pr, h'') := norm_expression max pr h' in
+ match norm_pr ! hashed_e with
+ | Some pr_op =>
+ (PTree.set hashed_e (pr_op ∨ p) norm_pr, h'')
+ | None =>
+ (PTree.set hashed_e p norm_pr, h'')
+ end
+ end.
+
+ Definition mk_pred_stmnt' (e: predicate) p_e := ¬ p_e ∨ Plit (true, e).
+
+ Definition mk_pred_stmnt t := PTree.fold (fun x a b => mk_pred_stmnt' a b ∧ x) t T.
+
+ Definition mk_pred_stmnt_l (t: list (predicate * pred_op)) :=
+ fold_left (fun x a => uncurry mk_pred_stmnt' a ∧ x) t T.
+
+ Definition encode_expression max pe h :=
+ let (tree, h) := norm_expression max pe h in
+ (mk_pred_stmnt tree, h).
+
+ Definition pred_expr_dec: forall pe1 pe2: predicated HS.t,
+ {pe1 = pe2} + {pe1 <> pe2}.
+ Proof.
+ pose proof HS.eq_dec as X.
+ pose proof (Predicate.eq_dec peq).
+ pose proof (NE.eq_dec _ X).
+ assert (forall a b: pred_op * HS.t, {a = b} + {a <> b})
+ by decide equality.
+ decide equality.
+ Defined.
+
+ Definition beq_pred_expr' (pe1 pe2: predicated HS.t) : bool :=
+ if pred_expr_dec pe1 pe2 then true else
+ let (p1, h) := encode_expression 1%positive pe1 (PTree.empty _) in
+ let (p2, h') := encode_expression 1%positive pe2 h in
+ equiv_check p1 p2.
+
+ Lemma mk_pred_stmnt_equiv' :
+ forall l l' e p,
+ mk_pred_stmnt_l l == mk_pred_stmnt_l l' ->
+ In (e, p) l ->
+ list_norepet (map fst l) ->
+ (exists p', In (e, p') l' /\ p == p')
+ \/ ~ In e (map fst l') /\ p == ⟂.
+ Proof. Abort.
+
+ Lemma mk_pred_stmnt_equiv :
+ forall tree tree',
+ mk_pred_stmnt tree == mk_pred_stmnt tree'.
+ Proof. Abort.
+
+ Definition tree_equiv_check_el (np2: PTree.t pred_op) (n: positive) (p: pred_op): bool :=
+ match np2 ! n with
+ | Some p' => equiv_check p p'
+ | None => equiv_check p ⟂
+ end.
+
+ Definition tree_equiv_check_None_el (np2: PTree.t pred_op) (n: positive) (p: pred_op): bool :=
+ match np2 ! n with
+ | Some p' => true
+ | None => equiv_check p ⟂
+ end.
+
+ Definition beq_pred_expr (pe1 pe2: predicated HS.t) : bool :=
+ if pred_expr_dec pe1 pe2 then true else
+ let (np1, h) := norm_expression 1 pe1 (PTree.empty _) in
+ let (np2, h') := norm_expression 1 pe2 h in
+ forall_ptree (tree_equiv_check_el np2) np1
+ && forall_ptree (tree_equiv_check_None_el np1) np2.
+
+ Lemma beq_pred_expr_correct:
+ forall np1 np2 e p p',
+ forall_ptree (tree_equiv_check_el np2) np1 = true ->
+ np1 ! e = Some p ->
+ np2 ! e = Some p' ->
+ p == p'.
+ Proof.
+ intros.
+ eapply forall_ptree_true in H; try eassumption.
+ unfold tree_equiv_check_el in H.
+ rewrite H1 in H. now apply equiv_check_correct.
+ Qed.
+
+ Lemma beq_pred_expr_correct2:
+ forall np1 np2 e p,
+ forall_ptree (tree_equiv_check_el np2) np1 = true ->
+ np1 ! e = Some p ->
+ np2 ! e = None ->
+ p == ⟂.
+ Proof.
+ intros.
+ eapply forall_ptree_true in H; try eassumption.
+ unfold tree_equiv_check_el in H.
+ rewrite H1 in H. now apply equiv_check_correct.
+ Qed.
+
+ Lemma beq_pred_expr_correct3:
+ forall np1 np2 e p,
+ forall_ptree (tree_equiv_check_None_el np1) np2 = true ->
+ np1 ! e = None ->
+ np2 ! e = Some p ->
+ p == ⟂.
+ Proof.
+ intros.
+ eapply forall_ptree_true in H; try eassumption.
+ unfold tree_equiv_check_None_el in H.
+ rewrite H0 in H. now apply equiv_check_correct.
+ Qed.
+
+ Section PRED_PROOFS.
+
+ Context {G B: Type}.
+ Context (f: PTree.t pred_pexpr).
+ Context (ps: PMap.t bool).
+ Context (a_sem: @Abstr.ctx G -> HS.t -> B -> Prop).
+ Context (ctx: @Abstr.ctx G).
+
+ Context (F_EVALULABLE: forall x, sem_pexpr ctx (get_forest_p' x f) ps !! x).
+
+ Variant sem_pred_tree: PTree.t HS.t -> PTree.t pred_op -> B -> Prop :=
+ | sem_pred_tree_intro :
+ forall expr e v et pt pr,
+ sem_pexpr ctx (from_pred_op f pr) true ->
+ a_sem ctx expr v ->
+ pt ! e = Some pr ->
+ et ! e = Some expr ->
+ sem_pred_tree et pt v.
+
+ Lemma norm_expression_in :
+ forall pe et pt h x y,
+ H.wf_hash_table h ->
+ norm_expression 1 pe h = (pt, et) ->
+ h ! x = Some y ->
+ et ! x = Some y.
+ Proof.
+ induction pe; crush; repeat (destruct_match; try discriminate; []).
+ - inv H0. eauto using H.hash_constant.
+ - destruct_match.
+ + inv H0. eapply IHpe.
+ eapply H.wf_hash_table_distr; eauto. eauto.
+ eauto using H.hash_constant.
+ + inv H0. eapply IHpe.
+ eapply H.wf_hash_table_distr; eauto. eauto.
+ eauto using H.hash_constant.
+ Qed.
+
+ Lemma norm_expression_exists :
+ forall pe et pt h x y,
+ H.wf_hash_table h ->
+ norm_expression 1 pe h = (pt, et) ->
+ pt ! x = Some y ->
+ exists z, et ! x = Some z.
+ Proof.
+ induction pe; crush; repeat (destruct_match; try discriminate; []).
+ - inv H0. destruct (peq x h0); subst; inv H1.
+ + eexists. eauto using H.hash_value_in.
+ + rewrite PTree.gso in H2 by auto. now rewrite PTree.gempty in H2.
+ - assert (H.wf_hash_table h1) by eauto using H.wf_hash_table_distr.
+ destruct_match; inv H0.
+ + destruct (peq h0 x); subst; eauto.
+ rewrite PTree.gso in H1 by auto. eauto.
+ + destruct (peq h0 x); subst; eauto.
+ * pose proof Heqp0 as X.
+ eapply H.hash_value_in in Heqp0.
+ eapply norm_expression_in in Heqn; eauto.
+ * rewrite PTree.gso in H1 by auto. eauto.
+ Qed.
+
+ Lemma norm_expression_wf :
+ forall pe et pt h,
+ H.wf_hash_table h ->
+ norm_expression 1 pe h = (pt, et) ->
+ H.wf_hash_table et.
+ Proof.
+ induction pe; crush; repeat (destruct_match; try discriminate; []).
+ - inv H0. eauto using H.wf_hash_table_distr.
+ - destruct_match.
+ + inv H0. eapply IHpe.
+ eapply H.wf_hash_table_distr; eauto. eauto.
+ + inv H0. eapply IHpe.
+ eapply H.wf_hash_table_distr; eauto. eauto.
+ Qed.
+
+ Lemma sem_pred_expr_in_true :
+ forall pe v,
+ sem_pred_expr f a_sem ctx pe v ->
+ exists p e, NE.In (p, e) pe
+ /\ sem_pexpr ctx (from_pred_op f p) true
+ /\ a_sem ctx e v.
+ Proof.
+ induction pe; crush.
+ - inv H. do 2 eexists; split; try split; eauto. constructor.
+ - inv H.
+ + do 2 eexists; split; try split; eauto. constructor; tauto.
+ + exploit IHpe; eauto. simplify.
+ do 2 eexists; split; try split; eauto. constructor; tauto.
+ Qed.
+
+ Definition pred_Ht_dec :
+ forall x y: pred_op * HS.t, {x = y} + {x <> y}.
+ Proof.
+ pose proof HS.eq_dec.
+ pose proof (@Predicate.eq_dec positive peq).
+ decide equality.
+ Defined.
+
+ Lemma sem_pred_mutexcl :
+ forall pe p t v,
+ predicated_mutexcl ((p, t) ::| pe) ->
+ sem_pred_expr f a_sem ctx pe v ->
+ sem_pexpr ctx (from_pred_op f p) false.
+ Proof.
+ intros. unfold predicated_mutexcl in H.
+ exploit sem_pred_expr_in_true; eauto; simplify.
+ unfold "⇒" in *. inv H5.
+ destruct (pred_Ht_dec (x, x0) (p, t)); subst.
+ { inv e; exfalso; apply H7; auto. }
+ assert (NE.In (x, x0) ((p, t) ::| pe)) by (constructor; tauto).
+ assert (NE.In (p, t) ((p, t) ::| pe)) by (constructor; tauto).
+ pose proof (H3 _ _ H H5 n).
+ assert (forall c, eval_predf c x = true -> eval_predf c (¬ p) = true)
+ by eauto.
+ eapply sem_pexpr_eval_inv in H1; eauto.
+ eapply sem_pexpr_eval; eauto. apply H9 in H1.
+ unfold eval_predf in *. rewrite negate_correct in H1.
+ symmetry in H1. apply negb_sym in H1. auto.
+ Qed.
+
+ Lemma exec_tree_exec_pe :
+ forall pe et pt v h
+ (MUTEXCL: predicated_mutexcl pe),
+ H.wf_hash_table h ->
+ norm_expression 1 pe h = (pt, et) ->
+ sem_pred_tree et pt v ->
+ sem_pred_expr f a_sem ctx pe v.
+ Proof.
+ induction pe; simplify; repeat (destruct_match; try discriminate; []).
+ - inv Heqp. inv H0. inv H1.
+ destruct (peq e h0); subst.
+ 2: { rewrite PTree.gso in H3 by auto.
+ rewrite PTree.gempty in H3. discriminate. }
+ assert (expr = t).
+ { apply H.hash_value_in in Heqp0. rewrite H4 in Heqp0. now inv Heqp0. }
+ subst. constructor; auto. rewrite PTree.gss in H3. inv H3; auto.
+ - inv Heqp. inv H1. destruct_match; inv H0; destruct (peq h0 e); subst.
+ + rewrite PTree.gss in H4. inv H4. inv H2. inv H1.
+ * exploit IHpe. eauto using predicated_cons.
+ eapply H.wf_hash_table_distr; eauto. eauto.
+ econstructor. eauto. eauto. eauto. eauto. intros.
+ assert (sem_pexpr ctx (from_pred_op f p) false)
+ by (eapply sem_pred_mutexcl; eauto).
+ eapply sem_pred_expr_cons_false; auto.
+ * assert (et ! e = Some t).
+ { eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
+ eauto. apply H.hash_value_in in Heqp0. auto. }
+ rewrite H1 in H5. inv H5.
+ constructor; auto.
+ + exploit IHpe. eauto using predicated_cons.
+ eapply H.wf_hash_table_distr; eauto. eauto.
+ econstructor. eauto. eauto. rewrite PTree.gso in H4; eauto. auto.
+ intros.
+ assert (sem_pexpr ctx (from_pred_op f p) false)
+ by (eapply sem_pred_mutexcl; eauto).
+ eapply sem_pred_expr_cons_false; auto.
+ + rewrite PTree.gss in H4. inv H4.
+ econstructor; auto.
+ assert (et ! e = Some t).
+ { eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
+ eauto. apply H.hash_value_in in Heqp0. auto. }
+ rewrite H0 in H5; inv H5. auto.
+ + rewrite PTree.gso in H4 by auto.
+ exploit IHpe. eauto using predicated_cons.
+ eapply H.wf_hash_table_distr; eauto. eauto.
+ econstructor. eauto. eauto. eauto. eauto. intros.
+ assert (sem_pexpr ctx (from_pred_op f p) false)
+ by (eapply sem_pred_mutexcl; eauto).
+ eapply sem_pred_expr_cons_false; auto.
+ Qed.
+
+ Lemma exec_pe_exec_tree :
+ forall pe et pt v h
+ (MUTEXCL: predicated_mutexcl pe),
+ H.wf_hash_table h ->
+ norm_expression 1 pe h = (pt, et) ->
+ sem_pred_expr f a_sem ctx pe v ->
+ sem_pred_tree et pt v.
+ Proof.
+ induction pe; simplify; repeat (destruct_match; try discriminate; []).
+ - inv H0. inv H1. econstructor; eauto. apply PTree.gss.
+ eapply H.hash_value_in; eauto.
+ - inv H1.
+ + destruct_match.
+ * inv H0. econstructor.
+ 2: { eauto. }
+ 2: { apply PTree.gss. }
+ constructor; tauto.
+ eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
+ eauto. eapply H.hash_value_in; eauto.
+ * inv H0. econstructor. eauto. eauto. apply PTree.gss.
+ eapply norm_expression_in. eapply H.wf_hash_table_distr; eauto.
+ eauto. eapply H.hash_value_in; eauto.
+ + destruct_match.
+ * inv H0. exploit IHpe.
+ eauto using predicated_cons.
+ eapply H.wf_hash_table_distr; eauto.
+ eauto. eauto. intros. inv H0.
+ destruct (peq e h0); subst.
+ -- rewrite H3 in Heqo. inv Heqo.
+ econstructor.
+ 3: { apply PTree.gss. }
+ constructor; tauto. eauto. auto.
+ -- econstructor. eauto. eauto. rewrite PTree.gso by eauto. auto.
+ auto.
+ * inv H0. exploit IHpe.
+ eauto using predicated_cons.
+ eapply H.wf_hash_table_distr; eauto.
+ eauto. eauto. intros. inv H0.
+ destruct (peq e h0); subst.
+ -- rewrite H3 in Heqo; discriminate.
+ -- econstructor; eauto. rewrite PTree.gso by auto. auto.
+ Qed.
+
+ Lemma beq_pred_expr_correct_top:
+ forall p1 p2 v
+ (MUTEXCL1: predicated_mutexcl p1)
+ (MUTEXCL2: predicated_mutexcl p2),
+ beq_pred_expr p1 p2 = true ->
+ sem_pred_expr f a_sem ctx p1 v ->
+ sem_pred_expr f a_sem ctx p2 v.
+ Proof.
+ unfold beq_pred_expr; intros.
+ destruct_match; subst; auto.
+ repeat (destruct_match; []).
+ symmetry in H. apply andb_true_eq in H. inv H.
+ symmetry in H1. symmetry in H2.
+ pose proof Heqn0. eapply norm_expression_wf in H.
+ 2: { unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3. }
+ eapply exec_tree_exec_pe; eauto.
+ eapply exec_pe_exec_tree in H0; auto.
+ 3: { eauto. }
+ 2: { unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3. }
+ inv H0. destruct (t0 ! e) eqn:?.
+ - assert (pr == p) by eauto using beq_pred_expr_correct.
+ assert (sem_pexpr ctx (from_pred_op f p) true).
+ { eapply sem_pexpr_eval; eauto. eapply sem_pexpr_eval_inv in H3; eauto. }
+ econstructor. apply H7. eauto. eauto.
+ eapply norm_expression_in; eauto.
+ - assert (pr == ⟂) by eauto using beq_pred_expr_correct2.
+ eapply sem_pexpr_eval_inv in H3; eauto. now rewrite H0 in H3.
+ Qed.
+
+ Lemma beq_pred_expr_correct_top2:
+ forall p1 p2 v
+ (MUTEXCL1: predicated_mutexcl p1)
+ (MUTEXCL2: predicated_mutexcl p2),
+ beq_pred_expr p1 p2 = true ->
+ sem_pred_expr f a_sem ctx p2 v ->
+ sem_pred_expr f a_sem ctx p1 v.
+ Proof.
+ unfold beq_pred_expr; intros.
+ destruct_match; subst; auto.
+ repeat (destruct_match; []).
+ symmetry in H. apply andb_true_eq in H. inv H.
+ symmetry in H1. symmetry in H2.
+ pose proof Heqn0. eapply norm_expression_wf in H.
+ 2: { unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3. }
+ eapply exec_tree_exec_pe; auto.
+ 2: { eauto. }
+ unfold H.wf_hash_table; intros. now rewrite PTree.gempty in H3.
+ eapply exec_pe_exec_tree in H0; auto.
+ 3: { eauto. }
+ 2: { auto. }
+ inv H0. destruct (t ! e) eqn:?.
+ - assert (p == pr) by eauto using beq_pred_expr_correct.
+ assert (sem_pexpr ctx (from_pred_op f p) true).
+ { eapply sem_pexpr_eval; eauto. eapply sem_pexpr_eval_inv in H3; eauto. }
+ econstructor. apply H7. eauto. eauto.
+ exploit norm_expression_exists.
+ 2: { eapply Heqn0. } unfold H.wf_hash_table; intros * YH.
+ now rewrite PTree.gempty in YH. eauto. simplify.
+ exploit norm_expression_in. 2: { eauto. } auto. eauto. intros.
+ crush.
+ - assert (pr == ⟂) by eauto using beq_pred_expr_correct3.
+ eapply sem_pexpr_eval_inv in H3; eauto. now rewrite H0 in H3.
+ Qed.
+
+ End PRED_PROOFS.
+
+End HashExprNorm.
+
+Module HN := HashExprNorm(HashExpr).
+Module EHN := HashExprNorm(EHashExpr).
+
+Definition check_mutexcl {A} (pe: predicated A) :=
+ let preds := map fst (NE.to_list pe) in
+ let pairs := map (fun x => x → or_list (remove (Predicate.eq_dec peq) x preds)) preds in
+ match sat_pred_simple (simplify (negate (and_list pairs))) with
+ | None => true
+ | _ => false
+ end.
+
+Lemma check_mutexcl_correct:
+ forall A (pe: predicated A),
+ check_mutexcl pe = true ->
+ predicated_mutexcl pe.
+Proof. Admitted.
+
+Definition check_mutexcl_tree {A} (f: PTree.t (predicated A)) :=
+ forall_ptree (fun _ => check_mutexcl) f.
+
+Lemma check_mutexcl_tree_correct:
+ forall A (f: PTree.t (predicated A)) i pe,
+ check_mutexcl_tree f = true ->
+ f ! i = Some pe ->
+ predicated_mutexcl pe.
+Proof.
+ unfold check_mutexcl_tree; intros.
+ eapply forall_ptree_true in H; eauto using check_mutexcl_correct.
+Qed.
+
+Definition check f1 f2 :=
+ RTree.beq HN.beq_pred_expr f1.(forest_regs) f2.(forest_regs)
+ && PTree.beq beq_pred_pexpr f1.(forest_preds) f2.(forest_preds)
+ && EHN.beq_pred_expr f1.(forest_exit) f2.(forest_exit)
+ && check_mutexcl_tree f1.(forest_regs)
+ && check_mutexcl_tree f2.(forest_regs)
+ && check_mutexcl f1.(forest_exit)
+ && check_mutexcl f2.(forest_exit).
+
+Lemma sem_pexpr_forward_eval1 :
+ forall A ctx f_p ps,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p,
+ @sem_pexpr A ctx (from_pred_op f_p p) false ->
+ eval_predf ps p = false.
+Proof.
+ induction p; crush.
+ - destruct_match. inv Heqp0. destruct b.
+ cbn. specialize (H p0).
+ eapply sem_pexpr_det; eauto. apply similar_refl.
+ specialize (H p0).
+ replace false with (negb true) in H0 by auto.
+ eapply sem_pexpr_negate2 in H0. cbn.
+ symmetry; apply negb_sym. cbn.
+ eapply sem_pexpr_det; eauto. apply similar_refl.
+ - inv H0.
+ - inv H0. inv H2. rewrite eval_predf_Pand. rewrite IHp1; eauto.
+ rewrite eval_predf_Pand. rewrite IHp2; eauto with bool.
+ - inv H0. rewrite eval_predf_Por. rewrite IHp1; eauto.
+Qed.
+
+Lemma sem_pexpr_forward_eval2 :
+ forall A ctx f_p ps,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p,
+ @sem_pexpr A ctx (from_pred_op f_p p) true ->
+ eval_predf ps p = true.
+Proof.
+ induction p; crush.
+ - destruct_match. inv Heqp0. destruct b.
+ cbn. specialize (H p0).
+ eapply sem_pexpr_det; eauto. apply similar_refl.
+ cbn. symmetry. apply negb_sym; cbn.
+ replace true with (negb false) in H0 by auto.
+ eapply sem_pexpr_negate2 in H0.
+ eapply sem_pexpr_det; eauto. apply similar_refl.
+ - inv H0.
+ - inv H0. rewrite eval_predf_Pand. rewrite IHp1; eauto.
+ - inv H0. inv H2. rewrite eval_predf_Por. rewrite IHp1; eauto.
+ rewrite eval_predf_Por. rewrite IHp2; eauto with bool.
+Qed.
+
+Lemma sem_pexpr_forward_eval :
+ forall A ctx f_p ps,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ forall p b,
+ @sem_pexpr A ctx (from_pred_op f_p p) b ->
+ eval_predf ps p = b.
+Proof.
+ intros; destruct b; eauto using sem_pexpr_forward_eval1, sem_pexpr_forward_eval2.
+Qed.
+
+Section BOOLEAN_EQUALITY.
+
+ Context {A B: Type}.
+ Context (beqA: A -> B -> bool).
+
+ Fixpoint beq2' (m1: PTree.tree' A) (m2: PTree.tree' B) {struct m1} : bool :=
+ match m1, m2 with
+ | PTree.Node001 r1, PTree.Node001 r2 => beq2' r1 r2
+ | PTree.Node010 x1, PTree.Node010 x2 => beqA x1 x2
+ | PTree.Node011 x1 r1, PTree.Node011 x2 r2 => beqA x1 x2 && beq2' r1 r2
+ | PTree.Node100 l1, PTree.Node100 l2 => beq2' l1 l2
+ | PTree.Node101 l1 r1, PTree.Node101 l2 r2 => beq2' l1 l2 && beq2' r1 r2
+ | PTree.Node110 l1 x1, PTree.Node110 l2 x2 => beqA x1 x2 && beq2' l1 l2
+ | PTree.Node111 l1 x1 r1, PTree.Node111 l2 x2 r2 => beqA x1 x2 && beq2' l1 l2 && beq2' r1 r2
+ | _, _ => false
+ end.
+
+ Definition beq2 (m1: PTree.t A) (m2 : PTree.t B) : bool :=
+ match m1, m2 with
+ | PTree.Empty, PTree.Empty => true
+ | PTree.Nodes m1', PTree.Nodes m2' => beq2' m1' m2'
+ | _, _ => false
+ end.
+
+ Let beq2_optA (o1: option A) (o2: option B) : bool :=
+ match o1, o2 with
+ | None, None => true
+ | Some a1, Some a2 => beqA a1 a2
+ | _, _ => false
+ end.
+
+ Lemma beq_correct_bool:
+ forall m1 m2,
+ beq2 m1 m2 = true <-> (forall x, beq2_optA (m1 ! x) (m2 ! x) = true).
+ Proof.
+ Local Transparent PTree.Node.
+ assert (beq_NN: forall l1 o1 r1 l2 o2 r2,
+ PTree.not_trivially_empty l1 o1 r1 ->
+ PTree.not_trivially_empty l2 o2 r2 ->
+ beq2 (PTree.Node l1 o1 r1) (PTree.Node l2 o2 r2) =
+ beq2 l1 l2 && beq2_optA o1 o2 && beq2 r1 r2).
+ { intros.
+ destruct l1, o1, r1; try contradiction; destruct l2, o2, r2; try contradiction;
+ simpl; rewrite ? andb_true_r, ? andb_false_r; auto.
+ rewrite andb_comm; auto.
+ f_equal; rewrite andb_comm; auto. }
+ induction m1 using PTree.tree_ind; [|induction m2 using PTree.tree_ind].
+ - intros. simpl; split; intros.
+ + destruct m2; try discriminate. simpl; auto.
+ + replace m2 with (@PTree.Empty B); auto. apply PTree.extensionality; intros x.
+ specialize (H x). destruct (m2 ! x); simpl; easy.
+ - split; intros.
+ + destruct (PTree.Node l o r); try discriminate. simpl; auto.
+ + replace (PTree.Node l o r) with (@PTree.Empty A); auto. apply PTree.extensionality; intros x.
+ specialize (H0 x). destruct ((PTree.Node l o r) ! x); simpl in *; easy.
+ - rewrite beq_NN by auto. split; intros.
+ + InvBooleans. rewrite ! PTree.gNode. destruct x.
+ * apply IHm0; auto.
+ * apply IHm1; auto.
+ * auto.
+ + apply andb_true_intro; split; [apply andb_true_intro; split|].
+ * apply IHm1. intros. specialize (H1 (xO x)); rewrite ! PTree.gNode in H1; auto.
+ * specialize (H1 xH); rewrite ! PTree.gNode in H1; auto.
+ * apply IHm0. intros. specialize (H1 (xI x)); rewrite ! PTree.gNode in H1; auto.
+ Qed.
+
+ Theorem beq2_correct:
+ forall m1 m2,
+ beq2 m1 m2 = true <->
+ (forall (x: PTree.elt),
+ match m1 ! x, m2 ! x with
+ | None, None => True
+ | Some y1, Some y2 => beqA y1 y2 = true
+ | _, _ => False
+ end).
+ Proof.
+ intros. rewrite beq_correct_bool. unfold beq2_optA. split; intros.
+ - specialize (H x). destruct (m1 ! x), (m2 ! x); intuition congruence.
+ - specialize (H x). destruct (m1 ! x), (m2 ! x); intuition auto.
+ Qed.
+
+End BOOLEAN_EQUALITY.
+
+Section GENERIC_CONTEXT.
+
+Context {A: Type}.
+Context (ctx: @ctx A).
+
+(*|
+Suitably restrict the predicate set so that one can evaluate a hashed predicate
+using that predicate set. However, one issue might be that we do not know that
+all the atoms of the original formula are actually evaluable.
+|*)
+
+Definition match_pred_states
+ (ht: PHT.hash_tree) (p_out: pred_op) (pred_set: predset) :=
+ forall (p: positive) (br: bool) (p_in: pred_expression),
+ PredIn p p_out ->
+ ht ! p = Some p_in ->
+ sem_pred ctx p_in (pred_set !! p).
+
+Lemma eval_hash_predicate :
+ forall max p_in ht p_out ht' br pred_set,
+ hash_predicate max p_in ht = (p_out, ht') ->
+ sem_pexpr ctx p_in br ->
+ match_pred_states ht' p_out pred_set ->
+ eval_predf pred_set p_out = br.
+Proof.
+ induction p_in; simplify.
+ + repeat destruct_match. inv H.
+ unfold eval_predf. cbn.
+ inv H0. inv H4. unfold match_pred_states in H1.
+ specialize (H1 h br).
+Abort.
+
+Lemma sem_pexpr_beq_correct :
+ forall p1 p2 b,
+ beq_pred_pexpr p1 p2 = true ->
+ sem_pexpr ctx p1 b ->
+ sem_pexpr ctx p2 b.
+Proof.
+ unfold beq_pred_pexpr.
+ induction p1; intros; destruct_match; crush.
+ Admitted.
+
+(*|
+This should only require a proof of sem_pexpr_beq_correct, the rest is
+straightforward.
+|*)
+
+Lemma pred_pexpr_beq_pred_pexpr :
+ forall pr a b br,
+ PTree.beq beq_pred_pexpr a b = true ->
+ sem_pexpr ctx (from_pred_op a pr) br ->
+ sem_pexpr ctx (from_pred_op b pr) br.
+Proof.
+ induction pr; crush.
+ - destruct_match. inv Heqp0. destruct b0.
+ + unfold get_forest_p' in *.
+ apply PTree.beq_correct with (x := p0) in H.
+ destruct a ! p0; destruct b ! p0; try (exfalso; assumption); auto.
+ eapply sem_pexpr_beq_correct; eauto.
+ + replace br with (negb (negb br)) by (now apply negb_involutive).
+ replace br with (negb (negb br)) in H0 by (now apply negb_involutive).
+ apply sem_pexpr_negate. apply sem_pexpr_negate2 in H0.
+ unfold get_forest_p' in *.
+ apply PTree.beq_correct with (x := p0) in H.
+ destruct a ! p0; destruct b ! p0; try (exfalso; assumption); auto.
+ eapply sem_pexpr_beq_correct; eauto.
+ - inv H0; try inv H4.
+ + eapply IHpr1 in H0; eauto. apply sem_pexpr_Pand_false; tauto.
+ + eapply IHpr2 in H0; eauto. apply sem_pexpr_Pand_false; tauto.
+ + eapply IHpr1 in H3; eauto. eapply IHpr2 in H5; eauto.
+ apply sem_pexpr_Pand_true; auto.
+ - inv H0; try inv H4.
+ + eapply IHpr1 in H0; eauto. apply sem_pexpr_Por_true; tauto.
+ + eapply IHpr2 in H0; eauto. apply sem_pexpr_Por_true; tauto.
+ + eapply IHpr1 in H3; eauto. eapply IHpr2 in H5; eauto.
+ apply sem_pexpr_Por_false; auto.
+Qed.
+
+(*|
+This lemma requires a theorem that equivalence of symbolic predicates means they
+will be. This further needs three-valued logic to be able to compare arbitrary
+predicates with each other, that will also show that the predicate will also be
+computable.
+|*)
+
+Lemma sem_pred_exec_beq_correct :
+ forall A B (sem: Abstr.ctx -> A -> B -> Prop) p a b r,
+ PTree.beq beq_pred_pexpr a b = true ->
+ sem_pred_expr a sem ctx p r ->
+ sem_pred_expr b sem ctx p r.
+Proof.
+ induction p; intros; inv H0.
+ - constructor; auto. eapply pred_pexpr_beq_pred_pexpr; eauto.
+ - constructor; auto. eapply pred_pexpr_beq_pred_pexpr; eauto.
+ - apply sem_pred_expr_cons_false; eauto.
+ eapply pred_pexpr_beq_pred_pexpr; eauto.
+Qed.
+
+End GENERIC_CONTEXT.
+
+Lemma tree_beq_pred_pexpr :
+ forall a b x,
+ RTree.beq beq_pred_pexpr (forest_preds a) (forest_preds b) = true ->
+ beq_pred_pexpr a #p x b #p x = true.
+Proof.
+ intros. unfold "#p". unfold get_forest_p'.
+ apply PTree.beq_correct with (x := x) in H.
+ destruct_match; destruct_match; auto.
+ unfold beq_pred_pexpr. destruct_match; auto.
+Qed.
+
+Lemma tree_beq_pred_expr :
+ forall a b x,
+ RTree.beq HN.beq_pred_expr (forest_regs a) (forest_regs b) = true ->
+ HN.beq_pred_expr a #r x b #r x = true.
+Proof.
+ intros. unfold "#r" in *.
+ apply PTree.beq_correct with (x := (R_indexed.index x)) in H.
+ unfold RTree.get in *.
+ unfold pred_expr in *.
+ destruct_match; destruct_match; auto.
+ unfold HN.beq_pred_expr. destruct_match; auto.
+Qed.
+
+Section CORRECT.
+
+Context {FUN TFUN: Type}.
+Context (ictx: @ctx FUN) (octx: @ctx TFUN) (HSIM: similar ictx octx).
+
+Lemma abstr_check_correct :
+ forall i' a b cf,
+ (exists ps, forall x, sem_pexpr ictx (get_forest_p' x (forest_preds a)) ps !! x) ->
+ check a b = true ->
+ sem ictx a (i', cf) ->
+ exists ti', sem octx b (ti', cf) /\ match_states i' ti'.
+Proof.
+ intros * EVALUABLE **. unfold check in H. simplify.
+ inv H0. inv H10. inv H11.
+ eexists; split; constructor; auto.
+ - constructor. intros.
+ eapply sem_pred_exec_beq_correct; eauto.
+ eapply sem_pred_expr_corr; eauto. now apply sem_value_corr.
+ eapply HN.beq_pred_expr_correct_top; eauto.
+ { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
+ { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
+ eapply tree_beq_pred_expr; eauto.
+ - (* This is where three-valued logic would be needed. *)
+ constructor. intros. apply sem_pexpr_beq_correct with (p1 := a #p x0).
+ apply tree_beq_pred_pexpr; auto.
+ apply sem_pexpr_corr with (ictx:=ictx); auto.
+ - eapply sem_pred_exec_beq_correct; eauto.
+ eapply sem_pred_expr_corr; eauto. now apply sem_mem_corr.
+ eapply HN.beq_pred_expr_correct_top; eauto.
+ { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
+ { unfold "#r"; destruct_match; eauto using check_mutexcl_tree_correct, predicated_singleton. }
+ eapply tree_beq_pred_expr; eauto.
+ - eapply sem_pred_exec_beq_correct; eauto.
+ eapply sem_pred_expr_corr; eauto. now apply sem_exit_corr.
+ eapply EHN.beq_pred_expr_correct_top; eauto using check_mutexcl_correct.
+Qed.
+
+(*|
+Proof Sketch:
+
+Two abstract states can be equivalent, without it being obvious that they can
+actually both be executed assuming one can be executed. One will therefore have
+to add a few more assumptions to makes it possible to execute the other.
+
+It currently assumes that all the predicates in the predicate tree are
+evaluable, which is actually something that can either be checked, or something
+that can be proven constructively. I believe that it should already be possible
+using the latter, so here it will only be assumed.
+
+Similarly, the current assumption is that mutual exclusivity of predicates is
+being checked within the ``check`` function, which could possibly also be proven
+constructively about the update function. This is a simpler short-term fix
+though.
+|*)
+
+End CORRECT.
diff --git a/src/hls/GiblePargenproofEvaluable.v b/src/hls/GiblePargenproofEvaluable.v
new file mode 100644
index 0000000..9013cd5
--- /dev/null
+++ b/src/hls/GiblePargenproofEvaluable.v
@@ -0,0 +1,395 @@
+Require Import compcert.backend.Registers.
+Require Import compcert.common.AST.
+Require Import compcert.common.Errors.
+Require Import compcert.common.Linking.
+Require Import compcert.common.Globalenvs.
+Require Import compcert.common.Memory.
+Require Import compcert.common.Values.
+Require Import compcert.lib.Maps.
+
+Require Import vericert.common.Vericertlib.
+Require Import vericert.hls.GibleSeq.
+Require Import vericert.hls.GiblePar.
+Require Import vericert.hls.Gible.
+Require Import vericert.hls.GiblePargenproofEquiv.
+Require Import vericert.hls.GiblePargen.
+Require Import vericert.hls.Predicate.
+Require Import vericert.hls.Abstr.
+Require Import vericert.common.Monad.
+
+#[local] Open Scope positive.
+#[local] Open Scope forest.
+#[local] Open Scope pred_op.
+
+#[local] Opaque simplify.
+#[local] Opaque deep_simplify.
+
+
+Definition evaluable {A B C} (sem: ctx -> B -> C -> Prop) (ctx: @ctx A) p := exists b, sem ctx p b.
+
+Definition p_evaluable {A} := @evaluable A _ _ sem_pexpr.
+
+Definition evaluable_expr {A} := @evaluable A _ _ sem_pred.
+
+Definition all_evaluable {A B} (ctx: @ctx A) f_p (l: predicated B) :=
+ forall p y, NE.In (p, y) l -> p_evaluable ctx (from_pred_op f_p p).
+
+Definition all_evaluable2 {A B C} (ctx: @ctx A) (sem: Abstr.ctx -> B -> C -> Prop) (l: predicated B) :=
+ forall p y, NE.In (p, y) l -> evaluable sem ctx y.
+
+Definition pred_forest_evaluable {A} (ctx: @ctx A) (ps: PTree.t pred_pexpr) :=
+ forall i p, ps ! i = Some p -> p_evaluable ctx p.
+
+Definition forest_evaluable {A} (ctx: @ctx A) (f: forest) :=
+ pred_forest_evaluable ctx f.(forest_preds).
+
+Lemma all_evaluable_cons :
+ forall A B pr ctx b a,
+ all_evaluable ctx pr (NE.cons a b) ->
+ @all_evaluable A B ctx pr b.
+Proof.
+ unfold all_evaluable; intros.
+ enough (NE.In (p, y) (NE.cons a b)); eauto.
+ constructor; tauto.
+Qed.
+
+Lemma all_evaluable2_cons :
+ forall A B C sem ctx b a,
+ all_evaluable2 ctx sem (NE.cons a b) ->
+ @all_evaluable2 A B C ctx sem b.
+Proof.
+ unfold all_evaluable2; intros.
+ enough (NE.In (p, y) (NE.cons a b)); eauto.
+ constructor; tauto.
+Qed.
+
+Lemma all_evaluable_cons2 :
+ forall A B pr ctx b a p,
+ @all_evaluable A B ctx pr (NE.cons (p, a) b) ->
+ p_evaluable ctx (from_pred_op pr p).
+Proof.
+ unfold all_evaluable; intros.
+ eapply H. constructor; eauto.
+Qed.
+
+Lemma all_evaluable2_cons2 :
+ forall A B C sem ctx b a p,
+ @all_evaluable2 A B C ctx sem (NE.cons (p, a) b) ->
+ evaluable sem ctx a.
+Proof.
+ unfold all_evaluable; intros.
+ eapply H. constructor; eauto.
+Qed.
+
+Lemma all_evaluable_cons3 :
+ forall A B pr ctx b p a,
+ all_evaluable ctx pr b ->
+ p_evaluable ctx (from_pred_op pr p) ->
+ @all_evaluable A B ctx pr (NE.cons (p, a) b).
+Proof.
+ unfold all_evaluable; intros. inv H1. inv H3. inv H1. auto.
+ eauto.
+Qed.
+
+Lemma all_evaluable_singleton :
+ forall A B pr ctx b p,
+ @all_evaluable A B ctx pr (NE.singleton (p, b)) ->
+ p_evaluable ctx (from_pred_op pr p).
+Proof.
+ unfold all_evaluable; intros. eapply H. constructor.
+Qed.
+
+Lemma get_forest_p_evaluable :
+ forall A (ctx: @ctx A) f r,
+ forest_evaluable ctx f ->
+ p_evaluable ctx (f #p r).
+Proof.
+ intros. unfold "#p", get_forest_p'.
+ destruct_match. unfold forest_evaluable in *.
+ unfold pred_forest_evaluable in *. eauto.
+ unfold p_evaluable, evaluable. eexists.
+ constructor. constructor.
+Qed.
+
+Lemma set_forest_p_evaluable :
+ forall A (ctx: @ctx A) f r p,
+ forest_evaluable ctx f ->
+ p_evaluable ctx p ->
+ forest_evaluable ctx (f #p r <- p).
+Proof.
+ unfold forest_evaluable, pred_forest_evaluable; intros.
+ destruct (peq i r); subst.
+ - rewrite forest_pred_gss2 in H1. now inv H1.
+ - rewrite forest_pred_gso2 in H1 by auto; eauto.
+Qed.
+
+ Lemma evaluable_singleton :
+ forall A B ctx fp r,
+ @all_evaluable A B ctx fp (NE.singleton (T, r)).
+ Proof.
+ unfold all_evaluable, evaluable; intros.
+ inv H. simpl. exists true. constructor.
+ Qed.
+
+Lemma evaluable_negate :
+ forall G (ctx: @ctx G) p,
+ p_evaluable ctx p ->
+ p_evaluable ctx (¬ p).
+Proof.
+ unfold p_evaluable, evaluable in *; intros.
+ inv H. eapply sem_pexpr_negate in H0. econstructor; eauto.
+Qed.
+
+Lemma predicated_evaluable :
+ forall G (ctx: @ctx G) ps (p: pred_op),
+ pred_forest_evaluable ctx ps ->
+ p_evaluable ctx (from_pred_op ps p).
+Proof.
+ unfold pred_forest_evaluable; intros. induction p; intros; cbn.
+ - repeat destruct_match; subst; unfold get_forest_p'.
+ destruct_match. eapply H; eauto. econstructor. constructor. constructor.
+ eapply evaluable_negate.
+ destruct_match. eapply H; eauto. econstructor. constructor. constructor.
+ - repeat econstructor.
+ - repeat econstructor.
+ - inv IHp1. inv IHp2. econstructor. apply sem_pexpr_Pand; eauto.
+ - inv IHp1. inv IHp2. econstructor. apply sem_pexpr_Por; eauto.
+Qed.
+
+Lemma predicated_evaluable_all :
+ forall A G (ctx: @ctx G) ps (p: predicated A),
+ pred_forest_evaluable ctx ps ->
+ all_evaluable ctx ps p.
+Proof.
+ unfold all_evaluable; intros.
+ eapply predicated_evaluable. eauto.
+Qed.
+
+Lemma predicated_evaluable_all_list :
+ forall A G (ctx: @ctx G) ps (p: list (predicated A)),
+ pred_forest_evaluable ctx ps ->
+ Forall (all_evaluable ctx ps) p.
+Proof.
+ induction p.
+ - intros. constructor.
+ - intros. constructor; eauto. apply predicated_evaluable_all; auto.
+Qed.
+
+#[local] Hint Resolve evaluable_singleton : core.
+#[local] Hint Resolve predicated_evaluable : core.
+#[local] Hint Resolve predicated_evaluable_all : core.
+#[local] Hint Resolve predicated_evaluable_all_list : core.
+
+Lemma evaluable_and_true :
+ forall G (ctx: @ctx G) ps p,
+ p_evaluable ctx (from_pred_op ps p) ->
+ p_evaluable ctx (from_pred_op ps (p ∧ T)).
+Proof.
+ intros. unfold evaluable in *. inv H. exists (x && true). cbn.
+ apply sem_pexpr_Pand; auto. constructor.
+Qed.
+
+Lemma all_evaluable_predicated_map :
+ forall A B G (ctx: @ctx G) ps (f: A -> B) p,
+ all_evaluable ctx ps p ->
+ all_evaluable ctx ps (predicated_map f p).
+Proof.
+ induction p.
+ - unfold all_evaluable; intros.
+ exploit NEin_map; eauto; intros. inv H1. inv H2.
+ eapply H; eauto.
+ - intros. cbn.
+ eapply all_evaluable_cons3. eapply IHp. eapply all_evaluable_cons; eauto.
+ cbn. destruct a. cbn in *. eapply all_evaluable_cons2; eauto.
+Qed.
+
+Lemma all_evaluable_predicated_map2 :
+ forall A B G (ctx: @ctx G) ps (f: A -> B) p,
+ all_evaluable ctx ps (predicated_map f p) ->
+ all_evaluable ctx ps p.
+Proof.
+ induction p.
+ - unfold all_evaluable in *; intros.
+ eapply H. eapply NEin_map2; eauto.
+ - intros. cbn. destruct a.
+ cbn in H. pose proof H. eapply all_evaluable_cons in H; eauto.
+ eapply all_evaluable_cons2 in H0; eauto.
+ unfold all_evaluable. specialize (IHp H).
+ unfold all_evaluable in IHp.
+ intros. inv H1. inv H3. inv H1; eauto.
+ specialize (IHp _ _ H1). eauto.
+Qed.
+
+Lemma all_evaluable_map_and :
+ forall A B G (ctx: @ctx G) ps (a: NE.non_empty ((pred_op * A) * (pred_op * B))),
+ (forall p1 x p2 y,
+ NE.In ((p1, x), (p2, y)) a ->
+ p_evaluable ctx (from_pred_op ps p1)
+ /\ p_evaluable ctx (from_pred_op ps p2)) ->
+ all_evaluable ctx ps (NE.map (fun x => match x with ((a, b), (c, d)) => (Pand a c, (b, d)) end) a).
+Proof.
+ induction a.
+ - intros. cbn. repeat destruct_match. inv Heqp.
+ unfold all_evaluable; intros. inv H0.
+ unfold evaluable.
+ exploit H. constructor.
+ intros [Ha Hb]. inv Ha. inv Hb.
+ exists (x && x0). apply sem_pexpr_Pand; auto.
+ - intros. unfold all_evaluable; cbn; intros. inv H0. inv H2.
+ + repeat destruct_match. inv Heqp0. inv H0.
+ specialize (H p2 a1 p3 b ltac:(constructor; eauto)).
+ inv H. inv H0. inv H1. exists (x && x0).
+ apply sem_pexpr_Pand; eauto.
+ + eapply IHa; intros. eapply H. econstructor. right. eauto.
+ eauto.
+Qed.
+
+Lemma all_evaluable_map_add :
+ forall A B G (ctx: @ctx G) ps (a: pred_op * A) (b: predicated B) p1 x p2 y,
+ p_evaluable ctx (from_pred_op ps (fst a)) ->
+ all_evaluable ctx ps b ->
+ NE.In (p1, x, (p2, y)) (NE.map (fun x : pred_op * B => (a, x)) b) ->
+ p_evaluable ctx (from_pred_op ps p1) /\ p_evaluable ctx (from_pred_op ps p2).
+Proof.
+ induction b; intros.
+ - cbn in *. inv H1. unfold all_evaluable in *. specialize (H0 _ _ ltac:(constructor)).
+ auto.
+ - cbn in *. inv H1. inv H3.
+ + inv H1. unfold all_evaluable in H0. specialize (H0 _ _ ltac:(constructor; eauto)); auto.
+ + eapply all_evaluable_cons in H0. specialize (IHb _ _ _ _ H H0 H1). auto.
+Qed.
+
+Lemma all_evaluable_non_empty_prod :
+ forall A B G (ctx: @ctx G) ps p1 x p2 y (a: predicated A) (b: predicated B),
+ all_evaluable ctx ps a ->
+ all_evaluable ctx ps b ->
+ NE.In ((p1, x), (p2, y)) (NE.non_empty_prod a b) ->
+ p_evaluable ctx (from_pred_op ps p1)
+ /\ p_evaluable ctx (from_pred_op ps p2).
+Proof.
+ induction a; intros.
+ - cbn in *. eapply all_evaluable_map_add; eauto. destruct a; cbn in *. eapply H; constructor.
+ - cbn in *. eapply NE.NEin_NEapp5 in H1. inv H1. eapply all_evaluable_map_add; eauto.
+ destruct a; cbn in *. eapply all_evaluable_cons2; eauto.
+ eapply all_evaluable_cons in H. eauto.
+Qed.
+
+Lemma all_evaluable_predicated_prod :
+ forall A B G (ctx: @ctx G) ps (a: predicated A) (b: predicated B),
+ all_evaluable ctx ps a ->
+ all_evaluable ctx ps b ->
+ all_evaluable ctx ps (predicated_prod a b).
+Proof.
+ intros. unfold all_evaluable; intros.
+ unfold predicated_prod in *. exploit all_evaluable_map_and.
+ 2: { eauto. }
+ intros. 2: { eauto. }
+Admitted. (* Requires simple lemma to prove, but this lemma is not needed. *)
+
+Lemma cons_pred_expr_evaluable :
+ forall G (ctx: @ctx G) ps a b,
+ all_evaluable ctx ps a ->
+ all_evaluable ctx ps b ->
+ all_evaluable ctx ps (cons_pred_expr a b).
+Proof.
+ unfold cons_pred_expr; intros.
+ apply all_evaluable_predicated_map.
+ now apply all_evaluable_predicated_prod.
+Qed.
+
+Lemma fold_right_all_evaluable :
+ forall G (ctx: @ctx G) ps n,
+ Forall (all_evaluable ctx ps) (NE.to_list n) ->
+ all_evaluable ctx ps (NE.fold_right cons_pred_expr (NE.singleton (T, Enil)) n).
+Proof.
+ induction n; cbn in *; intros.
+ - unfold cons_pred_expr. eapply all_evaluable_predicated_map. eapply all_evaluable_predicated_prod.
+ inv H. auto. unfold all_evaluable; intros. inv H0. exists true. constructor.
+ - inv H. specialize (IHn H3). now eapply cons_pred_expr_evaluable.
+Qed.
+
+Lemma merge_all_evaluable :
+ forall G (ctx: @ctx G) ps,
+ pred_forest_evaluable ctx ps ->
+ forall f args,
+ all_evaluable ctx ps (merge (list_translation args f)).
+Proof.
+ intros. eapply predicated_evaluable_all. eauto.
+Qed.
+
+ Lemma forest_evaluable_regset :
+ forall A f (ctx: @ctx A) n x,
+ forest_evaluable ctx f ->
+ forest_evaluable ctx f #r x <- n.
+ Proof.
+ unfold forest_evaluable, pred_forest_evaluable; intros.
+ eapply H. eauto.
+ Qed.
+
+ Lemma evaluable_impl :
+ forall A (ctx: @ctx A) a b,
+ p_evaluable ctx a ->
+ p_evaluable ctx b ->
+ p_evaluable ctx (a → b).
+ Proof.
+ unfold evaluable.
+ inversion_clear 1 as [b1 SEM1].
+ inversion_clear 1 as [b2 SEM2].
+ unfold Pimplies.
+ econstructor. apply sem_pexpr_Por;
+ eauto using sem_pexpr_negate.
+ Qed.
+
+ Lemma parts_evaluable :
+ forall A (ctx: @ctx A) b p0,
+ evaluable sem_pred ctx p0 ->
+ evaluable sem_pexpr ctx (Plit (b, p0)).
+ Proof.
+ intros. unfold evaluable in *. inv H.
+ destruct b. do 2 econstructor. eauto.
+ exists (negb x). constructor. rewrite negb_involutive. auto.
+ Qed.
+
+ Lemma p_evaluable_Pand :
+ forall A (ctx: @ctx A) a b,
+ p_evaluable ctx a ->
+ p_evaluable ctx b ->
+ p_evaluable ctx (a ∧ b).
+ Proof.
+ intros. inv H. inv H0.
+ econstructor. apply sem_pexpr_Pand; eauto.
+ Qed.
+
+ Lemma from_predicated_evaluable :
+ forall A (ctx: @ctx A) f b a,
+ pred_forest_evaluable ctx f ->
+ all_evaluable2 ctx sem_pred a ->
+ p_evaluable ctx (from_predicated b f a).
+ Proof.
+ induction a.
+ cbn. destruct_match; intros.
+ eapply evaluable_impl.
+ eauto using predicated_evaluable.
+ unfold all_evaluable2 in H0. unfold p_evaluable.
+ eapply parts_evaluable. eapply H0. econstructor.
+
+ intros. cbn. destruct_match.
+ eapply p_evaluable_Pand.
+ eapply all_evaluable2_cons2 in H0.
+ eapply evaluable_impl.
+ eauto using predicated_evaluable.
+ unfold all_evaluable2 in H0. unfold p_evaluable.
+ eapply parts_evaluable. eapply H0.
+ eapply all_evaluable2_cons in H0. eauto.
+ Qed.
+
+ Lemma p_evaluable_imp :
+ forall A (ctx: @ctx A) a b,
+ sem_pexpr ctx a false ->
+ p_evaluable ctx (a → b).
+ Proof.
+ intros. unfold "→".
+ unfold p_evaluable, evaluable. exists true.
+ constructor. replace true with (negb false) by trivial. left.
+ now apply sem_pexpr_negate.
+ Qed.
diff --git a/src/hls/GiblePargenproofForward.v b/src/hls/GiblePargenproofForward.v
new file mode 100644
index 0000000..8222513
--- /dev/null
+++ b/src/hls/GiblePargenproofForward.v
@@ -0,0 +1,1568 @@
+(*
+ * Vericert: Verified high-level synthesis.
+ * Copyright (C) 2023 Yann Herklotz <git@yannherklotz.com>
+ *
+ * This program is free software: you can redistribute it and/or modify
+ * it under the terms of the GNU General Public License as published by
+ * the Free Software Foundation, either version 3 of the License, or
+ * (at your option) any later version.
+ *
+ * This program is distributed in the hope that it will be useful,
+ * but WITHOUT ANY WARRANTY; without even the implied warranty of
+ * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+ * GNU General Public License for more details.
+ *
+ * You should have received a copy of the GNU General Public License
+ * along with this program. If not, see <https://www.gnu.org/licenses/>.
+ *)
+
+Require Import compcert.backend.Registers.
+Require Import compcert.common.AST.
+Require Import compcert.common.Errors.
+Require Import compcert.common.Linking.
+Require Import compcert.common.Globalenvs.
+Require Import compcert.common.Memory.
+Require Import compcert.common.Values.
+Require Import compcert.lib.Maps.
+
+Require Import vericert.common.Vericertlib.
+Require Import vericert.hls.GibleSeq.
+Require Import vericert.hls.GiblePar.
+Require Import vericert.hls.Gible.
+Require Import vericert.hls.GiblePargenproofEquiv.
+Require Import vericert.hls.GiblePargen.
+Require Import vericert.hls.Predicate.
+Require Import vericert.hls.Abstr.
+Require Import vericert.common.Monad.
+
+Module Import OptionExtra := MonadExtra(Option).
+
+#[local] Open Scope positive.
+#[local] Open Scope forest.
+#[local] Open Scope pred_op.
+
+#[local] Opaque simplify.
+#[local] Opaque deep_simplify.
+
+#[local] Ltac destr := destruct_match; try discriminate; [].
+
+Definition state_lessdef := GiblePargenproofEquiv.match_states.
+
+Definition is_regs i := match i with mk_instr_state rs _ _ => rs end.
+Definition is_mem i := match i with mk_instr_state _ _ m => m end.
+Definition is_ps i := match i with mk_instr_state _ p _ => p end.
+
+Definition check_dest i r' :=
+ match i with
+ | RBop p op rl r => (r =? r')%positive
+ | RBload p chunk addr rl r => (r =? r')%positive
+ | _ => false
+ end.
+
+Lemma check_dest_dec i r :
+ {check_dest i r = true} + {check_dest i r = false}.
+Proof. destruct (check_dest i r); tauto. Qed.
+
+Fixpoint check_dest_l l r :=
+ match l with
+ | nil => false
+ | a :: b => check_dest a r || check_dest_l b r
+ end.
+
+Lemma check_dest_l_forall :
+ forall l r,
+ check_dest_l l r = false ->
+ Forall (fun x => check_dest x r = false) l.
+Proof. induction l; crush. Qed.
+
+Lemma check_dest_l_dec i r :
+ {check_dest_l i r = true} + {check_dest_l i r = false}.
+Proof. destruct (check_dest_l i r); tauto. Qed.
+
+Lemma match_states_list :
+ forall A (rs: Regmap.t A) rs',
+ (forall r, rs !! r = rs' !! r) ->
+ forall l, rs ## l = rs' ## l.
+Proof. induction l; crush. Qed.
+
+Lemma PTree_matches :
+ forall A (v: A) res rs rs',
+ (forall r, rs !! r = rs' !! r) ->
+ forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x.
+Proof.
+ intros; destruct (Pos.eq_dec x res); subst;
+ [ repeat rewrite Regmap.gss by auto
+ | repeat rewrite Regmap.gso by auto ]; auto.
+Qed.
+
+Section CORRECTNESS.
+
+ Context (prog: GibleSeq.program) (tprog : GiblePar.program).
+
+ Let ge : GibleSeq.genv := Globalenvs.Genv.globalenv prog.
+ Let tge : GiblePar.genv := Globalenvs.Genv.globalenv tprog.
+
+ Lemma eval_predf_negate :
+ forall ps p,
+ eval_predf ps (negate p) = negb (eval_predf ps p).
+ Proof.
+ unfold eval_predf; intros. rewrite negate_correct. auto.
+ Qed.
+
+ Lemma is_truthy_negate :
+ forall ps p pred,
+ truthy ps p ->
+ falsy ps (combine_pred (Some (negate (Option.default T p))) pred).
+ Proof.
+ inversion 1; subst; simplify.
+ - destruct pred; constructor; auto.
+ - destruct pred; constructor.
+ rewrite eval_predf_Pand. rewrite eval_predf_negate. rewrite H0. auto.
+ rewrite eval_predf_negate. rewrite H0. auto.
+ Qed.
+
+ Lemma sem_pred_expr_NEapp :
+ forall A B C sem f_p ctx a b v,
+ sem_pred_expr f_p sem ctx a v ->
+ @sem_pred_expr A B C f_p sem ctx (NE.app a b) v.
+ Proof.
+ induction a; crush.
+ - inv H. constructor; auto.
+ - inv H. constructor; eauto.
+ eapply sem_pred_expr_cons_false; eauto.
+ Qed.
+
+ Lemma sem_pred_expr_NEapp2 :
+ forall A B C sem f_p ctx a b v ps,
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) (ps !! x)) ->
+ (forall x, NE.In x a -> eval_predf ps (fst x) = false) ->
+ sem_pred_expr f_p sem ctx b v ->
+ @sem_pred_expr A B C f_p sem ctx (NE.app a b) v.
+ Proof.
+ induction a; crush.
+ - assert (IN: NE.In a (NE.singleton a)) by constructor.
+ specialize (H0 _ IN). destruct a.
+ eapply sem_pred_expr_cons_false; eauto. cbn [fst] in *.
+ eapply sem_pexpr_eval; eauto.
+ - assert (NE.In a (NE.cons a a0)) by (constructor; tauto).
+ apply H0 in H2.
+ destruct a. cbn [fst] in *.
+ eapply sem_pred_expr_cons_false.
+ eapply sem_pexpr_eval; eauto. eapply IHa; eauto.
+ intros. eapply H0. constructor; tauto.
+ Qed.
+
+ Lemma sem_pred_expr_NEapp3 :
+ forall A B C sem f_p ctx (a b: predicated B) v,
+ (forall x, NE.In x a -> sem_pexpr ctx (from_pred_op f_p (fst x)) false) ->
+ sem_pred_expr f_p sem ctx b v ->
+ @sem_pred_expr A B C f_p sem ctx (NE.app a b) v.
+ Proof.
+ induction a; crush.
+ - assert (IN: NE.In a (NE.singleton a)) by constructor.
+ specialize (H _ IN). destruct a.
+ eapply sem_pred_expr_cons_false; eauto.
+ - assert (NE.In a (NE.cons a a0)) by (constructor; tauto).
+ apply H in H1.
+ destruct a. cbn [fst] in *.
+ eapply sem_pred_expr_cons_false; auto.
+ eapply IHa; eauto.
+ intros. eapply H. constructor; tauto.
+ Qed.
+
+ Lemma sem_pred_expr_map_not :
+ forall A p ps y x0,
+ eval_predf ps p = false ->
+ NE.In x0 (NE.map (fun x => (p ∧ fst x, snd x)) y) ->
+ eval_predf ps (@fst _ A x0) = false.
+ Proof.
+ induction y; crush.
+ - inv H0. simplify. rewrite eval_predf_Pand. rewrite H. auto.
+ - inv H0. inv H2. cbn -[eval_predf]. rewrite eval_predf_Pand.
+ rewrite H. auto. eauto.
+ Qed.
+
+ Lemma sem_pred_expr_map_Pand :
+ forall A B C sem ctx f_p ps x v p,
+ (forall x : positive, sem_pexpr ctx (get_forest_p' x f_p) ps !! x) ->
+ eval_predf ps p = true ->
+ sem_pred_expr f_p sem ctx x v ->
+ @sem_pred_expr A B C f_p sem ctx
+ (NE.map (fun x0 => (p ∧ fst x0, snd x0)) x) v.
+ Proof.
+ induction x; crush.
+ inv H1. simplify. constructor; eauto.
+ simplify. replace true with (true && true) by auto.
+ constructor; auto.
+ eapply sem_pexpr_eval; eauto.
+ inv H1. simplify. constructor; eauto.
+ simplify. replace true with (true && true) by auto.
+ constructor; auto.
+ eapply sem_pexpr_eval; eauto.
+ eapply sem_pred_expr_cons_false. cbn.
+ replace false with (true && false) by auto. apply sem_pexpr_Pand; auto.
+ eapply sem_pexpr_eval; eauto. eauto.
+ Qed.
+
+ Lemma sem_pred_expr_app_predicated :
+ forall A B C sem ctx f_p y x v p ps,
+ sem_pred_expr f_p sem ctx x v ->
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) (ps !! x)) ->
+ eval_predf ps p = true ->
+ @sem_pred_expr A B C f_p sem ctx (app_predicated p y x) v.
+ Proof.
+ intros * SEM PS EVAL. unfold app_predicated.
+ eapply sem_pred_expr_NEapp2; eauto.
+ intros. eapply sem_pred_expr_map_not; eauto.
+ rewrite eval_predf_negate. rewrite EVAL. auto.
+ eapply sem_pred_expr_map_Pand; eauto.
+ Qed.
+
+ Lemma sem_pred_expr_app_predicated_false :
+ forall A B C sem ctx f_p y x v p ps,
+ sem_pred_expr f_p sem ctx y v ->
+ (forall x, sem_pexpr ctx (get_forest_p' x f_p) (ps !! x)) ->
+ eval_predf ps p = false ->
+ @sem_pred_expr A B C f_p sem ctx (app_predicated p y x) v.
+ Admitted.
+
+ Lemma sem_pred_expr_prune_predicated :
+ forall A B C sem ctx f_p y x v,
+ sem_pred_expr f_p sem ctx x v ->
+ prune_predicated x = Some y ->
+ @sem_pred_expr A B C f_p sem ctx y v.
+ Proof.
+ intros * SEM PRUNE. unfold prune_predicated in *.
+ Admitted.
+
+ Inductive sem_ident {B A: Type} (c: @ctx B) (a: A): A -> Prop :=
+ | sem_ident_intro : sem_ident c a a.
+
+ Lemma sem_pred_expr_pred_ret :
+ forall G A (ctx: @Abstr.ctx G) (i: A) ps,
+ sem_pred_expr ps sem_ident ctx (pred_ret i) i.
+ Proof. intros; constructor; constructor. Qed.
+
+ Lemma sem_pred_expr_ident :
+ forall G A B ps (ctx: @Abstr.ctx G) (l: predicated A) (s: @Abstr.ctx G -> A -> B -> Prop) l_,
+ sem_pred_expr ps sem_ident ctx l l_ ->
+ forall (v: B),
+ s ctx l_ v ->
+ sem_pred_expr ps s ctx l v.
+ Proof.
+ induction 1.
+ - intros. constructor; auto. inv H0. auto.
+ - intros. apply sem_pred_expr_cons_false; auto.
+ - intros. inv H0. constructor; auto.
+ Qed.
+
+ Lemma sem_pred_expr_ident2 :
+ forall G A B ps (ctx: @Abstr.ctx G) (l: predicated A) (s: @Abstr.ctx G -> A -> B -> Prop) (v: B),
+ sem_pred_expr ps s ctx l v ->
+ exists l_, sem_pred_expr ps sem_ident ctx l l_ /\ s ctx l_ v.
+ Proof.
+ induction 1.
+ - exists e; split; auto. constructor; auto. constructor.
+ - inversion_clear IHsem_pred_expr as [x IH].
+ inversion_clear IH as [SEM EVALS].
+ exists x; split; auto. apply sem_pred_expr_cons_false; auto.
+ - exists e; split; auto; constructor; auto; constructor.
+ Qed.
+
+ Fixpoint of_elist (e: expression_list): list expression :=
+ match e with
+ | Econs a b => a :: of_elist b
+ | Enil => nil
+ end.
+
+ Fixpoint to_elist (e: list expression): expression_list :=
+ match e with
+ | a :: b => Econs a (to_elist b)
+ | nil => Enil
+ end.
+
+ Lemma sem_val_list_elist :
+ forall G (ctx: @Abstr.ctx G) l j,
+ sem_val_list ctx (to_elist l) j ->
+ Forall2 (sem_value ctx) l j.
+ Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
+
+ Lemma sem_val_list_elist2 :
+ forall G (ctx: @Abstr.ctx G) l j,
+ Forall2 (sem_value ctx) l j ->
+ sem_val_list ctx (to_elist l) j.
+ Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
+
+ Lemma sem_val_list_elist3 :
+ forall G (ctx: @Abstr.ctx G) l j,
+ sem_val_list ctx l j ->
+ Forall2 (sem_value ctx) (of_elist l) j.
+ Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
+
+ Lemma sem_val_list_elist4 :
+ forall G (ctx: @Abstr.ctx G) l j,
+ Forall2 (sem_value ctx) (of_elist l) j ->
+ sem_val_list ctx l j.
+ Proof. induction l; intros; cbn in *; inversion H; constructor; eauto. Qed.
+
+ Lemma sem_pred_expr_predicated_map :
+ forall A B C pr (f: C -> B) ctx (pred: predicated C) (pred': C),
+ sem_pred_expr pr sem_ident ctx pred pred' ->
+ @sem_pred_expr A _ _ pr sem_ident ctx (predicated_map f pred) (f pred').
+ Proof.
+ induction pred; unfold predicated_map; intros.
+ - inv H. inv H3. constructor; eauto. constructor.
+ - inv H. inv H5. constructor; eauto. constructor.
+ eapply sem_pred_expr_cons_false; eauto.
+ Qed.
+
+ Lemma NEapp_NEmap :
+ forall A B (f: A -> B) a b,
+ NE.map f (NE.app a b) = NE.app (NE.map f a) (NE.map f b).
+ Proof. induction a; crush. Qed.
+
+ Lemma sem_pred_expr_predicated_prod :
+ forall A B C pr ctx (a: predicated C) (b: predicated B) a' b',
+ sem_pred_expr pr sem_ident ctx a a' ->
+ sem_pred_expr pr sem_ident ctx b b' ->
+ @sem_pred_expr A _ _ pr sem_ident ctx (predicated_prod a b) (a', b').
+ Proof.
+ induction a; intros.
+ - inv H. inv H4. unfold predicated_prod.
+ generalize dependent b. induction b; intros.
+ + cbn. destruct_match. inv Heqp. inv H0. inv H6.
+ constructor. simplify. replace true with (true && true) by auto.
+ eapply sem_pexpr_Pand; eauto. constructor.
+ + inv H0. inv H6. cbn. constructor; cbn.
+ replace true with (true && true) by auto.
+ constructor; auto. constructor.
+ eapply sem_pred_expr_cons_false; eauto. cbn.
+ replace false with (true && false) by auto.
+ apply sem_pexpr_Pand; auto.
+ - unfold predicated_prod. simplify.
+ rewrite NEapp_NEmap.
+ inv H. eapply sem_pred_expr_NEapp.
+ { induction b; crush.
+ destruct a. inv H0. constructor.
+ replace true with (true && true) by auto.
+ eapply sem_pexpr_Pand; auto. inv H6. inv H7.
+ constructor.
+
+ destruct a. inv H0. constructor.
+ replace true with (true && true) by auto.
+ constructor; auto. inv H8. inv H6. constructor.
+
+ specialize (IHb H8). eapply sem_pred_expr_cons_false; auto.
+ replace false with (true && false) by auto.
+ apply sem_pexpr_Pand; auto.
+ }
+ { exploit IHa. eauto. eauto. intros.
+ unfold predicated_prod in *.
+ eapply sem_pred_expr_NEapp3; eauto; [].
+ clear H. clear H0.
+ induction b.
+ { intros. inv H. destruct a. simpl.
+ constructor; tauto. }
+ { intros. inv H. inv H1. destruct a. simpl.
+ constructor; tauto. eauto. } }
+ Qed.
+
+ Lemma sem_pred_expr_seq_app :
+ forall G A B (f: predicated (A -> B))
+ ps (ctx: @ctx G) l f_ l_,
+ sem_pred_expr ps sem_ident ctx l l_ ->
+ sem_pred_expr ps sem_ident ctx f f_ ->
+ sem_pred_expr ps sem_ident ctx (seq_app f l) (f_ l_).
+ Proof.
+ unfold seq_app; intros.
+ remember (fun x : (A -> B) * A => fst x (snd x)) as app.
+ replace (f_ l_) with (app (f_, l_)) by (rewrite Heqapp; auto).
+ eapply sem_pred_expr_predicated_map. eapply sem_pred_expr_predicated_prod; auto.
+ Qed.
+
+ Lemma sem_pred_expr_map :
+ forall A B C (ctx: @ctx A) ps (f: B -> C) y v,
+ sem_pred_expr ps sem_ident ctx y v ->
+ sem_pred_expr ps sem_ident ctx (NE.map (fun x => (fst x, f (snd x))) y) (f v).
+ Proof.
+ induction y; crush. inv H. constructor; crush. inv H3. constructor.
+ inv H. inv H5. constructor; eauto. constructor.
+ eapply sem_pred_expr_cons_false; eauto.
+ Qed.
+
+ Lemma sem_pred_expr_flap :
+ forall G A B C (f: predicated (A -> B -> C))
+ ps (ctx: @ctx G) l f_,
+ sem_pred_expr ps sem_ident ctx f f_ ->
+ sem_pred_expr ps sem_ident ctx (flap f l) (f_ l).
+ Proof.
+ induction f. unfold flap2; intros. inv H. inv H3.
+ constructor; eauto. constructor.
+ intros. inv H. cbn.
+ constructor; eauto. inv H5. constructor.
+ eapply sem_pred_expr_cons_false; eauto.
+ Qed.
+
+ Lemma sem_pred_expr_flap2 :
+ forall G A B C (f: predicated (A -> B -> C))
+ ps (ctx: @ctx G) l1 l2 f_,
+ sem_pred_expr ps sem_ident ctx f f_ ->
+ sem_pred_expr ps sem_ident ctx (flap2 f l1 l2) (f_ l1 l2).
+ Proof.
+ induction f. unfold flap2; intros. inv H. inv H3.
+ constructor; eauto. constructor.
+ intros. inv H. cbn.
+ constructor; eauto. inv H5. constructor.
+ eapply sem_pred_expr_cons_false; eauto.
+ Qed.
+
+ Lemma sem_pred_expr_seq_app_val :
+ forall G A B V (s: @Abstr.ctx G -> B -> V -> Prop)
+ (f: predicated (A -> B))
+ ps ctx l v f_ l_,
+ sem_pred_expr ps sem_ident ctx l l_ ->
+ sem_pred_expr ps sem_ident ctx f f_ ->
+ s ctx (f_ l_) v ->
+ sem_pred_expr ps s ctx (seq_app f l) v.
+ Proof.
+ intros. eapply sem_pred_expr_ident; [|eassumption].
+ eapply sem_pred_expr_seq_app; eauto.
+ Qed.
+
+ Fixpoint Eapp a b :=
+ match a with
+ | Enil => b
+ | Econs ax axs => Econs ax (Eapp axs b)
+ end.
+
+ Lemma Eapp_right_nil :
+ forall a, Eapp a Enil = a.
+ Proof. induction a; cbn; try rewrite IHa; auto. Qed.
+
+ Lemma Eapp_left_nil :
+ forall a, Eapp Enil a = a.
+ Proof. auto. Qed.
+
+ Lemma sem_pred_expr_cons_pred_expr :
+ forall A (ctx: @ctx A) pr s s' a e,
+ sem_pred_expr pr sem_ident ctx s s' ->
+ sem_pred_expr pr sem_ident ctx a e ->
+ sem_pred_expr pr sem_ident ctx (cons_pred_expr a s) (Econs e s').
+ Proof.
+ intros. unfold cons_pred_expr.
+ replace (Econs e s') with ((uncurry Econs) (e, s')) by auto.
+ eapply sem_pred_expr_predicated_map.
+ eapply sem_pred_expr_predicated_prod; eauto.
+ Qed.
+
+ Lemma sem_pred_expr_fold_right :
+ forall A pr ctx s a a' s',
+ sem_pred_expr pr sem_ident ctx s s' ->
+ Forall2 (sem_pred_expr pr sem_ident ctx) (NE.to_list a) (of_elist a') ->
+ @sem_pred_expr A _ _ pr sem_ident ctx (NE.fold_right cons_pred_expr s a) (Eapp a' s').
+ Proof.
+ induction a; crush. inv H0. inv H5. destruct a'; crush. destruct a'; crush.
+ inv H3. eapply sem_pred_expr_cons_pred_expr; eauto.
+ inv H0. destruct a'; crush. inv H3.
+ eapply sem_pred_expr_cons_pred_expr; eauto.
+ Qed.
+
+ Lemma sem_pred_expr_fold_right2 :
+ forall A pr ctx s a a' s',
+ sem_pred_expr pr sem_ident ctx s s' ->
+ @sem_pred_expr A _ _ pr sem_ident ctx (NE.fold_right cons_pred_expr s a) (Eapp a' s') ->
+ Forall2 (sem_pred_expr pr sem_ident ctx) (NE.to_list a) (of_elist a').
+ Proof.
+ induction a. Admitted.
+
+ Lemma NEof_list_some :
+ forall A a a' (e: A),
+ NE.of_list a = Some a' ->
+ NE.of_list (e :: a) = Some (NE.cons e a').
+ Proof.
+ induction a; [crush|].
+ intros.
+ cbn in H. destruct a0. inv H. auto.
+ destruct_match; [|discriminate].
+ inv H. specialize (IHa n a ltac:(trivial)).
+ cbn. destruct_match. unfold NE.of_list in IHa.
+ fold (@NE.of_list A) in IHa. rewrite IHa in Heqo0. inv Heqo0. auto.
+ unfold NE.of_list in IHa. fold (@NE.of_list A) in IHa. rewrite IHa in Heqo0. inv Heqo0.
+ Qed.
+
+ Lemma NEof_list_contra :
+ forall A b (a: A),
+ ~ NE.of_list (a :: b) = None.
+ Proof.
+ induction b; try solve [crush].
+ intros.
+ specialize (IHb a).
+ enough (X: exists x, NE.of_list (a :: b) = Some x).
+ inversion_clear X as [x X'].
+ erewrite NEof_list_some; eauto; discriminate.
+ destruct (NE.of_list (a :: b)) eqn:?; [eauto|contradiction].
+ Qed.
+
+ Lemma NEof_list_exists :
+ forall A b (a: A),
+ exists x, NE.of_list (a :: b) = Some x.
+ Proof.
+ intros. pose proof (NEof_list_contra _ b a).
+ destruct (NE.of_list (a :: b)); try contradiction.
+ eauto.
+ Qed.
+
+ Lemma NEof_list_exists2 :
+ forall A b (a c: A),
+ exists x,
+ NE.of_list (c :: a :: b) = Some (NE.cons c x)
+ /\ NE.of_list (a :: b) = Some x.
+ Proof.
+ intros. pose proof (NEof_list_exists _ b a).
+ inversion_clear H as [x B].
+ econstructor; split; eauto.
+ eapply NEof_list_some; eauto.
+ Qed.
+
+ Lemma NEof_list_to_list :
+ forall A (x: list A) y,
+ NE.of_list x = Some y ->
+ NE.to_list y = x.
+ Proof.
+ induction x; [crush|].
+ intros. destruct x; [crush|].
+ pose proof (NEof_list_exists2 _ x a0 a).
+ inversion_clear H0 as [x0 [HN1 HN2]]. rewrite HN1 in H. inv H.
+ cbn. f_equal. eauto.
+ Qed.
+
+ Lemma sem_pred_expr_merge :
+ forall G (ctx: @Abstr.ctx G) ps l args,
+ Forall2 (sem_pred_expr ps sem_ident ctx) args l ->
+ sem_pred_expr ps sem_ident ctx (merge args) (to_elist l).
+ Proof.
+ induction l; intros.
+ - inv H. cbn; repeat constructor.
+ - inv H. cbn. unfold merge. Admitted.
+
+ Lemma sem_pred_expr_merge2 :
+ forall A (ctx: @ctx A) pr l l',
+ sem_pred_expr pr sem_ident ctx (merge l) l' ->
+ Forall2 (sem_pred_expr pr sem_ident ctx) l (of_elist l').
+ Proof.
+ induction l; crush.
+ - unfold merge in *; cbn in *.
+ inv H. inv H5. constructor.
+ - unfold merge in H.
+ pose proof (NEof_list_exists _ l a) as Y.
+ inversion_clear Y as [x HNE]. rewrite HNE in H.
+ erewrite <- (NEof_list_to_list _ (a :: l)) by eassumption.
+ apply sem_pred_expr_fold_right2 with (s := (NE.singleton (T, Enil))) (s' := Enil).
+ repeat constructor.
+ rewrite Eapp_right_nil. auto.
+ Qed.
+
+ Lemma sem_merge_list :
+ forall A ctx f rs ps m args,
+ sem ctx f ((mk_instr_state rs ps m), None) ->
+ @sem_pred_expr A _ _ (forest_preds f) sem_val_list ctx
+ (merge (list_translation args f)) (rs ## args).
+ Proof.
+ induction args; crush. cbn. constructor; constructor.
+ unfold merge. specialize (IHargs H).
+ eapply sem_pred_expr_ident2 in IHargs.
+ inversion_clear IHargs as [l_ [HSEM HVAL]].
+ destruct_match; [|exfalso; eapply NEof_list_contra; eauto].
+ destruct args; cbn -[NE.of_list] in *.
+ - unfold merge in *. simplify.
+ inv H. inv H6. specialize (H a).
+ eapply sem_pred_expr_ident2 in H.
+ inversion_clear H as [l_2 [HSEM2 HVAL2]].
+ unfold cons_pred_expr.
+ eapply sem_pred_expr_ident.
+ eapply sem_pred_expr_predicated_map.
+ eapply sem_pred_expr_predicated_prod; [eassumption|repeat constructor].
+ repeat constructor; auto.
+ - pose proof (NEof_list_exists2 _ (list_translation args f) (f #r (Reg r)) (f #r (Reg a))) as Y.
+ inversion_clear Y as [x [Y1 Y2]]. rewrite Heqo in Y1. inv Y1.
+ inversion_clear H as [? ? ? ? ? ? REG PRED MEM EXIT].
+ inversion_clear REG as [? ? ? REG'].
+ inversion_clear PRED as [? ? ? PRED'].
+ pose proof (REG' a) as REGa. pose proof (REG' r) as REGr.
+ exploit sem_pred_expr_ident2; [exact REGa|].
+ intro REGI'; inversion_clear REGI' as [a' [SEMa SEMa']].
+ exploit sem_pred_expr_ident2; [exact REGr|].
+ intro REGI'; inversion_clear REGI' as [r' [SEMr SEMr']].
+ apply sem_pred_expr_ident with (l_ := Econs a' l_); [|constructor; auto].
+ replace (Econs a' l_) with (Eapp (Econs a' l_) Enil) by (now apply Eapp_right_nil).
+ apply sem_pred_expr_fold_right; eauto.
+ repeat constructor.
+ constructor; eauto.
+ erewrite NEof_list_to_list; eauto.
+ eapply sem_pred_expr_merge2; auto.
+ Qed.
+
+ Lemma sem_pred_expr_list_translation :
+ forall G ctx args f i,
+ @sem G ctx f (i, None) ->
+ exists l,
+ Forall2 (sem_pred_expr (forest_preds f) sem_ident ctx) (list_translation args f) l
+ /\ sem_val_list ctx (to_elist l) ((is_rs i)##args).
+ Proof.
+ induction args; intros.
+ - exists nil; cbn; split; auto; constructor.
+ - exploit IHargs; try eassumption; intro EX.
+ inversion_clear EX as [l [FOR SEM]].
+ destruct i as [rs' m' ps'].
+ inversion_clear H as [? ? ? ? ? ? REGSET PREDSET MEM EXIT].
+ inversion_clear REGSET as [? ? ? REG].
+ pose proof (REG a).
+ exploit sem_pred_expr_ident2; eauto; intro IDENT.
+ inversion_clear IDENT as [l_ [SEMP SV]].
+ exists (l_ :: l). split. constructor; auto.
+ cbn; constructor; auto.
+ Qed.
+
+(*|
+Here we can finally assume that everything in the forest is evaluable, which
+will allow us to prove that translating the list of register accesses will also
+all be evaluable.
+|*)
+
+ Lemma sem_update_Iop :
+ forall A op rs args m v f ps ctx,
+ Op.eval_operation (ctx_ge ctx) (ctx_sp ctx) op rs ## args (is_mem (ctx_is ctx)) = Some v ->
+ sem ctx f ((mk_instr_state rs ps m), None) ->
+ @sem_pred_expr A _ _ (forest_preds f) sem_value ctx
+ (seq_app (pred_ret (Eop op)) (merge (list_translation args f))) v.
+ Proof.
+ intros * EVAL SEM.
+ exploit sem_pred_expr_list_translation; try eassumption.
+ intro H; inversion_clear H as [x [HS HV]].
+ eapply sem_pred_expr_seq_app_val.
+ - cbn in *; eapply sem_pred_expr_merge; eauto.
+ - apply sem_pred_expr_pred_ret.
+ - econstructor; [eauto|]; auto.
+ Qed.
+
+ Lemma sem_update_Iload :
+ forall A rs args m v f ps ctx addr a0 chunk,
+ Op.eval_addressing (ctx_ge ctx) (ctx_sp ctx) addr rs ## args = Some a0 ->
+ Mem.loadv chunk m a0 = Some v ->
+ sem ctx f ((mk_instr_state rs ps m), None) ->
+ @sem_pred_expr A _ _ (forest_preds f) sem_value ctx
+ (seq_app (seq_app (pred_ret (Eload chunk addr)) (merge (list_translation args f))) (f #r Mem)) v.
+ Proof.
+ intros * EVAL MEM SEM.
+ exploit sem_pred_expr_list_translation; try eassumption.
+ intro H; inversion_clear H as [x [HS HV]].
+ inversion SEM as [? ? ? ? ? ? REG PRED HMEM EXIT]; subst.
+ exploit sem_pred_expr_ident2; [eapply HMEM|].
+ intros H; inversion_clear H as [x' [HS' HV']].
+ eapply sem_pred_expr_seq_app_val; eauto.
+ { eapply sem_pred_expr_seq_app; eauto.
+ - eapply sem_pred_expr_merge; eauto.
+ - apply sem_pred_expr_pred_ret.
+ }
+ econstructor; eauto.
+ Qed.
+
+ Lemma storev_valid_pointer1 :
+ forall chunk m m' s d b z,
+ Mem.storev chunk m s d = Some m' ->
+ Mem.valid_pointer m b z = true ->
+ Mem.valid_pointer m' b z = true.
+ Proof using.
+ intros. unfold Mem.storev in *. destruct_match; try discriminate; subst.
+ apply Mem.valid_pointer_nonempty_perm. apply Mem.valid_pointer_nonempty_perm in H0.
+ eapply Mem.perm_store_1; eauto.
+ Qed.
+
+ Lemma storev_valid_pointer2 :
+ forall chunk m m' s d b z,
+ Mem.storev chunk m s d = Some m' ->
+ Mem.valid_pointer m' b z = true ->
+ Mem.valid_pointer m b z = true.
+ Proof using.
+ intros. unfold Mem.storev in *. destruct_match; try discriminate; subst.
+ apply Mem.valid_pointer_nonempty_perm. apply Mem.valid_pointer_nonempty_perm in H0.
+ eapply Mem.perm_store_2; eauto.
+ Qed.
+
+ Definition valid_mem m m' :=
+ forall b z, Mem.valid_pointer m b z = Mem.valid_pointer m' b z.
+
+ #[global] Program Instance valid_mem_Equivalence : Equivalence valid_mem.
+ Next Obligation. unfold valid_mem; auto. Qed. (* Reflexivity *)
+ Next Obligation. unfold valid_mem; auto. Qed. (* Symmetry *)
+ Next Obligation. unfold valid_mem; eauto. Qed. (* Transitity *)
+
+ Lemma storev_valid_pointer :
+ forall chunk m m' s d,
+ Mem.storev chunk m s d = Some m' ->
+ valid_mem m m'.
+ Proof using.
+ unfold valid_mem. symmetry.
+ intros. destruct (Mem.valid_pointer m b z) eqn:?;
+ eauto using storev_valid_pointer1.
+ apply not_true_iff_false.
+ apply not_true_iff_false in Heqb0.
+ eauto using storev_valid_pointer2.
+ Qed.
+
+ Lemma storev_cmpu_bool :
+ forall m m' c v v0,
+ valid_mem m m' ->
+ Val.cmpu_bool (Mem.valid_pointer m) c v v0 = Val.cmpu_bool (Mem.valid_pointer m') c v v0.
+ Proof using.
+ unfold valid_mem.
+ intros. destruct v, v0; crush.
+ { destruct_match; crush.
+ destruct_match; crush.
+ destruct_match; crush.
+ apply orb_true_iff in H1.
+ inv H1. rewrite H in H2. rewrite H2 in Heqb1.
+ simplify. rewrite H0 in Heqb1. crush.
+ rewrite H in H2. rewrite H2 in Heqb1.
+ rewrite H0 in Heqb1. crush.
+ destruct_match; auto. simplify.
+ apply orb_true_iff in H1.
+ inv H1. rewrite <- H in H2. rewrite H2 in Heqb1.
+ simplify. rewrite H0 in Heqb1. crush.
+ rewrite <- H in H2. rewrite H2 in Heqb1.
+ rewrite H0 in Heqb1. crush. }
+
+ { destruct_match; crush.
+ destruct_match; crush.
+ destruct_match; crush.
+ apply orb_true_iff in H1.
+ inv H1. rewrite H in H2. rewrite H2 in Heqb1.
+ simplify. rewrite H0 in Heqb1. crush.
+ rewrite H in H2. rewrite H2 in Heqb1.
+ rewrite H0 in Heqb1. crush.
+ destruct_match; auto. simplify.
+ apply orb_true_iff in H1.
+ inv H1. rewrite <- H in H2. rewrite H2 in Heqb1.
+ simplify. rewrite H0 in Heqb1. crush.
+ rewrite <- H in H2. rewrite H2 in Heqb1.
+ rewrite H0 in Heqb1. crush. }
+
+ { destruct_match; auto. destruct_match; auto; crush.
+ { destruct_match; crush.
+ { destruct_match; crush.
+ setoid_rewrite H in H0; eauto.
+ setoid_rewrite H in H1; eauto.
+ rewrite H0 in Heqb. rewrite H1 in Heqb; crush.
+ }
+ { destruct_match; crush.
+ setoid_rewrite H in Heqb0; eauto.
+ rewrite H0 in Heqb0. rewrite H1 in Heqb0; crush. } }
+ { destruct_match; crush.
+ { destruct_match; crush.
+ setoid_rewrite H in H0; eauto.
+ setoid_rewrite H in H1; eauto.
+ rewrite H0 in Heqb0. rewrite H1 in Heqb0; crush.
+ }
+ { destruct_match; crush.
+ setoid_rewrite H in Heqb0; eauto.
+ rewrite H0 in Heqb0. rewrite H1 in Heqb0; crush. } } }
+ Qed.
+
+ Lemma storev_eval_condition :
+ forall m m' c rs,
+ valid_mem m m' ->
+ Op.eval_condition c rs m = Op.eval_condition c rs m'.
+ Proof using.
+ intros. destruct c; crush.
+ destruct rs; crush.
+ destruct rs; crush.
+ destruct rs; crush.
+ eapply storev_cmpu_bool; eauto.
+ destruct rs; crush.
+ destruct rs; crush.
+ eapply storev_cmpu_bool; eauto.
+ Qed.
+
+ Lemma eval_operation_valid_pointer :
+ forall A B (ge: Genv.t A B) sp op args m m' v,
+ valid_mem m m' ->
+ Op.eval_operation ge sp op args m' = Some v ->
+ Op.eval_operation ge sp op args m = Some v.
+ Proof.
+ intros * VALID OP. destruct op; auto.
+ - destruct cond; auto; cbn in *.
+ + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
+ + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
+ - destruct c; auto; cbn in *.
+ + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
+ + repeat destruct_match; auto. setoid_rewrite <- storev_cmpu_bool in OP; eauto.
+ Qed.
+
+ Lemma bb_memory_consistency_eval_operation :
+ forall A B (ge: Genv.t A B) sp state i state' args op v,
+ step_instr ge sp (Iexec state) i (Iexec state') ->
+ Op.eval_operation ge sp op args (is_mem state) = Some v ->
+ Op.eval_operation ge sp op args (is_mem state') = Some v.
+ Proof.
+ inversion_clear 1; intro EVAL; auto.
+ cbn in *.
+ eapply eval_operation_valid_pointer. unfold valid_mem. symmetry.
+ eapply storev_valid_pointer; eauto.
+ auto.
+ Qed.
+
+ Lemma truthy_dflt :
+ forall ps p,
+ truthy ps p -> eval_predf ps (dfltp p) = true.
+ Proof. intros. destruct p; cbn; inv H; auto. Qed.
+
+ Lemma sem_update_Istore :
+ forall A rs args m v f ps ctx addr a0 chunk m' v',
+ Op.eval_addressing (ctx_ge ctx) (ctx_sp ctx) addr rs ## args = Some a0 ->
+ Mem.storev chunk m a0 v' = Some m' ->
+ sem_value ctx v v' ->
+ sem ctx f ((mk_instr_state rs ps m), None) ->
+ @sem_pred_expr A _ _ (forest_preds f) sem_mem ctx
+ (seq_app (seq_app (pred_ret (Estore v chunk addr))
+ (merge (list_translation args f))) (f #r Mem)) m'.
+ Proof.
+ intros * EVAL STOREV SEMVAL SEM.
+ exploit sem_merge_list; try eassumption.
+ intro MERGE. exploit sem_pred_expr_ident2; eauto.
+ intro TMP; inversion_clear TMP as [x [HS HV]].
+ inversion_clear SEM as [? ? ? ? ? ? REG PRED HMEM EXIT].
+ exploit sem_pred_expr_ident2; [eapply HMEM|].
+ intros TMP; inversion_clear TMP as [x' [HS' HV']].
+ eapply sem_pred_expr_ident.
+ eapply sem_pred_expr_seq_app; eauto.
+ eapply sem_pred_expr_seq_app; eauto.
+ eapply sem_pred_expr_pred_ret.
+ econstructor; eauto.
+ Qed.
+
+ Lemma sem_update_Iop_top:
+ forall A f p p' f' rs m pr op res args p0 v state,
+ Op.eval_operation (ctx_ge state) (ctx_sp state) op rs ## args m = Some v ->
+ truthy pr p0 ->
+ valid_mem (is_mem (ctx_is state)) m ->
+ sem state f ((mk_instr_state rs pr m), None) ->
+ update (p, f) (RBop p0 op args res) = Some (p', f') ->
+ eval_predf pr p = true ->
+ @sem A state f' (mk_instr_state (rs # res <- v) pr m, None).
+ Proof.
+ intros * EVAL_OP TRUTHY VALID SEM UPD EVAL_PRED.
+ pose proof SEM as SEM2.
+ inversion UPD as [PRUNE]. unfold Option.bind in PRUNE.
+ destr. inversion_clear PRUNE.
+ rename Heqo into PRUNE.
+ inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
+ cbn [is_ps] in *. constructor.
+ + constructor; intro x. inversion_clear REG as [? ? ? REG']. specialize (REG' x).
+ destruct f as [fr fp fe]. cbn [forest_preds set_forest] in *.
+ destruct (peq x res); subst.
+ * rewrite forest_reg_gss in *. rewrite Regmap.gss in *.
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated; [| |eauto].
+ replace fp with (forest_preds {| forest_regs := fr; forest_preds := fp; forest_exit := fe |}) by auto.
+ eapply sem_update_Iop; eauto. cbn in *.
+ eapply eval_operation_valid_pointer; eauto.
+ inversion_clear SEM2 as [? ? ? ? ? ? REG2 PRED2 MEM2 EXIT2].
+ inversion_clear PRED2; eauto.
+ cbn -[eval_predf]. rewrite eval_predf_Pand.
+ rewrite EVAL_PRED. rewrite truthy_dflt; auto.
+ * rewrite forest_reg_gso. rewrite Regmap.gso; auto.
+ unfold not in *; intros. apply n0. inv H; auto.
+ + constructor; intros. inv PRED. rewrite forest_reg_pred. auto.
+ + rewrite forest_reg_gso; auto; discriminate.
+ + auto.
+ Qed.
+
+ Lemma sem_update_Iload_top:
+ forall A f p p' f' rs m pr res args p0 v state addr a chunk,
+ Op.eval_addressing (ctx_ge state) (ctx_sp state) addr rs ## args = Some a ->
+ Mem.loadv chunk m a = Some v ->
+ truthy pr p0 ->
+ valid_mem (is_mem (ctx_is state)) m ->
+ sem state f ((mk_instr_state rs pr m), None) ->
+ update (p, f) (RBload p0 chunk addr args res) = Some (p', f') ->
+ eval_predf pr p = true ->
+ @sem A state f' (mk_instr_state (rs # res <- v) pr m, None).
+ Proof.
+ intros * EVAL_OP LOAD TRUTHY VALID SEM UPD EVAL_PRED.
+ pose proof SEM as SEM2.
+ inversion UPD as [PRUNE]. unfold Option.bind in PRUNE. destr.
+ inversion_clear PRUNE.
+ rename Heqo into PRUNE.
+ inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
+ cbn [is_ps] in *. constructor.
+ + constructor; intro x. inversion_clear REG as [? ? ? REG']. specialize (REG' x).
+ destruct f as [fr fp fe]. cbn [forest_preds set_forest] in *.
+ destruct (peq x res); subst.
+ * rewrite forest_reg_gss in *. rewrite Regmap.gss in *.
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated; [| |eauto].
+ replace fp with (forest_preds {| forest_regs := fr; forest_preds := fp; forest_exit := fe |}) by auto.
+ eapply sem_update_Iload; eauto.
+ inversion_clear PRED; eauto.
+ cbn -[eval_predf]. rewrite eval_predf_Pand.
+ rewrite EVAL_PRED. rewrite truthy_dflt; auto.
+ * rewrite forest_reg_gso. rewrite Regmap.gso; auto.
+ unfold not in *; intros. apply n0. inv H; auto.
+ + constructor; intros. inv PRED. rewrite forest_reg_pred. auto.
+ + rewrite forest_reg_gso; auto; discriminate.
+ + auto.
+ Qed.
+
+ Lemma exists_sem_pred :
+ forall A B C (ctx: @ctx A) s pr r v,
+ @sem_pred_expr A B C pr s ctx r v ->
+ exists r',
+ NE.In r' r /\ s ctx (snd r') v.
+ Proof.
+ induction r; crush.
+ - inv H. econstructor. split. constructor. auto.
+ - inv H.
+ + econstructor. split. constructor. left; auto. auto.
+ + exploit IHr; eauto. intros HH. inversion_clear HH as [x HH']. inv HH'.
+ econstructor. split. constructor. right. eauto. auto.
+ Qed.
+
+ Lemma sem_update_Istore_top:
+ forall A f p p' f' rs m pr res args p0 state addr a chunk m',
+ Op.eval_addressing (ctx_ge state) (ctx_sp state) addr rs ## args = Some a ->
+ Mem.storev chunk m a rs !! res = Some m' ->
+ truthy pr p0 ->
+ valid_mem (is_mem (ctx_is state)) m ->
+ sem state f ((mk_instr_state rs pr m), None) ->
+ update (p, f) (RBstore p0 chunk addr args res) = Some (p', f') ->
+ eval_predf pr p = true ->
+ @sem A state f' (mk_instr_state rs pr m', None).
+ Proof.
+ intros * EVAL_OP STORE TRUTHY VALID SEM UPD EVAL_PRED.
+ pose proof SEM as SEM2.
+ inversion UPD as [PRUNE]. unfold Option.bind in PRUNE. destr.
+ inversion_clear PRUNE.
+ rename Heqo into PRUNE.
+ inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
+ cbn [is_ps] in *. constructor.
+ + constructor; intros. inv REG. rewrite forest_reg_gso; eauto. discriminate.
+ + constructor; intros. inv PRED. rewrite forest_reg_pred. auto.
+ + destruct f as [fr fp fm]; cbn -[seq_app] in *.
+ rewrite forest_reg_gss.
+ exploit sem_pred_expr_ident2; [exact MEM|]; intro HSEM_;
+ inversion_clear HSEM_ as [x [HSEM1 HSEM2]].
+ inv REG. specialize (H res).
+ pose proof H as HRES.
+ eapply sem_pred_expr_ident2 in HRES.
+ inversion_clear HRES as [r2 [HRES1 HRES2]].
+ apply exists_sem_pred in H. inversion_clear H as [r' [HNE HVAL]].
+ exploit sem_merge_list. eapply SEM2. instantiate (2 := args). intro HSPE. eapply sem_pred_expr_ident2 in HSPE.
+ inversion_clear HSPE as [l_ [HSPE1 HSPE2]].
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated.
+ eapply sem_pred_expr_seq_app_val; [solve [eauto]| |].
+ eapply sem_pred_expr_seq_app; [solve [eauto]|].
+ eapply sem_pred_expr_flap2.
+ eapply sem_pred_expr_seq_app; [solve [eauto]|].
+ eapply sem_pred_expr_pred_ret. econstructor. eauto. 3: { eauto. }
+ eauto. auto. eauto. inv PRED. eauto.
+ rewrite eval_predf_Pand. rewrite EVAL_PRED.
+ rewrite truthy_dflt. auto. auto.
+ + auto.
+ Qed.
+
+ Definition predicated_not_inP {A} (p: Gible.predicate) (l: predicated A) :=
+ forall op e, NE.In (op, e) l -> ~ PredIn p op.
+
+ Lemma predicated_not_inP_cons :
+ forall A p (a: (pred_op * A)) l,
+ predicated_not_inP p (NE.cons a l) ->
+ predicated_not_inP p l.
+ Proof.
+ unfold predicated_not_inP; intros. eapply H. econstructor. right; eauto.
+ Qed.
+
+ Lemma sem_pexpr_not_in :
+ forall G (ctx: @ctx G) p0 ps p e b,
+ ~ PredIn p p0 ->
+ sem_pexpr ctx (from_pred_op ps p0) b ->
+ sem_pexpr ctx (from_pred_op (PTree.set p e ps) p0) b.
+ Proof.
+ induction p0; auto; intros.
+ - cbn. destruct p. unfold get_forest_p'.
+ assert (p0 <> p) by
+ (unfold not; intros; apply H; subst; constructor).
+ rewrite PTree.gso; auto.
+ - cbn in *.
+ assert (X: ~ PredIn p p0_1 /\ ~ PredIn p p0_2) by
+ (split; unfold not; intros; apply H; constructor; tauto).
+ inversion_clear X as [X1 X2].
+ inv H0. inv H4.
+ specialize (IHp0_1 _ p e _ X1 H0). constructor. tauto.
+ specialize (IHp0_2 _ p e _ X2 H0). constructor. tauto.
+ constructor; auto.
+ - cbn in *.
+ assert (X: ~ PredIn p p0_1 /\ ~ PredIn p p0_2) by
+ (split; unfold not; intros; apply H; constructor; tauto).
+ inversion_clear X as [X1 X2].
+ inv H0. inv H4.
+ specialize (IHp0_1 _ p e _ X1 H0). constructor. tauto.
+ specialize (IHp0_2 _ p e _ X2 H0). constructor. tauto.
+ constructor; auto.
+ Qed.
+
+ Lemma sem_pred_not_in :
+ forall A B G (ctx: @ctx G) (s: @Abstr.ctx G -> A -> B -> Prop) l v p e ps,
+ sem_pred_expr ps s ctx l v ->
+ predicated_not_inP p l ->
+ sem_pred_expr (PTree.set p e ps) s ctx l v.
+ Proof.
+ induction l.
+ - intros. unfold predicated_not_inP in H0.
+ destruct a. constructor. apply sem_pexpr_not_in.
+ eapply H0. econstructor. inv H. auto. inv H. auto.
+ - intros. inv H. constructor. unfold predicated_not_inP in H0.
+ eapply sem_pexpr_not_in. eapply H0. constructor. left. eauto.
+ auto. auto.
+ apply sem_pred_expr_cons_false. apply sem_pexpr_not_in. eapply H0.
+ constructor. tauto. auto. auto.
+ eapply IHl. eauto. eapply predicated_not_inP_cons; eauto.
+ Qed.
+
+ Lemma pred_not_in_forestP :
+ forall pred f,
+ predicated_not_in_forest pred f = true ->
+ forall x, predicated_not_inP pred (f #r x).
+ Proof. Admitted.
+
+ Lemma pred_not_in_forest_exitP :
+ forall pred f,
+ predicated_not_in_forest pred f = true ->
+ predicated_not_inP pred (forest_exit f).
+ Proof. Admitted.
+
+ Lemma from_predicated_sem_pred_expr :
+ forall A (ctx: @ctx A) preds pe b,
+ sem_pred_expr preds sem_pred ctx pe b ->
+ sem_pexpr ctx (from_predicated true preds pe) b.
+ Proof. Admitted.
+
+ Lemma sem_update_Isetpred:
+ forall A (ctx: @ctx A) f pr p0 c args b rs m,
+ valid_mem (ctx_mem ctx) m ->
+ sem ctx f (mk_instr_state rs pr m, None) ->
+ Op.eval_condition c rs ## args m = Some b ->
+ truthy pr p0 ->
+ sem_pexpr ctx
+ (from_predicated true (forest_preds f) (seq_app (pred_ret (PEsetpred c)) (merge (list_translation args f)))) b.
+ Proof.
+ intros. eapply from_predicated_sem_pred_expr.
+ pose proof (sem_merge_list _ ctx f rs pr m args H0).
+ apply sem_pred_expr_ident2 in H3; simplify.
+ eapply sem_pred_expr_seq_app_val; [eauto| |].
+ constructor. constructor. constructor.
+ econstructor; eauto.
+ erewrite storev_eval_condition; eauto.
+ Qed.
+
+ Lemma sem_update_Isetpred_top:
+ forall A f p p' f' rs m pr args p0 state c pred b,
+ Op.eval_condition c rs ## args m = Some b ->
+ truthy pr p0 ->
+ valid_mem (is_mem (ctx_is state)) m ->
+ sem state f ((mk_instr_state rs pr m), None) ->
+ update (p, f) (RBsetpred p0 c args pred) = Some (p', f') ->
+ eval_predf pr p = true ->
+ @sem A state f' (mk_instr_state rs (pr # pred <- b) m, None).
+ Proof.
+ intros * EVAL_OP TRUTHY VALID SEM UPD EVAL_PRED.
+ pose proof SEM as SEM2.
+ inversion UPD as [PRUNE]. unfold Option.bind in PRUNE. destr. destr.
+ inversion_clear PRUNE.
+ rename Heqo into PRUNE.
+ inversion_clear SEM as [? ? ? ? ? ? REG PRED MEM EXIT].
+ cbn [is_ps] in *. constructor.
+ + constructor. intros. apply sem_pred_not_in. rewrite forest_pred_reg.
+ inv REG. auto. rewrite forest_pred_reg. apply pred_not_in_forestP.
+ unfold assert_ in *. repeat (destruct_match; try discriminate); auto.
+ + constructor; intros. destruct (peq x pred); subst.
+ * rewrite Regmap.gss.
+ rewrite forest_pred_gss.
+ cbn [update] in *. unfold Option.bind in *. destr. destr. inv UPD.
+ replace b with (b && true) by eauto with bool.
+ apply sem_pexpr_Pand.
+ destruct b. constructor. right.
+ eapply sem_update_Isetpred; eauto.
+ constructor. constructor. replace false with (negb true) by auto.
+ apply sem_pexpr_negate. inv TRUTHY. constructor.
+ cbn. eapply sem_pexpr_eval. inv PRED. eauto. auto.
+ replace false with (negb true) by auto.
+ apply sem_pexpr_negate.
+ eapply sem_pexpr_eval. inv PRED. eauto. auto.
+ eapply sem_update_Isetpred; eauto.
+ constructor. constructor. constructor.
+ replace true with (negb false) by auto. apply sem_pexpr_negate.
+ eapply sem_pexpr_eval. inv PRED. eauto. inv TRUTHY. auto. cbn -[eval_predf].
+ rewrite eval_predf_negate. rewrite H; auto.
+ replace true with (negb false) by auto. apply sem_pexpr_negate.
+ eapply sem_pexpr_eval. inv PRED. eauto. rewrite eval_predf_negate.
+ rewrite EVAL_PRED. auto.
+ * rewrite Regmap.gso by auto. inv PRED. specialize (H x).
+ rewrite forest_pred_gso by auto; auto.
+ + rewrite forest_pred_reg. apply sem_pred_not_in. auto. apply pred_not_in_forestP.
+ unfold assert_ in *. now repeat (destruct_match; try discriminate).
+ + cbn -[from_predicated from_pred_op seq_app]. apply sem_pred_not_in; auto.
+ apply pred_not_in_forest_exitP.
+ unfold assert_ in *. now repeat (destruct_match; try discriminate).
+ Qed.
+
+ Lemma sem_pexpr_impl :
+ forall A (state: @ctx A) a b res,
+ sem_pexpr state b res ->
+ sem_pexpr state a true ->
+ sem_pexpr state (a → b) res.
+ Proof.
+ intros. destruct res.
+ constructor; tauto.
+ constructor; auto. replace false with (negb true) by auto.
+ now apply sem_pexpr_negate.
+ Qed.
+
+ Lemma eval_predf_simplify :
+ forall ps x,
+ eval_predf ps (simplify x) = eval_predf ps x.
+ Proof.
+ unfold eval_predf; intros.
+ rewrite simplify_correct. auto.
+ Qed.
+
+ Lemma sem_update_falsy:
+ forall A state f f' rs ps m p a p',
+ instr_falsy ps a ->
+ update (p, f) a = Some (p', f') ->
+ sem state f (mk_instr_state rs ps m, None) ->
+ @sem A state f' (mk_instr_state rs ps m, None).
+ Proof.
+ inversion 1; cbn [update] in *; intros.
+ - unfold Option.bind in *. destr. inv H2.
+ constructor.
+ * constructor. intros. destruct (peq x res); subst.
+ rewrite forest_reg_gss. cbn.
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated_false. inv H3. inv H8. auto.
+ inv H3. inv H9. eauto. rewrite eval_predf_Pand. cbn -[eval_predf].
+ rewrite H0. auto.
+ rewrite forest_reg_gso. inv H3. inv H8. auto.
+ unfold not; intros; apply n0. now inv H1.
+ * constructor; intros. rewrite forest_reg_pred. inv H3. inv H9. auto.
+ * rewrite forest_reg_gso. inv H3. auto. unfold not; intros. inversion H1.
+ * inv H3. auto.
+ - unfold Option.bind in *. destr. inv H2.
+ constructor.
+ * constructor. intros. destruct (peq x dst); subst.
+ rewrite forest_reg_gss. cbn.
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated_false. inv H3. inv H8. auto.
+ inv H3. inv H9. eauto. rewrite eval_predf_Pand. cbn -[eval_predf].
+ rewrite H0. auto.
+ rewrite forest_reg_gso. inv H3. inv H8. auto.
+ unfold not; intros; apply n0. now inv H1.
+ * constructor; intros. rewrite forest_reg_pred. inv H3. inv H9. auto.
+ * rewrite forest_reg_gso. inv H3. auto. unfold not; intros. inversion H1.
+ * inv H3. auto.
+ - unfold Option.bind in *. destr. inv H2.
+ constructor.
+ * constructor. intros. rewrite forest_reg_gso by discriminate. inv H3. inv H8. auto.
+ * constructor; intros. rewrite forest_reg_pred. inv H3. inv H9. auto.
+ * rewrite forest_reg_gss. cbn. eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated_false. inv H3. auto. inv H3. inv H9. eauto.
+ rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H0. auto.
+ * inv H3. auto.
+ - unfold Option.bind in *. destr. destr. inv H2.
+ constructor.
+ * constructor; intros. rewrite forest_pred_reg. apply sem_pred_not_in.
+ inv H3. inv H8. auto. apply pred_not_in_forestP. unfold assert_ in Heqo. now destr.
+ * constructor. intros. destruct (peq x pred); subst.
+ rewrite forest_pred_gss. replace (ps !! pred) with (true && ps !! pred) by auto.
+ assert (sem_pexpr state0 (¬ (from_pred_op (forest_preds f) p0 ∧ from_pred_op (forest_preds f) p')) true).
+ { replace true with (negb false) by auto. apply sem_pexpr_negate.
+ constructor. left. eapply sem_pexpr_eval. inv H3. inv H9. eauto.
+ auto.
+ }
+ apply sem_pexpr_Pand. constructor; tauto.
+ apply sem_pexpr_impl. inv H3. inv H10. eauto.
+ { constructor. left. eapply sem_pexpr_eval. inv H3. inv H10. eauto.
+ rewrite eval_predf_negate. rewrite H0. auto.
+ }
+ rewrite forest_pred_gso by auto. inv H3. inv H9. auto.
+ * rewrite forest_pred_reg. apply sem_pred_not_in. inv H3. auto.
+ apply pred_not_in_forestP. unfold assert_ in Heqo. now destr.
+ * apply sem_pred_not_in. inv H3; auto. cbn.
+ apply pred_not_in_forest_exitP. unfold assert_ in Heqo. now destr.
+ - unfold Option.bind in *. destr. inv H2. inv H3. constructor.
+ * constructor. inv H8. auto.
+ * constructor. intros. inv H9. eauto.
+ * auto.
+ * cbn. eapply sem_pred_expr_prune_predicated; [|eauto].
+ eapply sem_pred_expr_app_predicated_false; auto.
+ inv H9. eauto.
+ rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H0. auto.
+ Qed.
+
+ Lemma sem_update_falsy_input:
+ forall A state f f' rs ps m p a p' exitcf,
+ eval_predf ps p = false ->
+ update (p, f) a = Some (p', f') ->
+ sem state f (mk_instr_state rs ps m, exitcf) ->
+ @sem A state f' (mk_instr_state rs ps m, exitcf)
+ /\ eval_predf ps p' = false.
+ Proof.
+ intros; destruct a; cbn [update] in *; intros.
+ - inv H0. auto.
+ - unfold Option.bind in *. destr. inv H0. split; [|solve [auto]].
+ constructor.
+ * constructor. intros. destruct (peq x r); subst.
+ rewrite forest_reg_gss. cbn.
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated_false. inv H1. inv H7. auto.
+ inv H1. inv H8. eauto. rewrite eval_predf_Pand.
+ rewrite H. eauto with bool.
+ rewrite forest_reg_gso. inv H1. inv H7. auto.
+ unfold not; intros; apply n0. now inv H0.
+ * constructor; intros. rewrite forest_reg_pred. inv H1. inv H8. auto.
+ * rewrite forest_reg_gso. inv H1. auto. unfold not; intros. inversion H0.
+ * inv H1. auto.
+ - unfold Option.bind in *. destr. inv H0. split; [|solve [auto]].
+ constructor.
+ * constructor. intros. destruct (peq x r); subst.
+ rewrite forest_reg_gss. cbn.
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated_false. inv H1. inv H7. auto.
+ inv H1. inv H8. eauto. rewrite eval_predf_Pand. cbn -[eval_predf].
+ rewrite H. eauto with bool.
+ rewrite forest_reg_gso. inv H1. inv H7. auto.
+ unfold not; intros; apply n0. now inv H0.
+ * constructor; intros. rewrite forest_reg_pred. inv H1. inv H8. auto.
+ * rewrite forest_reg_gso. inv H1. auto. unfold not; intros. inversion H0.
+ * inv H1. auto.
+ - unfold Option.bind in *. destr. inv H0. split; [|solve [auto]].
+ constructor.
+ * constructor. intros. rewrite forest_reg_gso by discriminate. inv H1. inv H7. auto.
+ * constructor; intros. rewrite forest_reg_pred. inv H1. inv H8. auto.
+ * rewrite forest_reg_gss. cbn. eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated_false. inv H1. auto. inv H1. inv H8. eauto.
+ rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H. eauto with bool.
+ * inv H1. auto.
+ - unfold Option.bind in *. destr. destr. inv H0. split; [|solve [auto]].
+ constructor.
+ * constructor; intros. rewrite forest_pred_reg. apply sem_pred_not_in.
+ inv H1. inv H7. auto. apply pred_not_in_forestP. unfold assert_ in Heqo0. now destr.
+ * constructor. intros. destruct (peq x p0); subst.
+ rewrite forest_pred_gss. replace (ps !! p0) with (true && ps !! p0) by auto.
+ assert (sem_pexpr state0 (¬ (from_pred_op (forest_preds f) (dfltp o) ∧ from_pred_op (forest_preds f) p')) true).
+ { replace true with (negb false) by auto. apply sem_pexpr_negate.
+ constructor. right. eapply sem_pexpr_eval. inv H1. inv H8. eauto.
+ auto.
+ }
+ apply sem_pexpr_Pand. constructor; tauto.
+ apply sem_pexpr_impl. inv H1. inv H9. eauto.
+ { constructor. right. eapply sem_pexpr_eval. inv H1. inv H9. eauto.
+ rewrite eval_predf_negate. rewrite H. auto.
+ }
+ rewrite forest_pred_gso by auto. inv H1. inv H8. auto.
+ * rewrite forest_pred_reg. apply sem_pred_not_in. inv H1. auto.
+ apply pred_not_in_forestP. unfold assert_ in Heqo0. now destr.
+ * apply sem_pred_not_in. inv H1; auto. cbn.
+ apply pred_not_in_forest_exitP. unfold assert_ in Heqo0. now destr.
+ - unfold Option.bind in *. destr. inv H0. inv H1. split.
+ -- constructor.
+ * constructor. inv H7. auto.
+ * constructor. intros. inv H8. eauto.
+ * auto.
+ * cbn. eapply sem_pred_expr_prune_predicated; [|eauto].
+ eapply sem_pred_expr_app_predicated_false; auto.
+ inv H8. eauto.
+ rewrite eval_predf_Pand. cbn -[eval_predf]. rewrite H. eauto with bool.
+ -- rewrite eval_predf_simplify. rewrite eval_predf_Pand. rewrite H. eauto with bool.
+ Qed.
+
+ Definition setpred_wf (i: instr): bool :=
+ match i with
+ | RBsetpred (Some op) _ _ p => negb (predin peq p op)
+ | _ => true
+ end.
+
+ Lemma sem_update_instr :
+ forall f i' i'' a sp p i p' f',
+ step_instr ge sp (Iexec i') a (Iexec i'') ->
+ valid_mem (is_mem i) (is_mem i') ->
+ sem (mk_ctx i sp ge) f (i', None) ->
+ update (p, f) a = Some (p', f') ->
+ eval_predf (is_ps i') p = true ->
+ sem (mk_ctx i sp ge) f' (i'', None).
+ Proof.
+ inversion 1; subst; clear H; intros VALID SEM UPD EVAL_PRED; pose proof SEM as SEM2.
+ - inv UPD; auto.
+ - eauto using sem_update_Iop_top.
+ - eauto using sem_update_Iload_top.
+ - eauto using sem_update_Istore_top.
+ - eauto using sem_update_Isetpred_top.
+ - destruct i''. eauto using sem_update_falsy.
+ Qed.
+
+ Lemma Pand_true_left :
+ forall ps a b,
+ eval_predf ps a = false ->
+ eval_predf ps (a ∧ b) = false.
+ Proof.
+ intros.
+ rewrite eval_predf_Pand. now rewrite H.
+ Qed.
+
+ Lemma Pand_true_right :
+ forall ps a b,
+ eval_predf ps b = false ->
+ eval_predf ps (a ∧ b) = false.
+ Proof.
+ intros.
+ rewrite eval_predf_Pand. rewrite H.
+ eauto with bool.
+ Qed.
+
+ Lemma sem_update_instr_term :
+ forall f i' i'' sp i cf p p' p'' f',
+ sem (mk_ctx i sp ge) f (i', None) ->
+ step_instr ge sp (Iexec i') (RBexit p cf) (Iterm i'' cf) ->
+ update (p', f) (RBexit p cf) = Some (p'', f') ->
+ eval_predf (is_ps i') p' = true ->
+ sem (mk_ctx i sp ge) f' (i'', Some cf)
+ /\ eval_predf (is_ps i') p'' = false.
+ Proof.
+ intros. inv H0. simpl in *.
+ unfold Option.bind in *. destruct_match; try discriminate.
+ apply truthy_dflt in H6. inv H1.
+ assert (eval_predf (Gible.is_ps i'') (¬ dfltp p) = false).
+ { rewrite eval_predf_negate. now rewrite negb_false_iff. }
+ apply Pand_true_left with (b := p') in H0.
+ rewrite <- eval_predf_simplify in H0. split; auto.
+ unfold "<-e". destruct i''.
+ inv H. constructor; auto.
+ constructor. inv H9. intros. cbn. eauto.
+ inv H10. constructor; intros. eauto.
+ cbn.
+ eapply sem_pred_expr_prune_predicated; eauto.
+ eapply sem_pred_expr_app_predicated.
+ constructor. constructor. constructor.
+ inv H10. eauto. cbn -[eval_predf] in *.
+ rewrite eval_predf_Pand. rewrite H2. now rewrite H6.
+ Qed.
+
+ Lemma step_instr_lessdef_term :
+ forall sp a i i' ti cf,
+ step_instr ge sp (Iexec i) a (Iterm i' cf) ->
+ state_lessdef i ti ->
+ exists ti', step_instr ge sp (Iexec ti) a (Iterm ti' cf) /\ state_lessdef i' ti'.
+ Proof.
+ inversion 1; intros; subst.
+ econstructor. split; eauto. constructor.
+ destruct p. constructor. erewrite eval_predf_pr_equiv. inv H4.
+ eauto. inv H6. eauto. constructor.
+ Qed.
+
+ Lemma combined_falsy :
+ forall i o1 o,
+ falsy i o1 ->
+ falsy i (combine_pred o o1).
+ Proof.
+ inversion 1; subst; crush. destruct o; simplify.
+ constructor. rewrite eval_predf_Pand. rewrite H0. crush.
+ constructor. auto.
+ Qed.
+
+ Lemma Abstr_match_states_sem :
+ forall i sp f i' ti cf,
+ sem (mk_ctx i sp ge) f (i', cf) ->
+ state_lessdef i ti ->
+ exists ti', sem (mk_ctx ti sp ge) f (ti', cf) /\ state_lessdef i' ti'.
+ Proof. Admitted. (* This needs a bit more in Abstr.v *)
+
+ Lemma mfold_left_update_Some :
+ forall xs x v,
+ mfold_left update xs x = Some v ->
+ exists y, x = Some y.
+ Proof.
+ induction xs; intros.
+ { cbn in *. inv H. eauto. }
+ cbn in *. unfold Option.bind in *. exploit IHxs; eauto.
+ intros. simplify. destruct x; crush.
+ eauto.
+ Qed.
+
+ Lemma step_instr_term_exists :
+ forall A B ge sp v x v2 cf,
+ @step_instr A B ge sp (Iexec v) x (Iterm v2 cf) ->
+ exists p, x = RBexit p cf.
+ Proof using. inversion 1; eauto. Qed.
+
+ Lemma eval_predf_update_true :
+ forall i i' curr_p next_p f f_next instr sp,
+ update (curr_p, f) instr = Some (next_p, f_next) ->
+ step_instr ge sp (Iexec i) instr (Iexec i') ->
+ eval_predf (is_ps i) curr_p = true ->
+ eval_predf (is_ps i') next_p = true.
+ Proof.
+ intros * UPD STEP EVAL. destruct instr; cbn [update] in UPD;
+ try solve [unfold Option.bind in *; try destr; inv UPD; inv STEP; auto].
+ - unfold Option.bind in *. destr. destr. inv UPD. inv STEP; auto. cbn [is_ps] in *.
+ unfold is_initial_pred_and_notin in Heqo1. unfold assert_ in Heqo1. destr. destr.
+ destr. destr. destr. destr. subst. assert (~ PredIn p2 next_p).
+ unfold not; intros. apply negb_true_iff in Heqb0. apply not_true_iff_false in Heqb0.
+ apply Heqb0. now apply predin_PredIn. rewrite eval_predf_not_PredIn; auto.
+ - unfold Option.bind in *. destr. inv UPD. inv STEP. inv H3. cbn.
+ rewrite eval_predf_simplify. rewrite eval_predf_Pand. rewrite eval_predf_negate.
+ destruct i'; cbn in *. rewrite H0. auto.
+ Qed.
+
+ Lemma seq_app_cons :
+ forall A B f a l p0 p1,
+ @seq_app A B (pred_ret f) (NE.cons a l) = NE.cons p0 p1 ->
+ @seq_app A B (pred_ret f) l = p1.
+ Proof. intros. cbn in *. inv H. eauto. Qed.
+
+ Lemma sem_update_valid_mem :
+ forall i i' i'' curr_p next_p f f_next instr sp,
+ step_instr ge sp (Iexec i') instr (Iexec i'') ->
+ update (curr_p, f) instr = Some (next_p, f_next) ->
+ sem (mk_ctx i sp ge) f (i', None) ->
+ valid_mem (is_mem i') (is_mem i'').
+ Proof.
+ inversion 1; crush.
+ unfold Option.bind in *. destr. inv H7.
+ eapply storev_valid_pointer; eauto.
+ Qed.
+
+ Lemma eval_predf_lessdef :
+ forall p a b,
+ state_lessdef a b ->
+ eval_predf (is_ps a) p = eval_predf (is_ps b) p.
+ Proof using.
+ induction p; crush.
+ - inv H. simpl. unfold eval_predf. simpl.
+ repeat destr. inv Heqp0. rewrite H1. auto.
+ - rewrite !eval_predf_Pand.
+ erewrite IHp1 by eassumption.
+ now erewrite IHp2 by eassumption.
+ - rewrite !eval_predf_Por.
+ erewrite IHp1 by eassumption.
+ now erewrite IHp2 by eassumption.
+ Qed.
+
+(*|
+``abstr_fold_falsy``: This lemma states that when we are at the end of an
+execution, the values in the register map do not continue to change.
+|*)
+
+ Lemma abstr_fold_falsy :
+ forall A ilist i sp ge f res p f' p',
+ @sem A (mk_ctx i sp ge) f res ->
+ mfold_left update ilist (Some (p, f)) = Some (p', f') ->
+ eval_predf (is_ps (fst res)) p = false ->
+ sem (mk_ctx i sp ge) f' res.
+ Proof.
+ induction ilist.
+ - intros. cbn in *. inv H0. auto.
+ - intros. cbn -[update] in H0.
+ exploit mfold_left_update_Some. eauto. intros. inv H2.
+ rewrite H3 in H0. destruct x.
+ destruct res. destruct i0.
+ exploit sem_update_falsy_input; try eassumption; intros.
+ inversion_clear H2.
+ eapply IHilist; eassumption.
+ Qed.
+
+ Lemma abstr_fold_correct :
+ forall sp x i i' i'' cf f p f' curr_p
+ (VALID: valid_mem (is_mem i) (is_mem i')),
+ SeqBB.step ge sp (Iexec i') x (Iterm i'' cf) ->
+ sem (mk_ctx i sp ge) f (i', None) ->
+ eval_predf (is_ps i') curr_p = true ->
+ mfold_left update x (Some (curr_p, f)) = Some (p, f') ->
+ forall ti,
+ state_lessdef i ti ->
+ exists ti', sem (mk_ctx ti sp ge) f' (ti', Some cf)
+ /\ state_lessdef i'' ti'
+ /\ valid_mem (is_mem i) (is_mem i'').
+ Proof.
+ induction x as [| x xs IHx ]; intros; cbn -[update] in *; inv H; cbn [fst snd] in *.
+ - (* inductive case *)
+ exploit mfold_left_update_Some; eauto; intros Hexists;
+ inversion Hexists as [[curr_p_inter f_inter] EXEQ]; clear Hexists.
+ exploit eval_predf_update_true;
+ eauto; intros EVALTRUE.
+ rewrite EXEQ in H2. eapply IHx in H2; cbn [fst snd] in *.
+ eauto.
+ transitivity (is_mem i'); auto.
+ eapply sem_update_valid_mem; eauto.
+ eauto.
+ eapply sem_update_instr; eauto. eauto. eauto.
+ - (* terminal case *)
+ exploit mfold_left_update_Some; eauto; intros Hexists;
+ inversion Hexists as [[curr_p_inter f_inter] EXEQ]; clear Hexists. rewrite EXEQ in H2.
+ exploit Abstr_match_states_sem; (* TODO *)
+ eauto; intros H; inversion H as [v [Hi LESSDEF]]; clear H.
+ exploit step_instr_lessdef_term;
+ eauto; intros H; inversion H as [v2 [Hi2 LESSDEF2]]; clear H.
+ exploit step_instr_term_exists; eauto; inversion 1 as [? ?]; subst; clear H.
+ erewrite eval_predf_lessdef in H1 by eassumption.
+ exploit sem_update_instr_term;
+ eauto; intros [A B].
+ exists v2.
+ exploit abstr_fold_falsy.
+ apply A.
+ eassumption. auto. cbn. inversion Hi2; subst. auto. intros.
+ split; auto. split; auto.
+ inversion H7; subst; auto.
+ Qed.
+
+ Lemma sem_regset_empty :
+ forall A ctx, @sem_regset A ctx empty (ctx_rs ctx).
+ Proof using.
+ intros; constructor; intros.
+ constructor; auto. constructor.
+ constructor.
+ Qed.
+
+ Lemma sem_predset_empty :
+ forall A ctx, @sem_predset A ctx empty (ctx_ps ctx).
+ Proof using.
+ intros; constructor; intros.
+ constructor; auto. constructor.
+ Qed.
+
+ Lemma sem_empty :
+ forall A ctx, @sem A ctx empty (ctx_is ctx, None).
+ Proof using.
+ intros. destruct ctx. destruct ctx_is.
+ constructor; try solve [constructor; constructor; crush].
+ eapply sem_regset_empty.
+ eapply sem_predset_empty.
+ Qed.
+
+ Lemma abstr_sequence_correct :
+ forall sp x i i'' cf x',
+ SeqBB.step ge sp (Iexec i) x (Iterm i'' cf) ->
+ abstract_sequence x = Some x' ->
+ forall ti,
+ state_lessdef i ti ->
+ exists ti', sem (mk_ctx ti sp ge) x' (ti', Some cf)
+ /\ state_lessdef i'' ti'.
+ Proof.
+ unfold abstract_sequence. intros. unfold Option.map in H0.
+ destruct_match; try easy.
+ destruct p as [p TMP]; simplify.
+ exploit abstr_fold_correct; eauto; crush.
+ { apply sem_empty. }
+ exists x0. auto.
+ Qed.
+
+End CORRECTNESS.
diff --git a/src/hls/Predicate.v b/src/hls/Predicate.v
index f99fa4f..92ad03f 100644
--- a/src/hls/Predicate.v
+++ b/src/hls/Predicate.v
@@ -879,3 +879,9 @@ Proof.
pose proof (sat_predicateP_det a p _ _ H1 H0).
rewrite H in H3. now rewrite H3 in H2.
Qed.
+
+Definition and_list {A} (p: list (@pred_op A)): @pred_op A :=
+ fold_left Pand p T.
+
+Definition or_list {A} (p: list (@pred_op A)): @pred_op A :=
+ fold_left Por p ⟂.