aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorChantal Keller <Chantal.Keller@inria.fr>2020-06-30 13:34:41 +0200
committerChantal Keller <Chantal.Keller@inria.fr>2020-06-30 13:34:41 +0200
commit88cfb1c2883a44a4295b4c5fe7f091f107913412 (patch)
treeca78f7b32ab84ed72b5f17f7b389844e4233dc6e
parent1300f0ab1aa346b4830d42e04693aba32d77d500 (diff)
parent5a48d8b315ba20a5d79b56f9f3fa0ab7a1e8e13a (diff)
downloadsmtcoq-88cfb1c2883a44a4295b4c5fe7f091f107913412.tar.gz
smtcoq-88cfb1c2883a44a4295b4c5fe7f091f107913412.zip
Merge branch 'master' of github.com:smtcoq/smtcoq into coq-8.10
-rw-r--r--src/Misc.v34
-rw-r--r--src/array/Array_checker.v4
-rw-r--r--src/bva/Bva_checker.v7
-rw-r--r--src/cnf/Cnf.v8
4 files changed, 34 insertions, 19 deletions
diff --git a/src/Misc.v b/src/Misc.v
index 952985f..8553172 100644
--- a/src/Misc.v
+++ b/src/Misc.v
@@ -495,6 +495,23 @@ Proof.
Qed.
+Lemma afold_left_and p a :
+ afold_left bool int true andb p a =
+ List.forallb p (to_list a).
+Proof.
+ rewrite afold_left_spec; auto.
+ rewrite fold_left_to_list.
+ assert (H:forall l acc, List.fold_left (fun (a0 : bool) (v : int) => a0 && p v) l acc =
+ acc && List.forallb p l).
+ {
+ clear a. induction l; simpl.
+ - intros; now rewrite andb_true_r.
+ - intro acc. rewrite IHl. now rewrite andb_assoc.
+ }
+ now apply H.
+Qed.
+
+
(* Case orb *)
Lemma afold_left_orb_true : forall A i a f,
@@ -541,6 +558,23 @@ Proof.
Qed.
+Lemma afold_left_or p a :
+ afold_left bool int false orb p a =
+ List.existsb p (to_list a).
+Proof.
+ rewrite afold_left_spec; auto.
+ rewrite fold_left_to_list.
+ assert (H:forall l acc, List.fold_left (fun (a0 : bool) (v : int) => a0 || p v) l acc =
+ acc || List.existsb p l).
+ {
+ clear a. induction l; simpl.
+ - intros; now rewrite orb_false_r.
+ - intro acc. rewrite IHl. now rewrite orb_assoc.
+ }
+ now apply H.
+Qed.
+
+
(* Case implb *)
Lemma afold_right_implb_false : forall A a f,
diff --git a/src/array/Array_checker.v b/src/array/Array_checker.v
index 3a982b2..0540855 100644
--- a/src/array/Array_checker.v
+++ b/src/array/Array_checker.v
@@ -893,10 +893,6 @@ Section certif.
apply read_over_other_write; now auto.
Qed.
- Axiom afold_left_or : forall a,
- afold_left bool int false orb (Lit.interp rho) a =
- C.interp rho (to_list a).
-
Lemma valid_check_ext lres : C.valid rho (check_ext lres).
unfold check_ext, eq_sel_sym.
case_eq (Lit.is_pos lres); intro Heq; simpl; try now apply C.interp_true.
diff --git a/src/bva/Bva_checker.v b/src/bva/Bva_checker.v
index 0f5bc1a..50cb0c0 100644
--- a/src/bva/Bva_checker.v
+++ b/src/bva/Bva_checker.v
@@ -1926,13 +1926,6 @@ Qed.
Lemma eq_head: forall {A: Type} a b (l: list A), (a :: l) = (b :: l) <-> a = b.
Proof. intros A a b l; split; [intros H; inversion H|intros ->]; auto. Qed.
-Axiom afold_left_and : forall a,
- afold_left bool int true andb (Lit.interp rho) a = List.forallb (Lit.interp rho) (to_list a).
-
- Axiom afold_left_or : forall a,
- afold_left bool int false orb (Lit.interp rho) a =
- C.interp rho (to_list a).
-
Axiom afold_left_xor : forall a,
afold_left bool int false xorb (Lit.interp rho) a =
C.interp rho (to_list a).
diff --git a/src/cnf/Cnf.v b/src/cnf/Cnf.v
index ec18c15..d6a6640 100644
--- a/src/cnf/Cnf.v
+++ b/src/cnf/Cnf.v
@@ -287,14 +287,6 @@ Section CHECKER.
match goal with |- context [Lit.interp rho ?x] =>
destruct (Lit.interp rho x);trivial end.
- Axiom afold_left_and : forall a,
- afold_left bool int true andb (Lit.interp rho) a =
- List.forallb (Lit.interp rho) (to_list a).
-
- Axiom afold_left_or : forall a,
- afold_left bool int false orb (Lit.interp rho) a =
- C.interp rho (to_list a).
-
Axiom afold_right_impb : forall a,
(afold_right bool int true implb (Lit.interp rho) a) =
C.interp rho (to_list (or_of_imp a)).