aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorChantal Keller <Chantal.Keller@lri.fr>2021-07-07 16:26:42 +0200
committerChantal Keller <Chantal.Keller@lri.fr>2021-07-07 16:26:42 +0200
commit0a5ce82066855bae205f9536c008715eebb470f2 (patch)
treed07ce5a99b4d680c31e2e43f635ad4d37dff8945
parentc7f38ea3b96ee0e52f0e1f20ccfabd3206ff6ad2 (diff)
parent0a0dc47da1916c2a850610cfb80ba04c17e64549 (diff)
downloadsmtcoq-0a5ce82066855bae205f9536c008715eebb470f2.tar.gz
smtcoq-0a5ce82066855bae205f9536c008715eebb470f2.zip
Merge remote-tracking branch 'remotes/origin/coq-8.12' into coq-8.13
-rw-r--r--src/Array/PArray.v477
-rw-r--r--src/Int63/Int63.v23
-rw-r--r--src/Int63/Int63Axioms.v313
-rw-r--r--src/Int63/Int63Native.v144
-rw-r--r--src/Int63/Int63Op.v334
-rw-r--r--src/Int63/Int63Properties.v2760
-rw-r--r--src/Makefile4
-rw-r--r--src/Misc.v1546
-rw-r--r--src/SMT_terms.v194
-rw-r--r--src/State.v35
-rw-r--r--src/Trace.v47
-rw-r--r--src/_CoqProject6
-rw-r--r--src/array/Array_checker.v200
-rw-r--r--src/bva/BVList.v32
-rw-r--r--src/bva/Bva_checker.v92
-rw-r--r--src/classes/SMT_classes_instances.v18
-rw-r--r--src/cnf/Cnf.v52
-rw-r--r--src/euf/Euf.v24
-rw-r--r--src/lia/Lia.v246
-rw-r--r--src/spl/Assumptions.v2
-rw-r--r--src/spl/Operators.v84
-rw-r--r--src/spl/Syntactic.v79
-rw-r--r--src/trace/coqInterface.ml26
-rw-r--r--src/trace/coqInterface.mli2
-rw-r--r--src/trace/coqTerms.ml2
-rw-r--r--src/trace/smtBtype.ml2
-rw-r--r--unit-tests/Tests_lfsc_tactics.v14
27 files changed, 1695 insertions, 5063 deletions
diff --git a/src/Array/PArray.v b/src/Array/PArray.v
index 25da052..3c98041 100644
--- a/src/Array/PArray.v
+++ b/src/Array/PArray.v
@@ -13,15 +13,58 @@
(* Software implementation of arrays, based on finite maps using AVL
trees *)
-
Declare Scope array_scope.
-Require Import Int31.
-Require Export Int63.
+Require Export Int63 Psatz.
Require FMapAVL.
+Require Import ZArith.
Local Open Scope int63_scope.
+Import OrderedType.
+
+Module IntOrderedType <: OrderedType.
+
+ Definition t := int.
+
+ Definition eq x y := (x == y) = true.
+
+ Definition lt x y := (x < y) = true.
+
+ Lemma eq_refl x : eq x x.
+ Proof. unfold eq. rewrite eqb_spec. reflexivity. Qed.
+
+ Lemma eq_sym x y : eq x y -> eq y x.
+ Proof. unfold eq. rewrite !eqb_spec. intros ->. reflexivity. Qed.
+
+ Lemma eq_trans x y z : eq x y -> eq y z -> eq x z.
+ Proof. unfold eq. rewrite !eqb_spec. intros -> ->. reflexivity. Qed.
+
+ Lemma lt_trans x y z : lt x y -> lt y z -> lt x z.
+ Proof. unfold lt; rewrite !ltb_spec; apply Z.lt_trans. Qed.
+
+ Lemma lt_not_eq x y : lt x y -> ~ eq x y.
+ Proof. unfold lt, eq. rewrite ltb_spec, eqb_spec. intros H1 H2. rewrite H2 in H1. lia. Qed.
+
+ Definition compare x y : Compare lt eq x y.
+ Proof.
+ case_eq (x < y); intro e.
+ exact (LT e).
+ case_eq (x == y); intro e2.
+ exact (EQ e2).
+ apply GT. unfold lt.
+ rewrite <- Bool.not_false_iff_true, <- Bool.not_true_iff_false, ltb_spec, eqb_spec in *; intro e3.
+ apply e2, to_Z_inj; lia.
+ Defined.
+
+ Definition eq_dec x y : { eq x y } + { ~ eq x y }.
+ Proof.
+ case_eq (x == y); intro e.
+ left; exact e.
+ right. intro H. rewrite H in e. discriminate.
+ Defined.
+
+End IntOrderedType.
Module Map := FMapAVL.Make(IntOrderedType).
@@ -33,12 +76,14 @@ Definition array (A:Type) : Type :=
Definition make {A:Type} (l:int) (d:A) : array A := (Map.empty A, d, l).
Definition get {A:Type} (t:array A) (i:int) : A :=
- let (td,_) := t in
- let (t,d) := td in
- match Map.find i t with
- | Some x => x
- | None => d
- end.
+ let (td, l) := t in
+ let (t, d) := td in
+ if i < l then
+ match Map.find i t with
+ | Some x => x
+ | None => d
+ end
+ else d.
Definition default {A:Type} (t:array A) : A :=
let (td,_) := t in let (_,d) := td in d.
@@ -56,340 +101,160 @@ Definition length {A:Type} (t:array A) : int :=
Definition copy {A:Type} (t:array A) : array A := t.
-Definition reroot : forall {A:Type}, array A -> array A := @copy.
-
-Definition init {A:Type} (l:int) (f:int -> A) (d:A) : array A :=
- let r :=
- if l == 0 then
- Map.empty A
- else
- foldi (fun j m => Map.add j (f j) m) 0 (l-1) (Map.empty A) in
- (r, d, l).
-
-Definition map {A B:Type} (f:A -> B) (t:array A) : array B :=
- let (td,l) := t in
- let (t,d) := td in
- (Map.map f t, f d, l).
-
Delimit Scope array_scope with array.
Notation "t '.[' i ']'" := (get t i) (at level 50) : array_scope.
Notation "t '.[' i '<-' a ']'" := (set t i a) (at level 50) : array_scope.
Local Open Scope array_scope.
-Definition max_array_length := 4194302%int31.
+Definition max_length := max_int.
(** Axioms *)
-Axiom get_outofbound : forall A (t:array A) i, (i < length t) = false -> t.[i] = default t.
-
-Axiom get_set_same : forall A t i (a:A), (i < length t) = true -> t.[i<-a].[i] = a.
-Axiom get_set_other : forall A t i j (a:A), i <> j -> t.[i<-a].[j] = t.[j].
-Axiom default_set : forall A t i (a:A), default (t.[i<-a]) = default t.
-
-
-Axiom get_make : forall A (a:A) size i, (make size a).[i] = a.
-Axiom default_make : forall A (a:A) size, (default (make size a)) = a.
-
-Axiom ltb_length : forall A (t:array A), length t <= max_array_length = true.
-
-Axiom length_make : forall A size (a:A),
- length (make size a) = if size <= max_array_length then size else max_array_length.
-Axiom length_set : forall A t i (a:A),
- length (t.[i<-a]) = length t.
-
-Axiom get_copy : forall A (t:array A) i, (copy t).[i] = t.[i].
-Axiom length_copy : forall A (t:array A), length (copy t) = length t.
-
-Axiom get_reroot : forall A (t:array A) i, (reroot t).[i] = t.[i].
-Axiom length_reroot : forall A (t:array A), length (reroot t) = length t.
-
-
-Axiom length_init : forall A f size (def:A),
- length (init size f def) = if size <= max_array_length then size else max_array_length.
-
-Axiom get_init : forall A f size (def:A) i,
- (init size f def).[i] = if i < length (init size f def) then f i else def.
-
-Axiom default_init : forall A f size (def:A), default (init size f def) = def.
-
-(* Not true in this implementation (see #71, many thanks to Andres Erbsen) *)
-(*
-Axiom get_ext : forall A (t1 t2:array A),
- length t1 = length t2 ->
- (forall i, i < length t1 = true -> t1.[i] = t2.[i]) ->
- default t1 = default t2 ->
- t1 = t2.
-*)
-
-(* Definition *)
-Definition to_list {A:Type} (t:array A) :=
- let len := length t in
- if 0 == len then nil
- else foldi_down (fun i l => t.[i] :: l)%list (len - 1) 0 nil.
-
-Definition forallbi {A:Type} (f:int-> A->bool) (t:array A) :=
- let len := length t in
- if 0 == len then true
- else forallb (fun i => f i (t.[i])) 0 (len - 1).
-
-Definition forallb {A:Type} (f: A->bool) (t:array A) :=
- let len := length t in
- if 0 == len then true
- else forallb (fun i => f (t.[i])) 0 (len - 1).
-
-Definition existsbi {A:Type} (f:int->A->bool) (t:array A) :=
- let len := length t in
- if 0 == len then false
- else existsb (fun i => f i (t.[i])) 0 (len - 1).
-
-Definition existsb {A:Type} (f:A->bool) (t:array A) :=
- let len := length t in
- if 0 == len then false
- else existsb (fun i => f (t.[i])) 0 (len - 1).
-
-(* TODO : We should add init as native and add it *)
-Definition mapi {A B:Type} (f:int->A->B) (t:array A) :=
- let size := length t in
- let def := f size (default t) in
- let tb := make size def in
- if size == 0 then tb
- else foldi (fun i tb => tb.[i<- f i (t.[i])]) 0 (size - 1) tb.
-
-Definition foldi_left {A B:Type} (f:int -> A -> B -> A) a (t:array B) :=
- let len := length t in
- if 0 == len then a
- else foldi (fun i a => f i a (t.[i])) 0 (len - 1) a.
-
-Definition fold_left {A B:Type} (f: A -> B -> A) a (t:array B) :=
- let len := length t in
- if 0 == len then a
- else foldi (fun i a => f a (t.[i])) 0 (length t - 1) a.
-
-Definition foldi_right {A B:Type} (f:int -> A -> B -> B) (t:array A) b :=
- let len := length t in
- if 0 == len then b
- else foldi_down (fun i b => f i (t.[i]) b) (len - 1) 0 b.
-
-Definition fold_right {A B:Type} (f: A -> B -> B) (t:array A) b :=
- let len := length t in
- if 0 == len then b
- else foldi_down (fun i b => f (t.[i]) b) (len - 1) 0 b.
-
-(* Lemmas *)
-
-Lemma default_copy : forall A (t:array A), default (copy t) = default t.
-Proof.
- intros A t;assert (length t < length t = false).
- apply Bool.not_true_is_false; apply leb_not_gtb; apply leb_refl.
- rewrite <- (get_outofbound _ (copy t) (length t)), <- (get_outofbound _ t (length t)), get_copy;trivial.
+Require FSets.FMapFacts.
+Module P := FSets.FMapFacts.WProperties_fun IntOrderedType Map.
+
+Lemma get_outofbound : forall A (t:array A) i, (i < length t) = false -> t.[i] = default t.
+intros A t i; unfold get.
+destruct t as ((t, d), l).
+unfold length; intro Hi; rewrite Hi; clear Hi.
+reflexivity.
Qed.
-Lemma reroot_default : forall A (t:array A), default (reroot t) = default t.
-Proof.
- intros A t;assert (length t < length t = false).
- apply Bool.not_true_is_false; apply leb_not_gtb; apply leb_refl.
- rewrite <- (get_outofbound _ (reroot t) (length t)), <- (get_outofbound _ t (length t)), get_reroot;trivial.
+Lemma get_set_same : forall A t i (a:A), (i < length t) = true -> t.[i<-a].[i] = a.
+intros A t i a.
+destruct t as ((t, d), l).
+unfold set, get, length.
+intro Hi; generalize Hi.
+rewrite ltb_spec.
+rewrite Z.lt_nge.
+rewrite <- leb_spec.
+rewrite Bool.not_true_iff_false.
+intro Hi'; rewrite Hi'; clear Hi'.
+rewrite Hi; clear Hi.
+rewrite P.F.add_eq_o.
+reflexivity.
+rewrite eqb_spec.
+reflexivity.
Qed.
-Lemma get_set_same_default :
- forall (A : Type) (t : array A) (i : int) ,
- (t .[ i <- default t]) .[ i] = default t.
-Proof.
- intros A t i;case_eq (i < (length t));intros.
- rewrite get_set_same;trivial.
- rewrite get_outofbound, default_set;trivial.
- rewrite length_set;trivial.
+Lemma get_set_other : forall A t i j (a:A), i <> j -> t.[i<-a].[j] = t.[j].
+intros A t i j a Hij.
+destruct t as ((t, d), l).
+unfold set, get, length.
+case (l <= i).
+reflexivity.
+rewrite P.F.add_neq_o.
+reflexivity.
+intro H; apply Hij; clear Hij.
+rewrite eqb_spec in H.
+assumption.
Qed.
-Lemma get_not_default_lt : forall A (t:array A) x,
- t.[x] <> default t -> x < length t = true.
-Proof.
- intros A t x Hd.
- case_eq (x < length t);intros Heq;[trivial | ].
- elim Hd;rewrite get_outofbound;trivial.
+Lemma default_set : forall A t i (a:A), default (t.[i<-a]) = default t.
+intros A t i a.
+destruct t as ((t, d), l).
+unfold default, set.
+case (l <= i); reflexivity.
Qed.
-Lemma foldi_left_Ind :
- forall A B (P : int -> A -> Prop) (f : int -> A -> B -> A) (t:array B),
- (forall a i, i < length t = true -> P i a -> P (i+1) (f i a (t.[i]))) ->
- forall a, P 0 a ->
- P (length t) (foldi_left f a t).
-Proof.
- intros;unfold foldi_left.
- destruct (reflect_eqb 0 (length t)).
- rewrite <- e;trivial.
- assert ((length t - 1) + 1 = length t) by ring.
- rewrite <- H1 at 1;apply foldi_Ind;auto.
- assert (W:= leb_max_int (length t));rewrite leb_spec in W.
- rewrite ltb_spec, to_Z_sub_1_diff;auto with zarith.
- intros Hlt;elim (ltb_0 _ Hlt).
- intros;apply H;trivial. rewrite ltb_leb_sub1;auto.
+Lemma get_make : forall A (a:A) size i, (make size a).[i] = a.
+intros A a size i.
+unfold make, get.
+rewrite P.F.empty_o.
+case (i < size); reflexivity.
Qed.
-Lemma fold_left_Ind :
- forall A B (P : int -> A -> Prop) (f : A -> B -> A) (t:array B),
- (forall a i, i < length t = true -> P i a -> P (i+1) (f a (t.[i]))) ->
- forall a, P 0 a ->
- P (length t) (fold_left f a t).
-Proof.
- intros.
- apply (foldi_left_Ind A B P (fun i => f));trivial.
+Lemma leb_length : forall A (t:array A), length t <= max_length = true.
+intros A t.
+generalize (length t); clear t.
+intro i.
+rewrite leb_spec.
+apply Z.lt_succ_r.
+change (Z.succ (to_Z max_length)) with wB.
+apply to_Z_bounded.
Qed.
-Lemma fold_left_ind :
- forall A B (P : A -> Prop) (f : A -> B -> A) (t:array B),
- (forall a i, i < length t = true -> P a -> P (f a (t.[i]))) ->
- forall a, P a ->
- P (fold_left f a t).
-Proof.
- intros;apply (fold_left_Ind A B (fun _ => P));trivial.
+Lemma length_make : forall A size (a:A),
+ length (make size a) = if size <= max_length then size else max_length.
+intros A size a.
+unfold length, make.
+replace (size <= max_length) with true.
+reflexivity.
+symmetry.
+rewrite leb_spec.
+apply Z.lt_succ_r.
+change (Z.succ (to_Z max_length)) with wB.
+apply to_Z_bounded.
Qed.
-Lemma foldi_right_Ind :
- forall A B (P : int -> A -> Prop) (f : int -> B -> A -> A) (t:array B),
- (forall a i, i < length t = true -> P (i+1) a -> P i (f i (t.[i]) a)) ->
- forall a, P (length t) a ->
- P 0 (foldi_right f t a).
-Proof.
- intros;unfold foldi_right.
- destruct (reflect_eqb 0 (length t)).
- rewrite e;trivial.
- set (P' z a := (*-1 <= z < [|length t|] ->*) P (of_Z (z + 1)) a).
- assert (P' ([|0|] - 1)%Z (foldi_down (fun (i : int) (b : A) => f i (t .[ i]) b) (length t - 1) 0 a)).
- apply foldi_down_ZInd;unfold P'.
- intros Hlt;elim (ltb_0 _ Hlt).
- rewrite to_Z_sub_1_diff;auto.
- ring_simplify ([|length t|] - 1 + 1)%Z;rewrite of_to_Z;trivial.
- intros;ring_simplify ([|i|] - 1 + 1)%Z;rewrite of_to_Z;auto.
- assert (i < length t = true).
- rewrite ltb_leb_sub1;auto.
- apply H;trivial.
- exact H1.
-Qed.
-
-Lemma fold_right_Ind :
- forall A B (P : int -> A -> Prop) (f : B -> A -> A) (t:array B),
- (forall a i, i < length t = true -> P (i+1) a -> P i (f (t.[i]) a)) ->
- forall a, P (length t) a ->
- P 0 (fold_right f t a).
-Proof.
- intros;apply (foldi_right_Ind A B P (fun i => f));trivial.
+Lemma length_set : forall A t i (a:A),
+ length (t.[i<-a]) = length t.
+intros A t i a.
+destruct t as ((t, d), l).
+unfold length, set.
+case (l <= i); reflexivity.
Qed.
-Lemma fold_right_ind :
- forall A B (P : A -> Prop) (f : B -> A -> A) (t:array B),
- (forall a i, i < length t = true -> P a -> P (f (t.[i]) a)) ->
- forall a, P a ->
- P (fold_right f t a).
-Proof.
- intros;apply (fold_right_Ind A B (fun i => P));trivial.
+Lemma get_copy : forall A (t:array A) i, (copy t).[i] = t.[i].
+intros A t i.
+unfold copy; reflexivity.
Qed.
-Lemma forallbi_spec : forall A (f : int -> A -> bool) t,
- forallbi f t = true <-> forall i, i < length t = true -> f i (t.[i]) = true.
-Proof.
- unfold forallbi;intros A f t.
- destruct (reflect_eqb 0 (length t)).
- split;[intros | trivial].
- elim (ltb_0 i);rewrite e;trivial.
- rewrite forallb_spec;split;intros Hi i;intros;apply Hi.
- apply leb_0. rewrite <- ltb_leb_sub1;auto. rewrite ltb_leb_sub1;auto.
+Lemma length_copy : forall A (t:array A), length (copy t) = length t.
+intros A t.
+unfold copy; reflexivity.
Qed.
-Lemma forallb_spec : forall A (f : A -> bool) t,
- forallb f t = true <-> forall i, i < length t = true -> f (t.[i]) = true.
-Proof.
- intros A f;apply (forallbi_spec A (fun i => f)).
-Qed.
+(* Not true in this implementation (see #71, many thanks to Andres Erbsen) *)
+(*
+Axiom array_ext : forall A (t1 t2:array A),
+ length t1 = length t2 ->
+ (forall i, i < length t1 = true -> t1.[i] = t2.[i]) ->
+ default t1 = default t2 ->
+ t1 = t2.
+*)
-Lemma existsbi_spec : forall A (f : int -> A -> bool) t,
- existsbi f t = true <-> exists i, i < length t = true /\ f i (t.[i]) = true.
-Proof.
- unfold existsbi;intros A f t.
- destruct (reflect_eqb 0 (length t)).
- split;[discriminate | intros [i [Hi _]];rewrite <- e in Hi;elim (ltb_0 _ Hi)].
- rewrite existsb_spec. repeat setoid_rewrite Bool.andb_true_iff.
- split;intros [i H];decompose [and] H;clear H;exists i;repeat split;trivial.
- rewrite ltb_leb_sub1;auto. apply leb_0. rewrite <- ltb_leb_sub1;auto.
-Qed.
+(* Lemmas *)
-Lemma existsb_spec : forall A (f : A -> bool) t,
- existsb f t = true <-> exists i, i < length t = true /\ f (t.[i]) = true.
-Proof.
- intros A f;apply (existsbi_spec A (fun i => f)).
+Lemma default_copy A (t:array A) : default (copy t) = default t.
+unfold copy; reflexivity.
Qed.
-Local Open Scope list_scope.
-
-Definition to_list_ntr A (t:array A) :=
- let len := length t in
- if 0 == len then nil
- else foldi_ntr _ (fun i l => t.[i] :: l) 0 (len - 1) nil.
-
-Lemma to_list_to_list_ntr : forall A (t:array A),
- to_list t = to_list_ntr _ t.
-Proof.
- unfold to_list, to_list_ntr; intros A t.
- destruct (reflect_eqb 0 (length t));trivial.
- rewrite foldi_ntr_foldi_down;trivial.
- apply leb_ltb_trans with max_array_length;[ | trivial].
- apply leb_trans with (length t);[ | apply ltb_length].
- rewrite leb_spec, sub_spec.
- rewrite to_Z_1, Zmod_small;try omega.
- generalize (to_Z_bounded (length t)).
- assert (0%Z <> [|length t|]);[ | omega].
- intros Heq;elim n;apply to_Z_inj;trivial.
+Lemma default_make A (a : A) size : default (make size a) = a.
+unfold default, make; reflexivity.
Qed.
-Lemma fold_left_to_list : forall (A B:Type) (t:array A) (f: B -> A -> B) b,
- fold_left f b t = List.fold_left f (to_list t) b.
-Proof.
- intros A B t f;rewrite to_list_to_list_ntr.
- unfold fold_left, to_list_ntr; destruct (reflect_eqb 0 (length t));[trivial | ].
- set (P1 := fun i => forall b,
- foldi (fun (i : int) (a : B) => f a (t .[ i])) i (length t - 1) b =
- List.fold_left f
- (foldi_ntr (list A) (fun (i : int) (l : list A) => t .[ i] :: l) i
- (length t - 1) nil) b).
- assert (W: P1 0);[ | trivial].
- apply int_ind_bounded with (max := length t - 1);unfold P1.
- apply leb_0.
- intros b;unfold foldi_ntr;rewrite foldi_eq, foldi_cont_eq;trivial.
- intros i _ Hlt Hrec b.
- unfold foldi_ntr;rewrite foldi_lt, foldi_cont_lt;trivial;simpl.
- apply Hrec.
+Lemma get_set_same_default A (t : array A) (i : int) : t.[i <- default t].[i] = default t.
+unfold default, get, set.
+destruct t as ((t, d), l).
+case_eq (i < l).
+intro H; generalize H.
+rewrite ltb_spec.
+rewrite Z.lt_nge.
+rewrite <- leb_spec.
+rewrite Bool.not_true_iff_false.
+intro H'; rewrite H'; clear H'.
+rewrite H; clear H.
+rewrite P.F.add_eq_o.
+reflexivity.
+rewrite eqb_spec.
+reflexivity.
+intro H; generalize H.
+rewrite <- Bool.not_true_iff_false.
+rewrite ltb_spec.
+rewrite <- Z.le_ngt.
+rewrite <- leb_spec.
+intro H'; rewrite H'; clear H'.
+rewrite H.
+reflexivity.
Qed.
-Require Import Bool.
-Local Open Scope bool_scope.
-
-Definition eqb {A:Type} (Aeqb: A->A->bool) (t1 t2:array A) :=
- (length t1 == length t2) &&
- Aeqb (default t1) (default t2) &&
- forallbi (fun i a1 => Aeqb a1 (t2.[i])) t1.
-
-(*
-Lemma reflect_eqb : forall (A:Type) (Aeqb:A->A->bool),
- (forall a1 a2, reflect (a1 = a2) (Aeqb a1 a2)) ->
- forall t1 t2, reflect (t1 = t2) (eqb Aeqb t1 t2).
-Proof.
- intros A Aeqb HA t1 t2.
- case_eq (eqb Aeqb t1 t2);unfold eqb;intros H;constructor.
- rewrite !andb_true_iff in H;destruct H as [[H1 H2] H3].
- apply get_ext.
- rewrite (reflect_iff _ _ (reflect_eqb _ _));trivial.
- rewrite forallbi_spec in H3.
- intros i Hlt;rewrite (reflect_iff _ _ (HA _ _));auto.
- rewrite (reflect_iff _ _ (HA _ _));trivial.
- intros Heq;rewrite Heq in H;clear Heq.
- revert H; rewrite Int63Axioms.eqb_refl;simpl.
- case_eq (Aeqb (default t2) (default t2));simpl;intros H0 H1.
- rewrite <- not_true_iff_false, forallbi_spec in H1;apply H1.
- intros i _; rewrite <- (reflect_iff _ _ (HA _ _));trivial.
- rewrite <- not_true_iff_false, <- (reflect_iff _ _ (HA _ _)) in H0;apply H0;trivial.
+Lemma get_not_default_lt A (t:array A) x :
+ t.[x] <> default t -> (x < length t) = true.
+unfold get, default, length.
+destruct t as ((t, d), l).
+case (x < l); tauto.
Qed.
-*)
-
(*
Local Variables:
diff --git a/src/Int63/Int63.v b/src/Int63/Int63.v
deleted file mode 100644
index acee305..0000000
--- a/src/Int63/Int63.v
+++ /dev/null
@@ -1,23 +0,0 @@
-(**************************************************************************)
-(* *)
-(* SMTCoq *)
-(* Copyright (C) 2011 - 2021 *)
-(* *)
-(* See file "AUTHORS" for the list of authors *)
-(* *)
-(* This file is distributed under the terms of the CeCILL-C licence *)
-(* *)
-(**************************************************************************)
-
-
-(** Glue with the Int31 library of standard coq, which is linked to
- native integers during VM computations.
-
- CAUTION: The name "Int63" is given for compatibility reasons, but
- int31 is used. **)
-
-Require Export Ring31.
-Require Export Int63Native.
-Require Export Int63Op.
-Require Export Int63Axioms.
-Require Export Int63Properties.
diff --git a/src/Int63/Int63Axioms.v b/src/Int63/Int63Axioms.v
deleted file mode 100644
index 9625bce..0000000
--- a/src/Int63/Int63Axioms.v
+++ /dev/null
@@ -1,313 +0,0 @@
-(**************************************************************************)
-(* *)
-(* SMTCoq *)
-(* Copyright (C) 2011 - 2021 *)
-(* *)
-(* See file "AUTHORS" for the list of authors *)
-(* *)
-(* This file is distributed under the terms of the CeCILL-C licence *)
-(* *)
-(**************************************************************************)
-
-
-Require Import Bvector.
-(* Require Export BigNumPrelude. *)
-Require Import Int31 Cyclic31.
-Require Export Int63Native.
-Require Export Int63Op.
-Require Import Psatz.
-
-Local Open Scope Z_scope.
-
-
-(* Taken from BigNumPrelude *)
-
- Lemma div_le_0 : forall p x, 0 <= x -> 0 <= x / 2 ^ p.
- Proof.
- intros p x Hle;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_le_lower_bound;auto with zarith.
- replace (2^p) with 0.
- destruct x;compute;intro;discriminate.
- destruct p;trivial;discriminate.
- Qed.
-
- Lemma div_lt : forall p x y, 0 <= x < y -> x / 2^p < y.
- Proof.
- intros p x y H;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_lt_upper_bound;auto with zarith.
- apply Z.lt_le_trans with y;auto with zarith.
- rewrite <- (Z.mul_1_r y);apply Z.mul_le_mono_nonneg;auto with zarith.
- assert (0 < 2^p);auto with zarith.
- replace (2^p) with 0.
- destruct x;change (0<y);auto with zarith.
- destruct p;trivial;discriminate.
- Qed.
-
-
-(* Int63Axioms *)
-
-Definition wB := (2^(Z_of_nat size)).
-
-Notation "[| x |]" := (to_Z x) (at level 0, x at level 99) : int63_scope.
-
-Notation "[+| c |]" :=
- (interp_carry 1 wB to_Z c) (at level 0, c at level 99) : int63_scope.
-
-Notation "[-| c |]" :=
- (interp_carry (-1) wB to_Z c) (at level 0, c at level 99) : int63_scope.
-
-Notation "[|| x ||]" :=
- (zn2z_to_Z wB to_Z x) (at level 0, x at level 99) : int63_scope.
-
-Local Open Scope int63_scope.
-Local Open Scope Z_scope.
-
-(* Bijection : int63 <-> Bvector size *)
-
-Theorem to_Z_inj : forall x y, [|x|] = [|y|] -> x = y.
-Proof Ring31.Int31_canonic.
-
-Theorem of_to_Z : forall x, of_Z ([|x|]) = x.
-Proof. exact phi_inv_phi. Qed.
-
-(* Comparisons *)
-Theorem eqb_refl x : (x == x)%int = true.
-Proof. now rewrite Ring31.eqb31_eq. Qed.
-
-Theorem ltb_spec x y : (x < y)%int = true <-> [|x|] < [|y|].
-Proof.
- unfold ltb. rewrite spec_compare, <- Z.compare_lt_iff.
- change (phi x ?= phi y) with ([|x|] ?= [|y|]).
- case ([|x|] ?= [|y|]); intuition; discriminate.
-Qed.
-
-Theorem leb_spec x y : (x <= y)%int = true <-> [|x|] <= [|y|].
-Proof.
- unfold leb. rewrite spec_compare, <- Z.compare_le_iff.
- change (phi x ?= phi y) with ([|x|] ?= [|y|]).
- case ([|x|] ?= [|y|]); intuition; discriminate.
-Qed.
-
-
-(** Specification of logical operations *)
-Lemma lsl_spec_alt p :
- forall x, [| addmuldiv31_alt p x 0 |] = ([|x|] * 2^(Z.of_nat p)) mod wB.
-Proof.
- induction p as [ |p IHp]; simpl; intro x.
- - rewrite Z.mul_1_r. symmetry. apply Zmod_small. apply phi_bounded.
- - rewrite IHp, phi_twice, Zmult_mod_idemp_l, Z.double_spec,
- Z.mul_comm, Z.mul_assoc, Z.mul_comm,
- Z.pow_pos_fold, Zpos_P_of_succ_nat, <- Z.add_1_r, Z.pow_add_r.
- * reflexivity.
- * apply Zle_0_nat.
- * exact Z.le_0_1.
-Qed.
-
-Theorem lsl_spec x p : [| x << p |] = ([|x|] * 2^[|p|]) mod wB.
-Proof.
- unfold lsl.
- rewrite addmuldiv31_equiv, lsl_spec_alt, Nat2Z.inj_abs_nat, Z.abs_eq.
- - reflexivity.
- - now destruct (phi_bounded p).
-Qed.
-
-
-Lemma div_greater (p x:int) (H:Z.of_nat Int31.size <= [|p|]) : [|x|] / 2 ^ [|p|] = 0.
-Proof.
- apply Z.div_small. destruct (phi_bounded x) as [H1 H2]. split; auto.
- apply (Z.lt_le_trans _ _ _ H2). apply Z.pow_le_mono_r; auto.
- exact Z.lt_0_2.
-Qed.
-
-Theorem lsr_spec x p : [|x >> p|] = [|x|] / 2 ^ [|p|].
-Proof.
- unfold lsr. case_eq (p < 31%int31)%int; intro Heq.
- - assert (H : [|31%int31 - p|] = 31 - [|p|]).
- * rewrite spec_sub. rewrite Zmod_small; auto.
- split.
- + rewrite ltb_spec in Heq. assert (forall x y, x < y -> 0 <= y - x) by (intros;lia); auto.
- + assert (H:forall x y z, 0 <= y /\ x < z -> x - y < z) by (intros;lia).
- apply H. destruct (phi_bounded p). destruct (phi_bounded (31%int31)). split; auto.
- * rewrite spec_add_mul_div.
- + rewrite Z.add_0_l. change (phi (31%int31 - p)) with [|31%int31 - p|]. rewrite H. replace (31 - (31 - [|p|])) with [|p|] by ring. apply Zmod_small. split.
- ++ apply div_le_0. now destruct (phi_bounded x).
- ++ apply div_lt. apply phi_bounded.
- + change (phi (31%int31 - p)) with [|31%int31 - p|]. rewrite H. assert (forall x y, 0 <= y -> x - y <= x) by (intros;lia). apply H0. now destruct (phi_bounded p).
- - rewrite div_greater; auto.
- destruct (Z.le_gt_cases [|31%int31|] [|p|]); auto.
- rewrite <- ltb_spec in H. rewrite H in Heq. discriminate.
-Qed.
-
-
-Lemma bit_testbit x i : bit x i = Z.testbit [|x|] [|i|].
-Admitted.
-(* Proof. *)
-(* case_eq [|i|]. *)
-(* - simpl. change 0 with [|0|]. intro Heq. apply Ring31.Int31_canonic in Heq. subst i. *)
-(* unfold bit. *)
-
-
-Lemma Z_pos_xO_pow i x (Hi:0 <= i) : Z.pos x < 2 ^ i <-> Z.pos x~0 < 2 ^ (i+1).
-Proof. rewrite Pos2Z.inj_xO, Z.pow_add_r; auto using Z.le_0_1; lia. Qed.
-
-Lemma Z_pos_xI_pow i x (Hi:0 <= i) : Z.pos x < 2 ^ i <-> Z.pos x~1 < 2 ^ (i+1).
-Proof. rewrite Pos2Z.inj_xI, Z.pow_add_r; auto using Z.le_0_1; lia. Qed.
-
-Lemma pow_nonneg i (Hi : 1 <= 2 ^ i) : 0 <= i.
-Proof.
- destruct (Z.le_gt_cases 0 i); auto.
- rewrite (Z.pow_neg_r _ _ H) in Hi. lia.
-Qed.
-
-Lemma pow_pos i (Hi : 1 < 2 ^ i) : 0 < i.
-Proof.
- destruct (Z.lt_trichotomy 0 i) as [H|[H|H]]; auto.
- - subst i. lia.
- - rewrite (Z.pow_neg_r _ _ H) in Hi. lia.
-Qed.
-
-Lemma pos_land_bounded : forall x y i,
- Z.pos x < 2 ^ i -> Z.pos y < 2 ^ i -> Z.of_N (Pos.land x y) < 2 ^ i.
-Proof.
- induction x as [x IHx|x IHx| ]; intros [y|y| ] i; auto.
- - simpl. intro H.
- assert (H4:0 <= i - 1) by (assert (H4:0 < i); try lia; apply pow_pos; apply (Z.le_lt_trans _ (Z.pos x~1)); auto; lia).
- generalize H. replace i with ((i-1)+1) at 1 2 by ring. rewrite <- !Z_pos_xI_pow; auto. intros H1 H2.
- assert (H3:=IHx _ _ H1 H2).
- unfold Pos.Nsucc_double. case_eq (Pos.land x y).
- * intros _. eapply Z.le_lt_trans; [ |exact H]. clear. lia.
- * intros p Hp. revert H3. rewrite Hp, N2Z.inj_pos, Z_pos_xI_pow; auto.
- replace (i - 1 + 1) with i by ring. clear. lia.
- - simpl. intro H.
- assert (H4:0 <= i - 1) by (assert (H4:0 < i); try lia; apply pow_pos; apply (Z.le_lt_trans _ (Z.pos x~1)); auto; lia).
- generalize H. replace i with ((i-1)+1) at 1 2 by ring. rewrite <- Z_pos_xI_pow, <- Z_pos_xO_pow; auto. intros H1 H2.
- assert (H3:=IHx _ _ H1 H2).
- unfold Pos.Ndouble. case_eq (Pos.land x y).
- * intros _. eapply Z.le_lt_trans; [ |exact H]. clear. lia.
- * intros p Hp. revert H3. rewrite Hp, N2Z.inj_pos, Z_pos_xO_pow; auto.
- replace (i - 1 + 1) with i by ring. clear. lia.
- - simpl. intro H.
- assert (H4:0 <= i - 1) by (assert (H4:0 < i); try lia; apply pow_pos; apply (Z.le_lt_trans _ (Z.pos x~0)); auto; lia).
- generalize H. replace i with ((i-1)+1) at 1 2 by ring. rewrite <- Z_pos_xI_pow, <- Z_pos_xO_pow; auto. intros H1 H2.
- assert (H3:=IHx _ _ H1 H2).
- unfold Pos.Ndouble. case_eq (Pos.land x y).
- * intros _. eapply Z.le_lt_trans; [ |exact H]. clear. lia.
- * intros p Hp. revert H3. rewrite Hp, N2Z.inj_pos, Z_pos_xO_pow; auto.
- replace (i - 1 + 1) with i by ring. clear. lia.
- - simpl. intro H.
- assert (H4:0 <= i - 1) by (assert (H4:0 < i); try lia; apply pow_pos; apply (Z.le_lt_trans _ (Z.pos x~0)); auto; lia).
- generalize H. replace i with ((i-1)+1) at 1 2 by ring. rewrite <- !Z_pos_xO_pow; auto. intros H1 H2.
- assert (H3:=IHx _ _ H1 H2).
- unfold Pos.Ndouble. case_eq (Pos.land x y).
- * intros _. eapply Z.le_lt_trans; [ |exact H]. clear. lia.
- * intros p Hp. revert H3. rewrite Hp, N2Z.inj_pos, Z_pos_xO_pow; auto.
- replace (i - 1 + 1) with i by ring. clear. lia.
- - simpl. intros H _. apply (Z.le_lt_trans _ (Z.pos x~0)); lia.
- - simpl. intros H _. apply (Z.le_lt_trans _ 1); lia.
-Qed.
-
-
-Lemma Z_land_bounded i : forall x y,
- 0 <= x < 2 ^ i -> 0 <= y < 2 ^ i -> 0 <= Z.land x y < 2 ^ i.
-Proof.
- intros [ |p|p] [ |q|q]; auto.
- - intros [_ H1] [_ H2]. simpl. split.
- * apply N2Z.is_nonneg.
- * now apply pos_land_bounded.
-Admitted.
-
-Theorem land_spec x y i : bit (x land y) i = bit x i && bit y i.
-Proof.
- rewrite !bit_testbit. change (x land y) with (land31 x y). unfold land31.
- rewrite phi_phi_inv. rewrite Zmod_small.
- - apply Z.land_spec.
- - split.
- * rewrite Z.land_nonneg. left. now destruct (phi_bounded x).
- * now destruct (Z_land_bounded _ _ _ (phi_bounded x) (phi_bounded y)).
-Qed.
-
-
-Axiom lor_spec: forall x y i, bit (x lor y) i = bit x i || bit y i.
-
-Axiom lxor_spec: forall x y i, bit (x lxor y) i = xorb (bit x i) (bit y i).
-
-(** Specification of basic opetations *)
-
-(* Arithmetic modulo operations *)
-
-(* Remarque : les axiomes seraient plus simple si on utilise of_Z a la place :
- exemple : add_spec : forall x y, of_Z (x + y) = of_Z x + of_Z y. *)
-
-Axiom add_spec : forall x y, [|x + y|] = ([|x|] + [|y|]) mod wB.
-
-Axiom sub_spec : forall x y, [|x - y|] = ([|x|] - [|y|]) mod wB.
-
-Axiom mul_spec : forall x y, [| x * y |] = [|x|] * [|y|] mod wB.
-
-Axiom mulc_spec : forall x y, [|x|] * [|y|] = [|fst (mulc x y)|] * wB + [|snd (mulc x y)|].
-
-Axiom div_spec : forall x y, [|x / y|] = [|x|] / [|y|].
-
-Axiom mod_spec : forall x y, [|x \% y|] = [|x|] mod [|y|].
-
-(** Iterators *)
-
-Axiom foldi_cont_gt : forall A B f from to cont,
- (to < from)%int = true -> foldi_cont (A:=A) (B:=B) f from to cont = cont.
-
-Axiom foldi_cont_eq : forall A B f from to cont,
- from = to -> foldi_cont (A:=A) (B:=B) f from to cont = f from cont.
-
-Axiom foldi_cont_lt : forall A B f from to cont,
- (from < to)%int = true->
- foldi_cont (A:=A) (B:=B) f from to cont =
- f from (fun a' => foldi_cont f (from + 1%int) to cont a').
-
-Axiom foldi_down_cont_lt : forall A B f from downto cont,
- (from < downto)%int = true -> foldi_down_cont (A:=A) (B:=B) f from downto cont = cont.
-
-Axiom foldi_down_cont_eq : forall A B f from downto cont,
- from = downto -> foldi_down_cont (A:=A) (B:=B) f from downto cont = f from cont.
-
-Axiom foldi_down_cont_gt : forall A B f from downto cont,
- (downto < from)%int = true->
- foldi_down_cont (A:=A) (B:=B) f from downto cont =
- f from (fun a' => foldi_down_cont f (from-1) downto cont a').
-
-(** Print *)
-
-Axiom print_int_spec : forall x, x = print_int x.
-
-(** Axioms on operations which are just short cut *)
-
-Axiom compare_def_spec : forall x y, compare x y = compare_def x y.
-
-Axiom head0_spec : forall x, 0 < [|x|] ->
- wB/ 2 <= 2 ^ ([|head0 x|]) * [|x|] < wB.
-
-Axiom tail0_spec : forall x, 0 < [|x|] ->
- (exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|tail0 x|]))%Z.
-
-Axiom addc_def_spec : forall x y, (x +c y)%int = addc_def x y.
-
-Axiom addcarryc_def_spec : forall x y, addcarryc x y = addcarryc_def x y.
-
-Axiom subc_def_spec : forall x y, (x -c y)%int = subc_def x y.
-
-Axiom subcarryc_def_spec : forall x y, subcarryc x y = subcarryc_def x y.
-
-Axiom diveucl_def_spec : forall x y, diveucl x y = diveucl_def x y.
-
-Axiom diveucl_21_spec : forall a1 a2 b,
- let (q,r) := diveucl_21 a1 a2 b in
- ([|q|],[|r|]) = Z.div_eucl ([|a1|] * wB + [|a2|]) [|b|].
-
-Axiom addmuldiv_def_spec : forall p x y,
- addmuldiv p x y = addmuldiv_def p x y.
-
-
-(*
- Local Variables:
- coq-load-path: ((rec "../../.." "SMTCoq"))
- End:
-*)
diff --git a/src/Int63/Int63Native.v b/src/Int63/Int63Native.v
deleted file mode 100644
index 0f9d6b7..0000000
--- a/src/Int63/Int63Native.v
+++ /dev/null
@@ -1,144 +0,0 @@
-(**************************************************************************)
-(* *)
-(* SMTCoq *)
-(* Copyright (C) 2011 - 2021 *)
-(* *)
-(* See file "AUTHORS" for the list of authors *)
-(* *)
-(* This file is distributed under the terms of the CeCILL-C licence *)
-(* *)
-(**************************************************************************)
-
-
-Require Export DoubleType.
-Require Import Int31 Cyclic31 Ring31.
-Require Import ZArith.
-Require Import Bool.
-
-
-Definition size := size.
-
-Notation int := int31.
-
-Declare Scope int63_scope.
-Delimit Scope int63_scope with int.
-Bind Scope int63_scope with int.
-
-(* Some constants *)
-Notation "0" := 0%int31 : int63_scope.
-Notation "1" := 1%int31 : int63_scope.
-Notation "2" := 2%int31 : int63_scope.
-Notation "3" := 3%int31 : int63_scope.
-
-(* Comparisons *)
-Definition eqb := eqb31.
-Notation "m '==' n" := (eqb m n) (at level 70, no associativity) : int63_scope.
-
-Definition ltb : int -> int -> bool :=
- fun i j => match compare31 i j with | Lt => true | _ => false end.
-Notation "m < n" := (ltb m n) : int63_scope.
-
-Definition leb : int -> int -> bool :=
- fun i j => match compare31 i j with | Gt => false | _ => true end.
-Notation "m <= n" := (leb m n) : int63_scope.
-
-
-Lemma eqb_correct : forall i j, (i==j)%int = true -> i = j.
-Proof. exact eqb31_correct. Qed.
-
-(* Logical operations *)
-Definition lsl : int -> int -> int :=
- fun i j => addmuldiv31 j i 0.
-Infix "<<" := lsl (at level 30, no associativity) : int63_scope.
-
-Definition lsr : int -> int -> int :=
- fun i j => if (j < 31%int31)%int then addmuldiv31 (31-j)%int31 0 i else 0%int31.
-Infix ">>" := lsr (at level 30, no associativity) : int63_scope.
-
-Definition land : int -> int -> int := land31.
-Global Arguments land i j : simpl never.
-Global Opaque land.
-Infix "land" := land (at level 40, left associativity) : int63_scope.
-
-Definition lor : int -> int -> int := lor31.
-Global Arguments lor i j : simpl never.
-Global Opaque lor.
-Infix "lor" := lor (at level 40, left associativity) : int63_scope.
-
-Definition lxor : int -> int -> int := lxor31.
-Global Arguments lxor i j : simpl never.
-Global Opaque lxor.
-Infix "lxor" := lxor (at level 40, left associativity) : int63_scope.
-
-(* Arithmetic modulo operations *)
-Notation "n + m" := (add31 n m) : int63_scope.
-Notation "n - m" := (sub31 n m) : int63_scope.
-Notation "n * m" := (mul31 n m) : int63_scope.
-
-Definition mulc : int -> int -> int * int :=
- fun i j => match mul31c i j with
- | W0 => (0%int, 0%int)
- | WW h l => (h, l)
- end.
-
-Definition div : int -> int -> int :=
- fun i j => let (q,_) := div31 i j in q.
-Notation "n / m" := (div n m) : int63_scope.
-
-Definition modulo : int -> int -> int :=
- fun i j => let (_,r) := div31 i j in r.
-Notation "n '\%' m" := (modulo n m) (at level 40, left associativity) : int63_scope.
-
-
-(* Iterators *)
-
-Definition firstr i := if ((i land 1) == 0)%int then D0 else D1.
-Fixpoint recr_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A)
- (i:int31) : A :=
- match n with
- | O => case0
- | S next =>
- if (i == 0)%int then
- case0
- else
- let si := (i >> 1)%int in
- caserec (firstr i) si (recr_aux next A case0 caserec si)
- end.
-Definition recr := recr_aux size.
-Definition iter_int31 i A f :=
- recr (A->A) (fun x => x)
- (fun b si rec => match b with
- | D0 => fun x => rec (rec x)
- | D1 => fun x => f (rec (rec x))
- end)
- i.
-
-Definition foldi_cont
- {A B : Type}
- (f : int -> (A -> B) -> A -> B)
- (from to : int)
- (cont : A -> B) : A -> B :=
- if ltb to from then
- cont
- else
- let (_,r) := iter_int31 (to - from) _ (fun (jy: (int * (A -> B))%type) =>
- let (j,y) := jy in ((j-1)%int, f j y)
- ) (to, cont) in
- f from r.
-
-Definition foldi_down_cont
- {A B : Type}
- (f : int -> (A -> B) -> A -> B)
- (from downto : int)
- (cont : A -> B) : A -> B :=
- if ltb from downto then
- cont
- else
- let (_,r) := iter_int31 (from - downto) _ (fun (jy: (int * (A -> B))%type) =>
- let (j,y) := jy in ((j+1)%int, f j y)
- ) (downto, cont) in
- f from r.
-
-(* Fake print *)
-
-Definition print_int : int -> int := fun i => i.
diff --git a/src/Int63/Int63Op.v b/src/Int63/Int63Op.v
deleted file mode 100644
index bb7d9a1..0000000
--- a/src/Int63/Int63Op.v
+++ /dev/null
@@ -1,334 +0,0 @@
-(**************************************************************************)
-(* *)
-(* SMTCoq *)
-(* Copyright (C) 2011 - 2021 *)
-(* *)
-(* See file "AUTHORS" for the list of authors *)
-(* *)
-(* This file is distributed under the terms of the CeCILL-C licence *)
-(* *)
-(**************************************************************************)
-
-
-Require Import Int31 Cyclic31.
-Require Export Int63Native.
-(* Require Import BigNumPrelude. *)
-Require Import Bvector.
-
-
-Local Open Scope int63_scope.
-
-(** The number of digits as a int *)
-Definition digits := 31%int31.
-
-(** The bigger int *)
-Definition max_int := Eval vm_compute in 0 - 1.
-
-(** Access to the nth digits *)
-Definition get_digit x p := (0 < (x land (1 << p))).
-
-Definition set_digit x p (b:bool) :=
- if (0 <= p) && (p < digits) then
- if b then x lor (1 << p)
- else x land (max_int lxor (1 << p))
- else x.
-
-(** Equality to 0 *)
-Definition is_zero (i:int) := i == 0.
-
-(** Parity *)
-Definition is_even (i:int) := is_zero (i land 1).
-
-(** Bit *)
-
-Definition bit i n := negb (is_zero ((i >> n) << (digits - 1))).
-(* Register bit as PrimInline. *)
-
-(** Extra modulo operations *)
-Definition opp (i:int) := 0 - i.
-Notation "- x" := (opp x) : int63_scope.
-
-Definition oppcarry i := max_int - i.
-
-Definition succ i := i + 1.
-
-Definition pred i := i - 1.
-
-Definition addcarry i j := i + j + 1.
-
-Definition subcarry i j := i - j - 1.
-
-(** Exact arithmetic operations *)
-
-Definition addc_def x y :=
- let r := x + y in
- if r < x then C1 r else C0 r.
-(* the same but direct implementation for efficiancy *)
-Definition addc : int -> int -> carry int := add31c.
-Notation "n '+c' m" := (addc n m) (at level 50, no associativity) : int63_scope.
-
-Definition addcarryc_def x y :=
- let r := addcarry x y in
- if r <= x then C1 r else C0 r.
-(* the same but direct implementation for efficiancy *)
-Definition addcarryc : int -> int -> carry int := add31carryc.
-
-Definition subc_def x y :=
- if y <= x then C0 (x - y) else C1 (x - y).
-(* the same but direct implementation for efficiancy *)
-Definition subc : int -> int -> carry int := sub31c.
-Notation "n '-c' m" := (subc n m) (at level 50, no associativity) : int63_scope.
-
-Definition subcarryc_def x y :=
- if y < x then C0 (x - y - 1) else C1 (x - y - 1).
-(* the same but direct implementation for efficiancy *)
-Definition subcarryc : int -> int -> carry int := sub31carryc.
-
-Definition diveucl_def x y := (x/y, x\%y).
-(* the same but direct implementation for efficiancy *)
-Definition diveucl : int -> int -> int * int := div31.
-
-Definition diveucl_21 : int -> int -> int -> int * int := div3121.
-
-Definition addmuldiv_def p x y :=
- (x << p) lor (y >> (digits - p)).
-(* the same but direct implementation for efficiancy *)
-Definition addmuldiv : int -> int -> int -> int := addmuldiv31.
-
-Definition oppc (i:int) := 0 -c i.
-
-Definition succc i := i +c 1.
-
-Definition predc i := i -c 1.
-
-(** Comparison *)
-Definition compare_def x y :=
- if x < y then Lt
- else if (x == y) then Eq else Gt.
-
-Definition compare : int -> int -> comparison := compare31.
-Notation "n ?= m" := (compare n m) (at level 70, no associativity) : int63_scope.
-
-(** Exotic operations *)
-
-(** I should add the definition (like for compare) *)
-Definition head0 : int -> int := head031.
-Definition tail0 : int -> int := tail031.
-
-(** Iterators *)
-
-Definition foldi {A} (f:int -> A -> A) from to :=
- foldi_cont (fun i cont a => cont (f i a)) from to (fun a => a).
-
-Definition fold {A} (f: A -> A) from to :=
- foldi_cont (fun i cont a => cont (f a)) from to (fun a => a).
-
-Definition foldi_down {A} (f:int -> A -> A) from downto :=
- foldi_down_cont (fun i cont a => cont (f i a)) from downto (fun a => a).
-
-Definition fold_down {A} (f:A -> A) from downto :=
- foldi_down_cont (fun i cont a => cont (f a)) from downto (fun a => a).
-
-Definition forallb (f:int -> bool) from to :=
- foldi_cont (fun i cont _ => if f i then cont tt else false) from to (fun _ => true) tt.
-
-Definition existsb (f:int -> bool) from to :=
- foldi_cont (fun i cont _ => if f i then true else cont tt) from to (fun _ => false) tt.
-
-(** Translation to Z *)
-
-(* Fixpoint to_Z_rec (n:nat) (i:int) := *)
-(* match n with *)
-(* | O => 0%Z *)
-(* | S n => *)
-(* (if is_even i then Zdouble else Zdouble_plus_one) (to_Z_rec n (i >> 1)) *)
-(* end. *)
-
-(* Definition to_Z := to_Z_rec size. *)
-
-Definition to_Z := phi.
-Definition of_Z := phi_inv.
-
-(* Fixpoint of_pos_rec (n:nat) (p:positive) := *)
-(* match n, p with *)
-(* | O, _ => 0 *)
-(* | S n, xH => 1 *)
-(* | S n, xO p => (of_pos_rec n p) << 1 *)
-(* | S n, xI p => (of_pos_rec n p) << 1 lor 1 *)
-(* end. *)
-
-(* Definition of_pos := of_pos_rec size. *)
-
-(* Definition of_Z z := *)
-(* match z with *)
-(* | Zpos p => of_pos p *)
-(* | Z0 => 0 *)
-(* | Zneg p => - (of_pos p) *)
-(* end. *)
-
-(** Gcd **)
-Fixpoint gcd_rec (guard:nat) (i j:int) {struct guard} :=
- match guard with
- | O => 1
- | S p => if j == 0 then i else gcd_rec p j (i \% j)
- end.
-
-Definition gcd := gcd_rec (2*size).
-
-(** Square root functions using newton iteration **)
-
-Definition sqrt_step (rec: int -> int -> int) (i j: int) :=
- let quo := i/j in
- if quo < j then rec i ((j + (i/j)%int) >> 1)
- else j.
-
-Definition iter_sqrt :=
- Eval lazy beta delta [sqrt_step] in
- fix iter_sqrt (n: nat) (rec: int -> int -> int)
- (i j: int) {struct n} : int :=
- sqrt_step
- (fun i j => match n with
- O => rec i j
- | S n => (iter_sqrt n (iter_sqrt n rec)) i j
- end) i j.
-
-Definition sqrt i :=
- match compare 1 i with
- Gt => 0
- | Eq => 1
- | Lt => iter_sqrt size (fun i j => j) i (i >> 1)
- end.
-
-Definition high_bit := 1 << (digits - 1).
-
-Definition sqrt2_step (rec: int -> int -> int -> int)
- (ih il j: int) :=
- if ih < j then
- let (quo,_) := diveucl_21 ih il j in
- if quo < j then
- match j +c quo with
- | C0 m1 => rec ih il (m1 >> 1)
- | C1 m1 => rec ih il ((m1 >> 1) + high_bit)
- end
- else j
- else j.
-
-Definition iter2_sqrt :=
- Eval lazy beta delta [sqrt2_step] in
- fix iter2_sqrt (n: nat)
- (rec: int -> int -> int -> int)
- (ih il j: int) {struct n} : int :=
- sqrt2_step
- (fun ih il j =>
- match n with
- | O => rec ih il j
- | S n => (iter2_sqrt n (iter2_sqrt n rec)) ih il j
- end) ih il j.
-
-Definition sqrt2 ih il :=
- let s := iter2_sqrt size (fun ih il j => j) ih il max_int in
- let (ih1, il1) := mulc s s in
- match il -c il1 with
- | C0 il2 =>
- if ih1 < ih then (s, C1 il2) else (s, C0 il2)
- | C1 il2 =>
- if ih1 < (ih - 1) then (s, C1 il2) else (s, C0 il2)
- end.
-
-(* Extra function on equality *)
-
-Definition cast_digit d1 d2 :
- option (forall P : Int31.digits -> Type, P d1 -> P d2) :=
- match d1, d2 with
- | D0, D0 | D1, D1 => Some (fun P h => h)
- | _, _ => None
- end.
-
-(* May need to improve this definition, but no reported efficiency problem for the moment *)
-Definition cast i j :
- option (forall P : int -> Type, P i -> P j) :=
- match i, j return option (forall P : int -> Type, P i -> P j) with
- | I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29 d30, I31 d'0 d'1 d'2 d'3 d'4 d'5 d'6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30 =>
- match
- cast_digit d0 d'0,
- cast_digit d1 d'1,
- cast_digit d2 d'2,
- cast_digit d3 d'3,
- cast_digit d4 d'4,
- cast_digit d5 d'5,
- cast_digit d6 d'6,
- cast_digit d7 d'7,
- cast_digit d8 d'8,
- cast_digit d9 d'9,
- cast_digit d10 d'10,
- cast_digit d11 d'11,
- cast_digit d12 d'12,
- cast_digit d13 d'13,
- cast_digit d14 d'14,
- cast_digit d15 d'15,
- cast_digit d16 d'16,
- cast_digit d17 d'17,
- cast_digit d18 d'18,
- cast_digit d19 d'19,
- cast_digit d20 d'20,
- cast_digit d21 d'21,
- cast_digit d22 d'22,
- cast_digit d23 d'23,
- cast_digit d24 d'24,
- cast_digit d25 d'25,
- cast_digit d26 d'26,
- cast_digit d27 d'27,
- cast_digit d28 d'28,
- cast_digit d29 d'29,
- cast_digit d30 d'30
- with
- | Some k0,
- Some k1,
- Some k2,
- Some k3,
- Some k4,
- Some k5,
- Some k6,
- Some k7,
- Some k8,
- Some k9,
- Some k10,
- Some k11,
- Some k12,
- Some k13,
- Some k14,
- Some k15,
- Some k16,
- Some k17,
- Some k18,
- Some k19,
- Some k20,
- Some k21,
- Some k22,
- Some k23,
- Some k24,
- Some k25,
- Some k26,
- Some k27,
- Some k28,
- Some k29,
- Some k30 =>
- Some (fun P h =>
- k0 (fun d0 => P (I31 d0 d'1 d'2 d'3 d'4 d'5 d'6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k1 (fun d1 => P (I31 d0 d1 d'2 d'3 d'4 d'5 d'6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k2 (fun d2 => P (I31 d0 d1 d2 d'3 d'4 d'5 d'6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k3 (fun d3 => P (I31 d0 d1 d2 d3 d'4 d'5 d'6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k4 (fun d4 => P (I31 d0 d1 d2 d3 d4 d'5 d'6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k5 (fun d5 => P (I31 d0 d1 d2 d3 d4 d5 d'6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k6 (fun d6 => P (I31 d0 d1 d2 d3 d4 d5 d6 d'7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k7 (fun d7 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d'8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k8 (fun d8 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d'9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k9 (fun d9 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d'10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k10 (fun d10 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d'11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k11 (fun d11 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d'12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k12 (fun d12 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d'13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k13 (fun d13 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d'14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k14 (fun d14 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d'15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k15 (fun d15 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d'16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k16 (fun d16 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d'17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k17 (fun d17 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d'18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k18 (fun d18 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d'19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k19 (fun d19 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d'20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k20 (fun d20 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d'21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k21 (fun d21 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d'22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k22 (fun d22 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d'23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k23 (fun d23 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d'24 d'25 d'26 d'27 d'28 d'29 d'30)) (k24 (fun d24 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d'25 d'26 d'27 d'28 d'29 d'30)) (k25 (fun d25 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d'26 d'27 d'28 d'29 d'30)) (k26 (fun d26 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d'27 d'28 d'29 d'30)) (k27 (fun d27 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d'28 d'29 d'30)) (k28 (fun d28 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d'29 d'30)) (k29 (fun d29 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29 d'30)) (k30 (fun d30 => P (I31 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29 d30)) h)))))))))))))))))))))))))))))))
- | _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _ => None
- end
- end.
-
-
-Definition eqo i j : option (i = j) :=
- match cast i j with
- | Some k => Some (k (fun j => i = j) (refl_equal i))
- | None => None
- end.
-
-
-(*
- Local Variables:
- coq-load-path: ((rec "../../.." "SMTCoq"))
- End:
-*)
diff --git a/src/Int63/Int63Properties.v b/src/Int63/Int63Properties.v
deleted file mode 100644
index 6ec5980..0000000
--- a/src/Int63/Int63Properties.v
+++ /dev/null
@@ -1,2760 +0,0 @@
-(**************************************************************************)
-(* *)
-(* SMTCoq *)
-(* Copyright (C) 2011 - 2021 *)
-(* *)
-(* See file "AUTHORS" for the list of authors *)
-(* *)
-(* This file is distributed under the terms of the CeCILL-C licence *)
-(* *)
-(**************************************************************************)
-
-
-Require Import Zgcd_alt.
-Require Import Bvector.
-Require Import Int31 Cyclic31.
-Require Export Int63Axioms.
-Require Import Eqdep_dec.
-Require Import Psatz.
-Require Import Znumtheory Zpow_facts.
-
-Local Open Scope int63_scope.
-Local Open Scope Z_scope.
-
-
-Notation Zpower_2 := Z.pow_2_r.
-Notation Zpower_Zsucc := Z.pow_succ_r.
-
-
-(* Taken from BigNumPrelude *)
-
-Lemma Zlt0_not_eq : forall n, 0<n -> n<>0.
-Proof.
- auto with zarith.
-Qed.
-
-Definition Z_div_plus_l a b c H := Zdiv.Z_div_plus_full_l a b c (Zlt0_not_eq _ H).
-
-Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
- Proof.
- intros a b H H1;case (Z_mod_lt a b);auto with zarith;intros H2 H3;split;auto.
- case (Z.le_gt_cases b a); intros H4; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- Qed.
-
-
-(** Trivial lemmas without axiom *)
-
-Lemma wB_diff_0 : wB <> 0.
-Proof. compute;discriminate. Qed.
-
-Lemma wB_pos : 0 < wB.
-Proof. reflexivity. Qed.
-
-Lemma to_Z_0 : [|0|] = 0.
-Proof. reflexivity. Qed.
-
-Lemma to_Z_1 : [|1|] = 1.
-Proof. reflexivity. Qed.
-
-(** equality *)
-Lemma eqb_complete : forall x y, x = y -> (x == y) = true.
-Proof.
- intros x y H;rewrite H, eqb_refl;trivial.
-Qed.
-
-Lemma eqb_spec : forall x y, (x == y) = true <-> x = y.
-Proof.
- split;auto using eqb_correct, eqb_complete.
-Qed.
-
-Lemma eqb_false_spec : forall x y, (x == y) = false <-> x <> y.
-Proof.
- intros;rewrite <- not_true_iff_false, eqb_spec;split;trivial.
-Qed.
-
-Lemma eqb_false_complete : forall x y, x <> y -> (x == y) = false.
-Proof.
- intros x y;rewrite eqb_false_spec;trivial.
-Qed.
-
-Lemma eqb_false_correct : forall x y, (x == y) = false -> x <> y.
-Proof.
- intros x y;rewrite eqb_false_spec;trivial.
-Qed.
-
-Definition eqs (i j : int) : {i = j} + { i <> j } :=
- (if i == j as b return ((b = true -> i = j) -> (b = false -> i <> j) -> {i=j} + {i <> j} )
- then fun (Heq : true = true -> i = j) _ => left _ (Heq (eq_refl true))
- else fun _ (Hdiff : false = false -> i <> j) => right _ (Hdiff (eq_refl false)))
- (eqb_correct i j)
- (eqb_false_correct i j).
-
-Lemma eq_dec : forall i j:int, i = j \/ i <> j.
-Proof.
- intros i j;destruct (eqs i j);auto.
-Qed.
-
-(* TODO: fill these proofs *)
-Lemma cast_refl : forall i, cast i i = Some (fun P H => H).
-Admitted.
-(* Proof. *)
-(* unfold cast;intros. *)
-(* generalize (eqb_correct i i). *)
-(* rewrite eqb_refl;intros. *)
-(* rewrite (eq_proofs_unicity eq_dec (e (eq_refl true)) (eq_refl i));trivial. *)
-(* Qed. *)
-
-Lemma cast_diff : forall i j, i == j = false -> cast i j = None.
-Admitted.
-(* Proof. *)
-(* intros;unfold cast;intros; generalize (eqb_correct i j). *)
-(* rewrite H;trivial. *)
-(* Qed. *)
-
-Lemma eqo_refl : forall i, eqo i i = Some (eq_refl i).
-Admitted.
-(* Proof. *)
-(* unfold eqo;intros. *)
-(* generalize (eqb_correct i i). *)
-(* rewrite eqb_refl;intros. *)
-(* rewrite (eq_proofs_unicity eq_dec (e (eq_refl true)) (eq_refl i));trivial. *)
-(* Qed. *)
-
-Lemma eqo_diff : forall i j, i == j = false -> eqo i j = None.
-Admitted.
-(* Proof. *)
-(* unfold eqo;intros; generalize (eqb_correct i j). *)
-(* rewrite H;trivial. *)
-(* Qed. *)
-
-(** translation with Z *)
-Require Import Ndigits.
-
-Lemma Z_of_N_double : forall n, Z_of_N (N.double n) = Z.double (Z_of_N n).
-Proof.
- destruct n;simpl;trivial.
-Qed.
-
-Lemma Z_of_N_double_plus_one : forall n, Z_of_N (Ndouble_plus_one n) = Zdouble_plus_one (Z_of_N n).
-Proof.
- destruct n;simpl;trivial.
-Qed.
-
-Lemma to_Z_bounded : forall x, 0 <= [|x|] < wB.
-Proof. apply phi_bounded. Qed.
-(* unfold to_Z, wB;induction size;intros. *)
-(* simpl;auto with zarith. *)
-(* rewrite inj_S;simpl;assert (W:= IHn (x >> 1)%int). *)
-(* rewrite Zpower_Zsucc;auto with zarith. *)
-(* destruct (is_even x). *)
-(* rewrite Z.double_mult;auto with zarith. *)
-(* rewrite Zdouble_plus_one_mult;auto with zarith. *)
-(* Qed. *)
-
-(* TODO: move_this *)
-(* Lemma orb_true_iff : forall b1 b2, b1 || b2 = true <-> b1 = true \/ b2 = true. *)
-(* Proof. *)
-(* split;intros;[apply orb_prop | apply orb_true_intro];trivial. *)
-(* Qed. *)
-
-Lemma to_Z_eq : forall x y, [|x|] = [|y|] <-> x = y.
-Proof.
- split;intros;subst;trivial.
- apply to_Z_inj;trivial.
-Qed.
-
-Lemma leb_ltb_eqb : forall x y, ((x <= y) = (x < y) || (x == y))%int.
-Proof.
- intros.
- apply eq_true_iff_eq.
- rewrite leb_spec, orb_true_iff, ltb_spec, eqb_spec, <- to_Z_eq;lia.
-Qed.
-
-
-(** Comparison *)
-
-Lemma compare_spec :
- forall x y, compare x y = ([|x|] ?= [|y|]).
-Proof.
- intros;rewrite compare_def_spec;unfold compare_def.
- case_eq (x < y)%int;intros Heq.
- rewrite ltb_spec in Heq.
- red in Heq;rewrite Heq;trivial.
- rewrite <- not_true_iff_false, ltb_spec in Heq.
- case_eq (x == y)%int;intros Heq1.
- rewrite eqb_spec in Heq1;rewrite Heq1, Z.compare_refl;trivial.
- rewrite <- not_true_iff_false, eqb_spec in Heq1.
- symmetry;change ([|x|] > [|y|]);rewrite <- to_Z_eq in Heq1;lia.
-Qed.
-
-Lemma is_zero_spec : forall x : int, is_zero x = true <-> x = 0%int.
-Proof.
- unfold is_zero;intros;apply eqb_spec.
-Qed.
-
-
-(** Addition *)
-
-Lemma addc_spec : forall x y, [+|x +c y|] = [|x|] + [|y|].
-Proof.
- intros;rewrite addc_def_spec;unfold addc_def.
- assert (W1 := to_Z_bounded x); assert (W2 := to_Z_bounded y).
- case_eq ((x + y < x)%int).
- rewrite ltb_spec;intros.
- change (wB + [|x+y|] = [|x|] + [|y|]).
- rewrite add_spec in H |- *.
- assert ([|x|] + [|y|] >= wB).
- destruct (Z_lt_ge_dec ([|x|] + [|y|]) wB);auto with zarith.
- elimtype False;rewrite Zmod_small in H;auto with zarith.
- assert (([|x|] + [|y|]) mod wB = [|x|] + [|y|] - wB).
- symmetry;apply Zmod_unique with 1;auto with zarith.
- rewrite H1;ring.
- rewrite <- not_true_iff_false, ltb_spec;intros.
- change ([|x+y|] = [|x|] + [|y|]).
- rewrite add_spec in *.
- assert ([|x|] + [|y|] < wB).
- destruct (Z_lt_ge_dec ([|x|] + [|y|]) wB);auto with zarith.
- assert (([|x|] + [|y|]) mod wB = [|x|] + [|y|] - wB).
- symmetry;apply Zmod_unique with 1;auto with zarith.
- elim H;lia.
- rewrite Zmod_small;auto with zarith.
-Qed.
-
-
-Lemma succc_spec : forall x, [+|succc x|] = [|x|] + 1.
-Proof. intros; unfold succc; apply addc_spec. Qed.
-
-Lemma addcarry_spec : forall x y, [|addcarry x y|] = ([|x|] + [|y|] + 1) mod wB.
-Proof.
- unfold addcarry;intros.
- rewrite add_spec,add_spec,Zplus_mod_idemp_l;trivial.
-Qed.
-
-Lemma addcarryc_spec : forall x y, [+|addcarryc x y|] = [|x|] + [|y|] + 1.
-Proof.
- intros;rewrite addcarryc_def_spec;unfold addcarryc_def.
- assert (W1 := to_Z_bounded x); assert (W2 := to_Z_bounded y).
- case_eq ((addcarry x y <= x)%int).
- rewrite leb_spec;intros.
- change (wB + [|(addcarry x y)|] = [|x|] + [|y|] + 1).
- rewrite addcarry_spec in H |- *.
- assert ([|x|] + [|y|] + 1 >= wB).
- destruct (Z_lt_ge_dec ([|x|] + [|y|] + 1) wB);auto with zarith.
- elimtype False;rewrite Zmod_small in H;auto with zarith.
- assert (([|x|] + [|y|] + 1) mod wB = [|x|] + [|y|] + 1 - wB).
- symmetry;apply Zmod_unique with 1;auto with zarith.
- rewrite H1;ring.
- rewrite <- not_true_iff_false, leb_spec;intros.
- change ([|addcarry x y|] = [|x|] + [|y|] + 1).
- rewrite addcarry_spec in *.
- assert ([|x|] + [|y|] + 1 < wB).
- destruct (Z_lt_ge_dec ([|x|] + [|y|] + 1) wB);auto with zarith.
- assert (([|x|] + [|y|] + 1) mod wB = [|x|] + [|y|] + 1 - wB).
- symmetry;apply Zmod_unique with 1;auto with zarith.
- elim H;lia.
- rewrite Zmod_small;auto with zarith.
-Qed.
-
-Lemma succ_spec : forall x, [|succ x|] = ([|x|] + 1) mod wB.
-Proof. intros; apply add_spec. Qed.
-
-(** Subtraction *)
-Lemma subc_spec : forall x y, [-|x -c y|] = [|x|] - [|y|].
-Proof.
- intros;rewrite subc_def_spec;unfold subc_def.
- assert (W1 := to_Z_bounded x); assert (W2 := to_Z_bounded y).
- case_eq (y <= x)%int.
- rewrite leb_spec;intros.
- change ([|x - y|] = [|x|] - [|y|]).
- rewrite sub_spec.
- rewrite Zmod_small;auto with zarith.
- rewrite <- not_true_iff_false, leb_spec;intros.
- change (-wB + [|x - y|] = [|x|] - [|y|]).
- rewrite sub_spec.
- assert (([|x|] - [|y|]) mod wB = [|x|] - [|y|] + wB).
- symmetry;apply Zmod_unique with (-1);auto with zarith.
- rewrite H0;ring.
-Qed.
-
-Lemma subcarry_spec :
- forall x y, [|subcarry x y|] = ([|x|] - [|y|] - 1) mod wB.
-Proof.
- unfold subcarry; intros.
- rewrite sub_spec,sub_spec,Zminus_mod_idemp_l;trivial.
-Qed.
-
-Lemma subcarryc_spec : forall x y, [-|subcarryc x y|] = [|x|] - [|y|] - 1.
- intros;rewrite subcarryc_def_spec;unfold subcarryc_def.
- assert (W1 := to_Z_bounded x); assert (W2 := to_Z_bounded y).
- (* fold (subcarry x y). *)
- replace ((x - y - 1)%int) with (subcarry x y) by reflexivity.
- case_eq (y < x)%int.
- rewrite ltb_spec;intros.
- change ([|subcarry x y|] = [|x|] - [|y|] - 1).
- rewrite subcarry_spec.
- rewrite Zmod_small;auto with zarith.
- rewrite <- not_true_iff_false, ltb_spec;intros.
- change (-wB + [|subcarry x y|] = [|x|] - [|y|] - 1).
- rewrite subcarry_spec.
- assert (([|x|] - [|y|] - 1) mod wB = [|x|] - [|y|] - 1 + wB).
- symmetry;apply Zmod_unique with (-1);auto with zarith.
- rewrite H0;ring.
-Qed.
-
-Lemma oppc_spec : forall x : int, [-|oppc x|] = - [|x|].
-Proof.
- unfold oppc;intros;rewrite subc_spec, to_Z_0;trivial.
-Qed.
-
-Lemma opp_spec : forall x : int, [|- x|] = - [|x|] mod wB.
-Proof.
- unfold opp;intros. rewrite sub_spec, to_Z_0;trivial.
-Qed.
-
-Lemma oppcarry_spec : forall x, [|oppcarry x|] = wB - [|x|] - 1.
-Proof.
- unfold oppcarry;intros.
- rewrite sub_spec.
- change [|max_int|] with (wB - 1).
- rewrite <- Zminus_plus_distr, Zplus_comm, Zminus_plus_distr.
- apply Zmod_small.
- generalize (to_Z_bounded x);auto with zarith.
-Qed.
-
-Lemma predc_spec : forall x, [-|predc x|] = [|x|] - 1.
-Proof. intros; unfold predc; apply subc_spec. Qed.
-
-Lemma pred_spec : forall x, [|pred x|] = ([|x|] - 1) mod wB.
-Proof. intros; unfold pred; apply sub_spec. Qed.
-
-Lemma diveucl_spec :
- forall x y,
- let (q,r) := diveucl x y in
- ([|q|],[|r|]) = Z.div_eucl [|x|] [|y|].
-Proof.
- intros;rewrite diveucl_def_spec.
- unfold diveucl_def;rewrite div_spec, mod_spec.
- unfold Z.div, Zmod;destruct (Z.div_eucl [|x|] [|y|]);trivial.
-Qed.
-
-(* Sqrt *)
-
- (* Direct transcription of an old proof
- of a fortran program in boyer-moore *)
-
-Lemma quotient_by_2 a: a - 1 <= (a/2) + (a/2).
-Proof.
- case (Z_mod_lt a 2); auto with zarith.
- intros H1; rewrite Zmod_eq_full; auto with zarith.
-Qed.
-
-Lemma sqrt_main_trick j k: 0 <= j -> 0 <= k ->
- (j * k) + j <= ((j + k)/2 + 1) ^ 2.
-Proof.
- intros Hj; generalize Hj k; pattern j; apply natlike_ind;
- auto; clear k j Hj.
- intros _ k Hk; repeat rewrite Zplus_0_l.
- apply Zmult_le_0_compat; generalize (Z_div_pos k 2); auto with zarith.
- intros j Hj Hrec _ k Hk; pattern k; apply natlike_ind; auto; clear k Hk.
- rewrite Zmult_0_r, Zplus_0_r, Zplus_0_l.
- generalize (sqr_pos (Z.succ j / 2)) (quotient_by_2 (Z.succ j));
- unfold Z.succ.
- rewrite Zpower_2, Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r.
- auto with zarith.
- intros k Hk _.
- replace ((Z.succ j + Z.succ k) / 2) with ((j + k)/2 + 1).
- generalize (Hrec Hj k Hk) (quotient_by_2 (j + k)).
- unfold Z.succ; repeat rewrite Zpower_2;
- repeat rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r.
- repeat rewrite Zmult_1_l; repeat rewrite Zmult_1_r.
- auto with zarith.
- rewrite Zplus_comm, <- Z_div_plus_full_l; auto with zarith.
- apply f_equal2 with (f := Z.div); auto with zarith.
-Qed.
-
-Lemma sqrt_main i j: 0 <= i -> 0 < j -> i < ((j + (i/j))/2 + 1) ^ 2.
-Proof.
- intros Hi Hj.
- assert (Hij: 0 <= i/j) by (apply Z_div_pos; auto with zarith).
- apply Z.lt_le_trans with (2 := sqrt_main_trick _ _ (Zlt_le_weak _ _ Hj) Hij).
- pattern i at 1; rewrite (Z_div_mod_eq i j); case (Z_mod_lt i j); auto with zarith.
-Qed.
-
-Lemma sqrt_init i: 1 < i -> i < (i/2 + 1) ^ 2.
-Proof.
- intros Hi.
- assert (H1: 0 <= i - 2) by auto with zarith.
- assert (H2: 1 <= (i / 2) ^ 2); auto with zarith.
- replace i with (1* 2 + (i - 2)); auto with zarith.
- rewrite Zpower_2, Z_div_plus_full_l; auto with zarith.
- generalize (sqr_pos ((i - 2)/ 2)) (Z_div_pos (i - 2) 2).
- rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r.
- auto with zarith.
- generalize (quotient_by_2 i).
- rewrite Zpower_2 in H2 |- *;
- repeat (rewrite Zmult_plus_distr_l ||
- rewrite Zmult_plus_distr_r ||
- rewrite Zmult_1_l || rewrite Zmult_1_r).
- auto with zarith.
-Qed.
-
-Lemma sqrt_test_true i j: 0 <= i -> 0 < j -> i/j >= j -> j ^ 2 <= i.
-Proof.
- intros Hi Hj Hd; rewrite Zpower_2.
- apply Z.le_trans with (j * (i/j)); auto with zarith.
- apply Z_mult_div_ge; auto with zarith.
-Qed.
-
-Lemma sqrt_test_false i j: 0 <= i -> 0 < j -> i/j < j -> (j + (i/j))/2 < j.
-Proof.
- intros Hi Hj H; case (Zle_or_lt j ((j + (i/j))/2)); auto.
- intros H1; contradict H; apply Zle_not_lt.
- assert (2 * j <= j + (i/j)); auto with zarith.
- apply Z.le_trans with (2 * ((j + (i/j))/2)); auto with zarith.
- apply Z_mult_div_ge; auto with zarith.
-Qed.
-
-
-Lemma sqrt_step_correct rec i j:
- 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 ->
- 2 * [|j|] < wB ->
- (forall j1 : int,
- 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 ->
- [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) ->
- [|sqrt_step rec i j|] ^ 2 <= [|i|] < ([|sqrt_step rec i j|] + 1) ^ 2.
-Proof.
- assert (Hp2: 0 < [|2|]) by exact (refl_equal Lt).
- intros Hi Hj Hij H31 Hrec.
- unfold sqrt_step.
- case_eq ((i / j < j)%int);[ | rewrite <- Bool.not_true_iff_false];
- rewrite ltb_spec, div_spec;intros.
- assert ([| j + (i / j)%int|] = [|j|] + [|i|]/[|j|]).
- {
- rewrite add_spec, Zmod_small;rewrite div_spec; auto with zarith.
- split.
- - apply Z.add_nonneg_nonneg.
- + apply Z.lt_le_incl; apply Z.le_lt_trans with (2 := H). apply Z_div_pos.
- * apply Z.lt_gt. abstract lia.
- * abstract lia.
- + apply Z_div_pos.
- * apply Z.lt_gt. assumption.
- * abstract lia.
- - abstract lia.
- }
- apply Hrec;rewrite lsr_spec, H0, to_Z_1;change (2^1) with 2.
- split; [ | apply sqrt_test_false;auto with zarith].
- replace ([|j|] + [|i|]/[|j|]) with
- (1 * 2 + (([|j|] - 2) + [|i|] / [|j|]));[ | ring].
- rewrite Z_div_plus_full_l; auto with zarith.
- assert (0 <= [|i|]/ [|j|]) by (apply Z_div_pos; auto with zarith).
- assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / 2) ; auto with zarith.
- case (Zle_lt_or_eq 1 [|j|]); auto with zarith.
- {
- intro. apply Z_div_pos.
- - apply Zgt_pos_0.
- - apply Z.add_nonneg_nonneg.
- + abstract lia.
- + assumption.
- }
- intros Hj1.
- rewrite <- Hj1, Zdiv_1_r.
- assert (0 <= ([|i|] - 1) /2)%Z;[ |apply Z_div_pos]; auto with zarith.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - abstract lia.
- }
- apply sqrt_main;auto with zarith.
- split;[apply sqrt_test_true | ];auto with zarith.
-Qed.
-
-Lemma iter_sqrt_correct n rec i j: 0 < [|i|] -> 0 < [|j|] ->
- [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB ->
- (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] ->
- [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB ->
- [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) ->
- [|iter_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter_sqrt n rec i j|] + 1) ^ 2.
-Proof.
- revert rec i j; elim n; unfold iter_sqrt; fold iter_sqrt; clear n.
- intros rec i j Hi Hj Hij H31 Hrec. replace (and (Z.le (Z.pow (to_Z match ltb (div i j) j return int with | true => rec i (lsr (add31 j (div i j)) In) | false => j end) (Zpos (xO xH))) (to_Z i)) (Z.lt (to_Z i) (Z.pow (Z.add (to_Z match ltb (div i j) j return int with | true => rec i (lsr (add31 j (div i j)) In) | false => j end) (Zpos xH)) (Zpos (xO xH))))) with ([|sqrt_step rec i j|] ^ 2 <= [|i|] < ([|sqrt_step rec i j|] + 1) ^ 2) by reflexivity. apply sqrt_step_correct; auto with zarith.
- intros n Hrec rec i j Hi Hj Hij H31 HHrec.
- replace (and (Z.le (Z.pow (to_Z match ltb (div i j) j return int with | true => iter_sqrt n (iter_sqrt n rec) i (lsr (add31 j (div i j)) In) | false => j end) (Zpos (xO xH))) (to_Z i)) (Z.lt (to_Z i) (Z.pow (Z.add (to_Z match ltb (div i j) j return int with | true => iter_sqrt n (iter_sqrt n rec) i (lsr (add31 j (div i j)) In) | false => j end) (Zpos xH)) (Zpos (xO xH))))) with ([|sqrt_step (iter_sqrt n (iter_sqrt n rec)) i j|] ^ 2 <= [|i|] < ([|sqrt_step (iter_sqrt n (iter_sqrt n rec)) i j|] + 1) ^ 2) by reflexivity.
- apply sqrt_step_correct; auto.
- intros j1 Hj1 Hjp1; apply Hrec; auto with zarith.
- intros j2 Hj2 H2j2 Hjp2 Hj31; apply Hrec; auto with zarith.
- intros j3 Hj3 Hpj3.
- apply HHrec; auto.
- rewrite inj_S, Zpower_Zsucc.
- apply Z.le_trans with (2 ^Z_of_nat n + [|j2|]); auto with zarith.
- apply Zle_0_nat.
-Qed.
-
-Lemma sqrt_spec : forall x,
- [|sqrt x|] ^ 2 <= [|x|] < ([|sqrt x|] + 1) ^ 2.
-Proof.
- intros i; unfold sqrt.
- rewrite compare_spec. case Z.compare_spec; rewrite to_Z_1;
- intros Hi; auto with zarith.
- apply iter_sqrt_correct; auto with zarith;
- rewrite lsr_spec, to_Z_1; change (2^1) with 2; auto with zarith.
- replace ([|i|]) with (1 * 2 + ([|i|] - 2))%Z; try ring.
- assert (0 <= ([|i|] - 2)/2)%Z by (apply Z_div_pos; auto with zarith).
- rewrite Z_div_plus_full_l; auto with zarith.
- apply sqrt_init; auto.
- assert (W:= Z_mult_div_ge [|i|] 2);assert (W':= to_Z_bounded i);auto with zarith.
- intros j2 H1 H2; contradict H2; apply Zlt_not_le.
- fold wB;assert (W:=to_Z_bounded i).
- apply Z.le_lt_trans with ([|i|]); auto with zarith.
- assert (0 <= [|i|]/2)%Z by (apply Z_div_pos; auto with zarith).
- apply Z.le_trans with (2 * ([|i|]/2)); auto with zarith.
- apply Z_mult_div_ge; auto with zarith.
- case (to_Z_bounded i); repeat rewrite Zpower_2; auto with zarith.
-Qed.
-
-Lemma sqrt2_step_def rec ih il j:
- sqrt2_step rec ih il j =
- if (ih < j)%int then
- let quo := fst (diveucl_21 ih il j) in
- if (quo < j)%int then
- let m :=
- match j +c quo with
- | C0 m1 => m1 >> 1
- | C1 m1 => (m1 >> 1 + 1 << (digits -1))%int
- end in
- rec ih il m
- else j
- else j.
-Proof.
- unfold sqrt2_step; case diveucl_21; intros;simpl.
- case (j +c i);trivial.
-Qed.
-
-Lemma sqrt2_lower_bound ih il j:
- [|| WW ih il||] < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|].
-Proof.
- intros H1.
- case (to_Z_bounded j); intros Hbj _.
- case (to_Z_bounded il); intros Hbil _.
- case (to_Z_bounded ih); intros Hbih Hbih1.
- assert (([|ih|] < [|j|] + 1)%Z); auto with zarith.
- apply Zlt_square_simpl; auto with zarith.
- simpl zn2z_to_Z in H1.
- repeat rewrite <-Zpower_2; apply Z.le_lt_trans with (2 := H1).
- apply Z.le_trans with ([|ih|] * wB)%Z;try rewrite Zpower_2; auto with zarith.
-Qed.
-
-
-Lemma div2_phi ih il j:
- [|fst (diveucl_21 ih il j)|] = [|| WW ih il||] /[|j|].
-Proof.
- generalize (diveucl_21_spec ih il j).
- case diveucl_21; intros q r Heq.
- simpl zn2z_to_Z;unfold Z.div;rewrite <- Heq;trivial.
-Qed.
-
-Lemma zn2z_to_Z_pos ih il : 0 <= [||WW ih il||].
-Proof.
- simpl zn2z_to_Z;destruct (to_Z_bounded ih);destruct (to_Z_bounded il);auto with zarith.
-Qed.
-
-
-Lemma sqrt2_step_correct rec ih il j:
- 2 ^ (Z_of_nat (size - 2)) <= [|ih|] ->
- 0 < [|j|] -> [|| WW ih il||] < ([|j|] + 1) ^ 2 ->
- (forall j1, 0 < [|j1|] < [|j|] -> [|| WW ih il||] < ([|j1|] + 1) ^ 2 ->
- [|rec ih il j1|] ^ 2 <= [||WW ih il||] < ([|rec ih il j1|] + 1) ^ 2) ->
- [|sqrt2_step rec ih il j|] ^ 2 <= [||WW ih il ||]
- < ([|sqrt2_step rec ih il j|] + 1) ^ 2.
-Proof.
- assert (Hp2: (0 < [|2|])%Z) by exact (refl_equal Lt).
- intros Hih Hj Hij Hrec; rewrite sqrt2_step_def.
- assert (H1: ([|ih|] <= [|j|])%Z) by (apply sqrt2_lower_bound with il; auto).
- case (to_Z_bounded ih); intros Hih1 _.
- case (to_Z_bounded il); intros Hil1 _.
- case (to_Z_bounded j); intros _ Hj1.
- assert (Hp3: (0 < [||WW ih il||])).
- simpl zn2z_to_Z;apply Z.lt_le_trans with ([|ih|] * wB)%Z; auto with zarith.
- apply Zmult_lt_0_compat; auto with zarith.
- apply Z.lt_le_trans with (2:= Hih); auto with zarith.
- cbv zeta.
- case_eq (ih < j)%int;intros Heq.
- rewrite ltb_spec in Heq.
- 2: rewrite <-not_true_iff_false, ltb_spec in Heq.
- 2: split; auto.
- 2: apply sqrt_test_true; auto with zarith.
- 2: unfold zn2z_to_Z; replace [|ih|] with [|j|]; auto with zarith.
- 2: assert (0 <= [|il|]/[|j|]) by (apply Z_div_pos; auto with zarith).
- 2: rewrite Zmult_comm, Z_div_plus_full_l; unfold base; auto with zarith.
- case (Zle_or_lt (2^(Z_of_nat size -1)) [|j|]); intros Hjj.
- case_eq (fst (diveucl_21 ih il j) < j)%int;intros Heq0.
- 2: rewrite <-not_true_iff_false, ltb_spec, div2_phi in Heq0.
- 2: split; auto; apply sqrt_test_true; auto with zarith.
- rewrite ltb_spec, div2_phi in Heq0.
- match goal with |- context[rec _ _ ?X] =>
- set (u := X)
- end.
- assert (H: [|u|] = ([|j|] + ([||WW ih il||])/([|j|]))/2).
- unfold u; generalize (addc_spec j (fst (diveucl_21 ih il j)));
- case addc;unfold interp_carry;rewrite div2_phi;simpl zn2z_to_Z.
- intros i H;rewrite lsr_spec, H;trivial.
- intros i H;rewrite <- H.
- case (to_Z_bounded i); intros H1i H2i.
- rewrite add_spec, Zmod_small, lsr_spec.
- change (1 * wB) with ([|(1 << (digits -1))|] * 2)%Z.
- rewrite Z_div_plus_full_l; auto with zarith.
- change wB with (2 * (wB/2))%Z; auto.
- replace [|(1 << (digits - 1))|] with (wB/2); auto.
- rewrite lsr_spec; auto.
- replace (2^[|1|]) with 2%Z; auto.
- split.
- {
- apply Z.add_nonneg_nonneg.
- - apply Z_div_pos.
- + apply Zgt_pos_0.
- + assumption.
- - apply Z_div_pos.
- + apply Zgt_pos_0.
- + abstract lia.
- }
- assert ([|i|]/2 < wB/2); auto with zarith.
- apply Zdiv_lt_upper_bound; auto with zarith.
- apply Hrec; rewrite H; clear u H.
- assert (Hf1: 0 <= [||WW ih il||]/ [|j|]) by (apply Z_div_pos; auto with zarith).
- case (Zle_lt_or_eq 1 ([|j|])); auto with zarith; intros Hf2.
- 2: contradict Heq0; apply Zle_not_lt; rewrite <- Hf2, Zdiv_1_r; assert (H10: forall (x:Z), 0 < x -> 1 <= x) by (intros; lia); auto.
- split.
- replace ([|j|] + [||WW ih il||]/ [|j|])%Z with
- (1 * 2 + (([|j|] - 2) + [||WW ih il||] / [|j|])); try ring.
- rewrite Z_div_plus_full_l; auto with zarith.
- assert (0 <= ([|j|] - 2 + [||WW ih il||] / [|j|]) / 2) ; auto with zarith.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - apply Z.add_nonneg_nonneg.
- + abstract lia.
- + assumption.
- }
- apply sqrt_test_false; auto with zarith.
- apply sqrt_main; auto with zarith.
- contradict Hij; apply Zle_not_lt.
- assert ((1 + [|j|]) <= 2 ^ (Z_of_nat size - 1)); auto with zarith.
- apply Z.le_trans with ((2 ^ (Z_of_nat size - 1)) ^2); auto with zarith.
- assert (0 <= 1 + [|j|]); auto with zarith.
- apply Zmult_le_compat; auto with zarith.
- change ((2 ^ (Z_of_nat size - 1))^2) with (2 ^ (Z_of_nat size - 2) * wB).
- apply Z.le_trans with ([|ih|] * wB); auto with zarith.
- unfold zn2z_to_Z, wB; auto with zarith.
-Qed.
-
-
-
-Lemma iter2_sqrt_correct n rec ih il j:
- 2^(Z_of_nat (size - 2)) <= [|ih|] -> 0 < [|j|] -> [||WW ih il||] < ([|j|] + 1) ^ 2 ->
- (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] ->
- [||WW ih il||] < ([|j1|] + 1) ^ 2 ->
- [|rec ih il j1|] ^ 2 <= [||WW ih il||] < ([|rec ih il j1|] + 1) ^ 2) ->
- [|iter2_sqrt n rec ih il j|] ^ 2 <= [||WW ih il||]
- < ([|iter2_sqrt n rec ih il j|] + 1) ^ 2.
-Proof.
- revert rec ih il j; elim n; unfold iter2_sqrt; fold iter2_sqrt; clear n.
- intros rec ih il j Hi Hj Hij Hrec; apply sqrt2_step_correct; auto with zarith.
- intros n Hrec rec ih il j Hi Hj Hij HHrec.
- apply sqrt2_step_correct; auto.
- intros j1 Hj1 Hjp1; apply Hrec; auto with zarith.
- intros j2 Hj2 H2j2 Hjp2; apply Hrec; auto with zarith.
- intros j3 Hj3 Hpj3.
- apply HHrec; auto.
- rewrite inj_S, Zpower_Zsucc.
- apply Z.le_trans with (2 ^Z_of_nat n + [|j2|])%Z; auto with zarith.
- apply Zle_0_nat.
-Qed.
-
-
-Lemma sqrt2_spec : forall x y,
- wB/ 4 <= [|x|] ->
- let (s,r) := sqrt2 x y in
- [||WW x y||] = [|s|] ^ 2 + [+|r|] /\
- [+|r|] <= 2 * [|s|].
- Proof.
- intros ih il Hih; unfold sqrt2.
- change [||WW ih il||] with ([||WW ih il||]).
- assert (Hbin: forall s, s * s + 2* s + 1 = (s + 1) ^ 2) by
- (intros s; ring).
- assert (Hb: 0 <= wB) by (red; intros HH; discriminate).
- assert (Hi2: [||WW ih il ||] < ([|max_int|] + 1) ^ 2).
- apply Z.le_lt_trans with ((wB - 1) * wB + (wB - 1)); auto with zarith.
- 2: apply refl_equal.
- case (to_Z_bounded ih); case (to_Z_bounded il); intros H1 H2 H3 H4.
- unfold zn2z_to_Z; auto with zarith.
- case (iter2_sqrt_correct size (fun _ _ j => j) ih il max_int); auto with zarith.
- apply refl_equal.
- intros j1 _ HH; contradict HH.
- apply Zlt_not_le.
- case (to_Z_bounded j1); auto with zarith.
- change (2 ^ Z_of_nat size) with ([|max_int|]+1)%Z; auto with zarith.
- set (s := iter2_sqrt size (fun _ _ j : int=> j) ih il max_int).
- intros Hs1 Hs2.
- generalize (mulc_spec s s); case mulc.
- simpl fst; simpl snd; intros ih1 il1 Hihl1.
- generalize (subc_spec il il1).
- case subc; intros il2 Hil2.
- simpl interp_carry in Hil2.
- case_eq (ih1 < ih)%int; [idtac | rewrite <- not_true_iff_false];
- rewrite ltb_spec; intros Heq.
- unfold interp_carry; rewrite Zmult_1_l.
- rewrite Zpower_2, Hihl1, Hil2.
- case (Zle_lt_or_eq ([|ih1|] + 1) ([|ih|])); auto with zarith.
- intros H2; contradict Hs2; apply Zle_not_lt.
- replace (([|s|] + 1) ^ 2) with ([||WW ih1 il1||] + 2 * [|s|] + 1).
- unfold zn2z_to_Z.
- case (to_Z_bounded il); intros Hpil _.
- assert (Hl1l: [|il1|] <= [|il|]).
- case (to_Z_bounded il2); rewrite Hil2; auto with zarith.
- assert ([|ih1|] * wB + 2 * [|s|] + 1 <= [|ih|] * wB); auto with zarith.
- case (to_Z_bounded s); intros _ Hps.
- case (to_Z_bounded ih1); intros Hpih1 _; auto with zarith.
- apply Z.le_trans with (([|ih1|] + 2) * wB); auto with zarith.
- unfold zn2z_to_Z; rewrite <-Hihl1, Hbin; auto.
- intros H2; split.
- unfold zn2z_to_Z; rewrite <- H2; ring.
- replace (wB + ([|il|] - [|il1|])) with ([||WW ih il||] - ([|s|] * [|s|])).
- rewrite <-Hbin in Hs2; assert (([||WW ih il||] < [|s|] * [|s|] + 2 * [|s|] + 1) -> ([||WW ih il||] - [|s|] * [|s|] <= 2 * [|s|])) by lia; auto.
- rewrite Hihl1; unfold zn2z_to_Z; rewrite <- H2; ring.
- unfold interp_carry.
- case (Zle_lt_or_eq [|ih|] [|ih1|]); auto with zarith; intros H.
- contradict Hs1.
- apply Zlt_not_le; rewrite Zpower_2, Hihl1.
- unfold zn2z_to_Z.
- case (to_Z_bounded il); intros _ H2.
- apply Z.lt_le_trans with (([|ih|] + 1) * wB + 0).
- rewrite Zmult_plus_distr_l, Zplus_0_r; auto with zarith.
- case (to_Z_bounded il1); intros H3 _.
- apply Zplus_le_compat; auto with zarith.
- split.
- rewrite Zpower_2, Hihl1.
- unfold zn2z_to_Z; ring[Hil2 H].
- replace [|il2|] with ([||WW ih il||] - [||WW ih1 il1||]).
- unfold zn2z_to_Z at 2; rewrite <-Hihl1.
- rewrite <-Hbin in Hs2; assert (([||WW ih il||] < [|s|] * [|s|] + 2 * [|s|] + 1) -> ([||WW ih il||] - [|s|] * [|s|] <= 2 * [|s|])) by lia; auto.
- unfold zn2z_to_Z; rewrite H, Hil2; ring.
- unfold interp_carry in Hil2 |- *.
- assert (Hsih: [|ih - 1|] = [|ih|] - 1).
- rewrite sub_spec, Zmod_small; auto; replace [|1|] with 1; auto.
- case (to_Z_bounded ih); intros H1 H2.
- split; auto with zarith.
- apply Z.le_trans with (wB/4 - 1); auto with zarith.
- case_eq (ih1 < ih - 1)%int; [idtac | rewrite <- not_true_iff_false];
- rewrite ltb_spec, Hsih; intros Heq.
- rewrite Zpower_2, Hihl1.
- case (Zle_lt_or_eq ([|ih1|] + 2) [|ih|]); auto with zarith.
- intros H2; contradict Hs2; apply Zle_not_lt.
- replace (([|s|] + 1) ^ 2) with ([||WW ih1 il1||] + 2 * [|s|] + 1).
- unfold zn2z_to_Z.
- assert ([|ih1|] * wB + 2 * [|s|] + 1 <= [|ih|] * wB + ([|il|] - [|il1|]));
- auto with zarith.
- rewrite <-Hil2.
- case (to_Z_bounded il2); intros Hpil2 _.
- apply Z.le_trans with ([|ih|] * wB + - wB); auto with zarith.
- case (to_Z_bounded s); intros _ Hps.
- assert (2 * [|s|] + 1 <= 2 * wB); auto with zarith.
- apply Z.le_trans with ([|ih1|] * wB + 2 * wB); auto with zarith.
- assert (Hi: ([|ih1|] + 3) * wB <= [|ih|] * wB); auto with zarith.
- unfold zn2z_to_Z; rewrite <-Hihl1, Hbin; auto.
- intros H2; unfold zn2z_to_Z; rewrite <-H2.
- split.
- replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring.
- rewrite <-Hil2; ring.
- replace (1 * wB + [|il2|]) with ([||WW ih il||] - [||WW ih1 il1||]).
- unfold zn2z_to_Z at 2; rewrite <-Hihl1.
- rewrite <-Hbin in Hs2; assert (([||WW ih il||] < [|s|] * [|s|] + 2 * [|s|] + 1) -> ([||WW ih il||] - [|s|] * [|s|] <= 2 * [|s|])) by lia; auto.
- unfold zn2z_to_Z; rewrite <-H2.
- replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring.
- rewrite <-Hil2; ring.
- case (Zle_lt_or_eq ([|ih|] - 1) ([|ih1|])); auto with zarith; intros H1.
- assert (He: [|ih|] = [|ih1|]).
- apply Zle_antisym; auto with zarith.
- case (Zle_or_lt [|ih1|] [|ih|]); auto; intros H2.
- contradict Hs1; apply Zlt_not_le; rewrite Zpower_2, Hihl1.
- unfold zn2z_to_Z.
- case (to_Z_bounded il); intros _ Hpil1.
- apply Z.lt_le_trans with (([|ih|] + 1) * wB).
- rewrite Zmult_plus_distr_l, Zmult_1_l; auto with zarith.
- case (to_Z_bounded il1); intros Hpil2 _.
- apply Z.le_trans with (([|ih1|]) * wB); auto with zarith.
- contradict Hs1; apply Zlt_not_le; rewrite Zpower_2, Hihl1.
- unfold zn2z_to_Z; rewrite He.
- assert ([|il|] - [|il1|] < 0); auto with zarith.
- rewrite <-Hil2.
- case (to_Z_bounded il2); auto with zarith.
- split.
- rewrite Zpower_2, Hihl1.
- unfold zn2z_to_Z; rewrite <-H1.
- apply trans_equal with ([|ih|] * wB + [|il1|] + ([|il|] - [|il1|])).
- ring.
- rewrite <-Hil2; ring.
- replace [|il2|] with ([||WW ih il||] - [||WW ih1 il1||]).
- unfold zn2z_to_Z at 2; rewrite <- Hihl1.
- rewrite <-Hbin in Hs2; assert (([||WW ih il||] < [|s|] * [|s|] + 2 * [|s|] + 1) -> ([||WW ih il||] - [|s|] * [|s|] <= 2 * [|s|])) by lia; auto.
- unfold zn2z_to_Z.
- rewrite <-H1.
- ring_simplify.
- apply trans_equal with (wB + ([|il|] - [|il1|])).
- ring.
- rewrite <-Hil2; ring.
-Qed.
-
-Lemma to_Z_gcd : forall i j,
- [|gcd i j|] = Zgcdn (2*size) [|j|] [|i|].
-Proof.
- unfold gcd.
- induction (2*size)%nat; intros.
- reflexivity.
- simpl.
- generalize (to_Z_bounded j)(to_Z_bounded i); intros.
- case_eq (j == 0)%int.
- rewrite eqb_spec;intros H1;rewrite H1.
- replace [|0|] with 0;trivial;rewrite Z.abs_eq;auto with zarith.
- rewrite <- not_true_iff_false, eqb_spec;intros.
- case_eq [|j|]; intros.
- elim H1;apply to_Z_inj;assumption.
- rewrite IHn, <- H2, mod_spec;trivial.
- rewrite H2 in H;destruct H as (H, _);elim H;trivial.
-Qed.
-
-Lemma gcd_spec : forall a b, Zis_gcd [|a|] [|b|] [|gcd a b|].
-Proof.
- intros.
- rewrite to_Z_gcd.
- apply Zis_gcd_sym.
- apply Zgcdn_is_gcd.
- unfold Zgcd_bound.
- generalize (to_Z_bounded b).
- destruct [|b|].
- unfold size; intros _; change Int31.size with 31%nat; lia.
- intros (_,H).
- cut (Psize p <= size)%nat; [ lia | rewrite <- Zpower2_Psize; auto].
- intros (H,_); compute in H; elim H; auto.
-Qed.
-
-Lemma head00_spec: forall x, [|x|] = 0 -> [|head0 x|] = [|digits|].
-Proof.
- change 0 with [|0|];intros x Heq.
- apply to_Z_inj in Heq;rewrite Heq;trivial.
-Qed.
-
-Lemma tail00_spec: forall x, [|x|] = 0 -> [|tail0 x|] = [|digits|].
-Proof.
- change 0 with [|0|];intros x Heq.
- apply to_Z_inj in Heq;rewrite Heq;trivial.
-Qed.
-
-(* lsr lsl *)
-Lemma lsl_0_l i: 0 << i = 0%int.
-Proof.
- apply to_Z_inj.
- generalize (lsl_spec 0 i).
- rewrite to_Z_0, Zmult_0_l, Zmod_0_l; auto.
-Qed.
-
-Lemma lsl_0_r i: i << 0 = i.
-Proof.
- apply to_Z_inj.
- rewrite lsl_spec, to_Z_0, Zmult_1_r.
- apply Zmod_small; apply (to_Z_bounded i).
-Qed.
-
-Lemma lsl_M_r x i (H: (digits <= i = true)%int) : x << i = 0%int.
-Proof.
- apply to_Z_inj.
- rewrite lsl_spec, to_Z_0.
- rewrite leb_spec in H.
- unfold wB; change (Z_of_nat size) with [|digits|].
- replace ([|i|]) with (([|i|] - [|digits|]) + [|digits|])%Z; try ring.
- rewrite Zpower_exp, Zmult_assoc, Z_mod_mult; auto with arith.
- apply Z.le_ge; auto with zarith.
- case (to_Z_bounded digits); auto with zarith.
-Qed.
-
-Lemma lsr_0_l i: 0 >> i = 0%int.
-Proof.
- apply to_Z_inj.
- generalize (lsr_spec 0 i).
- rewrite to_Z_0, Zdiv_0_l; auto.
-Qed.
-
-Lemma lsr_0_r i: i >> 0 = i.
-Proof.
- apply to_Z_inj.
- rewrite lsr_spec, to_Z_0, Zdiv_1_r; auto.
-Qed.
-
-Lemma lsr_M_r x i (H: (digits <= i = true)%int) : x >> i = 0%int.
-Proof.
- apply to_Z_inj.
- rewrite lsr_spec, to_Z_0.
- case (to_Z_bounded x); intros H1x H2x.
- case (to_Z_bounded digits); intros H1d H2d.
- rewrite leb_spec in H.
- apply Zdiv_small; split; auto.
- apply Z.lt_le_trans with (1 := H2x).
- unfold wB; change (Z_of_nat size) with [|digits|].
- apply Zpower_le_monotone; auto with zarith.
-Qed.
-
-Lemma add_le_r m n:
- if (n <= m + n)%int then ([|m|] + [|n|] < wB)%Z else (wB <= [|m|] + [|n|])%Z.
-Proof.
- case (to_Z_bounded m); intros H1m H2m.
- case (to_Z_bounded n); intros H1n H2n.
- case (Zle_or_lt wB ([|m|] + [|n|])); intros H.
- assert (H1: ([| m + n |] = [|m|] + [|n|] - wB)%Z).
- rewrite add_spec.
- replace (([|m|] + [|n|]) mod wB)%Z with (((([|m|] + [|n|]) - wB) + wB) mod wB)%Z.
- rewrite Zplus_mod, Z_mod_same_full, Zplus_0_r, !Zmod_small; auto with zarith.
- rewrite !Zmod_small; auto with zarith.
- apply f_equal2 with (f := Zmod); auto with zarith.
- case_eq (n <= m + n)%int; auto.
- rewrite leb_spec, H1; auto with zarith.
- assert (H1: ([| m + n |] = [|m|] + [|n|])%Z).
- rewrite add_spec, Zmod_small; auto with zarith.
- replace (n <= m + n)%int with true; auto.
- apply sym_equal; rewrite leb_spec, H1; auto with zarith.
-Qed.
-
-Lemma lsr_add i m n: ((i >> m) >> n = if n <= m + n then i >> (m + n) else 0)%int.
-Proof.
- case (to_Z_bounded m); intros H1m H2m.
- case (to_Z_bounded n); intros H1n H2n.
- case (to_Z_bounded i); intros H1i H2i.
- generalize (add_le_r m n); case (n <= m + n)%int; intros H.
- apply to_Z_inj; rewrite !lsr_spec, Zdiv_Zdiv, <- Zpower_exp; auto with zarith.
- rewrite add_spec, Zmod_small; auto with zarith.
- apply to_Z_inj; rewrite !lsr_spec, Zdiv_Zdiv, <- Zpower_exp; auto with zarith.
- apply Zdiv_small; split; auto with zarith.
- apply Z.lt_le_trans with (1 := H2i).
- apply Z.le_trans with (1 := H).
- apply Zpower2_le_lin; auto with zarith.
-Qed.
-
-Lemma lsl_add i m n: ((i << m) << n = if n <= m + n then i << (m + n) else 0)%int.
-Proof.
- case (to_Z_bounded m); intros H1m H2m.
- case (to_Z_bounded n); intros H1n H2n.
- case (to_Z_bounded i); intros H1i H2i.
- generalize (add_le_r m n); case (n <= m + n)%int; intros H.
- apply to_Z_inj; rewrite !lsl_spec, Zmult_mod, Zmod_mod, <- Zmult_mod.
- rewrite <-Zmult_assoc, <- Zpower_exp; auto with zarith.
- apply f_equal2 with (f := Zmod); auto.
- rewrite add_spec, Zmod_small; auto with zarith.
- apply to_Z_inj; rewrite !lsl_spec, Zmult_mod, Zmod_mod, <- Zmult_mod.
- rewrite <-Zmult_assoc, <- Zpower_exp; auto with zarith.
- unfold wB.
- replace ([|m|] + [|n|])%Z with
- ((([|m|] + [|n|]) - Z_of_nat size) + Z_of_nat size)%Z.
- 2: ring.
- rewrite Zpower_exp, Zmult_assoc, Z_mod_mult; auto with zarith.
- assert (Z_of_nat size < wB)%Z; auto with zarith.
- apply Zpower2_lt_lin; auto with zarith.
-Qed.
-
-
-Coercion b2i (b: bool) : int := if b then 1%int else 0%int.
-
-Lemma bit_0 n : bit 0 n = false.
-Proof. unfold bit; rewrite lsr_0_l; auto. Qed.
-
-Lemma lsr_1 n : 1 >> n = (n == 0).
-Proof.
- case_eq (n == 0).
- rewrite eqb_spec; intros H; rewrite H, lsr_0_r.
- apply refl_equal.
- intros Hn.
- assert (H1n : (1 >> n = 0)%int); auto.
- apply to_Z_inj; rewrite lsr_spec.
- apply Zdiv_small; rewrite to_Z_1; split; auto with zarith.
- change 1%Z with (2^0)%Z.
- apply Zpower_lt_monotone; split; auto with zarith.
- case (Zle_lt_or_eq 0 [|n|]); auto.
- case (to_Z_bounded n); auto.
- intros H1.
- assert ((n == 0) = true).
- rewrite eqb_spec; apply to_Z_inj; rewrite <-H1, to_Z_0; auto.
- generalize H; rewrite Hn; discriminate.
-Qed.
-
-Lemma bit_1 n : bit 1 n = (n == 0).
-Proof.
- unfold bit; rewrite lsr_1.
- case (n == 0).
- apply refl_equal.
- rewrite lsl_0_l; apply refl_equal.
-Qed.
-
-Lemma bit_M i n (H: (digits <= n = true)%int): bit i n = false.
-Proof. unfold bit; rewrite lsr_M_r; auto. Qed.
-
-Lemma bit_half i n (H: (n < digits = true)%int) : bit (i>>1) n = bit i (n+1).
-Proof.
- unfold bit.
- rewrite lsr_add.
- case_eq (n <= (1 + n))%int.
- replace (1+n)%int with (n+1)%int; [auto|idtac].
- apply to_Z_inj; rewrite !add_spec, Zplus_comm; auto.
- intros H1; assert (H2: n = max_int).
- 2: generalize H; rewrite H2; discriminate.
- case (to_Z_bounded n); intros H1n H2n.
- case (Zle_lt_or_eq [|n|] (wB - 1)); auto with zarith;
- intros H2; apply to_Z_inj; auto.
- generalize (add_le_r 1 n); rewrite H1.
- change [|max_int|] with (wB - 1)%Z.
- replace [|1|] with 1%Z; auto with zarith.
-Qed.
-
-Lemma bit_0_spec i: [|bit i 0|] = [|i|] mod 2.
-Proof.
- unfold bit, is_zero; rewrite lsr_0_r.
- assert (Hbi: ([|i|] mod 2 < 2)%Z).
- apply Z_mod_lt; auto with zarith.
- case (to_Z_bounded i); intros H1i H2i.
- case (Zmod_le_first [|i|] 2); auto with zarith; intros H3i H4i.
- assert (H2b: (0 < 2 ^ [|digits - 1|])%Z).
- apply Zpower_gt_0; auto with zarith.
- case (to_Z_bounded (digits -1)); auto with zarith.
- assert (H: [|i << (digits -1)|] = ([|i|] mod 2 * 2^ [|digits -1|])%Z).
- rewrite lsl_spec.
- rewrite (Z_div_mod_eq [|i|] 2) at 1; auto with zarith.
- rewrite Zmult_plus_distr_l, <-Zplus_mod_idemp_l.
- rewrite (Zmult_comm 2), <-Zmult_assoc.
- replace (2 * 2 ^ [|digits - 1|])%Z with wB; auto.
- rewrite Z_mod_mult, Zplus_0_l; apply Zmod_small.
- split; auto with zarith.
- replace wB with (2 * 2 ^ [|digits -1|])%Z; auto.
- apply Zmult_lt_compat_r; auto with zarith.
- case (Zle_lt_or_eq 0 ([|i|] mod 2)); auto with zarith; intros Hi.
- 2: generalize H; rewrite <-Hi, Zmult_0_l.
- 2: replace 0%Z with [|0|]; auto.
- 2: rewrite to_Z_eq, <-eqb_spec; intros H1; rewrite H1; auto.
- generalize H; replace ([|i|] mod 2) with 1%Z; auto with zarith.
- rewrite Zmult_1_l.
- intros H1.
- assert (H2: [|i << (digits - 1)|] <> [|0|]).
- replace [|0|] with 0%Z; auto with zarith.
- generalize (eqb_spec (i << (digits - 1)) 0).
- case (i << (digits - 1) == 0); auto.
- intros (H3,_); case H2.
- rewrite to_Z_eq; auto.
-Qed.
-
-Lemma bit_split i : (i = (i>>1)<<1 + bit i 0)%int.
-Proof.
- apply to_Z_inj.
- rewrite add_spec, lsl_spec, lsr_spec, bit_0_spec, Zplus_mod_idemp_l.
- replace (2 ^ [|1|]) with 2%Z; auto with zarith.
- rewrite Zmult_comm, <-Z_div_mod_eq; auto with zarith.
- rewrite Zmod_small; auto; case (to_Z_bounded i); auto.
-Qed.
-
-
-Lemma bit_eq i1 i2:
- i1 = i2 <-> forall i, bit i1 i = bit i2 i.
-Admitted. (* Too slow *)
-(* Proof. *)
-(* split; try (intros; subst; auto; fail). *)
-(* case (to_Z_bounded i2); case (to_Z_bounded i1). *)
-(* unfold wB; generalize i1 i2; elim size; clear i1 i2. *)
-(* replace (2^Z_of_nat 0) with 1%Z; auto with zarith. *)
-(* intros; apply to_Z_inj; auto with zarith. *)
-(* intros n IH i1 i2 H1i1 H2i1 H1i2 H2i2 H. *)
-(* rewrite (bit_split i1), (bit_split i2). *)
-(* rewrite H. *)
-(* apply f_equal2 with (f := add31); auto. *)
-(* apply f_equal2 with (f := lsl); auto. *)
-(* apply IH; try rewrite lsr_spec; *)
-(* replace (2^[|1|]) with 2%Z; auto with zarith. *)
-(* apply Zdiv_lt_upper_bound; auto with zarith. *)
-(* generalize H2i1; rewrite inj_S. *)
-(* unfold Z.succ; rewrite Zpower_exp; auto with zarith. *)
-(* apply Zdiv_lt_upper_bound; auto with zarith. *)
-(* generalize H2i2; rewrite inj_S. *)
-(* unfold Z.succ; rewrite Zpower_exp; auto with zarith. *)
-(* intros i. *)
-(* case (Zle_or_lt [|digits|] [|i|]); intros Hi. *)
-(* rewrite !bit_M; auto; rewrite leb_spec; auto. *)
-(* rewrite !bit_half; auto; rewrite ltb_spec; auto with zarith. *)
-(* Qed. *)
-
-Lemma bit_lsr x i j :
- (bit (x >> i) j = if j <= i + j then bit x (i + j) else false)%int.
-Proof.
- unfold bit; rewrite lsr_add; case leb; auto.
-Qed.
-
-Lemma bit_lsl x i j : bit (x << i) j =
-(if (j < i) || (digits <= j) then false else bit x (j - i))%int.
-Proof.
- assert (F1: 1 >= 0) by discriminate.
- case_eq (digits <= j)%int; intros H.
- rewrite orb_true_r, bit_M; auto.
- set (d := [|digits|]).
- case (Zle_or_lt d [|j|]); intros H1.
- case (leb_spec digits j); rewrite H; auto with zarith.
- intros _ HH; generalize (HH H1); discriminate.
- clear H.
- generalize (ltb_spec j i); case ltb; intros H2; unfold bit; [change (if true || false then false else negb (is_zero ((x >> (j - i)) << (digits - 1)))) with false | change (if false || false then false else negb (is_zero ((x >> (j - i)) << (digits - 1)))) with (negb (is_zero ((x >> (j - i)) << (digits - 1))))].
- assert (F2: ([|j|] < [|i|])%Z) by (case H2; auto); clear H2.
- replace (is_zero (((x << i) >> j) << (digits - 1))) with true; auto.
- case (to_Z_bounded j); intros H1j H2j.
- apply sym_equal; rewrite is_zero_spec; apply to_Z_inj.
- rewrite lsl_spec, lsr_spec, lsl_spec.
- replace wB with (2^d); auto.
- pattern d at 1; replace d with ((d - ([|j|] + 1)) + ([|j|] + 1))%Z.
- 2: ring.
- rewrite Zpower_exp; auto with zarith.
- replace [|i|] with (([|i|] - ([|j|] + 1)) + ([|j|] + 1))%Z.
- 2: ring.
- rewrite Zpower_exp, Zmult_assoc; auto with zarith.
- rewrite Zmult_mod_distr_r.
- rewrite Zplus_comm, Zpower_exp, !Zmult_assoc; auto with zarith.
- rewrite Z_div_mult_full; auto with zarith.
- 2: assert (0 < 2 ^ [|j|])%Z; auto with zarith.
- rewrite <-Zmult_assoc, <-Zpower_exp; auto with zarith.
- replace (1 + [|digits - 1|])%Z with d; auto with zarith.
- rewrite Z_mod_mult; auto.
- case H2; intros _ H3; case (Zle_or_lt [|i|] [|j|]); intros F2.
- 2: generalize (H3 F2); discriminate.
- clear H2 H3.
- apply f_equal with (f := negb).
- apply f_equal with (f := is_zero).
- apply to_Z_inj.
- rewrite !lsl_spec, !lsr_spec, !lsl_spec.
- pattern wB at 2 3; replace wB with (2^(1+ [|digits - 1|])); auto.
- rewrite Zpower_exp, Zpower_1_r; auto with zarith.
- rewrite !Zmult_mod_distr_r.
- apply f_equal2 with (f := Zmult); auto.
- replace wB with (2^ d); auto with zarith.
- replace d with ((d - [|i|]) + [|i|])%Z.
- 2: ring.
- case (to_Z_bounded i); intros H1i H2i.
- rewrite Zpower_exp; [ |apply Z.le_ge; lia|apply Z.le_ge; assumption].
- rewrite Zmult_mod_distr_r.
- case (to_Z_bounded j); intros H1j H2j.
- replace [|j - i|] with ([|j|] - [|i|])%Z.
- 2: rewrite sub_spec, Zmod_small; auto with zarith.
- set (d1 := (d - [|i|])%Z).
- set (d2 := ([|j|] - [|i|])%Z).
- pattern [|j|] at 1;
- replace [|j|] with (d2 + [|i|])%Z.
- 2: unfold d2; ring.
- rewrite Zpower_exp; auto with zarith.
- rewrite Zdiv_mult_cancel_r.
- 2: (apply Zlt0_not_eq; apply Z.pow_pos_nonneg; [apply Pos2Z.is_pos|assumption]).
- rewrite (Z_div_mod_eq [|x|] (2^d1)) at 2; auto with zarith.
- 2: apply Z.lt_gt; apply Zpower_gt_0; unfold d1; lia.
- pattern d1 at 2;
- replace d1 with (d2 + (1+ (d - [|j|] - 1)))%Z.
- 2: unfold d1, d2; ring.
- rewrite Zpower_exp; auto with zarith.
- rewrite <-Zmult_assoc, Zmult_comm.
- rewrite Z_div_plus_l; auto with zarith.
- rewrite Zpower_exp, Zpower_1_r; auto with zarith.
- rewrite <-Zplus_mod_idemp_l.
- rewrite <-!Zmult_assoc, Zmult_comm, Z_mod_mult, Zplus_0_l; auto.
-Qed.
-
-
-Lemma bit_b2i (b: bool) i : bit b i = (i == 0) && b.
-Proof.
- case b; unfold bit; simpl b2i.
- 2: rewrite lsr_0_l, lsl_0_l, andb_false_r; auto.
- rewrite lsr_1; case (i == 0); auto.
-Qed.
-
-Lemma bit_or_split i : (i = (i>>1)<<1 lor bit i 0)%int.
-Proof.
- rewrite bit_eq.
- intros n; rewrite lor_spec.
- rewrite bit_lsl, bit_lsr, bit_b2i.
- case (to_Z_bounded n); intros Hi _.
- case (Zle_lt_or_eq _ _ Hi).
- 2: replace 0%Z with [|0|]; auto; rewrite to_Z_eq.
- 2: intros H; rewrite <-H.
- 2: replace (0 < 1)%int with true; auto.
- intros H; clear Hi.
- case_eq (n == 0).
- rewrite eqb_spec; intros H1; generalize H; rewrite H1; discriminate.
- intros _; rewrite orb_false_r.
- case_eq (n < 1)%int.
- rewrite ltb_spec, to_Z_1; intros HH; contradict HH; auto with zarith.
- intros _.
- generalize (@bit_M i n); case leb.
- intros H1; rewrite H1; auto.
- intros _.
- case (to_Z_bounded n); intros H1n H2n.
- assert (F1: [|n - 1|] = ([|n|] - 1)%Z).
- rewrite sub_spec, Zmod_small; rewrite to_Z_1; auto with zarith.
- generalize (add_le_r 1 (n - 1)); case leb; rewrite F1, to_Z_1; intros HH.
- replace (1 + (n -1))%int with n. change (bit i n = bit i n). reflexivity.
- apply to_Z_inj; rewrite add_spec, F1, Zmod_small; rewrite to_Z_1;
- auto with zarith.
- rewrite bit_M; auto; rewrite leb_spec.
- replace [|n|] with wB; try discriminate; auto with zarith.
-Qed.
-
-(* is_zero *)
-Lemma is_zero_0: is_zero 0 = true.
-Proof. apply refl_equal. Qed.
-
-(* is_even *)
-Lemma is_even_bit i : is_even i = negb (bit i 0).
-Proof.
- unfold is_even.
- replace (i land 1) with (b2i (bit i 0)).
- case bit; auto.
- apply bit_eq; intros n.
- rewrite bit_b2i, land_spec, bit_1.
- generalize (eqb_spec n 0).
- case (n == 0); auto.
- intros(H,_); rewrite andb_true_r, H; auto.
- rewrite andb_false_r; auto.
-Qed.
-
-Lemma is_even_0: is_even 0 = true.
-Proof. apply refl_equal. Qed.
-
-Lemma is_even_lsl_1 i: is_even (i << 1) = true.
-Proof.
- rewrite is_even_bit, bit_lsl; auto.
-Qed.
-
-Lemma is_even_spec : forall x,
- if is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1.
-Proof.
-intros x; rewrite is_even_bit.
-generalize (bit_0_spec x); case bit; simpl; auto.
-Qed.
-
-(* More land *)
-
-Lemma land_0_l i: 0 land i = 0%int.
-Proof.
- apply bit_eq; intros n.
- rewrite land_spec, bit_0; auto.
-Qed.
-
-Lemma land_0_r i: i land 0 = 0%int.
-Proof.
- apply bit_eq; intros n.
- rewrite land_spec, bit_0, andb_false_r; auto.
-Qed.
-
-Lemma land_assoc i1 i2 i3 :
- i1 land (i2 land i3) = i1 land i2 land i3.
-Proof.
- apply bit_eq; intros n.
- rewrite !land_spec, andb_assoc; auto.
-Qed.
-
-
-Lemma land_comm i j : i land j = j land i.
-Proof.
- apply bit_eq; intros n.
- rewrite !land_spec, andb_comm; auto.
-Qed.
-
-Lemma lor_comm i1 i2 : i1 lor i2 = i2 lor i1.
-Proof.
- apply bit_eq; intros n.
- rewrite !lor_spec, orb_comm; auto.
-Qed.
-
-Lemma lor_assoc i1 i2 i3 :
- i1 lor (i2 lor i3) = i1 lor i2 lor i3.
-Proof.
- apply bit_eq; intros n.
- rewrite !lor_spec, orb_assoc; auto.
-Qed.
-
-Lemma land_lor_distrib_r i1 i2 i3 :
- i1 land (i2 lor i3) = (i1 land i2) lor (i1 land i3).
-Proof.
- apply bit_eq; intros n.
- rewrite !land_spec, !lor_spec, !land_spec, andb_orb_distrib_r; auto.
-Qed.
-
-Lemma land_lor_distrib_l i1 i2 i3 :
- (i1 lor i2) land i3 = (i1 land i3) lor (i2 land i3).
-Proof.
- apply bit_eq; intros n.
- rewrite !land_spec, !lor_spec, !land_spec, andb_orb_distrib_l; auto.
-Qed.
-
-Lemma lor_land_distrib_r i1 i2 i3:
- i1 lor (i2 land i3) = (i1 lor i2) land (i1 lor i3).
-Proof.
- apply bit_eq; intros n.
- rewrite !land_spec, !lor_spec, !land_spec, orb_andb_distrib_r; auto.
-Qed.
-
-Lemma lor_land_distrib_l i1 i2 i3:
- (i1 land i2) lor i3 = (i1 lor i3) land (i2 lor i3).
-Proof.
- apply bit_eq; intros n.
- rewrite !land_spec, !lor_spec, !land_spec, orb_andb_distrib_l; auto.
-Qed.
-
-Lemma absoption_land i1 i2 : i1 land (i1 lor i2) = i1.
-Proof.
- apply bit_eq; intros n.
- rewrite land_spec, lor_spec, absoption_andb; auto.
-Qed.
-
-Lemma absoption_lor i1 i2: i1 lor (i1 land i2) = i1.
-Proof.
- apply bit_eq; intros n.
- rewrite lor_spec, land_spec, absoption_orb; auto.
-Qed.
-
-Lemma land_lsl i1 i2 i: (i1 land i2) << i = (i1 << i) land (i2 << i).
-Proof.
- apply bit_eq; intros n.
- rewrite land_spec, !bit_lsl, land_spec.
- case (_ || _); auto.
-Qed.
-
-Lemma lor_lsl i1 i2 i: (i1 lor i2) << i = (i1 << i) lor (i2 << i).
-Proof.
- apply bit_eq; intros n.
- rewrite lor_spec, !bit_lsl, lor_spec.
- case (_ || _); auto.
-Qed.
-
-Lemma lxor_lsl i1 i2 i: (i1 lxor i2) << i = (i1 << i) lxor (i2 << i).
-Proof.
- apply bit_eq; intros n.
- rewrite lxor_spec, !bit_lsl, lxor_spec.
- case (_ || _); auto.
-Qed.
-
-Lemma land_lsr i1 i2 i: (i1 land i2) >> i = (i1 >> i) land (i2 >> i).
-Proof.
- apply bit_eq; intros n.
- rewrite land_spec, !bit_lsr, land_spec.
- case (_ <= _)%int; auto.
-Qed.
-
-Lemma lor_lsr i1 i2 i: (i1 lor i2) >> i = (i1 >> i) lor (i2 >> i).
-Proof.
- apply bit_eq; intros n.
- rewrite lor_spec, !bit_lsr, lor_spec.
- case (_ <= _)%int; auto.
-Qed.
-
-Lemma lxor_lsr i1 i2 i: (i1 lxor i2) >> i = (i1 >> i) lxor (i2 >> i).
-Proof.
- apply bit_eq; intros n.
- rewrite lxor_spec, !bit_lsr, lxor_spec.
- case (_ <= _)%int; auto.
-Qed.
-
-Lemma is_even_and i j : is_even (i land j) = is_even i || is_even j.
-Proof.
- rewrite !is_even_bit, land_spec; case bit; auto.
-Qed.
-
-Lemma is_even_or i j : is_even (i lor j) = is_even i && is_even j.
-Proof.
- rewrite !is_even_bit, lor_spec; case bit; auto.
-Qed.
-
-Lemma is_even_xor i j : is_even (i lxor j) = negb (xorb (is_even i) (is_even j)).
-Proof.
- rewrite !is_even_bit, lxor_spec; do 2 case bit; auto.
-Qed.
-
-Lemma lsl_add_distr x y n: (x + y) << n = ((x << n) + (y << n))%int.
-Proof.
- apply to_Z_inj; rewrite !lsl_spec, !add_spec, Zmult_mod_idemp_l.
- rewrite !lsl_spec, <-Zplus_mod.
- apply f_equal2 with (f := Zmod); auto with zarith.
-Qed.
-
-Lemma add_assoc x y z: (x + (y + z) = (x + y) + z)%int.
-Proof.
- apply to_Z_inj; rewrite !add_spec.
- rewrite Zplus_mod_idemp_l, Zplus_mod_idemp_r, Zplus_assoc; auto.
-Qed.
-
-Lemma add_comm x y: (x + y = y + x)%int.
-Proof.
- apply to_Z_inj; rewrite !add_spec, Zplus_comm; auto.
-Qed.
-
-Lemma lsr_add_distr x y n: (x + y) << n = ((x << n) + (y << n))%int.
-Proof.
- apply to_Z_inj.
- rewrite add_spec, !lsl_spec, add_spec.
- rewrite Zmult_mod_idemp_l, <-Zplus_mod.
- apply f_equal2 with (f := Zmod); auto with zarith.
-Qed.
-
-Lemma is_even_add x y :
- is_even (x + y) = negb (xorb (negb (is_even x)) (negb (is_even y))).
-Proof.
- assert (F : [|x + y|] mod 2 = ([|x|] mod 2 + [|y|] mod 2) mod 2).
- assert (F1: (2 | wB)) by (apply Zpower_divide; apply refl_equal).
- assert (F2: 0 < wB) by (apply refl_equal).
- case (to_Z_bounded x); intros H1x H2x.
- case (to_Z_bounded y); intros H1y H2y.
- rewrite add_spec, <-Zmod_div_mod; auto with zarith.
- rewrite (Z_div_mod_eq [|x|] 2) at 1; auto with zarith.
- rewrite (Z_div_mod_eq [|y|] 2) at 1; auto with zarith.
- rewrite Zplus_mod.
- rewrite Zmult_comm, (fun x => Zplus_comm (x * 2)), Z_mod_plus; auto with zarith.
- rewrite Zmult_comm, (fun x => Zplus_comm (x * 2)), Z_mod_plus; auto with zarith.
- rewrite !Zmod_mod, <-Zplus_mod; auto.
- generalize (is_even_spec (x + y)) (is_even_spec x) (is_even_spec y).
- do 3 case is_even; auto; rewrite F; intros H1 H2 H3;
- generalize H1; rewrite H2, H3; try discriminate.
-Qed.
-
-Lemma bit_add_0 x y: bit (x + y) 0 = xorb (bit x 0) (bit y 0).
-Proof.
- rewrite <-(fun x => (negb_involutive (bit x 0))).
- rewrite <-is_even_bit, is_even_add, !is_even_bit.
- do 2 case bit; auto.
-Qed.
-
-Lemma add_cancel_l x y z : (x + y = x + z)%int -> y = z.
-Proof.
- intros H; case (to_Z_bounded x); case (to_Z_bounded y); case (to_Z_bounded z);
- intros H1z H2z H1y H2y H1x H2x.
- generalize (add_le_r y x) (add_le_r z x); rewrite (add_comm y x), H, (add_comm z x).
- case_eq (x <= x + z)%int; intros H1 H2 H3.
- apply to_Z_inj; generalize H; rewrite <-to_Z_eq, !add_spec, !Zmod_small; auto with zarith.
- apply to_Z_inj; assert ([|x|] + [|y|] = [|x|] + [|z|]); auto with zarith.
- assert (F1: wB > 0) by apply refl_equal.
- rewrite (Z_div_mod_eq ([|x|] + [|y|]) wB), (Z_div_mod_eq ([|x|] + [|z|]) wB); auto.
- rewrite <-to_Z_eq, !add_spec in H; rewrite H.
- replace (([|x|] + [|y|])/wB) with 1.
- replace (([|x|] + [|z|])/wB) with 1; auto with zarith.
- apply Zle_antisym.
- apply Zdiv_le_lower_bound; auto with zarith.
- assert (F2: [|x|] + [|z|] < 2 * wB); auto with zarith.
- generalize (Zdiv_lt_upper_bound _ _ _ (Z.gt_lt _ _ F1) F2); auto with zarith.
- apply Zle_antisym.
- apply Zdiv_le_lower_bound; auto with zarith.
- assert (F2: [|x|] + [|y|] < 2 * wB); auto with zarith.
- generalize (Zdiv_lt_upper_bound _ _ _ (Z.gt_lt _ _ F1) F2); auto with zarith.
-Qed.
-
-Lemma add_cancel_r x y z : (y + x = z + x)%int -> y = z.
-Proof.
- rewrite !(fun t => add_comm t x); intros Hl; apply (add_cancel_l x); auto.
-Qed.
-
-Lemma to_Z_split x : [|x|] = [|(x >> 1)|] * 2 + [|bit x 0|].
-Proof.
- case (to_Z_bounded x); intros H1x H2x.
- case (to_Z_bounded (bit x 0)); intros H1b H2b.
- assert (F1: 0 <= [|x >> 1|] < wB/2).
- rewrite lsr_spec, to_Z_1, Zpower_1_r; split.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - assumption.
- }
- apply Zdiv_lt_upper_bound; auto with zarith.
- rewrite (bit_split x) at 1.
- rewrite add_spec, Zmod_small, lsl_spec, to_Z_1, Zpower_1_r, Zmod_small;
- split; auto with zarith.
- change wB with ((wB/2)*2); auto with zarith.
- rewrite lsl_spec, to_Z_1, Zpower_1_r, Zmod_small; auto with zarith.
- change wB with ((wB/2)*2); auto with zarith.
- rewrite lsl_spec, to_Z_1, Zpower_1_r, Zmod_small; auto with zarith.
- 2: change wB with ((wB/2)*2); auto with zarith.
- change wB with (((wB/2 - 1) * 2 + 1) + 1).
- assert ([|bit x 0|] <= 1); auto with zarith.
- case bit; discriminate.
-Qed.
-
-Lemma lor_le x y : (y <= x lor y)%int = true.
-Proof.
- generalize x y (to_Z_bounded x) (to_Z_bounded y); clear x y.
- unfold wB; elim size.
- replace (2^Z_of_nat 0) with 1%Z; auto with zarith.
- intros x y Hx Hy; replace x with 0%int.
- replace y with 0%int; auto.
- apply to_Z_inj; rewrite to_Z_0; auto with zarith.
- apply to_Z_inj; rewrite to_Z_0; auto with zarith.
- intros n IH x y; rewrite inj_S.
- unfold Z.succ; rewrite Zpower_exp, Zpower_1_r; auto with zarith.
- intros Hx Hy.
- rewrite leb_spec.
- rewrite (to_Z_split y) at 1; rewrite (to_Z_split (x lor y)).
- assert ([|y>>1|] <= [|(x lor y) >> 1|]).
- rewrite lor_lsr, <-leb_spec; apply IH.
- rewrite lsr_spec, to_Z_1, Zpower_1_r; split.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - abstract lia.
- }
- apply Zdiv_lt_upper_bound; auto with zarith.
- rewrite lsr_spec, to_Z_1, Zpower_1_r; split.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - abstract lia.
- }
- apply Zdiv_lt_upper_bound; auto with zarith.
- assert ([|bit y 0|] <= [|bit (x lor y) 0|]); auto with zarith.
- rewrite lor_spec; do 2 case bit; try discriminate.
-Qed.
-
-
-Lemma bit_add_or x y:
- (forall n, bit x n = true -> bit y n = true -> False) <-> (x + y)%int= x lor y.
-Proof.
- generalize x y (to_Z_bounded x) (to_Z_bounded y); clear x y.
- unfold wB; elim size.
- replace (2^Z_of_nat 0) with 1%Z; auto with zarith.
- intros x y Hx Hy; replace x with 0%int.
- replace y with 0%int.
- split; auto; intros _ n; rewrite !bit_0; discriminate.
- apply to_Z_inj; rewrite to_Z_0; auto with zarith.
- apply to_Z_inj; rewrite to_Z_0; auto with zarith.
- intros n IH x y; rewrite inj_S.
- unfold Z.succ; rewrite Zpower_exp, Zpower_1_r; auto with zarith.
- intros Hx Hy.
- split.
- intros Hn.
- assert (F1: ((x >> 1) + (y >> 1))%int = (x >> 1) lor (y >> 1)).
- apply IH.
- rewrite lsr_spec, Zpower_1_r; split.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - abstract lia.
- }
- apply Zdiv_lt_upper_bound; auto with zarith.
- rewrite lsr_spec, Zpower_1_r; split.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - abstract lia.
- }
- apply Zdiv_lt_upper_bound; auto with zarith.
- intros m H1 H2.
- case_eq (digits <= m)%int; [idtac | rewrite <- not_true_iff_false];
- intros Heq.
- rewrite bit_M in H1; auto; discriminate.
- rewrite leb_spec in Heq.
- apply (Hn (m + 1)%int);
- rewrite <-bit_half; auto; rewrite ltb_spec; auto with zarith.
- rewrite (bit_split (x lor y)), lor_lsr, <- F1, lor_spec.
- replace (b2i (bit x 0 || bit y 0)) with (bit x 0 + bit y 0)%int.
- 2: generalize (Hn 0%int); do 2 case bit; auto; intros [ ]; auto.
- rewrite lsl_add_distr.
- rewrite (bit_split x) at 1; rewrite (bit_split y) at 1.
- rewrite <-!add_assoc; apply f_equal2 with (f := add31); auto.
- rewrite add_comm, <-!add_assoc; apply f_equal2 with (f := add31); auto.
- rewrite add_comm; auto.
- intros Heq.
- generalize (add_le_r x y); rewrite Heq, lor_le; intro Hb.
- generalize Heq; rewrite (bit_split x) at 1; rewrite (bit_split y )at 1; clear Heq.
- rewrite (fun y => add_comm y (bit x 0)), <-!add_assoc, add_comm,
- <-!add_assoc, (add_comm (bit y 0)), add_assoc, <-lsr_add_distr.
- rewrite (bit_split (x lor y)), lor_spec.
- intros Heq.
- assert (F: (bit x 0 + bit y 0)%int = (bit x 0 || bit y 0)).
- assert (F1: (2 | wB)) by (apply Zpower_divide; apply refl_equal).
- assert (F2: 0 < wB) by (apply refl_equal).
- assert (F3: [|bit x 0 + bit y 0|] mod 2 = [|bit x 0 || bit y 0|] mod 2).
- apply trans_equal with (([|(x>>1 + y>>1) << 1|] + [|bit x 0 + bit y 0|]) mod 2).
- rewrite lsl_spec, Zplus_mod, <-Zmod_div_mod; auto with zarith.
- rewrite Zpower_1_r, Z_mod_mult, Zplus_0_l, Zmod_mod; auto with zarith.
- rewrite (Zmod_div_mod 2 wB), <-add_spec, Heq; auto with zarith.
- rewrite add_spec, <-Zmod_div_mod; auto with zarith.
- rewrite lsl_spec, Zplus_mod, <-Zmod_div_mod; auto with zarith.
- rewrite Zpower_1_r, Z_mod_mult, Zplus_0_l, Zmod_mod; auto with zarith.
- generalize F3; do 2 case bit; try discriminate; auto.
- case (IH (x >> 1) (y >> 1)).
- rewrite lsr_spec, to_Z_1, Zpower_1_r; split.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - abstract lia.
- }
- apply Zdiv_lt_upper_bound; auto with zarith.
- rewrite lsr_spec, to_Z_1, Zpower_1_r; split.
- {
- apply Z_div_pos.
- - apply Zgt_pos_0.
- - abstract lia.
- }
- apply Zdiv_lt_upper_bound; auto with zarith.
- intros _ HH m; case (to_Z_bounded m); intros H1m H2m.
- case_eq (digits <= m)%int.
- intros Hlm; rewrite bit_M; auto; discriminate.
- rewrite <- not_true_iff_false, leb_spec; intros Hlm.
- case (Zle_lt_or_eq 0 [|m|]); auto; intros Hm.
- replace m with ((m -1) + 1)%int.
- rewrite <-(bit_half x), <-(bit_half y); auto with zarith.
- apply HH.
- rewrite <-lor_lsr.
- assert (0 <= [|bit (x lor y) 0|] <= 1) by (case bit; split; discriminate).
- rewrite F in Heq; generalize (add_cancel_r _ _ _ Heq).
- intros Heq1; apply to_Z_inj.
- generalize Heq1; rewrite <-to_Z_eq, lsl_spec, to_Z_1, Zpower_1_r, Zmod_small.
- rewrite lsl_spec, to_Z_1, Zpower_1_r, Zmod_small; auto with zarith.
- case (to_Z_bounded (x lor y)); intros H1xy H2xy.
- rewrite lsr_spec, to_Z_1, Zpower_1_r; auto with zarith.
- change wB with ((wB/2)*2); split.
- {
- apply Z.mul_nonneg_nonneg.
- - apply Z_div_pos.
- + apply Zgt_pos_0.
- + assumption.
- - apply Pos2Z.is_nonneg.
- }
- assert ([|x lor y|] / 2 < wB / 2); auto with zarith.
- apply Zdiv_lt_upper_bound; auto with zarith.
- split.
- case (to_Z_bounded (x >> 1 + y >> 1)); auto with zarith.
- rewrite add_spec.
- apply Z.le_lt_trans with (([|x >> 1|] + [|y >> 1|]) * 2); auto with zarith.
- case (Zmod_le_first ([|x >> 1|] + [|y >> 1|]) wB); auto with zarith.
- case (to_Z_bounded (x >> 1)); case (to_Z_bounded (y >> 1)); auto with zarith.
- generalize Hb; rewrite (to_Z_split x) at 1; rewrite (to_Z_split y) at 1.
- case (to_Z_bounded (bit x 0)); case (to_Z_bounded (bit y 0)); auto with zarith.
- rewrite ltb_spec, sub_spec, to_Z_1, Zmod_small; auto with zarith.
- rewrite ltb_spec, sub_spec, to_Z_1, Zmod_small; auto with zarith.
- apply to_Z_inj.
- rewrite add_spec, sub_spec, Zplus_mod_idemp_l, to_Z_1, Zmod_small; auto with zarith.
- replace m with 0%int.
- intros Hbx Hby; generalize F; rewrite <-to_Z_eq, Hbx, Hby; discriminate.
- apply to_Z_inj; auto.
-Qed.
-
-Lemma addmuldiv_spec : forall x y p, [|p|] <= [|digits|] ->
- [| addmuldiv p x y |] =
- ([|x|] * (2 ^ [|p|]) + [|y|] / (2 ^ ([|digits|] - [|p|]))) mod wB.
-Proof.
- intros x y p H.
- assert (Fp := to_Z_bounded p); assert (Fd := to_Z_bounded digits).
- rewrite addmuldiv_def_spec; unfold addmuldiv_def.
- case (bit_add_or (x << p) (y >> (digits - p))); intros HH _.
- rewrite <-HH, add_spec, lsl_spec, lsr_spec, Zplus_mod_idemp_l, sub_spec.
- rewrite (fun x y => Zmod_small (x - y)); auto with zarith.
- intros n; rewrite bit_lsl, bit_lsr.
- generalize (add_le_r (digits - p) n).
- case leb; try discriminate.
- rewrite sub_spec, Zmod_small; auto with zarith; intros H1.
- case_eq (n < p)%int; try discriminate.
- rewrite <- not_true_iff_false, ltb_spec; intros H2.
- case leb; try discriminate.
- intros _; rewrite bit_M; try discriminate.
- rewrite leb_spec, add_spec, Zmod_small, sub_spec, Zmod_small; auto with zarith.
- rewrite sub_spec, Zmod_small; auto with zarith.
-Qed.
-
-Lemma lxor_comm: forall i1 i2 : int, i1 lxor i2 = i2 lxor i1.
-Proof.
- intros;apply bit_eq;intros.
- rewrite !lxor_spec;apply xorb_comm.
-Qed.
-
-Lemma lxor_assoc: forall i1 i2 i3 : int, i1 lxor (i2 lxor i3) = i1 lxor i2 lxor i3.
-Proof.
- intros;apply bit_eq;intros.
- rewrite !lxor_spec, xorb_assoc;trivial.
-Qed.
-
-Lemma lxor_0_l : forall i, 0 lxor i = i.
-Proof.
- intros;apply bit_eq;intros.
- rewrite lxor_spec, bit_0, xorb_false_l;trivial.
-Qed.
-
-Lemma lxor_0_r : forall i, i lxor 0 = i.
-Proof.
- intros;rewrite lxor_comm;apply lxor_0_l.
-Qed.
-
-Lemma lxor_nilpotent: forall i, i lxor i = 0%int.
-Proof.
- intros;apply bit_eq;intros.
- rewrite lxor_spec, xorb_nilpotent, bit_0;trivial.
-Qed.
-
-Lemma lor_0_l : forall i, 0 lor i = i.
-Proof.
- intros;apply bit_eq;intros.
- rewrite lor_spec, bit_0, orb_false_l;trivial.
-Qed.
-
-Lemma lor_0_r : forall i, i lor 0 = i.
-Proof.
- intros;rewrite lor_comm;apply lor_0_l.
-Qed.
-
-Lemma reflect_leb : forall i j, reflect ([|i|] <= [|j|])%Z (i <= j)%int.
-Proof.
- intros; apply iff_reflect.
- symmetry;apply leb_spec.
-Qed.
-
-Lemma reflect_eqb : forall i j, reflect (i = j)%Z (i == j).
-Proof.
- intros; apply iff_reflect.
- symmetry;apply eqb_spec.
-Qed.
-
-Lemma reflect_ltb : forall i j, reflect ([|i|] < [|j|])%Z (i < j)%int.
-Proof.
- intros; apply iff_reflect.
- symmetry;apply ltb_spec.
-Qed.
-
-Lemma lsr_is_even_eq : forall i j,
- i >> 1 = j >> 1 ->
- is_even i = is_even j ->
- i = j.
-Proof.
- intros;apply bit_eq.
- intros n;destruct (reflect_eqb n 0).
- rewrite <- (negb_involutive (bit i n)), <- (negb_involutive (bit j n)).
- rewrite e, <- !is_even_bit, H0;trivial.
- assert (W1 : [|n|] <> 0) by (intros Heq;apply n0;apply to_Z_inj;trivial).
- assert (W2 := to_Z_bounded n);clear n0.
- assert (W3 : [|n-1|] = [|n|] - 1).
- rewrite sub_spec, to_Z_1, Zmod_small;trivial;lia.
- assert (H1 : n = ((n-1)+1)%int).
- apply to_Z_inj;rewrite add_spec, W3.
- rewrite Zmod_small;rewrite to_Z_1; lia.
- destruct (reflect_ltb (n-1) digits).
- rewrite <- ltb_spec in l.
- rewrite H1, <- !bit_half, H;trivial.
- assert ((digits <= n)%int = true).
- rewrite leb_spec;lia.
- rewrite !bit_M;trivial.
-Qed.
-
-Lemma lsr1_bit : forall i k, (bit i k >> 1 = 0)%int.
-Proof.
- intros;destruct (bit i k);trivial.
-Qed.
-
-Lemma bit_xor_split: forall i : int, i = (i >> 1) << 1 lxor bit i 0.
-Proof.
- intros.
- rewrite bit_or_split at 1.
- apply lsr_is_even_eq.
- rewrite lxor_lsr, lor_lsr, lsr1_bit, lxor_0_r, lor_0_r;trivial.
- rewrite is_even_or, is_even_xor.
- rewrite is_even_lsl_1;trivial.
- rewrite (xorb_true_l (is_even (bit i 0))), negb_involutive;trivial.
-Qed.
-
-(** Order *)
-Local Open Scope int63_scope.
-
-Lemma succ_max_int : forall x,
- (x < max_int)%int = true -> (0 < x + 1)%int = true.
-Proof.
- intros x;rewrite ltb_spec, ltb_spec, add_spec.
- intros; assert (W:= to_Z_bounded x); assert (W1:= to_Z_bounded max_int).
- change [|0|] with 0%Z;change [|1|] with 1%Z.
- rewrite Zmod_small;lia.
-Qed.
-
-Lemma leb_max_int : forall x, (x <= max_int)%int = true.
-Proof.
- intros x;rewrite leb_spec;assert (W:= to_Z_bounded x).
- change [|max_int|] with (wB - 1)%Z;lia.
-Qed.
-
-Lemma leb_0 : forall x, 0 <= x = true.
-Proof.
- intros x;rewrite leb_spec;destruct (to_Z_bounded x);trivial.
-Qed.
-
-Lemma ltb_0 : forall x, ~ (x < 0 = true).
-Proof.
- intros x;rewrite ltb_spec, to_Z_0;destruct (to_Z_bounded x);lia.
-Qed.
-
-Lemma leb_trans : forall x y z, x <= y = true -> y <= z = true -> x <= z = true.
-Proof.
- intros x y z;rewrite !leb_spec;apply Z.le_trans.
-Qed.
-
-Lemma ltb_trans : forall x y z, x < y = true -> y < z = true -> x < z = true.
-Proof.
- intros x y z;rewrite !ltb_spec;apply Z.lt_trans.
-Qed.
-
-Lemma ltb_leb_trans : forall x y z, x < y = true -> y <= z = true -> x < z = true.
-Proof.
- intros x y z;rewrite leb_spec, !ltb_spec;apply Z.lt_le_trans.
-Qed.
-
-Lemma leb_ltb_trans : forall x y z, x <= y = true -> y < z = true -> x < z = true.
-Proof.
- intros x y z;rewrite leb_spec, !ltb_spec;apply Z.le_lt_trans.
-Qed.
-
-Lemma gtb_not_leb : forall n m, m < n = true -> ~(n <= m = true).
-Proof.
- intros n m; rewrite ltb_spec, leb_spec;lia.
-Qed.
-
-Lemma leb_not_gtb : forall n m, m <= n = true -> ~(n < m = true).
-Proof.
- intros n m; rewrite ltb_spec, leb_spec;lia.
-Qed.
-
-Lemma leb_refl : forall n, n <= n = true.
-Proof.
- intros n;rewrite leb_spec;apply Z.le_refl.
-Qed.
-
-Lemma leb_negb_gtb : forall x y, x <= y = negb (y < x).
-Proof.
- intros x y;apply Bool.eq_true_iff_eq;split;intros.
- apply Bool.eq_true_not_negb;apply leb_not_gtb;trivial.
- rewrite Bool.negb_true_iff, <- Bool.not_true_iff_false in H.
- rewrite leb_spec; rewrite ltb_spec in H;lia.
-Qed.
-
-Lemma ltb_negb_geb : forall x y, x < y = negb (y <= x).
-Proof.
- intros;rewrite leb_negb_gtb, Bool.negb_involutive;trivial.
-Qed.
-
-Lemma to_Z_sub_gt : forall x y, y <= x = true -> [|x - y|] = ([|x|] - [|y|])%Z.
-Proof.
- intros x y;assert (W:= to_Z_bounded x);assert (W0:= to_Z_bounded y);
- rewrite leb_spec;intros;rewrite sub_spec, Zmod_small;lia.
-Qed.
-
-Lemma not_0_ltb : forall x, x <> 0 <-> 0 < x = true.
-Proof.
- intros x;rewrite ltb_spec, to_Z_0;assert (W:=to_Z_bounded x);split.
- intros Hd;assert ([|x|] <> 0)%Z;[ | lia].
- intros Heq;elim Hd;apply to_Z_inj;trivial.
- intros Hlt Heq;elimtype False.
- assert ([|x|] = 0)%Z;[ rewrite Heq, to_Z_0;trivial | lia].
-Qed.
-
-Lemma not_ltb_refl : forall i, ~(i < i = true).
-Proof.
- intros;rewrite ltb_spec;lia.
-Qed.
-
-Lemma to_Z_sub_1 : forall x y, y < x = true -> ([| x - 1|] = [|x|] - 1)%Z.
-Proof.
- intros;apply to_Z_sub_gt.
- generalize (leb_ltb_trans _ _ _ (leb_0 y) H).
- rewrite ltb_spec, leb_spec, to_Z_0, to_Z_1;auto with zarith.
-Qed.
-
-Lemma to_Z_sub_1_diff : forall x, x <> 0 -> ([| x - 1|] = [|x|] - 1)%Z.
-Proof.
- intros x;rewrite not_0_ltb;apply to_Z_sub_1.
-Qed.
-
-Lemma to_Z_add_1 : forall x y, x < y = true -> [|x+1|] = ([|x|] + 1)%Z.
-Proof.
- intros x y;assert (W:= to_Z_bounded x);assert (W0:= to_Z_bounded y);
- rewrite ltb_spec;intros;rewrite add_spec, to_Z_1, Zmod_small;lia.
-Qed.
-
-Lemma ltb_leb_sub1 : forall x i, x <> 0 -> (i < x = true <-> i <= x - 1 = true).
-Proof.
- intros x i Hdiff.
- rewrite ltb_spec, leb_spec, to_Z_sub_1_diff;trivial.
- split;auto with zarith.
-Qed.
-
-Lemma ltb_leb_add1 : forall x y i, i < y = true -> (i < x = true <-> i + 1 <= x = true).
-Proof.
- intros x y i Hlt.
- rewrite ltb_spec, leb_spec.
- rewrite (to_Z_add_1 i y);trivial.
- split;auto with zarith.
-Qed.
-
-(** Iterators *)
-
-Lemma foldi_gt : forall A f from to (a:A),
- (to < from)%int = true -> foldi f from to a = a.
-Proof.
- intros;unfold foldi;rewrite foldi_cont_gt;trivial.
-Qed.
-
-Lemma foldi_eq : forall A f from to (a:A),
- from = to -> foldi f from to a = f from a.
-Proof.
- intros;unfold foldi;rewrite foldi_cont_eq;trivial.
-Qed.
-
-Lemma foldi_lt : forall A f from to (a:A),
- (from < to)%int = true -> foldi f from to a = foldi f (from + 1) to (f from a).
-Proof.
- intros;unfold foldi;rewrite foldi_cont_lt;trivial.
-Qed.
-
-Lemma fold_gt : forall A f from to (a:A),
- (to < from)%int = true -> fold f from to a = a.
-Proof.
- intros;apply foldi_gt;trivial.
-Qed.
-
-Lemma fold_eq : forall A f from to (a:A),
- from = to -> fold f from to a = f a.
-Proof.
- intros;apply foldi_eq;trivial.
-Qed.
-
-Lemma fold_lt : forall A f from to (a:A),
- (from < to)%int = true -> fold f from to a = fold f (from + 1) to (f a).
-Proof.
- intros;apply foldi_lt;trivial.
-Qed.
-
-Lemma foldi_down_lt : forall A f from downto (a:A),
- (from < downto)%int = true -> foldi_down f from downto a = a.
-Proof.
- intros;unfold foldi_down;rewrite foldi_down_cont_lt;trivial.
-Qed.
-
-Lemma foldi_down_eq : forall A f from downto (a:A),
- from = downto -> foldi_down f from downto a = f from a.
-Proof.
- intros;unfold foldi_down;rewrite foldi_down_cont_eq;trivial.
-Qed.
-
-Lemma foldi_down_gt : forall A f from downto (a:A),
- (downto < from)%int = true->
- foldi_down f from downto a =
- foldi_down f (from-1) downto (f from a).
-Proof.
- intros;unfold foldi_down;rewrite foldi_down_cont_gt;trivial.
-Qed.
-
-Lemma fold_down_lt : forall A f from downto (a:A),
- (from < downto)%int = true -> fold_down f from downto a = a.
-Proof.
- intros;apply foldi_down_lt;trivial.
-Qed.
-
-Lemma fold_down_eq : forall A f from downto (a:A),
- from = downto -> fold_down f from downto a = f a.
-Proof.
- intros;apply foldi_down_eq;trivial.
-Qed.
-
-Lemma fold_down_gt : forall A f from downto (a:A),
- (downto < from)%int = true->
- fold_down f from downto a =
- fold_down f (from-1) downto (f a).
-Proof.
- intros;apply foldi_down_gt;trivial.
-Qed.
-
-Require Import Wf_Z.
-
-Lemma int_ind : forall (P:int -> Type),
- P 0%int ->
- (forall i, (i < max_int)%int = true -> P i -> P (i + 1)%int) ->
- forall i, P i.
-Proof.
- intros P HP0 Hrec.
- assert (forall z, (0 <= z)%Z -> forall i, z = [|i|] -> P i).
- intros z H;pattern z;apply natlike_rec2;intros;trivial.
- rewrite <- (of_to_Z i), <- H0;exact HP0.
- assert (W:= to_Z_bounded i).
- assert ([|i - 1|] = [|i|] - 1)%Z.
- rewrite sub_spec, Zmod_small;rewrite to_Z_1;auto with zarith.
- assert (i = i - 1 + 1)%int.
- apply to_Z_inj.
- rewrite add_spec, H2.
- rewrite Zmod_small;rewrite to_Z_1;auto with zarith.
- rewrite H3;apply Hrec.
- rewrite ltb_spec, H2;change [|max_int|] with (wB - 1)%Z;auto with zarith.
- apply X;auto with zarith.
- intros;apply (X [|i|]);trivial.
- destruct (to_Z_bounded i);trivial.
-Qed.
-
-Lemma int_ind_bounded : forall (P:int-> Type) min max,
- min <= max =true ->
- P max ->
- (forall i, min <= i + 1 = true-> i < max =true-> P (i + 1) -> P i) ->
- P min.
-Proof.
- intros P min max Hle.
- intros Hmax Hrec.
- assert (W1:= to_Z_bounded max);assert (W2:= to_Z_bounded min).
- assert (forall z, (0 <= z)%Z -> (z <= [|max|] - [|min|])%Z -> forall i, z = [|i|] -> P (max - i)%int).
- intros z H1;pattern z;apply natlike_rec2;intros;trivial.
- assert (max - i = max)%int.
- apply to_Z_inj;rewrite sub_spec, <- H0, Zminus_0_r, Zmod_small;auto using to_Z_bounded.
- rewrite H2;trivial.
- assert (W3:= to_Z_bounded i);apply Hrec.
- rewrite leb_spec,add_spec, sub_spec, to_Z_1, (Zmod_small ([|max|] - [|i|])), Zmod_small;auto with zarith.
- rewrite ltb_spec, sub_spec, Zmod_small;auto with zarith.
- assert (max - i + 1 = max - (i - 1))%int.
- apply to_Z_inj;rewrite add_spec, !sub_spec, to_Z_1.
- rewrite (Zmod_small ([|max|] - [|i|]));auto with zarith.
- rewrite (Zmod_small ([|i|] - 1));auto with zarith.
- apply f_equal2;auto with zarith.
- rewrite H3;apply X;auto with zarith.
- rewrite sub_spec, to_Z_1, <- H2, Zmod_small;auto with zarith.
- rewrite leb_spec in Hle;assert (min = max - (max - min))%int.
- apply to_Z_inj.
- rewrite !sub_spec, !Zmod_small;auto with zarith.
- rewrite Zmod_small;auto with zarith.
- rewrite H;apply (X [| max - min |]);trivial;rewrite sub_spec, Zmod_small;auto with zarith.
-Qed.
-
-Lemma foldi_cont_ZInd : forall A B (P: Z -> (A -> B) -> Prop) (f:int -> (A -> B) -> (A -> B)) min max cont,
- (forall z, ([|max|] < z)%Z -> P z cont) ->
- (forall i cont, min <= i = true -> i <= max = true -> P ([|i|] + 1)%Z cont -> P [|i|] (f i cont)) ->
- P [|min|] (foldi_cont f min max cont).
-Proof.
- intros A B P f min max cont Ha Hf.
- assert (Bmax:= to_Z_bounded max);assert (Bmin:= to_Z_bounded min).
- case_eq (min <= max);intros Heq.
- generalize (leb_refl min).
- assert (P ([|max|] + 1)%Z cont) by (apply Ha;auto with zarith).
- clear Ha;revert cont H.
- pattern min at 2 3 4;apply int_ind_bounded with max;trivial.
- intros;rewrite foldi_cont_eq;auto using leb_refl.
- intros i Hle Hlt Hr cont Hcont Hle'.
- rewrite foldi_cont_lt;[ | trivial].
- apply Hf;trivial. rewrite leb_spec;rewrite ltb_spec in Hlt;auto with zarith.
- assert ([|i|] + 1 = [|i + 1|])%Z.
- rewrite ltb_spec in Hlt;assert (W:= to_Z_bounded i);rewrite add_spec, to_Z_1, Zmod_small;lia.
- rewrite H;apply Hr;trivial.
- assert (max < min = true) by (rewrite ltb_negb_geb,Heq;trivial).
- rewrite foldi_cont_gt;trivial;apply Ha;rewrite <- ltb_spec;trivial.
-Qed.
-
-
-(* Lemma of_pos_spec : forall p, [|of_pos p|] = Zpos p mod wB. *)
-(* Proof. *)
-(* unfold of_pos. *)
-(* unfold wB. *)
-(* assert (forall k, (k <= size)%nat -> *)
-(* forall p : positive, [|of_pos_rec k p|] = Zpos p mod 2 ^ Z_of_nat k). *)
-(* induction k. *)
-(* simpl;intros;rewrite to_Z_0,Zmod_1_r;trivial. *)
-(* Opaque Z_of_nat. *)
-(* destruct p;simpl. *)
-(* destruct (bit_add_or (of_pos_rec k p << 1) 1) as (H1, _). *)
-(* rewrite <- H1;clear H1. *)
-(* change (Zpos p~1) with (2*(Zpos p) + 1)%Z. *)
-(* rewrite add_spec,lsl_spec, IHk, to_Z_1. *)
-(* rewrite Zmult_comm, Zplus_mod_idemp_l, Zmod_small. *)
-(* change 2%Z with (2^1)%Z. *)
-(* rewrite Zmod_distr. *)
-(* rewrite inj_S, Zpower_Zsucc;[ | apply Zle_0_nat]. *)
-(* repeat change (2^1)%Z with 2%Z. *)
-(* rewrite Zmult_mod_distr_l;trivial. *)
-(* Transparent Z_of_nat. *)
-(* rewrite inj_S;lia. *)
-(* discriminate. *)
-(* split;[discriminate | trivial]. *)
-(* compute;trivial. *)
-(* assert (W:0 <= Zpos p mod 2 ^ Z_of_nat k < 2 ^ Z_of_nat k). *)
-(* apply Z.mod_pos_bound;auto with zarith. *)
-(* change (2^1)%Z with 2%Z;split;try lia. *)
-(* apply Z.lt_le_trans with (2 ^ Z_of_nat (S k)). *)
-(* rewrite inj_S, Zpower_Zsucc;lia. *)
-(* unfold wB;apply Zpower_le_monotone;auto with zarith. *)
-(* split;auto using inj_le with zarith. *)
-(* auto with zarith. *)
-(* intros n H1 H2. *)
-(* rewrite bit_1, eqb_spec in H2;subst. *)
-(* rewrite bit_lsl in H1;discriminate H1. *)
-
-(* change (Zpos p~0) with (2*(Zpos p))%Z. *)
-(* rewrite lsl_spec, IHk, to_Z_1. *)
-(* rewrite Zmult_comm, Zmod_small. *)
-(* rewrite inj_S, Zpower_Zsucc;[ | apply Zle_0_nat]. *)
-(* rewrite Zmult_mod_distr_l;trivial. *)
-(* assert (W:0 <= Zpos p mod 2 ^ Z_of_nat k < 2 ^ Z_of_nat k). *)
-(* apply Z.mod_pos_bound;auto with zarith. *)
-(* change (2^1)%Z with 2%Z;split;try lia. *)
-(* apply Z.lt_le_trans with (2 ^ Z_of_nat (S k)). *)
-(* rewrite inj_S, Zpower_Zsucc;lia. *)
-(* unfold wB;apply Zpower_le_monotone;auto with zarith. *)
-(* split;auto using inj_le with zarith. *)
-(* auto with zarith. *)
-
-(* rewrite to_Z_1, Zmod_small;trivial. *)
-(* split;auto with zarith. *)
-(* apply Zpower_gt_1;auto with zarith. *)
-(* rewrite inj_S;auto with zarith. *)
-
-(* apply H;auto with zarith. *)
-(* Qed. *)
-
-Lemma of_Z_spec : forall z, [|of_Z z|] = z mod wB.
-Admitted. (* no more of_pos *)
-(* Proof. *)
-(* unfold of_Z;destruct z. *)
-(* assert (W:= to_Z_bounded 0);rewrite Zmod_small;trivial. *)
-(* apply of_pos_spec. *)
-(* rewrite opp_spec, of_pos_spec. *)
-(* rewrite <- Zmod_opp_opp. *)
-(* change (- Zpos p)%Z with (Zneg p). *)
-(* destruct (Z_eq_dec (Zneg p mod wB) 0). *)
-(* rewrite e, Z_mod_zero_opp_r;trivial. *)
-(* rewrite Z_mod_nz_opp_r, Zminus_mod, Z_mod_same_full, Zmod_mod, Zminus_0_r, Zmod_mod;trivial. *)
-(* Qed. *)
-
-Lemma foldi_cont_Ind : forall A B (P: int -> (A -> B) -> Prop) (f:int -> (A -> B) -> (A -> B)) min max cont,
- max < max_int = true ->
- (forall z, max < z = true -> P z cont) ->
- (forall i cont, min <= i = true -> i <= max = true -> P (i + 1) cont -> P i (f i cont)) ->
- P min (foldi_cont f min max cont).
-Proof.
- intros.
- set (P' z cont := (0 <= z < wB)%Z -> P (of_Z z) cont).
- assert (P' [|min|] (foldi_cont f min max cont)).
- apply foldi_cont_ZInd;unfold P';intros.
- assert ([|(of_Z z)|] = z).
- rewrite of_Z_spec, Zmod_small;trivial.
- apply H0;rewrite ltb_spec, H4;trivial.
- rewrite of_to_Z;apply H1;trivial.
- assert (i < max_int = true).
- apply leb_ltb_trans with max;trivial.
- rewrite <- (to_Z_add_1 _ _ H6), of_to_Z in H4;apply H4.
- apply to_Z_bounded.
- unfold P' in H2;rewrite of_to_Z in H2;apply H2;apply to_Z_bounded.
-Qed.
-
-Lemma foldi_cont_ind : forall A B (P: (A -> B) -> Prop) (f:int -> (A -> B) -> (A -> B)) min max cont,
- P cont ->
- (forall i cont, min <= i = true -> i <= max = true -> P cont -> P (f i cont)) ->
- P (foldi_cont f min max cont).
-Proof.
- intros A B P f min max cont Ha Hf.
- set (P2 := fun (z:Z) b => P b);change (P2 [|min|] (foldi_cont f min max cont)).
- apply foldi_cont_ZInd;trivial.
-Qed.
-
-Lemma foldi_ZInd : forall A (P : Z -> A -> Prop) f min max a,
- (max < min = true -> P ([|max|] + 1)%Z a) ->
- P [|min|] a ->
- (forall i a, min <= i = true -> i <= max = true ->
- P [|i|] a -> P ([|i|] + 1)%Z (f i a)) ->
- P ([|max|]+1)%Z (foldi f min max a).
-Proof.
- unfold foldi;intros A P f min max a Hlt;intros.
- set (P' z cont :=
- if Zlt_bool [|max|] z then cont = (fun a0 : A => a0)
- else forall a, P z a -> P ([|max|]+1)%Z (cont a)).
- assert (P' [|min|] (foldi_cont (fun (i : int) (cont : A -> A) (a0 : A) => cont (f i a0)) min
- max (fun a0 : A => a0))).
- apply foldi_cont_ZInd;intros;red.
- rewrite Zlt_is_lt_bool in H1;rewrite H1;trivial.
- case_eq (Zlt_bool [|max|] [|i|]);intros.
- rewrite <- Zlt_is_lt_bool in H4;rewrite leb_spec in H2;elimtype False;lia.
- clear H4; revert H3;unfold P'.
- case_eq (Zlt_bool [|max|] ([|i|] + 1));intros;auto.
- rewrite <- Zlt_is_lt_bool in H3; assert ([|i|] = [|max|]) by (rewrite leb_spec in H2;lia).
- rewrite H4, <- H6;apply H0;trivial.
- revert H1;unfold P'.
- case_eq (Zlt_bool [|max|] [|min|]);auto.
- rewrite <- Zlt_is_lt_bool, <- ltb_spec;intros;rewrite foldi_cont_gt;auto.
-Qed.
-
-Lemma foldi_Ind : forall A (P : int -> A -> Prop) f min max a,
- (max < max_int = true) ->
- (max < min = true -> P (max + 1) a) ->
- P min a ->
- (forall i a, min <= i = true -> i <= max = true ->
- P i a -> P (i + 1) (f i a)) ->
- P (max+1) (foldi f min max a).
-Proof.
- intros.
- set (P' z a := (0 <= z < wB)%Z -> P (of_Z z) a).
- assert (W:= to_Z_add_1 _ _ H).
- assert (P' ([|max|]+1)%Z (foldi f min max a)).
- apply foldi_ZInd;unfold P';intros.
- rewrite <- W, of_to_Z;auto.
- rewrite of_to_Z;trivial.
- assert (i < max_int = true).
- apply leb_ltb_trans with max;trivial.
- rewrite <- (to_Z_add_1 _ _ H7), of_to_Z;apply H2;trivial.
- rewrite of_to_Z in H5;apply H5;apply to_Z_bounded.
- unfold P' in H3;rewrite <- W, of_to_Z in H3;apply H3;apply to_Z_bounded.
-Qed.
-
-Lemma foldi_ind : forall A (P: A -> Prop) (f:int -> A -> A) min max a,
- P a ->
- (forall i a, min <= i = true -> i <= max = true -> P a -> P (f i a)) ->
- P (foldi f min max a).
-Proof.
- unfold foldi;intros A P f min max a Ha Hr;revert a Ha.
- apply foldi_cont_ind;auto.
-Qed.
-
-Lemma fold_ind : forall A (P: A -> Prop) (f: A -> A) min max a,
- P a -> (forall a, P a -> P (f a)) -> P (fold f min max a).
-Proof.
- unfold fold;intros A P f min max a Ha Hr;revert a Ha.
- apply foldi_cont_ind;auto.
-Qed.
-
-Lemma foldi_down_cont_ZInd :
- forall A B (P: Z -> (A -> B) -> Prop) (f:int -> (A -> B) -> (A -> B)) max min cont,
- (forall z, (z < [|min|])%Z -> P z cont) ->
- (forall i cont, min <= i = true -> i <= max = true -> P ([|i|] - 1)%Z cont -> P [|i|] (f i cont)) ->
- P [|max|] (foldi_down_cont f max min cont).
-Proof.
- intros A B P f max min cont Ha Hf.
- assert (Bmax:= to_Z_bounded max);assert (Bmin:= to_Z_bounded min).
- case_eq (min <= max);intros Heq.
- generalize (leb_refl max).
- assert (P ([|min|] -1)%Z cont) by (apply Ha;auto with zarith).
- clear Ha;revert cont H Heq.
- pattern max at 1 2 4 5;apply int_ind;trivial.
- intros; assert (0 = min).
- apply to_Z_inj;revert Heq;rewrite leb_spec, to_Z_0;lia.
- rewrite foldi_down_cont_eq;subst;auto.
- intros i Hmaxi Hr cont Hcont Hmin Hmax.
- generalize Hmin;rewrite leb_ltb_eqb;case_eq (min < i+1);simpl;intros Hlt Hmin'.
- rewrite foldi_down_cont_gt;[ | trivial].
- apply Hf;trivial.
- assert ([|i|] + 1 = [|i + 1|])%Z.
- assert (W:= to_Z_bounded i);rewrite ltb_spec in Hmaxi;
- assert (W2 := to_Z_bounded max_int);rewrite add_spec, to_Z_1, Zmod_small;auto with zarith.
- assert (i + 1 - 1 = i).
- rewrite leb_spec in *;rewrite ltb_spec in *.
- assert (W1:= to_Z_bounded i); apply to_Z_inj;rewrite sub_spec,to_Z_1, Zmod_small;try lia.
- assert ([|i|] = [|i+1|]-1)%Z.
- rewrite <- H;ring.
- rewrite <- H1, H0;apply Hr;trivial.
- rewrite ltb_spec in Hlt;rewrite leb_spec;lia.
- rewrite leb_spec in Hmax |- *;lia.
- rewrite eqb_spec in Hmin';subst;rewrite foldi_down_cont_eq;auto.
- assert (max < min = true) by (rewrite ltb_negb_geb,Heq;trivial).
- rewrite foldi_down_cont_lt;trivial.
- apply Ha;rewrite <- ltb_spec;trivial.
-Qed.
-
-Lemma foldi_down_cont_ind : forall A B (P: (A -> B) -> Prop) (f:int -> (A -> B) -> (A -> B)) max min cont,
- P cont ->
- (forall i cont, min <= i = true -> i <= max = true -> P cont -> P (f i cont)) ->
- P (foldi_down_cont f max min cont).
-Proof.
- intros A B P f max min cont Ha Hf.
- set (P2 := fun (z:Z) b => P b);change (P2 [|max|] (foldi_down_cont f max min cont)).
- apply foldi_down_cont_ZInd;trivial.
-Qed.
-
-Lemma foldi_down_ZInd :
- forall A (P: Z -> A -> Prop) (f:int -> A -> A) max min a,
- (max < min = true -> P ([|min|] - 1)%Z a) ->
- (P [|max|] a) ->
- (forall i a, min <= i = true -> i <= max = true -> P [|i|]%Z a -> P ([|i|]-1)%Z (f i a)) ->
- P ([|min|] - 1)%Z (foldi_down f max min a).
-Proof.
- unfold foldi_down;intros A P f max min a Hlt;intros.
- set (P' z cont :=
- if Zlt_bool z [|min|] then cont = (fun a0 : A => a0)
- else forall a, P z a -> P ([|min|] - 1)%Z (cont a)).
- assert (P' [|max|] (foldi_down_cont (fun (i : int) (cont : A -> A) (a0 : A) => cont (f i a0)) max
- min (fun a0 : A => a0))).
- apply foldi_down_cont_ZInd;intros;red.
- rewrite Zlt_is_lt_bool in H1;rewrite H1;trivial.
- case_eq (Zlt_bool [|i|] [|min|]);intros.
- rewrite <- Zlt_is_lt_bool in H4;rewrite leb_spec in H1;elimtype False;lia.
- clear H4;revert H3;unfold P'.
- case_eq (Zlt_bool ([|i|] - 1) [|min|]);intros;auto.
- rewrite <- Zlt_is_lt_bool in H3; assert ([|i|] = [|min|]) by (rewrite leb_spec in H1;lia).
- rewrite H4, <- H6. apply H0;trivial.
- revert H1;unfold P'.
- case_eq (Zlt_bool [|max|] [|min|]);auto.
- rewrite <- Zlt_is_lt_bool, <- ltb_spec;intros;rewrite foldi_down_cont_lt;auto.
-Qed.
-
-Lemma foldi_down_ind : forall A (P: A -> Prop) (f:int -> A -> A) max min a,
- P a ->
- (forall i a, min <= i = true -> i <= max = true -> P a -> P (f i a)) ->
- P (foldi_down f max min a).
-Proof.
- unfold foldi_down;intros A P f max min a Ha Hr;revert a Ha.
- apply foldi_down_cont_ind;auto.
-Qed.
-
-Lemma fold_down_ind : forall A (P: A -> Prop) (f: A -> A) max min a,
- P a -> (forall a, P a -> P (f a)) -> P (fold_down f max min a).
-Proof.
- unfold fold_down;intros A P f max min a Ha Hr;revert a Ha.
- apply foldi_down_cont_ind;auto.
-Qed.
-
-Lemma foldi_down_Ind :
- forall A (P: int -> A -> Prop) (f:int -> A -> A) max min a,
- 0 < min = true ->
- (max < min = true -> P (min - 1) a) ->
- (P max a) ->
- (forall i a, min <= i = true -> i <= max = true -> P i a -> P (i - 1) (f i a)) ->
- P (min - 1) (foldi_down f max min a).
-Proof.
- intros.
- set (P' z a := (0 <= z < wB)%Z -> P (of_Z z) a).
- assert (W:= to_Z_sub_1 _ _ H).
- assert (P' ([|min|]-1)%Z (foldi_down f max min a)).
- apply foldi_down_ZInd;unfold P';intros.
- rewrite <- W, of_to_Z;auto.
- rewrite of_to_Z;trivial.
- assert (0 < i = true).
- apply ltb_leb_trans with min;trivial.
- rewrite <- (to_Z_sub_1 _ _ H7), of_to_Z;apply H2;trivial.
- rewrite of_to_Z in H5;apply H5;apply to_Z_bounded.
- unfold P' in H3;rewrite <- W, of_to_Z in H3;apply H3;apply to_Z_bounded.
-Qed.
-
-Lemma foldi_down_min :
- forall A f min max (a:A),
- min < max_int = true->
- (min <= max) = true ->
- foldi_down f max min a = f min (foldi_down f max (min + 1) a).
-Proof.
- intros.
- set (P:= fun i => i <= max - min = true ->
- forall a, foldi_down f (min + i) min a = f min (foldi_down f (min + i) (min + 1) a)).
- assert (min < min + 1 = true).
- rewrite ltb_leb_add1 with (y:=max_int), leb_refl;trivial.
- assert (P (max - min)).
- apply int_ind;unfold P.
- replace (min + 0) with min.
- intros _ a'; rewrite foldi_down_eq, foldi_down_lt;trivial.
- apply to_Z_inj;rewrite add_spec, to_Z_0, Zplus_0_r, Zmod_small;auto using to_Z_bounded.
- intros i Hi Hrec Hi1 a'.
- rewrite add_assoc.
- assert (Wi:= to_Z_add_1 _ _ Hi).
- assert (Wmin:= to_Z_add_1 _ _ H).
- assert ((min + 1) <= (min + i + 1) = true).
- assert (W1 := to_Z_bounded min); assert (W2:= to_Z_bounded max); assert (W3:= to_Z_bounded i).
- replace (min + i + 1) with (min + 1 + i).
- rewrite leb_spec, (add_spec (min+1)).
- unfold is_true in Hi1;rewrite leb_spec in *; rewrite ltb_spec in *.
- rewrite sub_spec in Hi1;rewrite Zmod_small in Hi1;[ | lia].
- rewrite Zmod_small;lia.
- rewrite <- !add_assoc, (add_comm 1 i);trivial.
- rewrite leb_ltb_eqb in H2;revert H2.
- case_eq (min + 1 < min + i + 1).
- intros Hlt _;rewrite foldi_down_gt.
- rewrite foldi_down_gt with (from := min + i + 1);trivial.
- replace (min + i + 1 - 1) with (min + i).
- apply Hrec.
- apply leb_trans with (i+1);[rewrite leb_spec;lia | trivial].
- apply to_Z_inj;rewrite sub_spec, (add_spec (min + i)), to_Z_1, Zminus_mod_idemp_l.
- assert (H100: forall (x:Z), (x + 1 - 1)%Z = x) by (intros; ring). rewrite H100.
- rewrite Zmod_small;auto using to_Z_bounded.
- apply leb_ltb_trans with (2:= Hlt).
- rewrite leb_spec;lia.
- simpl;rewrite eqb_spec;intros _ Heq.
- rewrite <- Heq.
- rewrite foldi_down_gt.
- replace (min + 1 - 1) with min.
- rewrite !foldi_down_eq;trivial.
- apply to_Z_inj;rewrite sub_spec, add_spec, to_Z_1, Zminus_mod_idemp_l.
- replace ([|min|] + 1 - 1)%Z with [|min|] by ring.
- rewrite Zmod_small;auto using to_Z_bounded.
- rewrite ltb_spec;lia.
- generalize (H2 (leb_refl _) a).
- replace (min + (max - min)) with max;trivial.
- apply to_Z_inj;rewrite add_spec, sub_spec, Zplus_mod_idemp_r.
- ring_simplify ([|min|] + ([|max|] - [|min|]))%Z.
- rewrite Zmod_small;auto using to_Z_bounded.
-Qed.
-
-Definition foldi_ntr A f min max (a:A) :=
- foldi_cont (fun i cont _ => f i (cont tt)) min max (fun _ => a) tt.
-
-Lemma foldi_ntr_foldi_down : forall A f min max (a:A),
- max < max_int = true ->
- foldi_down f max min a = foldi_ntr _ f min max a.
-Proof.
- intros;unfold foldi_ntr.
- apply foldi_cont_Ind;trivial.
- intros;apply foldi_down_lt;trivial.
- intros i cont Hmin Hmax Heq;rewrite <- Heq;clear Heq.
- apply foldi_down_min;trivial.
- apply leb_ltb_trans with (1:= Hmax);trivial.
-Qed.
-
-
-(** Two iterators *)
-
-Lemma foldi_cont_ZInd2 : forall A B C D (P: Z -> (A -> B) -> (C -> D) -> Prop) (f1 : int -> (A -> B) -> (A -> B)) (f2 : int -> (C -> D) -> (C -> D)) min max cont1 cont2,
- (forall z, ([|max|] < z)%Z -> P z cont1 cont2) ->
- (forall i cont1 cont2, min <= i = true -> i <= max = true -> P ([|i|] + 1)%Z cont1 cont2 ->
- P [|i|] (f1 i cont1) (f2 i cont2)) ->
- P [|min|] (foldi_cont f1 min max cont1) (foldi_cont f2 min max cont2).
-Proof.
- intros.
- set (P' z cont :=
- if Zlt_bool [|max|] z then cont = cont1
- else P z cont (foldi_cont f2 (of_Z z) max cont2)).
- assert (P' [|min|] (foldi_cont f1 min max cont1)).
- apply foldi_cont_ZInd;unfold P';intros.
- rewrite Zlt_is_lt_bool in H1;rewrite H1;trivial.
- case_eq (Zlt_bool [|max|] [|i|]);intros.
- rewrite <- Zlt_is_lt_bool, <- ltb_spec in H4.
- elim (not_ltb_refl max);apply ltb_leb_trans with i;trivial.
- rewrite of_to_Z;generalize H2;rewrite leb_ltb_eqb, orb_true_iff;intros [Hlt | Heq].
- rewrite foldi_cont_lt;[apply H0 | ];trivial.
- revert H3;case_eq (Zlt_bool [|max|] ([|i|] + 1)).
- rewrite <- Zlt_is_lt_bool;rewrite ltb_spec in Hlt;intros;elimtype False;lia.
- rewrite <- (to_Z_add_1 _ _ Hlt), of_to_Z; intros _ W;exact W.
- rewrite eqb_spec in Heq;subst.
- rewrite foldi_cont_eq;[apply H0 | ];trivial.
- assert ([|max|] < [|max|] + 1)%Z by auto with zarith.
- rewrite Zlt_is_lt_bool in H5;rewrite H5 in H3;rewrite H3.
- apply H;rewrite Zlt_is_lt_bool;trivial.
- revert H1;unfold P';case_eq (Zlt_bool [|max|] [|min|]).
- rewrite <- Zlt_is_lt_bool;intros.
- rewrite H2;rewrite foldi_cont_gt;[ | rewrite ltb_spec];auto.
- rewrite of_to_Z;auto.
-Qed.
-
-
-Lemma foldi_cont_ind2 : forall A B C D (P: (A -> B) -> (C -> D) -> Prop) (f:int -> (A -> B) -> (A -> B)) (g:int -> (C -> D) -> (C -> D)) min max cont1 cont2,
- P cont1 cont2 ->
- (forall i cont1 cont2, min <= i = true -> i <= max = true -> P cont1 cont2 -> P (f i cont1) (g i cont2)) ->
- P (foldi_cont f min max cont1) (foldi_cont g min max cont2).
-Proof.
- intros A B C D P f g min max cont1 cont2 Ha Hf.
- set (P2 := fun (z:Z) b c => P b c);change (P2 [|min|] (foldi_cont f min max cont1) (foldi_cont g min max cont2)).
- apply foldi_cont_ZInd2;trivial.
-Qed.
-
-
-Lemma foldi_ZInd2 : forall A B (P : Z -> A -> B -> Prop) f g min max a b,
- (max < min = true -> P ([|max|] + 1)%Z a b) ->
- P [|min|] a b ->
- (forall i a b, min <= i = true -> i <= max = true ->
- P [|i|] a b -> P ([|i|] + 1)%Z (f i a) (g i b)) ->
- P ([|max|]+1)%Z (foldi f min max a) (foldi g min max b).
-Proof.
- unfold foldi;intros A B P f g min max a b Hlt;intros.
- set (P' z cont1 cont2 :=
- if Zlt_bool [|max|] z then cont1 = (fun a : A => a) /\ cont2 = (fun b : B => b)
- else forall a b, P z a b -> P ([|max|]+1)%Z (cont1 a) (cont2 b)).
- assert (P' [|min|] (foldi_cont (fun (i : int) (cont : A -> A) (a : A) => cont (f i a)) min
- max (fun a : A => a))
- (foldi_cont (fun (i : int) (cont : B -> B) (b : B) => cont (g i b)) min
- max (fun b : B => b))).
- apply foldi_cont_ZInd2;intros;red.
- rewrite Zlt_is_lt_bool in H1;rewrite H1;auto.
- case_eq (Zlt_bool [|max|] [|i|]);intros.
- rewrite <- Zlt_is_lt_bool in H4;rewrite leb_spec in H2;elimtype False;lia.
- clear H4; revert H3;unfold P'.
- case_eq (Zlt_bool [|max|] ([|i|] + 1));intros;auto.
- rewrite <- Zlt_is_lt_bool in H3; assert ([|i|] = [|max|]) by (rewrite leb_spec in H2;lia).
- destruct H4;subst;rewrite <- H6;apply H0;trivial.
- revert H1;unfold P'.
- case_eq (Zlt_bool [|max|] [|min|]);auto.
- rewrite <- Zlt_is_lt_bool, <- ltb_spec;intros;rewrite !foldi_cont_gt;auto.
-Qed.
-
-
-Lemma foldi_Ind2 : forall A B (P : int -> A -> B -> Prop) f g min max a b,
- (max < max_int = true) ->
- (max < min = true -> P (max + 1) a b) ->
- P min a b ->
- (forall i a b, min <= i = true -> i <= max = true ->
- P i a b -> P (i + 1) (f i a) (g i b)) ->
- P (max+1) (foldi f min max a) (foldi g min max b).
-Proof.
- intros.
- set (P' z a b := (0 <= z < wB)%Z -> P (of_Z z) a b).
- assert (W:= to_Z_add_1 _ _ H).
- assert (P' ([|max|]+1)%Z (foldi f min max a) (foldi g min max b)).
- apply foldi_ZInd2;unfold P';intros.
- rewrite <- W, of_to_Z;auto.
- rewrite of_to_Z;trivial.
- assert (i < max_int = true).
- apply leb_ltb_trans with max;trivial.
- rewrite <- (to_Z_add_1 _ _ H7), of_to_Z;apply H2;trivial.
- rewrite of_to_Z in H5;apply H5;apply to_Z_bounded.
- unfold P' in H3;rewrite <- W, of_to_Z in H3;apply H3;apply to_Z_bounded.
-Qed.
-
-
-Lemma foldi_ind2 : forall A B (P: A -> B -> Prop) (f:int -> A -> A) (g:int -> B -> B) min max a b,
- P a b ->
- (forall i a b, min <= i = true -> i <= max = true -> P a b -> P (f i a) (g i b)) ->
- P (foldi f min max a) (foldi g min max b).
-Proof.
- unfold foldi;intros A B P f g min max a b Ha Hr; revert a b Ha.
- apply (foldi_cont_ind2 _ _ _ _ (fun cont1 cont2 => forall a b, P a b -> P (cont1 a) (cont2 b))); auto.
-Qed.
-
-
-Lemma fold_ind2 : forall A B (P: A -> B -> Prop) (f: A -> A) (g: B -> B) min max a b,
- P a b -> (forall a b, P a b -> P (f a) (g b)) -> P (fold f min max a) (fold g min max b).
-Proof.
- unfold fold;intros A B P f g min max a b Ha Hr;revert a b Ha.
- apply (foldi_cont_ind2 _ _ _ _ (fun cont1 cont2 => forall a b, P a b -> P (cont1 a) (cont2 b)));auto.
-Qed.
-
-Lemma foldi_eq_compat : forall A (f1 f2:int -> A -> A) min max a,
- (forall i a, min <= i = true -> i <= max = true -> f1 i a = f2 i a) ->
- foldi f1 min max a = foldi f2 min max a.
-Proof.
- intros; set (P' (z:Z) (a1 a2:A) := a1 = a2).
- assert (P' ([|max|] + 1)%Z (foldi f1 min max a) (foldi f2 min max a)).
- apply foldi_ZInd2;unfold P';intros;subst;auto.
- apply H0.
-Qed.
-
-Lemma foldi_down_cont_ZInd2 :
- forall A B C D (P: Z -> (A -> B) -> (C -> D) -> Prop) (f1:int -> (A -> B) -> (A -> B)) (f2:int -> (C -> D) -> (C -> D)) max min cont1 cont2,
- (forall z, (z < [|min|])%Z -> P z cont1 cont2) ->
- (forall i cont1 cont2, min <= i = true -> i <= max = true -> P ([|i|] - 1)%Z cont1 cont2 ->
- P [|i|] (f1 i cont1) (f2 i cont2)) ->
- P [|max|] (foldi_down_cont f1 max min cont1) (foldi_down_cont f2 max min cont2).
-Proof.
- intros.
- set (P' z cont :=
- if Zlt_bool z [|min|] then cont = cont1
- else P z cont (foldi_down_cont f2 (of_Z z) min cont2)).
- assert (P' [|max|] (foldi_down_cont f1 max min cont1)).
- apply foldi_down_cont_ZInd;unfold P';intros.
- rewrite Zlt_is_lt_bool in H1;rewrite H1;trivial.
- case_eq (Zlt_bool [|i|] [|min|]);intros.
- rewrite <- Zlt_is_lt_bool, <- ltb_spec in H4.
- elim (not_ltb_refl min);apply leb_ltb_trans with i;trivial.
- rewrite of_to_Z;generalize H1;rewrite leb_ltb_eqb, orb_true_iff;intros [Hlt | Heq].
- rewrite foldi_down_cont_gt;[apply H0 | ];trivial.
- revert H3;case_eq (Zlt_bool ([|i|] - 1) [|min|]).
- rewrite <- Zlt_is_lt_bool;rewrite ltb_spec in Hlt;intros;elimtype False;lia.
- rewrite <- (to_Z_sub_1 _ _ Hlt), of_to_Z; intros _ W;exact W.
- rewrite eqb_spec in Heq;subst.
- rewrite foldi_down_cont_eq;[apply H0 | ];trivial.
- assert ([|i|] - 1 < [|i|])%Z by auto with zarith.
- rewrite Zlt_is_lt_bool in H5;rewrite H5 in H3;rewrite H3.
- apply H;rewrite Zlt_is_lt_bool;trivial.
- revert H1;unfold P';case_eq (Zlt_bool [|max|] [|min|]).
- rewrite <- Zlt_is_lt_bool;intros.
- rewrite H2;rewrite foldi_down_cont_lt;[ | rewrite ltb_spec];auto.
- rewrite of_to_Z;auto.
-Qed.
-
-
-Lemma foldi_down_cont_ind2 : forall A B C D (P: (A -> B) -> (C -> D) -> Prop) (f:int -> (A -> B) -> (A -> B)) (g:int -> (C -> D) -> (C -> D)) max min cont1 cont2,
- P cont1 cont2 ->
- (forall i cont1 cont2, min <= i = true -> i <= max = true -> P cont1 cont2 -> P (f i cont1) (g i cont2)) ->
- P (foldi_down_cont f max min cont1) (foldi_down_cont g max min cont2).
-Proof.
- intros A B C D P f g max min cont1 cont2 Ha Hf.
- set (P2 := fun (z:Z) b c => P b c);change (P2 [|max|] (foldi_down_cont f max min cont1) (foldi_down_cont g max min cont2)).
- apply foldi_down_cont_ZInd2;trivial.
-Qed.
-
-
-Lemma foldi_down_ZInd2 :
- forall A B (P: Z -> A -> B -> Prop) (f1:int -> A -> A) (f2:int -> B -> B) max min a1 a2,
- (max < min = true -> P ([|min|] - 1)%Z a1 a2) ->
- (P [|max|] a1 a2) ->
- (forall z, (z < [|min|])%Z -> P z a1 a2) ->
- (forall i a1 a2, min <= i = true -> i <= max = true -> P [|i|] a1 a2 ->
- P ([|i|] - 1)%Z (f1 i a1) (f2 i a2)) ->
- P ([|min|] - 1)%Z (foldi_down f1 max min a1) (foldi_down f2 max min a2).
-Proof.
- unfold foldi_down;intros A B P f1 f2 max min a1 a2 Hlt;intros.
- set (P' z cont1 cont2 :=
- if Zlt_bool z [|min|] then cont1 = (fun a0 : A => a0) /\ cont2 = (fun a0 : B => a0)
- else forall a1 a2, P z a1 a2 -> P ([|min|] - 1)%Z (cont1 a1) (cont2 a2)).
- assert (P' [|max|] (foldi_down_cont (fun (i : int) (cont : A -> A) (a0 : A) => cont (f1 i a0)) max
- min (fun a0 : A => a0))
- (foldi_down_cont (fun (i : int) (cont : B -> B) (a0 : B) => cont (f2 i a0)) max
- min (fun a0 : B => a0))).
- apply foldi_down_cont_ZInd2;intros;red.
- rewrite Zlt_is_lt_bool in H2;rewrite H2;auto.
- case_eq (Zlt_bool [|i|] [|min|]);intros.
- rewrite <- Zlt_is_lt_bool in H5;rewrite leb_spec in H2;elimtype False;lia.
- clear H5;revert H4;unfold P'.
- case_eq (Zlt_bool ([|i|] - 1) [|min|]);intros;auto.
- rewrite <- Zlt_is_lt_bool in H4; assert ([|i|] = [|min|]) by (rewrite leb_spec in H2;lia).
- destruct H5;subst;rewrite <- H7;apply H1;trivial.
- revert H2;unfold P'.
- case_eq (Zlt_bool [|max|] [|min|]);auto.
- rewrite <- Zlt_is_lt_bool, <- ltb_spec;intros;rewrite foldi_down_cont_lt;auto.
- destruct H3. rewrite H4;auto.
-Qed.
-
-
-Lemma foldi_down_ind2 : forall A B (P: A -> B -> Prop) (f:int -> A -> A) (g:int -> B -> B) max min a b,
- P a b ->
- (forall i a b, min <= i = true -> i <= max = true -> P a b -> P (f i a) (g i b)) ->
- P (foldi_down f max min a) (foldi_down g max min b).
-Proof.
- unfold foldi_down;intros A B P f g max min a b Ha Hr;revert a b Ha.
- apply (foldi_down_cont_ind2 _ _ _ _ (fun cont1 cont2 => forall a b, P a b -> P (cont1 a) (cont2 b)));auto.
-Qed.
-
-
-Lemma fold_down_ind2 : forall A B (P: A -> B -> Prop) (f: A -> A) (g: B -> B) max min a b,
- P a b -> (forall a b, P a b -> P (f a) (g b)) -> P (fold_down f max min a) (fold_down g max min b).
-Proof.
- unfold fold_down;intros A B P f g max min a b Ha Hr;revert a b Ha.
- apply (foldi_down_cont_ind2 _ _ _ _ (fun cont1 cont2 => forall a b, P a b -> P (cont1 a) (cont2 b)));auto.
-Qed.
-
-Lemma foldi_down_eq_compat : forall A (f1 f2:int -> A -> A) max min a,
- (forall i a, min <= i = true -> i <= max = true -> f1 i a = f2 i a) ->
- foldi_down f1 max min a = foldi_down f2 max min a.
-Proof.
- intros; set (P' (z:Z) (a1 a2:A) := a1 = a2).
- assert (P' ([|min|] - 1)%Z (foldi_down f1 max min a) (foldi_down f2 max min a)).
- apply foldi_down_ZInd2;unfold P';intros;subst;auto.
- apply H0.
-Qed.
-
-
-Lemma forallb_spec : forall f from to,
- forallb f from to = true <->
- forall i, from <= i = true -> i <= to = true -> f i = true.
-Proof.
- unfold forallb;intros f from to.
- setoid_rewrite leb_spec.
- apply foldi_cont_ZInd.
- intros;split;[intros;elimtype False;lia | trivial].
- intros i cont Hfr Hto Hcont.
- case_eq (f i);intros Heq.
- rewrite Hcont;clear Hcont;split;auto with zarith;intros.
- assert (H2 : ([|i0|] = [|i|] \/ [|i|] + 1 <= [|i0|])%Z) by lia; destruct H2;auto with zarith.
- apply to_Z_inj in H2;rewrite H2;trivial.
- split;[discriminate | intros].
- rewrite leb_spec in Hto;rewrite <- Heq;auto with zarith.
-Qed.
-
-Lemma forallb_eq_compat : forall f1 f2 from to,
- (forall i, from <= i = true -> i <= to = true -> f1 i = f2 i) ->
- forallb f1 from to = forallb f2 from to.
-Proof.
- unfold forallb;intros.
- set (P' (z:Z) (cont1 cont2:unit -> bool) := cont1 tt = cont2 tt).
- refine (foldi_cont_ZInd2 _ _ _ _ P' _ _ from to _ _ _ _);unfold P';intros;trivial.
- rewrite H2, H;trivial.
-Qed.
-
-Lemma existsb_spec : forall f from to,
- existsb f from to = true <->
- exists i, ((from <= i) && (i <= to) && (f i)) = true .
-Proof.
- unfold existsb;intros.
- repeat setoid_rewrite andb_true_iff;setoid_rewrite leb_spec.
- apply foldi_cont_ZInd.
- intros;split;[discriminate | intros [i [W1 W2]];elimtype False;lia].
- intros i cont Hfr Hto Hcont.
- case_eq (f i);intros Heq.
- split;trivial.
- exists i;rewrite leb_spec in Hto;auto with zarith.
- rewrite Hcont;clear Hcont;split;intros [i0 [W1 W2]];
- exists i0;split;auto with zarith.
- assert (~ [|i|] = [|i0|]);[ | auto with zarith].
- intros W;apply to_Z_inj in W;rewrite W in Heq;rewrite Heq in W2;discriminate.
-Qed.
-
-Lemma existsb_eq_compat : forall f1 f2 from to,
- (forall i, from <= i = true -> i <= to = true -> f1 i = f2 i) ->
- existsb f1 from to = existsb f2 from to.
-Proof.
- unfold existsb;intros.
- set (P' (z:Z) (cont1 cont2:unit -> bool) := cont1 tt = cont2 tt).
- refine (foldi_cont_ZInd2 _ _ _ _ P' _ _ from to _ _ _ _);unfold P';intros;trivial.
- rewrite H2, H;trivial.
-Qed.
-
-
-Lemma bit_max_int : forall i, (i < digits)%int = true -> bit max_int i = true.
-Proof.
- intros;apply (forallb_spec (bit max_int) 0 (digits - 1)).
- vm_compute;trivial.
- apply leb_0.
- rewrite ltb_spec in H.
- destruct (to_Z_bounded i);rewrite leb_spec.
- change [|digits - 1 |] with ([|digits|] - 1)%Z;lia.
-Qed.
-
-Lemma land_max_int_l : forall i, max_int land i = i.
-Proof.
- intros;apply bit_eq;intros.
- rewrite land_spec.
- destruct (reflect_leb digits i0).
- rewrite <- leb_spec in l.
- rewrite !bit_M;trivial.
- rewrite bit_max_int;trivial.
- rewrite ltb_spec;lia.
-Qed.
-
-Lemma land_max_int_r : forall i, i land max_int = i.
-Proof.
- intros;rewrite land_comm;apply land_max_int_l.
-Qed.
-
-
-(* int is an OrderedType *)
-
-Require Import OrderedType.
-
-Module IntOrderedType <: OrderedType.
-
- Definition t := int.
-
- Definition eq x y := (x == y) = true.
-
- Definition lt x y := (x < y) = true.
-
- Lemma eq_refl x : eq x x.
- Proof. unfold eq. rewrite eqb_spec. reflexivity. Qed.
-
- Lemma eq_sym x y : eq x y -> eq y x.
- Proof. unfold eq. rewrite !eqb_spec. intros ->. reflexivity. Qed.
-
- Lemma eq_trans x y z : eq x y -> eq y z -> eq x z.
- Proof. unfold eq. rewrite !eqb_spec. intros -> ->. reflexivity. Qed.
-
- Lemma lt_trans x y z : lt x y -> lt y z -> lt x z.
- Proof. apply ltb_trans. Qed.
-
- Lemma lt_not_eq x y : lt x y -> ~ eq x y.
- Proof. unfold lt, eq. rewrite ltb_negb_geb, eqb_spec. intros H1 H2. rewrite H2, leb_refl in H1. discriminate. Qed.
-
- Definition compare x y : Compare lt eq x y.
- Proof.
- case_eq (x < y); intro e.
- exact (LT e).
- case_eq (x == y); intro e2.
- exact (EQ e2). apply GT. unfold lt. rewrite ltb_negb_geb, leb_ltb_eqb, e, e2. reflexivity.
- Defined.
-
- Definition eq_dec x y : { eq x y } + { ~ eq x y }.
- Proof.
- case_eq (x == y); intro e.
- left; exact e.
- right. intro H. rewrite H in e. discriminate.
- Defined.
-
-End IntOrderedType.
-
-
-(*
- Local Variables:
- coq-load-path: ((rec "../../.." "SMTCoq"))
- End:
-*)
diff --git a/src/Makefile b/src/Makefile
index 95955bb..3efcfc3 100644
--- a/src/Makefile
+++ b/src/Makefile
@@ -7,7 +7,7 @@
## # GNU Lesser General Public License Version 2.1 ##
## # (see LICENSE file for the text of the license) ##
##########################################################################
-## GNUMakefile for Coq 8.13.1
+## GNUMakefile for Coq 8.13.2
# For debugging purposes (must stay here, don't move below)
INITIAL_VARS := $(.VARIABLES)
@@ -218,7 +218,7 @@ COQDOCLIBS?=$(COQLIBS_NOML)
# The version of Coq being run and the version of coq_makefile that
# generated this makefile
COQ_VERSION:=$(shell $(COQC) --print-version | cut -d " " -f 1)
-COQMAKEFILE_VERSION:=8.13.1
+COQMAKEFILE_VERSION:=8.13.2
COQSRCLIBS?= $(foreach d,$(COQ_SRC_SUBDIRS), -I "$(COQLIB)/$(d)")
diff --git a/src/Misc.v b/src/Misc.v
index 520c41c..e558ce3 100644
--- a/src/Misc.v
+++ b/src/Misc.v
@@ -10,10 +10,12 @@
(**************************************************************************)
-Require Import Bool List PArray Int63 ZArith Psatz.
+Require Import Bool List PArray Int63 Ring63 ZArith Psatz.
Local Open Scope int63_scope.
Local Open Scope array_scope.
+Global Notation "[| x |]" := (φ x).
+
(** Lemmas about Bool *)
@@ -23,6 +25,23 @@ Proof. intros [ | ]; reflexivity. Qed.
(** Lemmas about Int63 *)
+Lemma reflect_eqb : forall i j, reflect (i = j)%Z (i == j).
+Proof.
+ intros; apply iff_reflect.
+ symmetry;apply eqb_spec.
+Qed.
+
+Lemma to_Z_eq : forall x y, [|x|] = [|y|] <-> x = y.
+Proof.
+ split;intros;subst;trivial.
+ apply to_Z_inj;trivial.
+Qed.
+
+Lemma max_int_wB : [|max_int|] = (wB - 1)%Z.
+Proof.
+ reflexivity.
+Qed.
+
Lemma le_eq : forall i j,
(j <= i) = true -> (j + 1 <= i) = false -> i = j.
Proof.
@@ -32,7 +51,7 @@ Proof.
assert (H2: (([|j|] + 1)%Z < wB)%Z \/ ([|j|] + 1)%Z = wB).
pose (H3 := to_Z_bounded j); lia.
destruct H2 as [H2|H2].
- rewrite Zmod_small in H.
+ rewrite Z.mod_small in H.
lia.
split.
pose (H3 := to_Z_bounded j); lia.
@@ -40,6 +59,15 @@ Proof.
rewrite H2, Z_mod_same_full in H; elim H; destruct (to_Z_bounded i) as [H3 _]; assumption.
Qed.
+Lemma leb_0 : forall x, 0 <= x = true.
+Proof.
+ intros x;rewrite leb_spec;destruct (to_Z_bounded x);trivial.
+Qed.
+
+Lemma leb_refl : forall n, n <= n = true.
+Proof.
+ intros n;rewrite leb_spec;apply Z.le_refl.
+Qed.
Lemma lt_eq : forall i j,
(i < j + 1) = true -> (i < j) = false -> i = j.
@@ -50,7 +78,7 @@ Proof.
assert (H2: (([|j|] + 1)%Z < wB)%Z \/ ([|j|] + 1)%Z = wB).
pose (H3 := to_Z_bounded j); lia.
destruct H2 as [H2|H2].
- rewrite Zmod_small in H1.
+ rewrite Z.mod_small in H1.
lia.
split.
pose (H3 := to_Z_bounded j); lia.
@@ -58,6 +86,100 @@ Proof.
rewrite H2, Z_mod_same_full in H1; elimtype False. destruct (to_Z_bounded i) as [H3 _]. lia.
Qed.
+Lemma not_0_ltb : forall x, x <> 0 <-> 0 < x = true.
+Proof.
+ intros x;rewrite ltb_spec, to_Z_0;assert (W:=to_Z_bounded x);split.
+ intros Hd;assert ([|x|] <> 0)%Z;[ | omega].
+ intros Heq;elim Hd;apply to_Z_inj;trivial.
+ intros Hlt Heq;elimtype False.
+ assert ([|x|] = 0)%Z;[ rewrite Heq, to_Z_0;trivial | omega].
+Qed.
+
+Lemma ltb_0 : forall x, ~ (x < 0 = true).
+Proof.
+ intros x;rewrite ltb_spec, to_Z_0;destruct (to_Z_bounded x);omega.
+Qed.
+
+Lemma not_ltb_refl : forall i, ~(i < i = true).
+Proof.
+ intros;rewrite ltb_spec;omega.
+Qed.
+
+Lemma ltb_trans : forall x y z, x < y = true -> y < z = true -> x < z = true.
+Proof.
+ intros x y z;rewrite !ltb_spec;apply Z.lt_trans.
+Qed.
+
+Lemma leb_ltb_eqb : forall x y, ((x <= y) = (x < y) || (x == y)).
+Proof.
+ intros.
+ apply eq_true_iff_eq.
+ rewrite leb_spec, orb_true_iff, ltb_spec, eqb_spec, <- to_Z_eq;omega.
+Qed.
+
+Lemma leb_ltb_trans : forall x y z, x <= y = true -> y < z = true -> x < z = true.
+Proof.
+ intros x y z;rewrite leb_spec, !ltb_spec;apply Z.le_lt_trans.
+Qed.
+
+Lemma to_Z_add_1 : forall x y, x < y = true -> [|x+1|] = ([|x|] + 1)%Z.
+Proof.
+ intros x y;assert (W:= to_Z_bounded x);assert (W0:= to_Z_bounded y);
+ rewrite ltb_spec;intros;rewrite add_spec, to_Z_1, Z.mod_small;omega.
+Qed.
+
+Lemma to_Z_add_1_wB : forall x, ([|x|] < wB - 1)%Z -> [|x + 1|] = ([|x|] + 1)%Z.
+Proof.
+ intros; assert (Bx := to_Z_bounded x); rewrite add_spec, to_Z_1, Z.mod_small; lia.
+Qed.
+
+Lemma leb_not_gtb : forall n m, m <= n = true -> ~(n < m = true).
+Proof.
+ intros n m; rewrite ltb_spec, leb_spec;omega.
+Qed.
+
+Lemma leb_negb_gtb : forall x y, x <= y = negb (y < x).
+Proof.
+ intros x y;apply Bool.eq_true_iff_eq;split;intros.
+ apply Bool.eq_true_not_negb;apply leb_not_gtb;trivial.
+ rewrite Bool.negb_true_iff, <- Bool.not_true_iff_false in H.
+ rewrite leb_spec; rewrite ltb_spec in H;omega.
+Qed.
+
+Lemma ltb_negb_geb : forall x y, x < y = negb (y <= x).
+Proof.
+ intros;rewrite leb_negb_gtb, Bool.negb_involutive;trivial.
+Qed.
+
+Lemma to_Z_sub_gt : forall x y, y <= x = true -> [|x - y|] = ([|x|] - [|y|])%Z.
+Proof.
+ intros x y;assert (W:= to_Z_bounded x);assert (W0:= to_Z_bounded y);
+ rewrite leb_spec;intros;rewrite sub_spec, Zmod_small;omega.
+Qed.
+
+Lemma to_Z_sub_1 : forall x y, y < x = true -> ([| x - 1|] = [|x|] - 1)%Z.
+Proof.
+ intros;apply to_Z_sub_gt.
+ generalize (leb_ltb_trans _ _ _ (leb_0 y) H).
+ rewrite ltb_spec, leb_spec, to_Z_0, to_Z_1;auto with zarith.
+Qed.
+
+Lemma to_Z_sub_1_diff : forall x, x <> 0 -> ([| x - 1|] = [|x|] - 1)%Z.
+Proof.
+ intros x;rewrite not_0_ltb;apply to_Z_sub_1.
+Qed.
+
+Lemma to_Z_sub_1_0 : forall x, (0 < [|x|])%Z -> [|x - 1|] = ([|x|] - 1)%Z.
+Proof.
+ intros; apply (to_Z_sub_1 _ 0); rewrite ltb_spec; assumption.
+Qed.
+
+Lemma ltb_leb_sub1 : forall x i, x <> 0 -> (i < x = true <-> i <= x - 1 = true).
+Proof.
+ intros x i Hdiff.
+ rewrite ltb_spec, leb_spec, to_Z_sub_1_diff;trivial.
+ split;auto with zarith.
+Qed.
Lemma minus_1_lt i : (i == 0) = false -> i - 1 < i = true.
Proof.
@@ -70,689 +192,908 @@ Proof.
clear -H H1. change [|0|] with 0%Z. lia.
Qed.
+Lemma lsr0_l i: 0 >> i = 0.
+Proof.
+ apply to_Z_inj.
+ generalize (lsr_spec 0 i).
+ rewrite to_Z_0, Zdiv_0_l; auto.
+Qed.
-Lemma foldi_down_ZInd2 :
- forall A (P: Z -> A -> Prop) (f:int -> A -> A) max min a,
- (max < min = true -> P ([|min|])%Z a) ->
- (P ([|max|]+1)%Z a) ->
- (forall i a, min <= i = true -> i <= max = true -> P ([|i|]+1)%Z a -> P [|i|] (f i a)) ->
- P [|min|] (foldi_down f max min a).
-Proof.
- unfold foldi_down;intros A P f max min a Hlt;intros.
- set (P' z cont :=
- if Zlt_bool z ([|min|]+1)%Z then cont = (fun a0 : A => a0)
- else forall a, P z a -> P [|min|] (cont a)).
- assert (H1: P' ([|max|]+1)%Z (foldi_down_cont (fun (i : int) (cont : A -> A) (a0 : A) => cont (f i a0)) max
- min (fun a0 : A => a0))).
- apply foldi_down_cont_ZInd;intros;red.
- assert (H20: (z+1 < [|min|]+1)%Z).
- lia.
- rewrite Zlt_is_lt_bool in H20; rewrite H20;trivial.
- case_eq (Zlt_bool ([|i|]+1) ([|min|]+1));intros.
- rewrite <- Zlt_is_lt_bool in H4;rewrite leb_spec in H1;elimtype False;lia.
- clear H4;revert H3;unfold P'.
- case_eq (Zlt_bool ([|i|] - 1 + 1) ([|min|]+1));intros;auto.
- rewrite <- Zlt_is_lt_bool in H3; assert ([|i|] = [|min|]) by (rewrite leb_spec in H1;lia).
- rewrite H4, <- H6. apply H0;trivial.
- apply H4. replace ([|i|] - 1 + 1)%Z with [|i|] by lia. auto.
- revert H1;unfold P'.
- case_eq (Zlt_bool ([|max|]+1)%Z ([|min|]+1)%Z);auto.
- rewrite <- Zlt_is_lt_bool.
- intro H22; assert (H21: ([|max|] < [|min|])%Z). lia.
- rewrite <- ltb_spec in H21. intros;rewrite foldi_down_cont_lt;auto.
-Qed.
-
-
-Lemma foldi_down_Ind2 : forall A (P : int -> A -> Prop) f max min a,
- (max < max_int = true) ->
- (max < min = true -> P min a) ->
- P (max+1) a ->
- (forall i a, min <= i = true -> i <= max = true ->
- P (i+1) a -> P i (f i a)) ->
- P min (foldi_down f max min a).
-Proof.
- intros A P f max min a H H0 H1 H2.
- set (P' z a := (0 <= z < wB)%Z -> P (of_Z z) a).
- assert (W:= to_Z_add_1 _ _ H).
- assert (P' ([|min|])%Z (foldi_down f max min a)).
- apply foldi_down_ZInd2;unfold P';intros.
- rewrite of_to_Z;auto.
- rewrite <- W, of_to_Z;auto.
- rewrite of_to_Z. apply H2; trivial.
- assert (i < max_int = true).
- apply leb_ltb_trans with max; trivial.
- rewrite <- (to_Z_add_1 _ _ H7) in H5. rewrite of_to_Z in H5. apply H5. apply to_Z_bounded.
- unfold P' in H3; rewrite of_to_Z in H3;apply H3;apply to_Z_bounded.
-Qed.
-
-
-(** Lemmas about PArray.to_list *)
+Lemma lxor_lsr i1 i2 i: (i1 lxor i2) >> i = (i1 >> i) lxor (i2 >> i).
+Proof.
+ apply bit_ext; intros n.
+ rewrite lxor_spec, !bit_lsr, lxor_spec.
+ case (_ <= _); auto.
+Qed.
-Lemma to_list_In : forall {A} (t: array A) i,
- i < length t = true -> In (t.[i]) (to_list t).
+Lemma bit_or_split i : i = (i>>1)<<1 lor bit i 0.
Proof.
- intros A t i H; unfold to_list; case_eq (0 == length t); intro Heq.
- unfold is_true in H; rewrite ltb_spec in H; rewrite eqb_spec in Heq; rewrite <- Heq in H; rewrite to_Z_0 in H; pose (H1 := to_Z_bounded i); elimtype False; lia.
- pose (P := fun j a => i < j = true \/ In (t .[ i]) a).
- pose (H1:= foldi_down_Ind2 _ P); unfold P in H1.
- assert (H2: i < 0 = true \/ In (t .[ i]) (foldi_down (fun (i0 : int) (l : list A) => t .[ i0] :: l) (length t - 1) 0 nil)).
- apply H1.
- rewrite ltb_spec; erewrite to_Z_sub_1; try eassumption.
- pose (H2 := to_Z_bounded (length t)); change [|max_int|] with (wB-1)%Z; lia.
- intro H2; elimtype False; rewrite ltb_spec, to_Z_0 in H2; pose (H3 := to_Z_bounded (length t - 1)); lia.
- left; unfold is_true; rewrite ltb_spec; rewrite (to_Z_add_1 _ max_int).
- erewrite to_Z_sub_1; try eassumption.
- unfold is_true in H; rewrite ltb_spec in H; lia.
- rewrite ltb_spec; erewrite to_Z_sub_1; try eassumption.
- pose (H2 := to_Z_bounded (length t)); change [|max_int|] with (wB-1)%Z; lia.
- intros j a H2 H3 [H4|H4].
- case_eq (i < j); intro Heq2.
- left; reflexivity.
- right; rewrite (lt_eq _ _ H4 Heq2); constructor; reflexivity.
- right; constructor 2; assumption.
- destruct H2 as [H2|H2]; try assumption.
- unfold is_true in H2; rewrite ltb_spec, to_Z_0 in H2; pose (H3 := to_Z_bounded i); elimtype False; lia.
+ apply bit_ext.
+ intros n; rewrite lor_spec.
+ rewrite bit_lsl, bit_lsr, bit_b2i.
+ case (to_Z_bounded n); intros Hi _.
+ case (Zle_lt_or_eq _ _ Hi).
+ 2: replace 0%Z with [|0|]; auto; rewrite to_Z_eq.
+ 2: intros H; rewrite <-H.
+ 2: replace (0 < 1) with true; auto.
+ intros H; clear Hi.
+ case_eq (n == 0).
+ rewrite eqb_spec; intros H1; generalize H; rewrite H1; discriminate.
+ intros _; rewrite orb_false_r.
+ case_eq (n < 1).
+ rewrite ltb_spec, to_Z_1; intros HH; contradict HH; auto with zarith.
+ intros _.
+ generalize (@bit_M i n); case (_ <= _).
+ intros H1; rewrite H1; auto.
+ intros _.
+ case (to_Z_bounded n); intros H1n H2n.
+ assert (F1: [|n - 1|] = ([|n|] - 1)%Z).
+ rewrite sub_spec, Zmod_small; rewrite to_Z_1; auto with zarith.
+ generalize (add_le_r 1 (n - 1)); case (_ <= _); rewrite F1, to_Z_1; intros HH.
+ replace (1 + (n -1)) with n. change (bit i n = bit i n). reflexivity.
+ apply to_Z_inj; rewrite add_spec, F1, Zmod_small; rewrite to_Z_1;
+ auto with zarith.
+ rewrite bit_M; auto; rewrite leb_spec.
+ replace [|n|] with wB; try discriminate; auto with zarith.
Qed.
-Lemma to_list_In_eq : forall {A} (t: array A) i x,
- i < length t = true -> x = t.[i] -> In x (to_list t).
+Lemma lsr_is_even_eq : forall i j,
+ i >> 1 = j >> 1 ->
+ is_even i = is_even j ->
+ i = j.
Proof.
- intros A t i x Hi ->. now apply to_list_In.
+ intros;apply bit_ext.
+ intros n;destruct (reflect_eqb n 0).
+ rewrite <- (negb_involutive (bit i n)), <- (negb_involutive (bit j n)).
+ rewrite e, <- !is_even_bit, H0;trivial.
+ assert (W1 : [|n|] <> 0%Z) by (intros Heq;apply n0;apply to_Z_inj;trivial).
+ assert (W2 := to_Z_bounded n);clear n0.
+ assert (W3 : [|n-1|] = ([|n|] - 1)%Z).
+ rewrite sub_spec, to_Z_1, Zmod_small;trivial;omega.
+ assert (H1 : n = (n-1)+1).
+ apply to_Z_inj;rewrite add_spec, W3.
+ rewrite Zmod_small;rewrite to_Z_1; omega.
+ case_eq ((n-1) < digits); intro l.
+ rewrite ltb_spec in l.
+ rewrite H1, <- !bit_half, H; trivial; rewrite ltb_spec; trivial.
+ assert ((digits <= n) = true).
+ rewrite <- Bool.not_true_iff_false, ltb_spec in l; rewrite leb_spec;omega.
+ rewrite !bit_M;trivial.
Qed.
-Lemma In_to_list : forall {A} (t: array A) x,
- In x (to_list t) -> exists i, i < length t = true /\ x = t.[i].
+Lemma lsr1_bit : forall i k, bit i k >> 1 = 0.
Proof.
- intros A t x; unfold to_list; case_eq (0 == length t); intro Heq.
- intro H; inversion H.
- rewrite eqb_false_spec in Heq.
- pose (P (_:int) l := In x l ->
- exists i : int, (i < length t) = true /\ x = t .[ i]).
- pose (H1 := foldi_down_Ind2 _ P (fun (i : int) (l : list A) => t .[ i] :: l) (length t - 1) 0); unfold P in H1; apply H1.
- rewrite ltb_spec, to_Z_sub_1_diff; auto; change [|max_int|] with (wB-1)%Z; pose (H2 := to_Z_bounded (length t)); lia.
- intros _ H; inversion H.
- intro H; inversion H.
- simpl; intros i a _ H2 IH [H3|H3].
- exists i; split; auto; rewrite ltb_spec; rewrite leb_spec, to_Z_sub_1_diff in H2; auto; lia.
- destruct (IH H3) as [j [H4 H5]]; exists j; auto.
+ intros;destruct (bit i k);trivial.
Qed.
+Lemma is_even_or i j : is_even (i lor j) = is_even i && is_even j.
+Proof.
+ rewrite !is_even_bit, lor_spec; case bit; auto.
+Qed.
-(** Lemmas about PArray.mapi *)
+Lemma is_even_xor i j : is_even (i lxor j) = negb (xorb (is_even i) (is_even j)).
+Proof.
+ rewrite !is_even_bit, lxor_spec; do 2 case bit; auto.
+Qed.
-Lemma length_mapi : forall {A B} (f:int -> A -> B) t,
- length (mapi f t) = length t.
+Lemma bit_xor_split: forall i : int, i = (i >> 1) << 1 lxor bit i 0.
Proof.
- unfold mapi; intros A B f t; case_eq (length t == 0).
- rewrite Int63Properties.eqb_spec; intro Heq; rewrite Heq, length_make; auto.
- rewrite eqb_false_spec; intro Heq; apply foldi_ind.
- rewrite length_make, ltb_length; auto.
- intros i a _ H1 H2; rewrite length_set; auto.
+ intros.
+ rewrite bit_or_split at 1.
+ apply lsr_is_even_eq.
+ rewrite lxor_lsr, lor_lsr, lsr1_bit, lxor0_r, lor0_r;trivial.
+ rewrite is_even_or, is_even_xor.
+ rewrite is_even_lsl_1;trivial.
+ rewrite (xorb_true_l (is_even (bit i 0))), negb_involutive;trivial.
Qed.
+Lemma lxor_nilpotent: forall i, i lxor i = 0.
+Proof.
+ intros;apply bit_ext;intros.
+ rewrite lxor_spec, xorb_nilpotent, bit_0;trivial.
+Qed.
-Lemma default_mapi : forall {A B} (f:int -> A -> B) t,
- default (mapi f t) = f (length t) (default t).
+Lemma int_ind : forall (P : int -> Prop),
+ P 0 ->
+ (forall i, (i < max_int) = true -> P i -> P (i + 1)) ->
+ forall i, P i.
Proof.
- unfold mapi; intros A B f t; case (length t == 0).
- rewrite default_make; auto.
- apply foldi_ind.
- rewrite default_make; auto.
- intros; rewrite default_set; auto.
+ intros P HP0 Hrec i.
+ assert (Bi := to_Z_bounded i).
+ destruct Bi as [ Bi0 Bi ].
+ rewrite <- of_to_Z.
+ rewrite Z2Nat.inj_lt in Bi; [ | exact Bi0 | lia ]; clear Bi0.
+ rewrite <- (Z2Nat.id (to_Z i)); [ | apply to_Z_bounded ].
+ revert Bi.
+ induction (Z.to_nat (to_Z i)); clear i.
+ intro; apply HP0.
+ rewrite Nat2Z.inj_lt.
+ rewrite Z2Nat.id; [ | generalize wB_pos; clear IHn; lia ].
+ rewrite Nat2Z.inj_succ.
+ rewrite <- Z.add_1_r.
+ rewrite <- (Nat2Z.id n) in IHn at 1.
+ rewrite <- Z2Nat.inj_lt in IHn; [ | clear IHn; lia | clear IHn; generalize wB_pos; lia ].
+ generalize (Z.of_nat n) IHn (Nat2Z.is_nonneg n); clear n IHn; intros z IHz Hz1 Hz2.
+ replace (of_Z (z + 1)) with (of_Z z + 1).
+ apply Hrec.
+ apply ltb_spec.
+ rewrite of_Z_spec, Z.mod_small, max_int_wB; lia.
+ apply IHz; lia.
+ apply to_Z_inj.
+ rewrite of_Z_spec, Z.mod_small by lia.
+ rewrite to_Z_add_1_wB, of_Z_spec.
+ rewrite Z.mod_small; lia.
+ rewrite of_Z_spec, Z.mod_small; lia.
Qed.
-Lemma get_mapi : forall {A B} (f:int -> A -> B) t i,
- i < length t = true -> (mapi f t).[i] = f i (t.[i]).
+Lemma int_ind_bounded : forall (P : int -> Prop) min max,
+ min <= max = true ->
+ P min ->
+ (forall i, min <= i = true -> i < max = true -> P i -> P (i + 1)) ->
+ P max.
Proof.
- intros A B f t i Hi; generalize (length_mapi f t); unfold mapi; case_eq (length t == 0).
- rewrite Int63Properties.eqb_spec; intro Heq; rewrite Heq in Hi; eelim ltb_0; eassumption.
- rewrite eqb_false_spec; intro Heq; pose (Hi':=Hi); replace (length t) with ((length t - 1) + 1) in Hi'.
- generalize Hi'; apply (foldi_Ind _ (fun j a => (i < j) = true -> length a = length t -> a.[i] = f i (t.[i]))).
- rewrite ltb_spec, (to_Z_sub_1 _ i); auto; destruct (to_Z_bounded (length t)) as [_ H]; change [|max_int|] with (wB-1)%Z; lia.
- intros H _; eelim ltb_0; eassumption.
- intros H; eelim ltb_0; eassumption.
- intros j a _ H1 IH H2 H3; rewrite length_set in H3; case_eq (j == i).
- rewrite Int63Properties.eqb_spec; intro Heq2; subst i; rewrite get_set_same; auto; rewrite H3; auto.
- rewrite eqb_false_spec; intro Heq2; rewrite get_set_other; auto; apply IH; auto; rewrite ltb_spec; rewrite ltb_spec, (to_Z_add_1 _ (length t)) in H2.
- assert (H4: [|i|] <> [|j|]) by (intro H4; apply Heq2, to_Z_inj; auto); lia.
- rewrite ltb_spec; rewrite leb_spec, (to_Z_sub_1 _ _ Hi) in H1; lia.
- apply to_Z_inj; rewrite (to_Z_add_1 _ max_int).
- rewrite to_Z_sub_1_diff; auto; lia.
- rewrite ltb_spec, to_Z_sub_1_diff; auto; destruct (to_Z_bounded (length t)) as [_ H]; change [|max_int|] with (wB-1)%Z; lia.
+ intros P min max Hle Hmin Hrec.
+ rewrite leb_spec in Hle.
+ assert (Bmin := to_Z_bounded min);assert (Bmax := to_Z_bounded max).
+ replace max with (min + (max - min)) by ring.
+ generalize (leb_refl (max - min)).
+ pattern (max - min) at 1 3.
+ apply int_ind.
+ intros _; replace (min + 0) with min by ring; exact Hmin.
+ intros i Hi1 IH; revert Hi1.
+ rewrite ltb_spec, leb_spec.
+ assert (Bi := to_Z_bounded i).
+ rewrite max_int_wB; intro Hi1.
+ replace (min + (i + 1)) with (min + i + 1) by ring.
+ rewrite to_Z_add_1_wB, sub_spec, Z.mod_small by lia.
+ intro Hi2; apply Hrec.
+ rewrite leb_spec, add_spec, Z.mod_small; lia.
+ rewrite ltb_spec, add_spec, Z.mod_small; lia.
+ apply IH.
+ rewrite leb_spec, sub_spec, Z.mod_small; lia.
Qed.
+Definition foldi {A : Type} (f : int -> A -> A) (from to : int) (a : A) : A :=
+ let fix foldi_rec (n : nat) (a : A) : A :=
+ match n with
+ | O => a
+ | S m => foldi_rec m (f (to - of_Z (Z.of_nat n)) a)
+ end in
+ foldi_rec (Z.to_nat (to_Z to) - Z.to_nat (to_Z from))%nat a.
-Lemma get_mapi_outofbound : forall {A B} (f:int -> A -> B) t i,
- i < length t = false -> (mapi f t).[i] = f (length t) (default t).
+Lemma foldi_ge : forall A f from to (a:A),
+ to <= from = true -> foldi f from to a = a.
Proof.
- intros A B f t i H1; rewrite get_outofbound.
- apply default_mapi.
- rewrite length_mapi; auto.
+ intros A f from to a; unfold foldi.
+ assert (Bfrom := to_Z_bounded from); assert (Bto := to_Z_bounded to).
+ rewrite leb_spec, Z2Nat.inj_le by lia; intro Hle.
+ replace (_ - _)%nat with O by lia; tauto.
+Qed.
+
+Lemma foldi_lt_l : forall A f from to (a:A),
+ from < to = true -> foldi f from to a = foldi f (from + 1) to (f from a).
+Proof.
+ intros A f from to a; rewrite ltb_spec; intro Hlt; unfold foldi.
+ assert (Bfrom := to_Z_bounded from); assert (Bto := to_Z_bounded to).
+ rewrite <- Nat.sub_succ.
+ rewrite Nat.sub_succ_l by (rewrite Nat.le_succ_l, <- Z2Nat.inj_lt; lia); f_equal.
+ rewrite to_Z_add_1_wB, Z.add_1_r, Z2Nat.inj_succ by lia; reflexivity.
+ f_equal.
+ rewrite <- Nat.sub_succ_l by (rewrite Nat.le_succ_l, <- Z2Nat.inj_lt; lia).
+ rewrite Nat.sub_succ, <- Z2Nat.inj_sub, Z2Nat.id by lia.
+ apply to_Z_inj; rewrite sub_spec, of_Z_spec, <- 2!sub_spec; f_equal; ring.
+Qed.
+
+Lemma foldi_lt_r : forall A f from to (a:A),
+ from < to = true -> foldi f from to a = f (to - 1) (foldi f from (to - 1) a).
+Proof.
+ intros A f from to a; rewrite ltb_spec; intro Hlt.
+ assert (Bfrom := to_Z_bounded from); assert (Bto := to_Z_bounded to).
+ replace from with (max_int - (max_int - from)) by ring.
+ revert a; pattern (max_int - from).
+ apply (int_ind_bounded _ (max_int - (to - 1))).
+ rewrite leb_spec, sub_spec, to_Z_sub_1_0, sub_spec, max_int_wB, 2!Z.mod_small by lia; lia.
+ intro a; replace (max_int - (max_int - (to - 1))) with (to - 1) by ring.
+ rewrite foldi_lt_l by (rewrite ltb_spec, to_Z_sub_1_0; lia).
+ ring_simplify (to - 1 + 1).
+ rewrite 2!foldi_ge by (rewrite leb_spec; lia); reflexivity.
+ intro i; rewrite leb_spec, ltb_spec, sub_spec, to_Z_sub_1_0, sub_spec, max_int_wB, 2!Z.mod_small by lia.
+ intros Hi1 Hi2 IH a.
+ rewrite foldi_lt_l by (rewrite ltb_spec, sub_spec, to_Z_add_1_wB, max_int_wB, Z.mod_small by lia; lia).
+ rewrite (foldi_lt_l _ _ (max_int - (i + 1))) by (rewrite ltb_spec, sub_spec, to_Z_add_1_wB, to_Z_sub_1_0, max_int_wB, Z.mod_small by lia; lia).
+ replace (max_int - (i + 1) + 1) with (max_int - i) by ring.
+ apply IH.
+Qed.
+
+Lemma foldi_ind : forall A (P : int -> A -> Prop) f from to a,
+ from <= to = true -> P from a ->
+ (forall i a, from <= i = true -> i < to = true -> P i a -> P (i + 1) (f i a)) ->
+ P to (foldi f from to a).
+Proof.
+ intros A P f from to a Hle Hfrom IH.
+ assert (Bfrom := to_Z_bounded from); assert (Bto := to_Z_bounded to).
+ pattern to; apply (int_ind_bounded _ from).
+ exact Hle.
+ rewrite foldi_ge by (rewrite leb_spec; lia).
+ exact Hfrom.
+ intro i; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, ltb_spec; intros Hi1 Hi2; rewrite (foldi_lt_r _ _ _ (i + 1)) by (rewrite ltb_spec, to_Z_add_1_wB; lia).
+ ring_simplify (i + 1 - 1); apply IH; [ rewrite leb_spec; exact Hi1 | rewrite ltb_spec; exact Hi2 ].
+Qed.
+
+Lemma foldi_ind2 : forall A B (P : int -> A -> B -> Prop) f1 f2 from to a1 a2,
+ from <= to = true -> P from a1 a2 ->
+ (forall i a1 a2, from <= i = true -> i < to = true -> P i a1 a2 -> P (i + 1) (f1 i a1) (f2 i a2)) ->
+ P to (foldi f1 from to a1) (foldi f2 from to a2).
+Proof.
+ intros A B P f1 f2 from to a1 a2 Hle Hfrom IH.
+ assert (Bfrom := to_Z_bounded from); assert (Bto := to_Z_bounded to).
+ pattern to; apply (int_ind_bounded _ from).
+ exact Hle.
+ rewrite 2!foldi_ge by (rewrite leb_spec; lia).
+ exact Hfrom.
+ intro i; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, ltb_spec; intros Hi1 Hi2; rewrite 2!(foldi_lt_r _ _ _ (i + 1)) by (rewrite ltb_spec, to_Z_add_1_wB; lia).
+ ring_simplify (i + 1 - 1); apply IH; [ rewrite leb_spec; exact Hi1 | rewrite ltb_spec; exact Hi2 ].
Qed.
+Lemma foldi_eq_compat : forall A (f1 f2:int -> A -> A) min max a,
+ (forall i a, min <= i = true -> i < max = true -> f1 i a = f2 i a) ->
+ foldi f1 min max a = foldi f2 min max a.
+Proof.
+ intros A f1 f2 min max a Hf.
+ assert (Bmin := to_Z_bounded min); assert (Bmax := to_Z_bounded max).
+ case (Z.lt_ge_cases [|min|] [|max|]); [ intro Hlt | intro Hle ].
+ apply (foldi_ind2 _ _ (fun _ a b => a = b)); [ rewrite leb_spec; lia | reflexivity | ].
+ intros i a1 a2 Hi1 Hi2 Heq; rewrite Heq; apply Hf; assumption.
+ rewrite 2!foldi_ge by (rewrite leb_spec; lia); reflexivity.
+Qed.
-(** Custom fold_left and fold_right *)
+(** Lemmas about to_list *)
-Definition afold_left A B default (OP : A -> A -> A) (F : B -> A) (V : array B) :=
- let n := PArray.length V in
- if n == 0 then default
- else foldi (fun i a => OP a (F (V.[i]))) 1 (n-1) (F (V.[0])).
+Definition to_list {A : Type} (t : array A) :=
+ List.rev (foldi (fun i l => t.[i] :: l)%list 0 (length t) nil).
+Lemma foldi_to_list : forall A B (f : A -> B -> A) a e,
+ foldi (fun i x => f x (a.[i])) 0 (length a) e = fold_left f (to_list a) e.
+Proof.
+ intros A B f a e; unfold to_list.
+ rewrite <- fold_left_rev_right, rev_involutive.
+ apply (foldi_ind2 _ _ (fun _ a b => a = fold_right (fun y x => f x y) e b)).
+ apply leb_0.
+ reflexivity.
+ intros i x l _ Hi IH.
+ simpl; f_equal; exact IH.
+Qed.
-Definition afold_right A B default (OP : A -> A -> A) (F : B -> A) (V : array B) :=
- let n := PArray.length V in
- if n == 0 then default else
- if n <= 1 then F (V.[0])
- else foldi_down (fun i b => OP (F (V.[i])) b) (n-2) 0 (F (V.[n-1])).
+Lemma to_list_In : forall {A} (t: array A) i,
+ i < length t = true -> In (t.[i]) (to_list t).
+ intros A t i; assert (Bt := to_Z_bounded (length t)); assert (Bi := to_Z_bounded i); rewrite ltb_spec; unfold to_list.
+ rewrite <- in_rev.
+ apply foldi_ind.
+ rewrite leb_spec, to_Z_0; lia.
+ rewrite to_Z_0; lia.
+ intros j l _; assert (Bj := to_Z_bounded j).
+ rewrite ltb_spec; intros Hj IH.
+ rewrite to_Z_add_1_wB by lia; intro Hij.
+ case (reflect_eqb j i); [ intro Heq; rewrite Heq; clear Heq | rewrite <- to_Z_eq; intro Hneq ].
+ apply in_eq.
+ apply in_cons.
+ apply IH.
+ lia.
+Qed.
+Lemma to_list_In_eq : forall {A} (t: array A) i x,
+ i < length t = true -> x = t.[i] -> In x (to_list t).
+Proof.
+ intros A t i x Hi ->. now apply to_list_In.
+Qed.
-(** Some properties about afold_left *)
+Lemma In_to_list : forall {A} (t: array A) x,
+ In x (to_list t) -> exists i, i < length t = true /\ x = t.[i].
+Proof.
+ intros A t x; assert (Bt := to_Z_bounded (length t)); unfold to_list.
+ rewrite <- in_rev.
+ set (a := foldi _ _ _ _); pattern (length t) at 0, a; subst a; apply foldi_ind.
+ rewrite leb_spec, to_Z_0; lia.
+ intro H; elim (in_nil H).
+ intros i a _; assert (Bi := to_Z_bounded i); rewrite ltb_spec; intros Hi IH.
+ intro Hin; case (in_inv Hin); clear Hin; [ | exact IH ].
+ intro H; rewrite <- H; clear H.
+ exists i.
+ split; [ rewrite ltb_spec; lia | reflexivity ].
+Qed.
-Lemma afold_left_eq :
- forall A B OP def (F1 F2 : A -> B) V1 V2,
- length V1 = length V2 ->
- (forall i, i < length V1 = true -> F1 (V1.[i]) = F2 (V2.[i])) ->
- afold_left _ _ def OP F1 V1 = afold_left _ _ def OP F2 V2.
+(** Lemmas about amapi/amap *)
+
+Definition amapi {A B:Type} (f:int->A->B) (t:array A) :=
+ let l := length t in
+ foldi (fun i tb => tb.[i <- f i (t.[i])]) 0 l (make l (f l (default t))).
+
+Definition amap {A B:Type} (f:A->B) := amapi (fun _ => f).
+
+Lemma length_amapi : forall {A B} (f:int -> A -> B) t,
+ length (amapi f t) = length t.
Proof.
- unfold afold_left;intros. rewrite <- H.
- destruct (Int63Properties.reflect_eqb (length V1) 0);trivial.
- rewrite (H0 0); [ | unfold is_true;rewrite <- not_0_ltb;trivial].
- apply foldi_eq_compat;intros;rewrite H0;trivial.
- unfold is_true;rewrite ltb_leb_sub1;trivial.
+ unfold amapi; intros A B f t.
+ assert (Bt := to_Z_bounded (length t)).
+ apply (foldi_ind _ (fun _ a => length a = length t)).
+ apply leb_0.
+ rewrite length_make, leb_length by reflexivity; reflexivity.
+ intros i a _; assert (Bi := to_Z_bounded i).
+ rewrite ltb_spec; intros Hi IH.
+ rewrite length_set; exact IH.
Qed.
+Lemma length_amap : forall {A B} (f:A -> B) t,
+ length (amap f t) = length t.
+Proof.
+ intros; unfold amap; apply length_amapi.
+Qed.
-Definition afoldi_left {A B:Type} default (OP : int -> A -> A -> A) (F : B -> A) (V : array B) :=
- let n := PArray.length V in
- if n == 0 then default
- else foldi (fun i a => OP i a (F (V.[i]))) 1 (n-1) (F (V.[0])).
+Lemma default_amapi : forall {A B} (f:int -> A -> B) t,
+ default (amapi f t) = f (length t) (default t).
+Proof.
+ unfold amapi; intros A B f t.
+ assert (Bt := to_Z_bounded (length t)).
+ apply (foldi_ind _ (fun i a => default a = f (length t) (default t))).
+ apply leb_0.
+ rewrite default_make by reflexivity; reflexivity.
+ intros i a _; assert (Bi := to_Z_bounded i).
+ rewrite ltb_spec; intros Hi IH.
+ rewrite default_set; exact IH.
+Qed.
+Lemma default_amap : forall {A B} (f:A -> B) t,
+ default (amap f t) = f (default t).
+Proof.
+ intros; unfold amap; apply default_amapi.
+Qed.
-Lemma afoldi_left_Ind :
- forall {A B: Type} (P : int -> A -> Prop) default (OP : int -> A -> A -> A) (F : B -> A) (t:array B),
- if length t == 0 then
- True
- else
- (forall a i, i < length t = true -> P i a -> P (i+1) (OP i a (F (t.[i])))) ->
- P 1 (F (t.[0])) ->
- P (length t) (afoldi_left default OP F t).
+Lemma get_amapi : forall {A B} (f:int -> A -> B) t i,
+ i < length t = true -> (amapi f t).[i] = f i (t.[i]).
Proof.
- intros A B P default OP F t; case_eq (length t == 0).
- intros; exact I.
- intros Heq H1 H2; unfold afoldi_left; rewrite Heq;
- assert (H: (length t - 1) + 1 = length t) by ring.
- rewrite <- H at 1; apply foldi_Ind; auto.
- assert (W:= leb_max_int (length t)); rewrite leb_spec in W.
- rewrite ltb_spec, to_Z_sub_1_diff; auto with zarith.
- intro H3; rewrite H3 in Heq; discriminate.
- intro Hlt; assert (H3: length t - 1 = 0).
- rewrite ltb_spec, to_Z_1 in Hlt; apply to_Z_inj; rewrite to_Z_0; pose (H3 := to_Z_bounded (length t - 1)); lia.
- rewrite H3; assumption.
- intros i a H3 H4; apply H1; trivial.
- rewrite ltb_leb_sub1; auto.
- intro H5; rewrite H5 in Heq; discriminate.
+ intros A B f t.
+ assert (Bt := to_Z_bounded (length t)).
+ intro i; assert (Bi := to_Z_bounded i).
+ rewrite ltb_spec; intro Hi.
+ generalize (length_amapi f t); unfold amapi; revert Hi.
+ set (a := foldi _ _ _ _); pattern (length t) at 1, a; subst a; apply foldi_ind.
+ rewrite leb_spec, to_Z_0; lia.
+ rewrite to_Z_0; lia.
+ intros j a _; assert (Bj := to_Z_bounded j).
+ rewrite ltb_spec; intros Hj IH.
+ rewrite to_Z_add_1_wB by lia; intro Hij.
+ rewrite length_set; case (reflect_eqb j i); [ intro Heq; rewrite Heq | intro Hneq ]; intro Hlength.
+ rewrite get_set_same by (rewrite Hlength, ltb_spec; lia); reflexivity.
+ rewrite get_set_other by exact Hneq.
+ apply IH; [ rewrite <- to_Z_eq in Hneq; lia | exact Hlength ].
Qed.
+Lemma get_amap : forall {A B} (f:A -> B) t i,
+ i < length t = true -> (amap f t).[i] = f (t.[i]).
+Proof.
+ intros; unfold amap; apply get_amapi; assumption.
+Qed.
-Lemma afold_left_Ind :
- forall A B (P : int -> A -> Prop) default (OP : A -> A -> A) (F : B -> A) (t:array B),
- if length t == 0 then
- True
- else
- (forall a i, i < length t = true -> P i a -> P (i+1) (OP a (F (t.[i])))) ->
- P 1 (F (t.[0])) ->
- P (length t) (afold_left A B default OP F t).
+Lemma get_amapi_outofbound : forall {A B} (f:int -> A -> B) t i,
+ i < length t = false -> (amapi f t).[i] = f (length t) (default t).
Proof.
- intros A B P default OP F t;
- apply (afoldi_left_Ind P default (fun _ => OP)); trivial.
+ intros A B f t i H1; rewrite get_outofbound.
+ apply default_amapi.
+ rewrite length_amapi; auto.
Qed.
+Lemma get_amap_outofbound : forall {A B} (f:A -> B) t i,
+ i < length t = false -> (amap f t).[i] = f (default t).
+Proof.
+ intros; unfold amap; apply get_amapi_outofbound; assumption.
+Qed.
-Lemma afold_left_ind :
- forall A B (P : A -> Prop) default (OP : A -> A -> A) (F : B -> A) (t:array B),
- (forall a i, i < length t = true -> P a -> P (OP a (F (t.[i])))) ->
- P default -> P (F (t.[0])) ->
- P (afold_left _ _ default OP F t).
+Lemma to_list_amap : forall A B (f : A -> B) t, to_list (amap f t) = List.map f (to_list t).
Proof.
- intros A B P default OP F t H1 H2 H4.
- pose (H3 := afold_left_Ind A B (fun _ => P) default OP F t).
- case_eq (length t == 0); intro Heq.
- unfold afold_left; rewrite Heq; assumption.
- rewrite Heq in H3; apply H3; trivial.
+ intros A B f t.
+ assert (Bt := to_Z_bounded (length t)).
+ unfold to_list; rewrite length_amap.
+ rewrite map_rev; f_equal.
+ apply (foldi_ind2 _ _ (fun i a b => a = map f b)).
+ apply leb_0.
+ reflexivity.
+ intros i a1 a2 _; assert (Bi := to_Z_bounded i).
+ rewrite ltb_spec; intros Hi IH.
+ simpl; f_equal.
+ apply get_amap.
+ rewrite ltb_spec; lia.
+ apply IH.
Qed.
+(** Some properties about afold_left *)
+
+Definition afold_left A default (OP : A -> A -> A) (V : array A) :=
+ if length V == 0 then default
+ else foldi (fun i a => OP a (V.[i])) 1 (length V) (V.[0]).
-Lemma afold_left_spec : forall A B (f:B -> A) args op e,
+Lemma afold_left_spec : forall A args op (e : A),
(forall a, op e a = a) ->
- afold_left _ _ e op f args =
- fold_left (fun a v => op a (f v)) e args.
+ afold_left _ e op args =
+ foldi (fun i a => op a (args.[i])) 0 (length args) e.
Proof.
- unfold afold_left, fold_left;intros A B f args op neu H10.
+ unfold afold_left;intros A args op neu H10.
destruct (reflect_eqb (length args) 0) as [e|n].
- rewrite e, eqb_refl;trivial.
- apply not_eq_sym in n;rewrite (eqb_false_complete _ _ n).
- case_eq (0 < (length args - 1));intros H.
- rewrite foldi_lt with (from := 0);trivial.
- rewrite H10; auto.
- assert (H0: (0 <> [|length args|])%Z).
- intros Heq;apply n;apply to_Z_inj;trivial.
- assert (H1: length args = 1).
- generalize (to_Z_bounded (length args)).
- rewrite <- not_true_iff_false, ltb_spec, to_Z_0, to_Z_sub_1_diff in H;auto.
- intros;apply to_Z_inj;rewrite to_Z_1;lia.
- rewrite H1; change (1 - 1) with 0; rewrite (foldi_eq _ _ 0 0); auto.
+ rewrite e, foldi_ge by reflexivity;trivial.
+ rewrite (foldi_lt_l _ _ 0) by (apply not_0_ltb; assumption).
+ f_equal; rewrite H10; reflexivity.
Qed.
+Lemma afold_left_eq :
+ forall A OP (def : A) V1 V2,
+ length V1 = length V2 ->
+ (forall i, i < length V1 = true -> V1.[i] = V2.[i]) ->
+ afold_left _ def OP V1 = afold_left _ def OP V2.
+Proof.
+ intros A OP def V1 V2 Heqlength HeqV.
+ assert (BV1 := to_Z_bounded (length V1)).
+ unfold afold_left.
+ rewrite <- Heqlength.
+ case (reflect_eqb (length V1) 0).
+ reflexivity.
+ rewrite <- to_Z_eq, to_Z_0; intro Hneq.
+ rewrite <- HeqV by (rewrite ltb_spec, to_Z_0; lia).
+ apply (foldi_ind2 _ _ (fun i a b => a = b)).
+ rewrite leb_spec, to_Z_1; lia.
+ reflexivity.
+ intros i a1 a2; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, to_Z_1; intro Hi1.
+ rewrite ltb_spec by lia; intros Hi2 IH.
+ f_equal;[ exact IH | apply HeqV; rewrite ltb_spec; lia ].
+Qed.
+
+Lemma afold_left_ind : forall A OP def V (P : int -> A -> Prop),
+ (length V = 0 -> P 0 def) ->
+ (0 < length V = true -> P 1 (V.[0])) ->
+ (forall a i, 0 < i = true -> i < length V = true -> P i a -> P (i + 1) (OP a (V.[i]))) ->
+ P (length V) (afold_left A def OP V).
+Proof.
+ intros A OP def V P HP0 HP1 HPOP.
+ assert (BV := to_Z_bounded (length V)).
+ unfold afold_left.
+ case (reflect_eqb (length V) 0); [ intro Heq; rewrite Heq; tauto | intro Hneq ].
+ rewrite <- to_Z_eq, to_Z_0 in Hneq.
+ apply foldi_ind.
+ rewrite leb_spec, to_Z_1; lia.
+ apply HP1; rewrite ltb_spec, to_Z_0; lia.
+ intros i a; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, to_Z_1, ltb_spec; intros Hi1 Hi2 IH.
+ apply HPOP; [ rewrite ltb_spec, to_Z_0; lia | rewrite ltb_spec; lia | exact IH ].
+Qed.
(** Some properties about afold_right *)
-Lemma afold_right_eq :
- forall A B OP def (F1 F2 : A -> B) V1 V2,
- length V1 = length V2 ->
- (forall i, i < length V1 = true -> F1 (V1.[i]) = F2 (V2.[i])) ->
- afold_right _ _ def OP F1 V1 = afold_right _ _ def OP F2 V2.
-Proof.
- unfold afold_right;intros.
- rewrite <- H.
- destruct (Int63Properties.reflect_eqb (length V1) 0);trivial.
- destruct (reflect_leb (length V1) 1);intros.
- apply H0;unfold is_true;rewrite ltb_leb_sub1;trivial. apply leb_0.
- assert (length V1 - 1 < length V1 = true).
- rewrite ltb_leb_sub1;auto using leb_refl.
- rewrite (H0 (length V1 - 1));trivial.
- apply foldi_down_eq_compat;intros;rewrite H0;trivial.
- unfold is_true;rewrite ltb_leb_sub1;[ | trivial].
- apply ltb_leb_sub1;trivial.
- revert n0 H3;rewrite ltb_spec, leb_spec, to_Z_1, sub_spec.
- change [|2|] with 2%Z.
- intros HH;assert (W:= to_Z_bounded (length V1));rewrite Zmod_small;lia.
-Qed.
-
-
-Definition afoldi_right {A B:Type} default (OP : int -> A -> A -> A) (F : B -> A) (V : array B) :=
- let n := PArray.length V in
- if n == 0 then default
- else if n <= 1 then F (V .[ 0])
- else foldi_down (fun i a => OP i (F (V.[i])) a) (n-2) 0 (F (V.[n-1])).
-
-
-Lemma afoldi_right_Ind :
- forall {A B: Type} (P : int -> A -> Prop) default (OP : int -> A -> A -> A) (F : B -> A) (t:array B),
- if length t <= 1 then
- True
- else
- (forall a i, i < length t - 1 = true -> P (i+1) a -> P i (OP i (F (t.[i])) a)) ->
- P ((length t)-1) (F (t.[(length t)-1])) ->
- P 0 (afoldi_right default OP F t).
-Proof.
- intros A B P default OP F t; case_eq (length t <= 1).
- intros; exact I.
- intros Heq H1 H2; unfold afoldi_right. replace (length t == 0) with false.
- rewrite Heq.
- set (P' z a := P (of_Z (z + 1)) a).
- change (P' ([|0|] - 1)%Z (foldi_down (fun (i : int) (a : A) => OP i (F (t .[ i])) a) (length t - 2) 0 (F (t .[ length t - 1])))).
- apply foldi_down_ZInd;unfold P'.
- intros Hlt;elim (ltb_0 _ Hlt).
- replace (length t - 2) with (length t - 1 - 1) by ring.
- rewrite to_Z_sub_1_diff.
- ring_simplify ([|length t - 1|] - 1 + 1)%Z;rewrite of_to_Z;trivial.
- assert (H10: (1 < length t) = true) by (rewrite ltb_negb_geb, Heq; auto).
- intro H11. rewrite ltb_spec in H10. assert (H12: [|length t - 1|] = 0%Z) by (rewrite H11; auto). change [|1|] with (1%Z) in H10. rewrite to_Z_sub_1_diff in H12; [lia| ]. intro H13. assert (H14: [|length t|] = 0%Z) by (rewrite H13; auto). lia.
- intros;ring_simplify ([|i|] - 1 + 1)%Z;rewrite of_to_Z;auto.
- assert (i < length t - 1 = true).
- rewrite ltb_spec. rewrite leb_spec in H0. replace (length t - 2) with (length t - 1 - 1) in H0 by ring. rewrite to_Z_sub_1_diff in H0; [lia| ]. intro H4. assert (H5: [|length t - 1|] = 0%Z) by (rewrite H4; auto). assert (H6: 1 < length t = true) by (rewrite ltb_negb_geb, Heq; auto). rewrite ltb_spec in H6. change ([|1|]) with (1%Z) in H6. rewrite to_Z_sub_1_diff in H5; [lia| ]. intro H7. assert (H8: [|length t|] = 0%Z) by (rewrite H7; auto). lia.
- apply H1; [trivial| ].
- rewrite <-(to_Z_add_1 _ _ H4), of_to_Z in H3;auto.
- symmetry. rewrite eqb_false_spec. intro H. rewrite H in Heq. discriminate.
-Qed.
-
-
-Lemma afold_right_Ind :
- forall A B (P : int -> A -> Prop) default (OP : A -> A -> A) (F : B -> A) (t:array B),
- if length t <= 1 then
- True
- else
- (forall a i, i < length t - 1 = true -> P (i+1) a -> P i (OP (F (t.[i])) a)) ->
- P ((length t)-1) (F (t.[(length t)-1])) ->
- P 0 (afold_right A B default OP F t).
-Proof.
- intros A B P default OP F t;
- apply (afoldi_right_Ind P default (fun _ => OP) F).
-Qed.
-
-
-Lemma afold_right_ind :
- forall A B (P : A -> Prop) default (OP : A -> A -> A) (F : B -> A) (t:array B),
- (forall a i, i < length t - 1 = true -> P a -> P (OP (F (t.[i])) a)) ->
- P default -> P (F (t.[length t - 1])) ->
- P (afold_right _ _ default OP F t).
-Proof.
- intros A B P default OP F t H1 H2 H4.
- pose (H3 := afold_right_Ind A B (fun _ => P) default OP F t).
- unfold afold_right. case_eq (length t == 0); auto. intro H10. assert (H := H10). rewrite eqb_false_spec in H. case_eq (length t <= 1); intro Heq.
- replace 0 with (length t - 1); auto. apply to_Z_inj. rewrite to_Z_sub_1_diff; auto. rewrite leb_spec in Heq. assert (H5 := leb_0 (length t)). rewrite leb_spec in H5. change [|0|] with 0%Z in *. change [|1|] with 1%Z in Heq. assert (H6 : [|length t|] <> 0%Z) by (intro H6; elim H; apply to_Z_inj; auto). lia. rewrite Heq in H3. unfold afold_right in H3. rewrite H10, Heq in H3. apply H3; auto.
-Qed.
-
-
-Lemma afold_right_ind_nonempty :
- forall A B (P : A -> Prop) default (OP : A -> A -> A) (F : B -> A) (t:array B),
- (forall a i, i < length t - 1 = true -> P a -> P (OP (F (t.[i])) a)) ->
- 0 < length t = true -> P (F (t.[length t - 1])) ->
- P (afold_right _ _ default OP F t).
-Proof.
- intros A B P default OP F t H1 H2 H4.
- pose (H3 := afold_right_Ind A B (fun _ => P) default OP F t).
- unfold afold_right. assert (H10 : length t == 0 = false) by (rewrite eqb_false_spec; intro H; rewrite H in H2; discriminate). rewrite H10. assert (H := H10). rewrite eqb_false_spec in H. case_eq (length t <= 1); intro Heq.
- replace 0 with (length t - 1); auto. apply to_Z_inj. rewrite to_Z_sub_1_diff; auto. rewrite leb_spec in Heq. assert (H5 := leb_0 (length t)). rewrite leb_spec in H5. change [|0|] with 0%Z in *. change [|1|] with 1%Z in Heq. assert (H6 : [|length t|] <> 0%Z) by (intro H6; elim H; apply to_Z_inj; auto). lia. rewrite Heq in H3. unfold afold_right in H3. rewrite H10, Heq in H3. apply H3; auto.
-Qed.
-
-
-Lemma afold_right_spec : forall A B (f:B -> A) args op e,
+Definition afold_right A default (OP : A -> A -> A) (V : array A) :=
+ if length V == 0 then default
+ else foldi (fun i => OP (V.[length V - 1 - i])) 1 (length V) (V.[length V - 1]).
+
+Lemma afold_right_spec : forall A args op (e : A),
(forall a, op a e = a) ->
- afold_right _ _ e op f args =
- fold_right (fun v a => op (f v) a) args e.
+ afold_right _ e op args =
+ foldi (fun i a => op (args.[length args - 1 - i]) a) 0 (length args) e.
Proof.
- unfold afold_right, fold_right;intros A B f args op neu H10.
+ unfold afold_right;intros A args op neu H10.
+ assert (Bargs := to_Z_bounded (length args)).
destruct (reflect_eqb (length args) 0) as [e|n].
- rewrite e, eqb_refl;trivial.
- apply not_eq_sym in n;rewrite (eqb_false_complete _ _ n).
- case_eq (length args <= 1); intro Heq.
- assert (H11: length args = 1).
- apply to_Z_inj. rewrite leb_spec in Heq. assert (H11: 0%Z <> [|length args|]) by (intro H; elim n; apply to_Z_inj; auto). change [|1|] with (1%Z) in *. assert (H12 := leb_0 (length args)). rewrite leb_spec in H12. change [|0|] with 0%Z in H12. lia.
- rewrite H11, foldi_down_eq; auto.
- assert (H11: 1 < length args = true) by (rewrite ltb_negb_geb, Heq; auto). replace (foldi_down (fun (i : int) (b : A) => op (f (args .[ i])) b) (length args - 1) 0 neu) with (foldi_down (fun (i : int) (b : A) => op (f (args .[ i])) b) (length args - 1 - 1) 0 (op (f (args .[ length args - 1])) neu)).
- replace (length args - 1 - 1) with (length args - 2) by ring. rewrite H10. auto.
- symmetry. apply foldi_down_gt. rewrite ltb_spec. change [|0|] with 0%Z. rewrite to_Z_sub_1_diff; auto. rewrite ltb_spec in H11. change [|1|] with 1%Z in H11. lia.
+ rewrite e, foldi_ge by reflexivity;trivial.
+ change 1 with (0 + 1) at 2.
+ replace (length args - 1) with (length args - 1 - 0) at 1 by ring.
+ rewrite <- (H10 (args.[length args - 1 - 0])).
+ rewrite <- (foldi_lt_l _ (fun i => op (args.[length args - 1 - i]))) by (apply not_0_ltb; assumption).
+ apply foldi_eq_compat; intros; reflexivity.
Qed.
+Lemma afold_right_eq :
+ forall A OP (def : A) V1 V2,
+ length V1 = length V2 ->
+ (forall i, i < length V1 = true -> V1.[i] = V2.[i]) ->
+ afold_right _ def OP V1 = afold_right _ def OP V2.
+Proof.
+ intros A OP def V1 V2 Heqlength HeqV.
+ assert (BV1 := to_Z_bounded (length V1)).
+ unfold afold_right.
+ rewrite <- Heqlength.
+ case (reflect_eqb (length V1) 0); [ reflexivity | intro Hneq ].
+ rewrite <- to_Z_eq, to_Z_0 in Hneq.
+ rewrite <- HeqV by (rewrite ltb_spec, to_Z_sub_1_0; lia).
+ apply (foldi_ind2 _ _ (fun i a b => a = b)).
+ rewrite leb_spec, to_Z_1; lia.
+ reflexivity.
+ intros i a1 a2; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, to_Z_1; intro Hi1.
+ rewrite ltb_spec by lia; intros Hi2 IH.
+ f_equal;[ apply HeqV; rewrite ltb_spec, sub_spec, to_Z_sub_1_0, Z.mod_small; lia | exact IH ].
+Qed.
+
+Lemma afold_right_ind : forall A OP def V (P : int -> A -> Prop),
+ (length V = 0 -> P 0 def) ->
+ (0 < length V = true -> P (length V - 1) (V.[length V - 1])) ->
+ (forall a i, 0 < i = true -> i < length V = true -> P i a -> P (i - 1) (OP (V.[i - 1]) a)) ->
+ P 0 (afold_right A def OP V).
+Proof.
+ intros A OP def V P HP0 HP1 HPOP.
+ assert (BV := to_Z_bounded (length V)).
+ unfold afold_right.
+ case (reflect_eqb (length V) 0); [ intro Heq; apply HP0; exact Heq | intro Hneq ].
+ rewrite <- to_Z_eq, to_Z_0 in Hneq.
+ replace 0 with (length V - length V) at 1 by ring.
+ apply (foldi_ind _ (fun i a => P (length V - i) a)).
+ rewrite leb_spec, to_Z_1; lia.
+ apply HP1; rewrite ltb_spec, to_Z_0; lia.
+ intros i a; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, to_Z_1, ltb_spec; intros Hi1 Hi2 IH.
+ replace (length V - (i + 1)) with (length V - i - 1) by ring.
+ replace (length V - 1 - i) with (length V - i - 1) by ring.
+ apply HPOP; [ rewrite ltb_spec, to_Z_0, sub_spec, Z.mod_small; lia | rewrite ltb_spec, sub_spec, Z.mod_small; lia | exact IH ].
+Qed.
(** Application to our uses of afold_left and afold_right *)
(* Case andb *)
-Lemma afold_left_andb_false : forall A i a f,
+Lemma afold_left_andb_false : forall i a,
i < length a = true ->
- f (a .[ i]) = false ->
- afold_left bool A true andb f a = false.
+ a .[ i] = false ->
+ afold_left bool true andb a = false.
Proof.
- intros A i a f; rewrite afold_left_spec; auto; apply (fold_left_Ind _ _ (fun j t => (i < j) = true -> f (a .[ i]) = false -> t = false)).
- intros b j H1 H2 H3 H4; case_eq (i == j).
- rewrite Int63Properties.eqb_spec; intro; subst i; rewrite H4, andb_false_r; auto.
- rewrite eqb_false_spec; intro Heq; rewrite H2; auto; rewrite ltb_spec; rewrite ltb_spec in H3; rewrite (to_Z_add_1 _ (length a)) in H3; auto; assert (H5: [|i|] <> [|j|]) by (intro H5; apply Heq, to_Z_inj; auto); lia.
- intro H; eelim ltb_0; eassumption.
+ intros i a; assert (Ba := to_Z_bounded (length a)); assert (Bi := to_Z_bounded i).
+ rewrite afold_left_spec by apply andb_true_l; apply foldi_ind.
+ apply leb_0.
+ rewrite ltb_spec, to_Z_0; lia.
+ intros j b _; assert (Bj := to_Z_bounded j).
+ rewrite 2!ltb_spec; intros Hj IH.
+ rewrite ltb_spec, to_Z_add_1_wB by lia; intro Hij.
+ case (reflect_eqb i j).
+ intros Heq Hai; rewrite <- Heq, Hai; apply andb_false_r.
+ rewrite <- to_Z_eq; intros Hneq Hai.
+ rewrite IH; [ apply andb_false_l | lia | exact Hai ].
Qed.
-
-Lemma afold_left_andb_false_inv : forall A a f,
- afold_left bool A true andb f a = false ->
- exists i, (i < length a = true) /\ (f (a .[ i]) = false).
+Lemma afold_left_andb_false_inv : forall a,
+ afold_left bool true andb a = false ->
+ exists i, (i < length a = true) /\ (a .[ i] = false).
Proof.
- intros A a f; rewrite afold_left_spec; auto; apply fold_left_ind; try discriminate.
- intros b i H1; case b; simpl.
- intros _ H2; exists i; auto.
- intros H2 _; destruct (H2 (refl_equal false)) as [j [H3 H4]]; exists j; auto.
+ intro a; assert (Ba := to_Z_bounded (length a)).
+ rewrite afold_left_spec by apply andb_true_l; apply foldi_ind.
+ apply leb_0.
+ discriminate.
+ intros i b _; assert (Bi := to_Z_bounded i).
+ rewrite ltb_spec; intros Hj IH.
+ destruct b.
+ rewrite andb_true_l; intro H; exists i; rewrite H.
+ split; [ rewrite ltb_spec, to_Z_add_1_wB; lia | reflexivity ].
+ generalize (IH eq_refl); clear IH; intros [ j [ Hji Haj ] ] _.
+ rewrite ltb_spec in Hji; exists j.
+ split; [ rewrite ltb_spec, to_Z_add_1_wB; lia | exact Haj ].
Qed.
-
-Lemma afold_left_andb_true : forall A a f,
- (forall i, i < length a = true -> f (a.[i]) = true) ->
- afold_left bool A true andb f a = true.
+Lemma afold_left_andb_true : forall a,
+ (forall i, i < length a = true -> a.[i] = true) ->
+ afold_left bool true andb a = true.
Proof.
- intros A a f H; rewrite afold_left_spec; auto; apply fold_left_ind; trivial; intros b j H1 H2; rewrite H2; simpl; rewrite H; trivial.
+ intros a H.
+ rewrite afold_left_spec by apply andb_true_l; apply foldi_ind.
+ apply leb_0.
+ reflexivity.
+ intros b j _ H1 H2; rewrite H2; simpl; rewrite H; trivial.
Qed.
-
-Lemma afold_left_andb_true_inv : forall A a f,
- afold_left bool A true andb f a = true ->
- forall i, i < length a = true -> f (a.[i]) = true.
+Lemma afold_left_andb_true_inv : forall a,
+ afold_left bool true andb a = true ->
+ forall i, i < length a = true -> a.[i] = true.
Proof.
- intros A a f; rewrite afold_left_spec; auto; apply (fold_left_Ind _ _ (fun j t => t = true -> forall i, (i < j) = true -> f (a .[ i]) = true)).
- intros b i H1; case b; simpl; try discriminate; intros H2 H3 j Hj; case_eq (j == i); intro Heq.
- rewrite Int63Properties.eqb_spec in Heq; subst j; auto.
- apply H2; auto; rewrite eqb_false_spec in Heq; rewrite ltb_spec; rewrite ltb_spec in Hj; assert (H4: [|j|] <> [|i|]) by (intro H; apply Heq, to_Z_inj; auto); rewrite (to_Z_add_1 _ (length a)) in Hj; auto; lia.
- intros _ i H; eelim ltb_0; eassumption.
+ intros a H i; assert (Ba := to_Z_bounded (length a)); assert (Bi := to_Z_bounded i).
+ revert H; rewrite afold_left_spec by apply andb_true_l; apply foldi_ind.
+ apply leb_0.
+ rewrite ltb_spec, to_Z_0; lia.
+ intros j b _; assert (Bj := to_Z_bounded j).
+ rewrite 2!ltb_spec; intros Hj IH.
+ rewrite ltb_spec, to_Z_add_1_wB by lia.
+ rewrite andb_true_iff.
+ case (reflect_eqb i j).
+ intro Heq; rewrite <- Heq; tauto.
+ rewrite <- to_Z_eq; intros Hneq [ Hb Haj ] Hij.
+ apply IH; [ exact Hb | lia ].
Qed.
-
-Lemma afold_left_and p a :
- afold_left bool int true andb p a =
+Lemma afold_left_and A (p : A -> bool) a :
+ afold_left bool true andb (amap 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.
+ rewrite afold_left_spec, foldi_to_list, to_list_amap by exact andb_true_l.
+ rewrite <- andb_true_r.
+ generalize true.
+ induction (to_list a) as [ | x l ]; clear a; intro b.
+ reflexivity.
+ simpl; rewrite IHl.
+ rewrite (andb_comm b (p x)), (andb_comm (p x) (forallb p l)); apply andb_assoc.
Qed.
-
(* Case orb *)
-Lemma afold_left_orb_true : forall A i a f,
+Lemma afold_left_orb_true : forall i a,
i < length a = true ->
- f (a .[ i]) = true ->
- afold_left bool A false orb f a = true.
+ a .[ i] = true ->
+ afold_left bool false orb a = true.
Proof.
- intros A i a f; rewrite afold_left_spec; auto; apply (fold_left_Ind _ _ (fun j t => (i < j) = true -> f (a .[ i]) = true -> t = true)).
- intros b j H1 H2 H3 H4; case_eq (i == j).
- rewrite Int63Properties.eqb_spec; intro; subst i; rewrite H4, orb_true_r; auto.
- rewrite eqb_false_spec; intro Heq; rewrite H2; auto; rewrite ltb_spec; rewrite ltb_spec in H3; rewrite (to_Z_add_1 _ (length a)) in H3; auto; assert (H5: [|i|] <> [|j|]) by (intro H5; apply Heq, to_Z_inj; auto); lia.
- intro H; eelim ltb_0; eassumption.
+ intros i a; assert (Ba := to_Z_bounded (length a)); assert (Bi := to_Z_bounded i).
+ rewrite afold_left_spec by apply orb_false_l; apply foldi_ind.
+ apply leb_0.
+ rewrite ltb_spec, to_Z_0; lia.
+ intros j b _; assert (Bj := to_Z_bounded j).
+ rewrite 2!ltb_spec; intros Hj IH.
+ rewrite ltb_spec, to_Z_add_1_wB by lia; intro Hij.
+ case (reflect_eqb i j).
+ intros Heq Hai; rewrite <- Heq, Hai; apply orb_true_r.
+ rewrite <- to_Z_eq; intros Hneq Hai.
+ rewrite IH; [ apply orb_true_l | lia | exact Hai ].
Qed.
-
-Lemma afold_left_orb_true_inv : forall A a f,
- afold_left bool A false orb f a = true ->
- exists i, (i < length a = true) /\ (f (a .[ i]) = true).
+Lemma afold_left_orb_true_inv : forall a,
+ afold_left bool false orb a = true ->
+ exists i, i < length a = true /\ a .[ i] = true.
Proof.
- intros A a f; rewrite afold_left_spec; auto; apply fold_left_ind; try discriminate.
- intros b i H1; case b; simpl.
- intros H2 _; destruct (H2 (refl_equal true)) as [j [H3 H4]]; exists j; auto.
- intros _ H2; exists i; auto.
+ intro a; assert (Ba := to_Z_bounded (length a)).
+ rewrite afold_left_spec by apply andb_true_l; apply foldi_ind.
+ apply leb_0.
+ discriminate.
+ intros i b _; assert (Bi := to_Z_bounded i).
+ rewrite ltb_spec; intros Hj IH.
+ destruct b.
+ generalize (IH eq_refl); clear IH; intros [ j [ Hji Haj ] ] _.
+ rewrite ltb_spec in Hji; exists j.
+ split; [ rewrite ltb_spec, to_Z_add_1_wB; lia | exact Haj ].
+ rewrite orb_false_l; intro H; exists i; rewrite H.
+ split; [ rewrite ltb_spec, to_Z_add_1_wB; lia | reflexivity ].
Qed.
-
-Lemma afold_left_orb_false : forall A a f,
- (forall i, i < length a = true -> f (a.[i]) = false) ->
- afold_left bool A false orb f a = false.
+Lemma afold_left_orb_false : forall a,
+ (forall i, i < length a = true -> a.[i] = false) ->
+ afold_left bool false orb a = false.
Proof.
- intros A a f H; rewrite afold_left_spec; auto; apply fold_left_ind; trivial; intros b j H1 H2; rewrite H2; simpl; rewrite H; trivial.
+ intros a H.
+ rewrite afold_left_spec by apply andb_true_l; apply foldi_ind.
+ apply leb_0.
+ reflexivity.
+ intros b j _ H1 H2; rewrite H2; simpl; rewrite H; trivial.
Qed.
-
-Lemma afold_left_orb_false_inv : forall A a f,
- afold_left bool A false orb f a = false ->
- forall i, i < length a = true -> f (a.[i]) = false.
+Lemma afold_left_orb_false_inv : forall a,
+ afold_left bool false orb a = false ->
+ forall i, i < length a = true -> a.[i] = false.
Proof.
- intros A a f; rewrite afold_left_spec; auto; apply (fold_left_Ind _ _ (fun j t => t = false -> forall i, (i < j) = true -> f (a .[ i]) = false)).
- intros b i H1; case b; simpl; try discriminate; intros H2 H3 j Hj; case_eq (j == i); intro Heq.
- rewrite Int63Properties.eqb_spec in Heq; subst j; auto.
- apply H2; auto; rewrite eqb_false_spec in Heq; rewrite ltb_spec; rewrite ltb_spec in Hj; assert (H4: [|j|] <> [|i|]) by (intro H; apply Heq, to_Z_inj; auto); rewrite (to_Z_add_1 _ (length a)) in Hj; auto; lia.
- intros _ i H; eelim ltb_0; eassumption.
+ intros a H i; assert (Ba := to_Z_bounded (length a)); assert (Bi := to_Z_bounded i).
+ revert H; rewrite afold_left_spec by apply andb_true_l; apply foldi_ind.
+ apply leb_0.
+ rewrite ltb_spec, to_Z_0; lia.
+ intros j b _; assert (Bj := to_Z_bounded j).
+ rewrite 2!ltb_spec; intros Hj IH.
+ rewrite ltb_spec, to_Z_add_1_wB by lia.
+ rewrite orb_false_iff.
+ case (reflect_eqb i j).
+ intro Heq; rewrite <- Heq; tauto.
+ rewrite <- to_Z_eq; intros Hneq [ Hb Haj ] Hij.
+ apply IH; [ exact Hb | lia ].
Qed.
-
-Lemma afold_left_or p a :
- afold_left bool int false orb p a =
+Lemma afold_left_or A (p : A -> bool) a :
+ afold_left bool false orb (amap 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.
+ rewrite afold_left_spec, foldi_to_list, to_list_amap by exact andb_true_l.
+ rewrite <- orb_false_r.
+ generalize false.
+ induction (to_list a) as [ | x l ]; clear a; intro b.
+ reflexivity.
+ simpl; rewrite IHl.
+ rewrite (orb_comm b (p x)), (orb_comm (p x) (existsb p l)); apply orb_assoc.
Qed.
-
(* Case implb *)
-Lemma afold_right_implb_false : forall A a f,
+Lemma afold_right_implb_false : forall a,
0 < length a = true /\
- (forall i, i < length a - 1 = true -> f (a .[ i]) = true) /\
- f (a.[length a - 1]) = false ->
- afold_right bool A true implb f a = false.
+ (forall i, i < length a - 1 = true -> a .[ i] = true) /\
+ a.[length a - 1] = false ->
+ afold_right bool true implb a = false.
Proof.
- intros A a f; intros [H1 [H2 H3]]; apply afold_right_ind_nonempty; auto; intros b i H4 H5; rewrite H5, H2; auto.
+ intros a; intros [H1 [H2 H3]].
+ pattern 0; apply afold_right_ind.
+ intro H; rewrite H in H1; discriminate.
+ intros _; exact H3.
+ intros b i H4 H5 H6; rewrite H6; clear H6.
+ rewrite H2; [ reflexivity | ].
+ assert (Ba := to_Z_bounded (length a)); assert (Bi := to_Z_bounded i).
+ revert H1 H4 H5; rewrite 4!ltb_spec, to_Z_0; intros H1 H4 H5.
+ rewrite 2!to_Z_sub_1_0; lia.
Qed.
-
-Lemma afold_right_implb_false_inv : forall A a f,
- afold_right bool A true implb f a = false ->
+Lemma afold_right_implb_false_inv : forall a,
+ afold_right bool true implb a = false ->
0 < length a = true /\
- (forall i, i < length a - 1 = true -> f (a .[ i]) = true) /\
- f (a.[length a - 1]) = false.
-Proof.
- intros A a f; case_eq (length a == 0); intro Heq1.
- unfold afold_right; rewrite Heq1; discriminate.
- intro H; split.
- rewrite eqb_false_spec in Heq1; rewrite <- not_0_ltb; auto.
- generalize H; clear H; case_eq (length a <= 1); intro Heq2.
- unfold afold_right; rewrite Heq1, Heq2; intro H; replace (length a - 1) with 0.
- split; auto; intros i Hi; elim (ltb_0 i); auto.
- rewrite eqb_false_spec in Heq1; apply to_Z_inj; rewrite to_Z_sub_1_diff; auto; rewrite leb_spec in Heq2; change [|1|] with 1%Z in Heq2; assert ([|length a|] <> 0%Z) by (intro H1; apply Heq1, to_Z_inj; auto); generalize (leb_0 (length a)); rewrite leb_spec; change [|0|] with 0%Z; lia.
- pose (P j k := k = false -> (forall i : int, (j <= i) = true -> (i < length a - 1) = true -> f (a .[ i]) = true) /\ f (a .[ length a - 1]) = false); assert (H: P 0 (afold_right bool A true implb f a)).
- generalize (afold_right_Ind _ _ P true implb f a); rewrite Heq2; intro IH; apply IH; clear IH; unfold P.
- intros b i H1 H2 H3; case_eq b; intro Heq3.
- rewrite Heq3 in H3; generalize H3; case (f (a .[ i])); discriminate.
- destruct (H2 Heq3) as [H4 H5]; split; auto; intros j H6 H7; case_eq (i == j); intro Heq4.
- rewrite eqb_spec in Heq4; subst j b; generalize H3; case (f (a .[ i])); auto; discriminate.
- apply H4; auto; rewrite leb_spec in *; rewrite (to_Z_add_1 _ _ H1); rewrite eqb_false_spec in Heq4; assert ([|i|] <> [|j|]) by (intro H; apply Heq4, to_Z_inj; auto); lia.
- intro H; split; auto; intros i H1 H2; elimtype False; rewrite leb_spec in H1; rewrite ltb_spec in H2; lia.
- unfold P in H; intro H1; destruct (H H1) as [H2 H3]; split; auto; intro i; apply H2, leb_0.
-Qed.
-
-
-Lemma afold_right_implb_true_aux : forall A a f,
- (exists i, i < length a - 1 = true /\ f (a.[i]) = false) ->
- afold_right bool A true implb f a = true.
-Proof.
- intros A a f; case_eq (length a == 0); intro Heq1.
- intros _; unfold afold_right; rewrite Heq1; auto.
- case_eq (length a <= 1); intro Heq2.
- intros [i [Hi _]]; elim (ltb_0 i); replace 0 with (length a - 1); auto; rewrite eqb_false_spec in Heq1; apply to_Z_inj; rewrite to_Z_sub_1_diff; auto; assert (H1: [|length a|] <> 0%Z) by (intro H; apply Heq1, to_Z_inj; auto); rewrite leb_spec in Heq2; generalize (leb_0 (length a)); rewrite leb_spec; change [|0|] with 0%Z; change [|1|] with 1%Z in Heq2; lia.
- pose (P j k := (exists i : int, (j <= i) = true /\ (i < length a - 1) = true /\ f (a .[ i]) = false) -> k = true); assert (H: P 0 (afold_right bool A true implb f a)).
- generalize (afold_right_Ind _ _ P true implb f a); rewrite Heq2; intro IH; apply IH; clear IH; unfold P.
- intros b i H1 H2 [j [H3 [H4 H5]]]; case_eq (i == j); intro Heq3.
- rewrite eqb_spec in Heq3; subst i; rewrite H5; case b; auto.
- rewrite H2.
- case (f (a .[ i])); auto.
- exists j; repeat split; auto; assert (H: i < j = true).
- rewrite ltb_spec; rewrite leb_spec in H3; rewrite eqb_false_spec in Heq3; assert (H: [|i|] <> [|j|]) by (intro H; apply Heq3, to_Z_inj; auto); lia.
- rewrite leb_spec, (to_Z_add_1 _ _ H); rewrite ltb_spec in H; lia.
- intros [i [H1 [H2 _]]]; elimtype False; rewrite leb_spec in H1; rewrite ltb_spec in H2; lia.
- unfold P in H; intros [i Hi]; apply H; exists i; split; auto; apply leb_0.
-Qed.
-
-
-Lemma afold_right_implb_true : forall A a f,
- length a = 0 \/ (exists i, i < length a - 1 = true /\ f (a.[i]) = false) \/
- (forall i, i < length a = true -> f (a.[i]) = true) ->
- afold_right bool A true implb f a = true.
-Proof.
- intros A a f; case_eq (length a == 0).
- intros H _; unfold afold_right; rewrite H; auto.
- intros Heq [H1|[H1|H1]].
- rewrite H1 in Heq; discriminate.
+ (forall i, i < length a - 1 = true -> a .[ i] = true) /\
+ a.[length a - 1] = false.
+Proof.
+ intros a H; assert (Ba := to_Z_bounded (length a)); split; [ | split ].
+ revert H; unfold afold_right.
+ case (reflect_eqb (length a) 0).
+ intro Heq; rewrite Heq; discriminate.
+ rewrite <- to_Z_eq, to_Z_0; intros Hlength _.
+ rewrite ltb_spec, to_Z_0; lia.
+ intro i; generalize (leb_0 i); revert H i; apply afold_right_ind.
+ discriminate.
+ intros _ _ i; rewrite leb_spec, ltb_spec; lia.
+ intros b j; assert (Bj := to_Z_bounded j).
+ rewrite 2!ltb_spec, to_Z_0; intros Hj1 Hj2 IH.
+ unfold implb; case_eq (a.[j - 1]); [ | discriminate ]; intros Ha H; subst b; intro i.
+ case (reflect_eqb i (j - 1)).
+ intro Heq; subst i; intros; exact Ha.
+ rewrite <- to_Z_eq, to_Z_sub_1_0 by lia; intro Hneq.
+ rewrite leb_spec, to_Z_sub_1_0 by lia; intro Hji.
+ apply IH; [ reflexivity | rewrite leb_spec; lia ].
+ revert H; unfold afold_right.
+ case (reflect_eqb (length a) 0).
+ discriminate.
+ rewrite <- to_Z_eq, to_Z_0; intro Hlength.
+ apply (foldi_ind _ (fun i b => b = false -> a.[length a - 1] = false)).
+ rewrite leb_spec, to_Z_1; lia.
+ tauto.
+ intros i b; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, to_Z_1; intro Hi1.
+ rewrite ltb_spec; intros Hi2 IH.
+ unfold implb at 1; case (a.[length a - 1 - i]); [ exact IH | discriminate ].
+Qed.
+
+Lemma afold_right_implb_true_aux : forall a,
+ (exists i, i < length a - 1 = true /\ a.[i] = false) ->
+ afold_right bool true implb a = true.
+Proof.
+ intros a [ i [ Hi Hai ] ].
+ assert (Bi := to_Z_bounded i).
+ generalize (leb_0 i); apply afold_right_ind.
+ reflexivity.
+ intros _; revert Hi; rewrite ltb_spec, leb_spec; lia.
+ intros b j.
+ assert (Bj := to_Z_bounded j).
+ rewrite ltb_spec, to_Z_0; intro Hj1.
+ rewrite ltb_spec; intro Hj2.
+ rewrite leb_spec; intro IH.
+ rewrite leb_spec, to_Z_sub_1_0 by lia; intro Hji.
+ case (reflect_eqb i (j - 1)).
+ intro Heq; rewrite Heq in Hai; rewrite Hai; reflexivity.
+ rewrite <- to_Z_eq, to_Z_sub_1_0 by lia; intro Hneq.
+ rewrite IH by lia; case (a.[j - 1]); reflexivity.
+Qed.
+
+Lemma afold_right_implb_true : forall a,
+ length a = 0 \/ (exists i, i < length a - 1 = true /\ a.[i] = false) \/
+ (forall i, i < length a = true -> a.[i] = true) ->
+ afold_right bool true implb a = true.
+Proof.
+ intro a; assert (Ba := to_Z_bounded (length a)); case (reflect_eqb (length a) 0).
+ intros H _; unfold afold_right; rewrite H; reflexivity.
+ intro Hneq.
+ intros [H1|[H1|H1]].
+ elim (Hneq H1).
apply afold_right_implb_true_aux; auto.
- apply afold_right_ind_nonempty.
- intros b i H2 H3; subst b; case (f (a .[ i])); auto.
- apply not_0_ltb; intro H; rewrite H in Heq; discriminate.
- apply H1; rewrite ltb_spec, to_Z_sub_1_diff; [lia| ]; intro H; rewrite H in Heq; discriminate.
-Qed.
-
-
-Lemma afold_right_implb_true_inv : forall A a f,
- afold_right bool A true implb f a = true ->
- length a = 0 \/ (exists i, i < length a - 1 = true /\ f (a.[i]) = false) \/
- (forall i, i < length a = true -> f (a.[i]) = true).
-Proof.
- intros A a f; case_eq (length a == 0); intro Heq1.
- intros _; left; rewrite eqb_spec in Heq1; auto.
- case_eq (length a <= 1); intro Heq2.
- unfold afold_right; rewrite Heq1, Heq2; intro H; right; right; intros i Hi; replace i with 0; auto; apply to_Z_inj; rewrite ltb_spec in Hi; rewrite eqb_false_spec in Heq1; assert (H1: [|length a|] <> 0%Z) by (intro H1; apply Heq1, to_Z_inj; auto); rewrite leb_spec in Heq2; change [|1|] with 1%Z in Heq2; generalize (leb_0 (length a)) (leb_0 i); rewrite !leb_spec; change [|0|] with 0%Z; lia.
- pose (P j k := k = true -> (exists i : int, (j <= i) = true /\ (i < length a - 1) = true /\ f (a .[ i]) = false) \/ (forall i : int, (j <= i) = true -> (i < length a) = true -> f (a .[ i]) = true)); assert (H: P 0 (afold_right bool A true implb f a)).
- generalize (afold_right_Ind _ _ P true implb f a); rewrite Heq2; intro IH; apply IH; clear IH; unfold P.
- intros b i H1 H2 H3; case_eq b; intro Heq3.
- destruct (H2 Heq3) as [[j [H4 [H5 H6]]]|H4].
- left; exists j; repeat split; auto; rewrite leb_spec in *; rewrite (to_Z_add_1 _ _ H1) in H4; lia.
- case_eq (f (a.[i])); intro Heq4.
- right; intros j H5 H6; case_eq (i == j); intro Heq5.
- rewrite eqb_spec in Heq5; subst j; auto.
- apply H4; auto; rewrite leb_spec in *; rewrite (to_Z_add_1 _ _ H1); rewrite eqb_false_spec in Heq5; assert ([|i|] <> [|j|]) by (intro H; apply Heq5, to_Z_inj; auto); lia.
- left; exists i; repeat split; auto; apply leb_refl.
- rewrite Heq3 in H3; case_eq (f (a .[ i])); intro Heq4; rewrite Heq4 in H3; try discriminate; left; exists i; repeat split; auto; apply leb_refl.
- intros H1; right; intros i H2 H3; replace i with (length a - 1); auto; apply to_Z_inj; rewrite leb_spec in H2; rewrite (to_Z_sub_1 _ _ H3) in *; rewrite ltb_spec in H3; lia.
- unfold P in H; intro H1; right; destruct (H H1) as [[i [_ H2]]|H2].
- left; exists i; auto.
- right; intro i; apply H2, leb_0.
+ assert (Heq : length a == 0 = false) by (rewrite <- not_true_iff_false, eqb_spec; exact Hneq).
+ unfold afold_right; rewrite Heq.
+ revert Hneq; rewrite <- to_Z_eq, to_Z_0; intro Hneq.
+ apply (foldi_ind _ (fun i a => a = true)).
+ rewrite leb_spec, to_Z_1; lia.
+ apply H1; rewrite ltb_spec, to_Z_sub_1_0; lia.
+ intros i b; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, to_Z_1; intro Hi1.
+ rewrite ltb_spec; intros Hi2 IH.
+ rewrite IH; case (a.[length a - 1 - i]); reflexivity.
Qed.
+Lemma afold_right_implb_true_inv : forall a,
+ afold_right bool true implb a = true ->
+ length a = 0 \/ (exists i, i < length a - 1 = true /\ a.[i] = false) \/
+ (forall i, i < length a = true -> a.[i] = true).
+Proof.
+ intros a H; cut (length a = 0
+ \/ (exists i, 0 <= i = true /\ i < length a - 1 = true /\ a.[i] = false)
+ \/ (forall i, 0 <= i = true -> i < length a = true -> a.[i] = true)).
+ clear H; intro H; destruct H as [ H | H ].
+ left; tauto.
+ destruct H as [ H | H ].
+ destruct H as [ i [ Hi1 Hi2 ] ].
+ right; left; exists i; tauto.
+ right; right; intro i; apply H; apply leb_0.
+ assert (Ba := to_Z_bounded (length a)).
+ revert H; apply afold_right_ind.
+ left; tauto.
+ rewrite ltb_spec, to_Z_0; intro Hlength.
+ intro Ha; right; right.
+ intro i; assert (Bi := to_Z_bounded i).
+ rewrite leb_spec, to_Z_sub_1_0 by lia; intro Hi1.
+ rewrite ltb_spec; intro Hi2.
+ replace i with (length a - 1) by (rewrite <- to_Z_eq, to_Z_sub_1_0; lia); exact Ha.
+ intros b i; assert (Bi := to_Z_bounded i).
+ rewrite ltb_spec, to_Z_0; intro Hi1.
+ rewrite ltb_spec; intros Hi2 IH.
+ case_eq (a.[i - 1]); unfold implb.
+ intros Ha Hb; destruct (IH Hb) as [ Heq | H ]; clear IH.
+ rewrite Heq in Hi2; lia.
+ destruct H as [ [ j [ Hij Hj ] ] | H ].
+ right; left; exists j.
+ split; [ | exact Hj ].
+ revert Hij; rewrite 2!leb_spec, to_Z_sub_1_0; lia.
+ right; right; intro j.
+ rewrite leb_spec, sub_spec, to_Z_1, Z.mod_small by lia; intro Hij.
+ rewrite ltb_spec; intro Hj.
+ case (reflect_eqb j (i - 1)).
+ intro Heq; rewrite Heq; exact Ha.
+ rewrite <- to_Z_eq, to_Z_sub_1_0 by lia; intro Hneq.
+ apply H; [ rewrite leb_spec; lia | rewrite ltb_spec; lia ].
+ intros Ha _; right; left; exists (i - 1).
+ split; [ rewrite leb_spec; lia | ].
+ split; [ rewrite ltb_spec, 2!to_Z_sub_1_0; lia | exact Ha ].
+Qed.
(* Other cases *)
-Lemma afold_left_length_2 : forall A B default OP F t,
+Lemma afold_left_length_2 : forall A default OP t,
(length t == 2) = true ->
- afold_left A B default OP F t = OP (F (t.[0])) (F (t.[1])).
+ afold_left A default OP t = OP (t.[0]) (t.[1]).
Proof.
- intros A B default OP F t H; unfold afold_left; rewrite eqb_spec in H; rewrite H; change (2 == 0) with false; simpl; change (2-1) with 1; rewrite foldi_eq; trivial.
+ intros A default OP t H; unfold afold_left; rewrite eqb_spec in H; rewrite H; change (2 == 0) with false; reflexivity.
Qed.
-Lemma afold_right_length_2 : forall A B default OP F t,
+Lemma afold_right_length_2 : forall A default OP t,
(length t == 2) = true ->
- afold_right A B default OP F t = OP (F (t.[0])) (F (t.[1])).
+ afold_right A default OP t = OP (t.[0]) (t.[1]).
Proof.
- intros A B default OP F t H; unfold afold_right; rewrite eqb_spec in H; rewrite H; change (2 == 0) with false; simpl; change (2<=1) with false; simpl; change (2-2) with 0; rewrite foldi_down_eq; trivial.
+ intros A default OP t H; unfold afold_right; rewrite eqb_spec in H; rewrite H; change (2 == 0) with false; reflexivity.
Qed.
Ltac tac_left :=
- intros t f H H1 H2; rewrite afold_left_length_2;
+ intros t H H1 H2; rewrite afold_left_length_2;
[rewrite H1, H2| ]; trivial.
Ltac tac_right :=
- try (intros t f H H1 H2; rewrite afold_right_length_2;
+ try (intros t H H1 H2; rewrite afold_right_length_2;
[rewrite H1, H2| ]; trivial);
- try (intros t f H H1; rewrite afold_right_length_2;
+ try (intros t H H1; rewrite afold_right_length_2;
[rewrite H1| ]; trivial);
try (rewrite implb_true_r; trivial).
-Lemma afold_left_xorb_false1 : forall t f,
+Lemma afold_left_xorb_false1 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = false -> f (t .[ 1]) = false ->
- afold_left bool int false xorb f t = false.
+ t .[ 0] = false -> t .[ 1] = false ->
+ afold_left bool false xorb t = false.
Proof. tac_left. Qed.
-Lemma afold_left_xorb_false2 : forall t f,
+Lemma afold_left_xorb_false2 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = true -> f (t .[ 1]) = true ->
- afold_left bool int false xorb f t = false.
+ t .[ 0] = true -> t .[ 1] = true ->
+ afold_left bool false xorb t = false.
Proof. tac_left. Qed.
-Lemma afold_left_xorb_true1 : forall t f,
+Lemma afold_left_xorb_true1 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = false -> f (t .[ 1]) = true ->
- afold_left bool int false xorb f t = true.
+ t .[ 0] = false -> t .[ 1] = true ->
+ afold_left bool false xorb t = true.
Proof. tac_left. Qed.
-Lemma afold_left_xorb_true2 : forall t f,
+Lemma afold_left_xorb_true2 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = true -> f (t .[ 1]) = false ->
- afold_left bool int false xorb f t = true.
+ t .[ 0] = true -> t .[ 1] = false ->
+ afold_left bool false xorb t = true.
Proof. tac_left. Qed.
@@ -777,31 +1118,31 @@ Proof. tac_left. Qed.
(* Proof. tac_right. Qed. *)
-Lemma afold_left_eqb_false1 : forall t f,
+Lemma afold_left_eqb_false1 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = false -> f (t .[ 1]) = true ->
- afold_left bool int true eqb f t = false.
+ t .[ 0] = false -> t .[ 1] = true ->
+ afold_left bool true eqb t = false.
Proof. tac_left. Qed.
-Lemma afold_left_eqb_false2 : forall t f,
+Lemma afold_left_eqb_false2 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = true -> f (t .[ 1]) = false ->
- afold_left bool int true eqb f t = false.
+ t .[ 0] = true -> t .[ 1] = false ->
+ afold_left bool true eqb t = false.
Proof. tac_left. Qed.
-Lemma afold_left_eqb_true1 : forall t f,
+Lemma afold_left_eqb_true1 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = true -> f (t .[ 1]) = true ->
- afold_left bool int true eqb f t = true.
+ t .[ 0] = true -> t .[ 1] = true ->
+ afold_left bool true eqb t = true.
Proof. tac_left. Qed.
-Lemma afold_left_eqb_true2 : forall t f,
+Lemma afold_left_eqb_true2 : forall t,
(PArray.length t == 2) = true ->
- f (t .[ 0]) = false -> f (t .[ 1]) = false ->
- afold_left bool int true eqb f t = true.
+ t .[ 0] = false -> t .[ 1] = false ->
+ afold_left bool true eqb t = true.
Proof. tac_left. Qed.
@@ -990,48 +1331,67 @@ End Distinct.
Arguments distinct [A] eq l.
+(** Specification of aexistsbi and aforallbi *)
-(** Specification of existsb *)
+Definition aexistsbi {A:Type} (f:int->A->bool) (t:array A) :=
+ afold_left _ false orb (amapi f t).
-Lemma existsb_false_spec : forall f from to,
- existsb f from to = false <->
- forall i, ((from <= i) = true /\ (i <= to) = true) -> f i = false.
+Lemma aexistsbi_false_spec : forall A (f : int -> A -> bool) t,
+ aexistsbi f t = false <->
+ forall i, i < length t = true -> f i (t.[i]) = false.
Proof.
- unfold existsb;intros; setoid_rewrite leb_spec; apply foldi_cont_ZInd.
- intros z Hz; split; auto; intros _ i [H1 H2]; assert (H3 := Z.le_trans _ _ _ H1 H2); elimtype False; lia.
- intros i cont H1 H2 H3; case_eq (f i); intro Heq.
- split; try discriminate; intro H; rewrite <- Heq; apply H; split; try lia; rewrite leb_spec in H2; auto.
- rewrite H3; split; intros H j [Hj1 Hj2].
- case_eq (i == j); intro Heq2.
- rewrite eqb_spec in Heq2; subst j; auto.
- apply H; split; auto; rewrite eqb_false_spec in Heq2; assert ([|i|] <> [|j|]) by (intro; apply Heq2, to_Z_inj; auto); lia.
- apply H; lia.
+ intros A f t; unfold aexistsbi.
+ split.
+ intro H; generalize (afold_left_orb_false_inv _ H); clear H.
+ rewrite length_amapi; intros H i Hi.
+ rewrite <- get_amapi by exact Hi.
+ apply H; exact Hi.
+ intro H; apply afold_left_orb_false.
+ intro i; rewrite length_amapi; intro Hi; rewrite get_amapi by exact Hi; apply H; exact Hi.
Qed.
-
-Lemma array_existsbi_false_spec : forall A (f : int -> A -> bool) t,
- existsbi f t = false <->
- forall i, i < length t = true -> f i (t.[i]) = false.
+Lemma aexistsbi_spec : forall A (f : int -> A -> bool) t,
+ aexistsbi f t = true <-> exists i, i < length t = true /\ f i (t.[i]) = true.
Proof.
- unfold existsbi;intros A f t; destruct (reflect_eqb 0 (length t)).
- split; auto. intros _ i Hi. elim (ltb_0 i). rewrite e. auto.
- rewrite existsb_false_spec. split.
- intros H i Hi. apply H. split; [apply leb_0| ]. rewrite leb_spec. rewrite (to_Z_sub_1 _ _ Hi). rewrite ltb_spec in Hi. lia.
- intros H i [_ Hi]. apply H. rewrite ltb_spec. rewrite leb_spec in Hi. rewrite to_Z_sub_1_diff in Hi; auto; lia.
+ intros A f t; unfold aexistsbi.
+ split.
+ intro H; generalize (afold_left_orb_true_inv _ H); clear H.
+ intros [ i [ Hi Hf ] ]; exists i.
+ rewrite length_amapi in Hi; rewrite get_amapi in Hf by exact Hi.
+ split; [ exact Hi | exact Hf ].
+ intros [ i [ Hi Hf ] ].
+ apply (afold_left_orb_true i); [ rewrite length_amapi; exact Hi | rewrite get_amapi by exact Hi; exact Hf ].
Qed.
+Definition aforallbi {A:Type} (f:int->A->bool) (t:array A) :=
+ afold_left _ true andb (amapi f t).
-Lemma array_existsb_false_spec : forall A (f : A -> bool) t,
- PArray.existsb f t = false <->
- forall i, i < length t = true -> f (t.[i]) = false.
+Lemma aforallbi_false_spec : forall A (f : int -> A -> bool) t,
+ aforallbi f t = false <-> exists i, i < length t = true /\ f i (t.[i]) = false.
Proof.
- intros A f t; unfold PArray.existsb; case_eq (0 == length t).
- rewrite eqb_spec; intro H; split; auto; intros _ i Hi; elim (ltb_0 i); rewrite H; auto.
- intro H; rewrite existsb_false_spec; split.
- intros H1 i Hi; apply H1; split; [apply leb_0| ]; rewrite leb_spec, (to_Z_sub_1 _ _ Hi); rewrite ltb_spec in Hi; lia.
- intros H1 i [_ H2]; apply H1; rewrite ltb_spec; rewrite leb_spec in H2; rewrite to_Z_sub_1_diff in H2; [lia| ]; intro H3; rewrite H3 in H; discriminate.
+ intros A f t; unfold aforallbi.
+ split.
+ intro H; generalize (afold_left_andb_false_inv _ H); clear H.
+ intros [ i [ Hi Hf ] ]; exists i.
+ rewrite length_amapi in Hi; rewrite get_amapi in Hf by exact Hi.
+ split; [ exact Hi | exact Hf ].
+ intros [ i [ Hi Hf ] ].
+ apply (afold_left_andb_false i); [ rewrite length_amapi; exact Hi | rewrite get_amapi by exact Hi; exact Hf ].
Qed.
+Lemma aforallbi_spec : forall A (f : int -> A -> bool) t,
+ aforallbi f t = true <->
+ forall i, i < length t = true -> f i (t.[i]) = true.
+Proof.
+ intros A f t; unfold aforallbi.
+ split.
+ intro H; generalize (afold_left_andb_true_inv _ H); clear H.
+ rewrite length_amapi; intros H i Hi.
+ rewrite <- get_amapi by exact Hi.
+ apply H; exact Hi.
+ intro H; apply afold_left_andb_true.
+ intro i; rewrite length_amapi; intro Hi; rewrite get_amapi by exact Hi; apply H; exact Hi.
+Qed.
(** Forall of two lists at the same time *)
@@ -1051,38 +1411,6 @@ End Forall2.
Arguments forallb2 {A B} f l1 l2.
-(* Compatibility between native-coq and Coq 8.5 *)
-
-Definition Nat_eqb :=
- fix eqb (n m : nat) {struct n} : bool :=
- match n with
- | O => match m with
- | O => true
- | S _ => false
- end
- | S n' => match m with
- | O => false
- | S m' => eqb n' m'
- end
- end.
-
-Definition List_map_ext_in
- : forall (A B : Type) (f g : A -> B) (l : list A),
- (forall a : A, In a l -> f a = g a) -> List.map f l = List.map g l :=
- fun (A B : Type) (f g : A -> B) (l : list A) =>
- list_ind
- (fun l0 : list A =>
- (forall a : A, In a l0 -> f a = g a) -> List.map f l0 = List.map g l0)
- (fun _ : forall a : A, False -> f a = g a => eq_refl)
- (fun (a : A) (l0 : list A)
- (IHl : (forall a0 : A, In a0 l0 -> f a0 = g a0) -> List.map f l0 = List.map g l0)
- (H : forall a0 : A, a = a0 \/ In a0 l0 -> f a0 = g a0) =>
- eq_ind_r (fun b : B => b :: List.map f l0 = g a :: List.map g l0)
- (eq_ind_r (fun l1 : list B => g a :: l1 = g a :: List.map g l0) eq_refl
- (IHl (fun (a0 : A) (H0 : In a0 l0) => H a0 (or_intror H0))))
- (H a (or_introl eq_refl))) l.
-
-
(* Misc lemmas *)
Lemma neg_eq_true_eq_false b : b = false <-> b <> true.
diff --git a/src/SMT_terms.v b/src/SMT_terms.v
index b3df896..aac66be 100644
--- a/src/SMT_terms.v
+++ b/src/SMT_terms.v
@@ -10,7 +10,7 @@
(**************************************************************************)
-Require Import Bool Int63 PArray BinNat BinPos ZArith SMT_classes_instances.
+Require Import Bool Int63 Psatz PArray BinNat BinPos ZArith SMT_classes_instances.
Require Import Misc State BVList. (* FArray Equalities DecidableTypeEx. *)
Require FArray.
Require List .
@@ -79,10 +79,10 @@ Module Form.
| Fatom a => interp_atom a
| Ftrue => true
| Ffalse => false
- | Fnot2 i l => fold (fun b => negb (negb b)) 1 i (Lit.interp interp_var l)
- | Fand args => afold_left _ _ true andb (Lit.interp interp_var) args
- | For args => afold_left _ _ false orb (Lit.interp interp_var) args
- | Fimp args => afold_right _ _ true implb (Lit.interp interp_var) args
+ | Fnot2 i l => foldi (fun _ b => negb (negb b)) 0 i (Lit.interp interp_var l)
+ | Fand args => afold_left _ true andb (amap (Lit.interp interp_var) args)
+ | For args => afold_left _ false orb (amap (Lit.interp interp_var) args)
+ | Fimp args => afold_right _ true implb (amap (Lit.interp interp_var) args)
| Fxor a b => xorb (Lit.interp interp_var a) (Lit.interp interp_var b)
| Fiff a b => Bool.eqb (Lit.interp interp_var a) (Lit.interp interp_var b)
| Fite a b c =>
@@ -98,19 +98,19 @@ Module Form.
Section Interp_get.
- Variable t_form : PArray.array form.
+ Variable t_form : array form.
- Definition t_interp : PArray.array bool :=
- PArray.foldi_left (fun i t_b hf =>
- t_b.[i <- interp_aux (PArray.get t_b) hf])
- (PArray.make (PArray.length t_form) true) t_form.
+ Definition t_interp : array bool :=
+ foldi (fun i t_b =>
+ t_b.[i <- interp_aux (get t_b) (t_form.[i])]) 0 (length t_form)
+ (make (length t_form) true).
Definition lt_form i h :=
match h with
| Fatom _ | Ftrue | Ffalse => true
| Fnot2 _ l => Lit.blit l < i
| Fand args | For args | Fimp args =>
- PArray.forallb (fun l => Lit.blit l < i) args
+ aforallbi (fun _ l => Lit.blit l < i) args
| Fxor a b | Fiff a b => (Lit.blit a < i) && (Lit.blit b < i)
| Fite a b c => (Lit.blit a < i) && (Lit.blit b < i) && (Lit.blit c < i)
| FbbT _ ls => List.forallb (fun l => Lit.blit l < i) ls
@@ -123,41 +123,42 @@ Module Form.
interp_aux f1 h = interp_aux f2 h.
Proof.
destruct h;simpl;intros;trivial;
- try (apply afold_left_eq;unfold is_true in H0;
- rewrite PArray.forallb_spec in H0;intros;
+ try (try apply afold_left_eq; try apply afold_right_eq;unfold is_true in H0;
+ rewrite ?forallb_spec, ?aforallbi_spec, ?andb_true_iff in H0;intros;
+ rewrite ?length_amap; try rewrite length_amap in H1;
+ rewrite ?get_amap by assumption;
auto using Lit.interp_eq_compat with smtcoq_core).
- f_equal;auto using Lit.interp_eq_compat.
- - apply afold_right_eq;unfold is_true in H0;
- rewrite PArray.forallb_spec in H0;intros;
- auto using Lit.interp_eq_compat with smtcoq_core.
- - unfold is_true in H0;rewrite !andb_true_iff in H0;decompose [and] H0;
- rewrite !(Lit.interp_eq_compat f1 f2);auto.
- - unfold is_true in H0;rewrite !andb_true_iff in H0;decompose [and] H0;
- rewrite !(Lit.interp_eq_compat f1 f2);auto.
- - unfold is_true in H0;rewrite !andb_true_iff in H0;decompose [and] H0;
- rewrite !(Lit.interp_eq_compat f1 f2);auto.
+ - decompose [and] H0; rewrite !(Lit.interp_eq_compat f1 f2);auto.
+ - decompose [and] H0; rewrite !(Lit.interp_eq_compat f1 f2);auto.
+ - decompose [and] H0; rewrite !(Lit.interp_eq_compat f1 f2);auto.
- replace (List.map (Lit.interp f2) l) with (List.map (Lit.interp f1) l); auto.
unfold is_true in H0. rewrite List.forallb_forall in H0.
- apply List_map_ext_in. intros x Hx. apply Lit.interp_eq_compat; auto with smtcoq_core.
+ apply List.map_ext_in. intros x Hx. apply Lit.interp_eq_compat; auto with smtcoq_core.
Qed.
- Definition wf := PArray.forallbi lt_form t_form.
+ Definition wf := aforallbi lt_form t_form.
Hypothesis wf_t_i : wf.
Lemma length_t_interp : length t_interp = length t_form.
Proof.
- unfold t_interp;apply PArray.foldi_left_Ind with (P := fun i a => length a = length t_form).
+ unfold t_interp.
+ assert (Bt := to_Z_bounded (length t_form)).
+ apply (foldi_ind _ (fun i a => length a = length t_form)).
+ rewrite leb_spec, to_Z_0; lia.
+ rewrite length_make, leb_length; reflexivity.
intros;rewrite length_set;trivial.
- rewrite length_make, ltb_length;trivial.
Qed.
Lemma default_t_interp : default t_interp = true.
Proof.
- unfold t_interp;apply PArray.foldi_left_Ind with
- (P := fun i a => default a = true).
- intros;rewrite default_set;trivial.
+ unfold t_interp.
+ assert (Bt := to_Z_bounded (length t_form)).
+ apply (foldi_ind _ (fun i a => default a = true)).
+ rewrite leb_spec, to_Z_0; lia.
apply default_make.
+ intros;rewrite default_set;trivial.
Qed.
Lemma t_interp_wf : forall i, i < PArray.length t_form ->
@@ -167,26 +168,27 @@ Module Form.
forall j, j < i ->
t.[j] = interp_aux (PArray.get t) (t_form.[j])).
assert (P' (length t_form) t_interp).
- unfold is_true, wf in wf_t_i;rewrite PArray.forallbi_spec in wf_t_i.
- unfold t_interp;apply foldi_left_Ind;unfold P';intros.
- rewrite length_set in H1.
- destruct (Int63Properties.reflect_eqb j i).
- rewrite e, PArray.get_set_same.
- apply lt_form_interp_form_aux with (2:= wf_t_i i H).
+ unfold is_true, wf in wf_t_i;rewrite aforallbi_spec in wf_t_i.
+ unfold t_interp;apply foldi_ind;unfold P';intros.
+ apply leb_0.
+ elim (ltb_0 _ H0).
+ rewrite length_set in H2.
+ destruct (reflect_eqb j i).
+ rewrite e, get_set_same.
+ apply lt_form_interp_form_aux with (2:= wf_t_i i H0).
intros;rewrite get_set_other;trivial.
intros Heq;elim (not_ltb_refl i);rewrite Heq at 1;trivial.
- rewrite H1;trivial.
+ rewrite H2;trivial.
assert (j < i).
assert ([|j|] <> [|i|]) by (intros Heq1;elim n;apply to_Z_inj;trivial).
- generalize H2;unfold is_true;rewrite !ltb_spec, (to_Z_add_1 _ _ H);
+ generalize H3;unfold is_true;rewrite !ltb_spec, (to_Z_add_1 _ _ H0);
auto with zarith.
- rewrite get_set_other, H0;auto.
+ rewrite get_set_other, H1;auto.
apply lt_form_interp_form_aux with
- (2:= wf_t_i j (ltb_trans _ _ _ H3 H)).
+ (2:= wf_t_i j (ltb_trans _ _ _ H4 H0)).
intros;rewrite get_set_other;trivial.
intros Heq;elim (not_ltb_refl i);apply ltb_trans with j;
[ rewrite Heq| ];trivial.
- elim (ltb_0 _ H0).
apply H;apply length_t_interp.
Qed.
@@ -252,7 +254,7 @@ Module Typ.
Import FArray.
- Notation index := int (only parsing).
+ Notation index := N (only parsing).
Inductive type :=
| TFArray : type -> type -> type
@@ -283,7 +285,7 @@ Module Typ.
(type_compdec (projT2 (interp_compdec_aux te))))
| Tindex i =>
existT (fun ty : Type => CompDec ty)
- (te_carrier (t_i .[ i])) (te_compdec (t_i .[ i]))
+ (te_carrier (t_i .[of_Z (Z.of_N i)])) (te_compdec (t_i .[of_Z (Z.of_N i)]))
| TZ => existT (fun ty : Type => CompDec ty) Z Z_compdec
| Tbool => existT (fun ty : Type => CompDec ty) bool bool_compdec
| Tpositive => existT (fun ty : Type => CompDec ty) positive Positive_compdec
@@ -379,7 +381,7 @@ Module Typ.
reflect (x = y) (eqb_of_compdec c x y).
Proof.
intros x y.
- apply reflect_eqb.
+ apply SMT_classes.reflect_eqb.
Qed.
@@ -396,7 +398,7 @@ Module Typ.
Definition i_eqb_eqb (t:type) : interp t -> interp t -> bool :=
match t with
- | Tindex i => eqb_of_compdec (t_i.[i]).(te_compdec)
+ | Tindex i => eqb_of_compdec (t_i.[of_Z (Z.of_N i)]).(te_compdec)
| TZ => Z.eqb (* Zeq_bool *)
| Tbool => Bool.eqb
| Tpositive => Pos.eqb
@@ -505,7 +507,7 @@ Module Typ.
Fixpoint cast (A B: type) : cast_result A B :=
match A as C, B as D return cast_result C D with
| Tindex i, Tindex j =>
- match Int63Op.cast i j with
+ match N_cast i j with
| Some k => Cast (fun P => k (fun y => P (Tindex y)))
| None => NoCast
end
@@ -543,7 +545,7 @@ Module Typ.
Proof.
destruct A;simpl;trivial.
do 2 rewrite cast_refl. easy.
- rewrite Int63Properties.cast_refl;trivial.
+ rewrite N_cast_refl;trivial.
rewrite N_cast_refl;trivial.
Qed.
@@ -551,7 +553,7 @@ Module Typ.
(* Remark : I use this definition because eqb will not be used only in the interpretation *)
Fixpoint eqb (A B: type) : bool :=
match A, B with
- | Tindex i, Tindex j => i == j
+ | Tindex i, Tindex j => N.eqb i j
| TZ, TZ => true
| Tbool, Tbool => true
| Tpositive, Tpositive => true
@@ -599,7 +601,7 @@ Module Typ.
rewrite andb_false_iff in H.
destruct H; apply cast_diff in H; rewrite H; auto.
case (cast A1 B1); auto.
- intros H. rewrite (Int63Properties.cast_diff _ _ H);trivial.
+ rewrite N.eqb_neq. intro Heq. now rewrite N_cast_diff.
rewrite N.eqb_neq. intro Heq. now rewrite N_cast_diff.
Qed.
@@ -634,7 +636,7 @@ Module Typ.
apply (reflect_iff _ _ (reflect_eqb x1 y1)) in H.
apply (reflect_iff _ _ (reflect_eqb x2 y2)) in H0.
subst; auto.
- apply iff_reflect;rewrite Int63Properties.eqb_spec;split;intros H;[inversion H | subst]; trivial.
+ apply iff_reflect. rewrite N.eqb_eq. split;intros H;[inversion H | subst]; trivial.
apply iff_reflect. rewrite N.eqb_eq. split;intros H;[inversion H | subst]; trivial.
Qed.
@@ -804,7 +806,7 @@ Module Atom.
| UO_Zpos, UO_Zpos
| UO_Zneg, UO_Zneg
| UO_Zopp, UO_Zopp => true
- | UO_BVbitOf s1 n, UO_BVbitOf s2 m => Nat_eqb n m && N.eqb s1 s2
+ | UO_BVbitOf s1 n, UO_BVbitOf s2 m => Nat.eqb n m && N.eqb s1 s2
| UO_BVnot s1, UO_BVnot s2 => N.eqb s1 s2
| UO_BVneg s1, UO_BVneg s2 => N.eqb s1 s2
| UO_BVextr i0 n00 n01, UO_BVextr i1 n10 n11 => N.eqb i0 i1 && N.eqb n00 n10 && N.eqb n01 n11
@@ -853,10 +855,10 @@ Module Atom.
| Acop o, Acop o' => cop_eqb o o'
| Auop o t, Auop o' t' => uop_eqb o o' && (t == t')
| Abop o t1 t2, Abop o' t1' t2' => bop_eqb o o' && (t1 == t1') && (t2 == t2')
- | Anop o t, Anop o' t' => nop_eqb o o' && list_beq Int63Native.eqb t t'
+ | Anop o t, Anop o' t' => nop_eqb o o' && list_beq Int63.eqb t t'
| Atop o t1 t2 t3, Atop o' t1' t2' t3' =>
top_eqb o o' && (t1 == t1') && (t2 == t2') && (t3 == t3')
- | Aapp a la, Aapp b lb => (a == b) && list_beq Int63Native.eqb la lb
+ | Aapp a la, Aapp b lb => (a == b) && list_beq Int63.eqb la lb
| _, _ => false
end.
@@ -880,7 +882,7 @@ Module Atom.
Lemma reflect_uop_eqb : forall o1 o2, reflect (o1 = o2) (uop_eqb o1 o2).
Proof.
intros [ | | | | | s1 n1 | s1 | s1 | s1 | s1 | s1 ] [ | | | | |s2 n2 | s2 | s2 | s2 | s2 | s2 ];simpl; try constructor;trivial; try discriminate.
- - apply iff_reflect. case_eq (Nat_eqb n1 n2).
+ - apply iff_reflect. case_eq (Nat.eqb n1 n2).
+ case_eq ((s1 =? s2)%N).
* rewrite N.eqb_eq, beq_nat_true_iff.
intros -> ->. split; reflexivity.
@@ -982,26 +984,26 @@ Qed.
(* Constants *)
preflect (reflect_cop_eqb c c0);constructor;subst;trivial.
(* Unary operators *)
- preflect (reflect_uop_eqb u u0); preflect (Int63Properties.reflect_eqb i i0);
+ preflect (reflect_uop_eqb u u0); preflect (Misc.reflect_eqb i i0);
constructor;subst;trivial.
(* Binary operators *)
preflect (reflect_bop_eqb b b0);
- preflect (Int63Properties.reflect_eqb i i1);
- preflect (Int63Properties.reflect_eqb i0 i2);
+ preflect (Misc.reflect_eqb i i1);
+ preflect (Misc.reflect_eqb i0 i2);
constructor;subst;trivial.
(* Ternary operators *)
preflect (reflect_top_eqb t t0).
- preflect (Int63Properties.reflect_eqb i i2).
- preflect (Int63Properties.reflect_eqb i0 i3).
- preflect (Int63Properties.reflect_eqb i1 i4).
+ preflect (Misc.reflect_eqb i i2).
+ preflect (Misc.reflect_eqb i0 i3).
+ preflect (Misc.reflect_eqb i1 i4).
constructor;subst;trivial.
(* N-ary operators *)
preflect (reflect_nop_eqb n n0);
- preflect (reflect_list_beq _ _ Int63Properties.reflect_eqb l l0);
+ preflect (reflect_list_beq _ _ Misc.reflect_eqb l l0);
constructor; subst; reflexivity.
(* Application *)
- preflect (Int63Properties.reflect_eqb i i0);
- preflect (reflect_list_beq _ _ Int63Properties.reflect_eqb l l0);
+ preflect (Misc.reflect_eqb i i0);
+ preflect (reflect_list_beq _ _ Misc.reflect_eqb l l0);
constructor;subst;trivial.
Qed.
@@ -2228,9 +2230,9 @@ Qed.
Variable t_atom : PArray.array atom.
- Definition t_interp : PArray.array bval :=
- PArray.foldi_left (fun i t_a a => t_a.[i <- interp_aux (PArray.get t_a) a])
- (PArray.make (PArray.length t_atom) (interp_cop CO_xH)) t_atom.
+ Definition t_interp : array bval :=
+ foldi (fun i t_a => t_a.[i <- interp_aux (get t_a) (t_atom.[i])]) 0 (length t_atom)
+ (make (length t_atom) (interp_cop CO_xH)).
Definition lt_atom i a :=
match a with
@@ -2262,52 +2264,58 @@ Qed.
unfold is_true in H;rewrite andb_true_iff in H;destruct H;rewrite Hf, IHl;trivial.
Qed.
- Definition wf := PArray.forallbi lt_atom t_atom.
+ Definition wf := aforallbi lt_atom t_atom.
Hypothesis wf_t_i : wf.
Lemma length_t_interp : length t_interp = length t_atom.
Proof.
- unfold t_interp;apply PArray.foldi_left_Ind with
- (P := fun i a => length a = length t_atom).
+ unfold t_interp.
+ assert (Bt := to_Z_bounded (length t_atom)).
+ apply (foldi_ind _ (fun i a => length a = length t_atom)).
+ rewrite leb_spec, to_Z_0; lia.
+ rewrite length_make, leb_length;trivial.
intros;rewrite length_set;trivial.
- rewrite length_make, ltb_length;trivial.
Qed.
Lemma default_t_interp : default t_interp = interp_cop CO_xH.
Proof.
- unfold t_interp;apply PArray.foldi_left_Ind with
- (P := fun i a => default a = interp_cop CO_xH).
- intros;rewrite default_set;trivial.
+ unfold t_interp.
+ assert (Bt := to_Z_bounded (length t_atom)).
+ apply (foldi_ind _ (fun i a => default a = interp_cop CO_xH)).
+ rewrite leb_spec, to_Z_0; lia.
apply default_make.
+ intros;rewrite default_set;trivial.
Qed.
Lemma t_interp_wf_lt : forall i, i < PArray.length t_atom ->
t_interp.[i] = interp_aux (PArray.get t_interp) (t_atom.[i]).
Proof.
+ assert (Bt := to_Z_bounded (length t_atom)).
set (P' i t := length t = length t_atom ->
forall j, j < i ->
t.[j] = interp_aux (PArray.get t) (t_atom.[j])).
assert (P' (length t_atom) t_interp).
- unfold is_true, wf in wf_t_i;rewrite PArray.forallbi_spec in wf_t_i.
- unfold t_interp;apply foldi_left_Ind;unfold P';intros.
- rewrite length_set in H1.
- destruct (Int63Properties.reflect_eqb j i).
+ unfold is_true, wf in wf_t_i;rewrite aforallbi_spec in wf_t_i.
+ unfold t_interp;apply foldi_ind;unfold P';intros.
+ rewrite leb_spec, to_Z_0; lia.
+ elim (ltb_0 _ H0).
+ rewrite length_set in H2.
+ destruct (Misc.reflect_eqb j i).
rewrite e, PArray.get_set_same.
- apply lt_interp_aux with (2:= wf_t_i i H).
+ apply lt_interp_aux with (2:= wf_t_i i H0).
intros;rewrite get_set_other;trivial.
intros Heq;elim (not_ltb_refl i);rewrite Heq at 1;trivial.
- rewrite H1;trivial.
+ rewrite H2;trivial.
assert (j < i).
assert ([|j|] <> [|i|]) by(intros Heq1;elim n;apply to_Z_inj;trivial).
- generalize H2;unfold is_true;rewrite !ltb_spec,
- (to_Z_add_1 _ _ H);auto with zarith.
- rewrite get_set_other, H0;auto.
- apply lt_interp_aux with (2:= wf_t_i j (ltb_trans _ _ _ H3 H)).
+ generalize H3;unfold is_true;rewrite !ltb_spec,
+ (to_Z_add_1 _ _ H0);auto with zarith.
+ rewrite get_set_other, H1;auto.
+ apply lt_interp_aux with (2:= wf_t_i j (ltb_trans _ _ _ H4 H0)).
intros;rewrite get_set_other;trivial.
intros Heq;elim (not_ltb_refl i);apply ltb_trans with j;
[ rewrite Heq| ];trivial.
- elim (ltb_0 _ H0).
apply H;apply length_t_interp.
Qed.
@@ -2366,7 +2374,7 @@ Qed.
exists v, interp_aux (get a) (t_atom.[h]) =
Bval (v_type _ _ (interp_aux (get a) (t_atom.[h]))) v.
Proof.
- unfold wf, is_true in wf_t_i; rewrite forallbi_spec in wf_t_i.
+ unfold wf, is_true in wf_t_i; rewrite aforallbi_spec in wf_t_i.
intros h Hh a IH; generalize (wf_t_i h Hh).
case (t_atom.[h]); simpl.
(* Constants *)
@@ -2503,22 +2511,24 @@ Qed.
Lemma check_aux_interp_hatom_lt : forall h, h < length t_atom ->
exists v, t_interp.[h] = Bval (get_type h) v.
Proof.
+ assert (Bt := to_Z_bounded (length t_atom)).
set (P' i t := length t = length t_atom ->
forall j, j < i ->
exists v, t.[j] = Bval (v_type Typ.type interp_t (t.[j])) v).
assert (P' (length t_atom) t_interp).
- unfold t_interp;apply foldi_left_Ind;unfold P';intros.
- rewrite length_set in H1.
- destruct (Int63Properties.reflect_eqb j i).
+ unfold t_interp;apply foldi_ind;unfold P';intros.
+ rewrite leb_spec, to_Z_0; lia.
+ elim (ltb_0 _ H0).
+ rewrite length_set in H2.
+ destruct (Misc.reflect_eqb j i).
rewrite e, PArray.get_set_same.
apply check_aux_interp_aux_lt; auto.
- rewrite H1; auto.
+ rewrite H2; auto.
assert (j < i).
assert ([|j|] <> [|i|]) by(intros Heq1;elim n;apply to_Z_inj;trivial).
- generalize H2;unfold is_true;rewrite !ltb_spec,
- (to_Z_add_1 _ _ H);auto with zarith.
+ generalize H3;unfold is_true;rewrite !ltb_spec,
+ (to_Z_add_1 _ _ H0);auto with zarith.
rewrite get_set_other;auto.
- elim (ltb_0 _ H0).
apply H;apply length_t_interp.
Qed.
@@ -2557,7 +2567,7 @@ Qed.
Definition wt t_atom :=
let t_interp := t_interp t_atom in
let get_type := get_type' t_interp in
- PArray.forallbi (fun i h => check_aux get_type h (get_type i)) t_atom.
+ aforallbi (fun i h => check_aux get_type h (get_type i)) t_atom.
Definition interp_hatom (t_atom : PArray.array atom) :=
let t_a := t_interp t_atom in
diff --git a/src/State.v b/src/State.v
index 518007a..ae69f37 100644
--- a/src/State.v
+++ b/src/State.v
@@ -10,7 +10,7 @@
(**************************************************************************)
-Require Import List Bool Int63 PArray Psatz.
+Require Import List Bool Int63 Psatz Ring63 PArray Omega Misc Ring.
(* Require Import AxiomesInt. *)
@@ -88,13 +88,13 @@ Module Lit.
Lemma neg_involutive : forall l, neg (neg l) = l.
Proof.
- unfold neg;intros; rewrite <- lxor_assoc;change (1 lxor 1) with 0;rewrite lxor_0_r;trivial.
+ unfold neg;intros; rewrite <- lxorA;change (1 lxor 1) with 0;rewrite lxor0_r;trivial.
Qed.
Lemma blit_neg : forall l, blit (neg l) = blit l.
Proof.
unfold blit, neg;intros l.
- rewrite lxor_lsr, lxor_0_r;trivial.
+ rewrite lxor_lsr, lxor0_r;trivial.
Qed.
Lemma lit_blit: forall l,
@@ -103,7 +103,7 @@ Module Lit.
unfold is_pos, lit, blit;intros.
rewrite (bit_xor_split l) at 2.
rewrite is_even_bit, negb_true_iff in H;rewrite H.
- symmetry;apply lxor_0_r.
+ symmetry;apply lxor0_r.
Qed.
Lemma lit_blit_neg: forall l,
@@ -220,7 +220,7 @@ Module Lit.
Lemma lxor_neg : forall l1 l2, (l1 lxor l2 == 1) = true -> l1 = Lit.neg l2.
Proof.
unfold Lit.neg; intros l1 l2;rewrite eqb_spec;intros Heq;rewrite <- Heq.
- rewrite lxor_comm, <- lxor_assoc, lxor_nilpotent, lxor_0_r;trivial.
+ rewrite lxorC, <- lxorA, lxor_nilpotent, lxor0_r;trivial.
Qed.
End Lit.
@@ -484,7 +484,7 @@ Module S.
forall id, valid rho (set s id c).
Proof.
unfold valid, get;simpl;intros.
- destruct (Int63Properties.reflect_eqb id id0);subst.
+ destruct (reflect_eqb id id0);subst.
case_eq (id0 < length s);intros.
rewrite PArray.get_set_same;trivial.
rewrite PArray.get_outofbound.
@@ -622,7 +622,7 @@ Module S.
C.valid rho c -> valid rho (set_clause s pos c).
Proof.
unfold valid, get, set_clause. intros rho s Hrho Hs pos c Hc id.
- destruct (Int63Properties.reflect_eqb pos id);subst.
+ destruct (reflect_eqb pos id);subst.
case_eq (id < length s); intro H.
unfold get;rewrite PArray.get_set_same; trivial.
unfold C.valid;rewrite sort_correct;trivial.
@@ -640,7 +640,7 @@ Module S.
C.valid rho c -> valid rho (set_clause_keep s pos c).
Proof.
unfold valid, get, set_clause_keep. intros rho s Hrho Hs pos c Hc id.
- destruct (Int63Properties.reflect_eqb pos id);subst.
+ destruct (reflect_eqb pos id);subst.
case_eq (id < length s); intro H.
unfold get;rewrite PArray.get_set_same; trivial.
unfold C.valid;rewrite sort_keep_correct;trivial.
@@ -657,7 +657,7 @@ Module S.
let len := PArray.length r in
if len == 0 then s
else
- let c := foldi (fun i c' => (C.resolve (get s (r.[i])) c')) 1 (len - 1) (get s (r.[0])) in
+ let c := foldi (fun i c' => (C.resolve (get s (r.[i])) c')) 1 len (get s (r.[0])) in
(* S.set_clause *) internal_set s pos c.
Lemma valid_set_resolve :
@@ -665,13 +665,18 @@ Module S.
forall pos r, valid rho (set_resolve s pos r).
Proof.
unfold set_resolve; intros rho s Hrho Hv pos r.
- destruct (Int63Properties.reflect_eqb (length r) 0);[trivial | ].
+ destruct (reflect_eqb (length r) 0);[trivial | ].
apply valid_internal_set;trivial.
- (* apply S.valid_set_clause; auto. *)
- apply foldi_ind;auto.
- intros i c _ _ Hc. apply C.resolve_correct;auto;apply Hv.
- Qed.
-
+ pattern (length r); apply (int_ind_bounded _ 1).
+ generalize (to_Z_bounded (length r)); rewrite <- to_Z_eq, to_Z_0 in n; rewrite leb_spec, to_Z_1; lia.
+ rewrite foldi_ge; [ apply Hv | reflexivity ].
+ intros i Hi1 Hi2 Hc.
+ rewrite foldi_lt_r.
+ apply C.resolve_correct; [ apply Hv | ring_simplify (i + 1 - 1); exact Hc ].
+ rewrite ltb_spec, to_Z_add_1_wB, to_Z_1.
+ rewrite leb_spec, to_Z_1 in Hi1; generalize (to_Z_bounded i); omega.
+ rewrite ltb_spec in Hi2; generalize (to_Z_bounded (length r)); omega.
+ Qed.
(* Weakening *)
diff --git a/src/Trace.v b/src/Trace.v
index b6715ab..3f61dc7 100644
--- a/src/Trace.v
+++ b/src/Trace.v
@@ -120,37 +120,37 @@ Module Sat_Checker.
Definition dimacs := PArray.array (PArray.array _lit).
Definition C_interp_or rho c :=
- afold_left _ _ false orb (Lit.interp rho) c.
+ afold_left _ false orb (amap (Lit.interp rho) c).
Lemma C_interp_or_spec : forall rho c,
C_interp_or rho c = C.interp rho (to_list c).
Proof.
intros rho c; unfold C_interp_or; case_eq (C.interp rho (to_list c)).
- unfold C.interp; rewrite List.existsb_exists; intros [x [H1 H2]]; destruct (In_to_list _ _ H1) as [i [H3 H4]]; subst x; apply (afold_left_orb_true _ i); auto.
- unfold C.interp; intro H; apply afold_left_orb_false; intros i H1; case_eq (Lit.interp rho (c .[ i])); auto; intro Heq; assert (H2: exists x, List.In x (to_list c) /\ Lit.interp rho x = true).
+ unfold C.interp; rewrite List.existsb_exists; intros [x [H1 H2]]; destruct (In_to_list _ _ H1) as [i [H3 H4]]; subst x; apply (afold_left_orb_true i); rewrite ?length_amap,?get_amap;auto.
+ unfold C.interp; intro H; apply afold_left_orb_false; rewrite length_amap; intros i H1; rewrite get_amap; case_eq (Lit.interp rho (c .[ i])); auto; intro Heq; assert (H2: exists x, List.In x (to_list c) /\ Lit.interp rho x = true).
exists (c.[i]); split; auto; apply to_list_In; auto.
rewrite <- List.existsb_exists in H2; rewrite H2 in H; auto.
Qed.
Definition valid rho (d:dimacs) :=
- afold_left _ _ true andb (C_interp_or rho) d.
+ afold_left _ true andb (amap (C_interp_or rho) d).
Lemma valid_spec : forall rho d,
valid rho d <->
- (forall i : int, i < length d -> C.interp rho (PArray.to_list (d.[i]))).
+ (forall i : int, i < length d -> C.interp rho (to_list (d.[i]))).
Proof.
unfold valid; intros rho d; split; intro H.
intros i Hi; case_eq (C.interp rho (to_list (d .[ i]))); try reflexivity.
- intro Heq; erewrite afold_left_andb_false in H; try eassumption.
+ intro Heq; erewrite afold_left_andb_false in H; rewrite ?length_amap, ?get_amap; try eassumption.
rewrite C_interp_or_spec; auto.
- apply afold_left_andb_true; try assumption; intros i Hi; rewrite C_interp_or_spec; apply H; auto.
+ apply afold_left_andb_true; rewrite length_amap; intros i Hi; rewrite get_amap, C_interp_or_spec by assumption; apply H; auto.
Qed.
Inductive certif :=
| Certif : int -> _trace_ step -> clause_id -> certif.
Definition add_roots s (d:dimacs) :=
- PArray.foldi_right (fun i c s => S.set_clause s i (PArray.to_list c)) d s.
+ foldi (fun i s => S.set_clause s i (to_list (d.[i]))) 0 (length d) s.
Definition checker (d:dimacs) (c:certif) :=
let (nclauses, t, confl_id) := c in
@@ -160,7 +160,8 @@ Qed.
forall d s, valid rho d -> S.valid rho s ->
S.valid rho (add_roots s d).
Proof.
- intros rho Hwr d s Hd Hs; unfold add_roots; apply (PArray.foldi_right_Ind _ _ (fun _ a => S.valid rho a)); auto; intros a i Hlt Hv; apply S.valid_set_clause; auto; rewrite valid_spec in Hd; apply Hd; auto.
+ intros rho Hwr d s Hd Hs; unfold add_roots.
+apply (foldi_ind _ (fun _ a => S.valid rho a)); [ apply leb_0 | | ]; auto; intros i a _ Hle Hv; apply S.valid_set_clause; auto; rewrite valid_spec in Hd; apply Hd; auto.
Qed.
Lemma checker_correct : forall d c,
@@ -233,7 +234,7 @@ Module Cnf_Checker.
Proof.
intros rho rhobv t_form Ht s H; destruct (Form.check_form_correct rho rhobv _ Ht) as [[Ht1 Ht2] Ht3]; intros [pos res|pos cid lf|pos|pos|pos l|pos l|pos l i|pos cid|pos cid|pos cid i]; simpl; try apply S.valid_set_clause; auto.
apply S.valid_set_resolve; auto.
- apply valid_check_flatten; auto; try discriminate; intros a1 a2; unfold is_true; rewrite Int63Properties.eqb_spec; intro; subst a1; auto.
+ apply valid_check_flatten; auto; try discriminate; intros a1 a2; unfold is_true; rewrite Int63.eqb_spec; intro; subst a1; auto.
apply valid_check_True; auto.
apply valid_check_False; auto.
apply valid_check_BuildDef; auto.
@@ -299,7 +300,7 @@ Module Cnf_Checker.
Lit.interp (Form.interp_state_var (PArray.get t_var) (fun _ s => BITVECTOR_LIST.zeros s) t_form) l1 =
Lit.interp (Form.interp_state_var (PArray.get t_var) (fun _ s => BITVECTOR_LIST.zeros s) t_form) l2.
Proof.
- unfold checker_eq; intros t_var t_form l1 l2 l c; rewrite !andb_true_iff; case_eq (t_form .[ Lit.blit l]); [intros _ _|intros _|intros _|intros _ _ _|intros _ _|intros _ _|intros _ _|intros _ _ _|intros l1' l2' Heq|intros _ _ _ _|intros a ls Heq]; intros [[H1 H2] H3]; try discriminate; rewrite andb_true_iff in H2; rewrite !Int63Properties.eqb_spec in H2; destruct H2 as [H2 H4]; subst l1' l2'; case_eq (Lit.is_pos l); intro Heq'; rewrite Heq' in H1; try discriminate; clear H1; assert (H:PArray.default t_form = Form.Ftrue /\ Form.wf t_form).
+ unfold checker_eq; intros t_var t_form l1 l2 l c; rewrite !andb_true_iff; case_eq (t_form .[ Lit.blit l]); [intros _ _|intros _|intros _|intros _ _ _|intros _ _|intros _ _|intros _ _|intros _ _ _|intros l1' l2' Heq|intros _ _ _ _|intros a ls Heq]; intros [[H1 H2] H3]; try discriminate; rewrite andb_true_iff in H2; rewrite !Int63.eqb_spec in H2; destruct H2 as [H2 H4]; subst l1' l2'; case_eq (Lit.is_pos l); intro Heq'; rewrite Heq' in H1; try discriminate; clear H1; assert (H:PArray.default t_form = Form.Ftrue /\ Form.wf t_form).
unfold checker in H3; destruct c as (nclauses, t, confl); rewrite andb_true_iff in H3; destruct H3 as [H3 _]; destruct (Form.check_form_correct (get t_var) (fun _ s => BITVECTOR_LIST.zeros s) _ H3) as [[Ht1 Ht2] Ht3]; split; auto with smtcoq_core.
destruct H as [H1 H2]; case_eq (Lit.interp (Form.interp_state_var (get t_var) (fun _ s => BITVECTOR_LIST.zeros s) t_form) l1); intro Heq1; case_eq (Lit.interp (Form.interp_state_var (get t_var) (fun _ s => BITVECTOR_LIST.zeros s) t_form) l2); intro Heq2; auto with smtcoq_core; elim (checker_correct H3 (rho:=get t_var) (rhobv:=fun _ s => BITVECTOR_LIST.zeros s)); unfold Lit.interp; rewrite Heq'; unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_core; rewrite Heq; simpl; rewrite Heq1, Heq2; auto with smtcoq_core.
Qed.
@@ -502,15 +503,15 @@ Inductive step :=
Definition add_roots s d used_roots :=
match used_roots with
- | Some ur => PArray.foldi_right (fun i c_index s =>
- let c := if c_index < length d then (d.[c_index])::nil else C._true in
- S.set_clause s i c) ur s
- | None => PArray.foldi_right (fun i c s => S.set_clause s i (c::nil)) d s
+ | Some ur => foldi (fun i s =>
+ let c := if (ur.[i]) < length d then (d.[ur.[i]])::nil else C._true in
+ S.set_clause s i c) 0 (length ur) s
+ | None => foldi (fun i s => S.set_clause s i (d.[i]::nil)) 0 (length d) s
end.
Definition valid t_i t_func t_atom t_form d :=
let rho := Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form in
- afold_left _ _ true andb (Lit.interp rho) d.
+ afold_left _ true andb (amap (Lit.interp rho) d).
Lemma add_roots_correct : (* forall t_i t_func t_atom t_form, *)
let rho := Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form in
@@ -519,11 +520,11 @@ Inductive step :=
forall s d used_roots, S.valid rho s -> valid t_func t_atom t_form d ->
S.valid rho (add_roots s d used_roots).
Proof.
- intros (* t_i t_func t_atom t_form *) rho H1 H2 H10 s d used_roots H3; unfold valid; intro H4; pose (H5 := (afold_left_andb_true_inv _ _ _ H4)); unfold add_roots; assert (Valuation.wf rho) by (destruct (Form.check_form_correct (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) _ H1) as [_ H]; auto with smtcoq_core); case used_roots.
- intro ur; apply (foldi_right_Ind _ _ (fun _ a => S.valid rho a)); auto with smtcoq_core; intros a i H6 Ha; apply S.valid_set_clause; auto with smtcoq_core; case_eq (ur .[ i] < length d).
- intro; unfold C.valid; simpl; rewrite H5; auto with smtcoq_core.
+ intros (* t_i t_func t_atom t_form *) rho H1 H2 H10 s d used_roots H3; unfold valid; intro H4; pose (H5 := (afold_left_andb_true_inv _ H4)); unfold add_roots; assert (Valuation.wf rho) by (destruct (Form.check_form_correct (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) _ H1) as [_ H]; auto with smtcoq_core); case used_roots.
+ intro ur; apply (foldi_ind _ (fun _ a => S.valid rho a)); [ apply leb_0 | | ]; auto with smtcoq_core; intros i a _ H6 Ha; apply S.valid_set_clause; auto with smtcoq_core; case_eq (ur .[ i] < length d).
+ intro; unfold C.valid; simpl; specialize (H5 (ur.[i])); rewrite length_amap, get_amap in H5 by assumption; unfold rho; rewrite H5; auto with smtcoq_core.
intros; apply C.interp_true; auto with smtcoq_core.
- apply (foldi_right_Ind _ _ (fun _ a => S.valid rho a)); auto with smtcoq_core; intros a i H6 Ha; apply S.valid_set_clause; auto with smtcoq_core; unfold C.valid; simpl; rewrite H5; auto with smtcoq_core.
+ apply (foldi_ind _ (fun _ a => S.valid rho a)); [ apply leb_0 | | ]; auto with smtcoq_core; intros i a _ H6 Ha; apply S.valid_set_clause; auto with smtcoq_core; unfold C.valid; simpl; specialize (H5 i); rewrite length_amap, get_amap in H5 by assumption; unfold rho; rewrite H5; auto with smtcoq_core.
Qed.
Definition checker (* t_i t_func t_atom t_form *) d used_roots (c:certif) :=
@@ -726,7 +727,7 @@ Inductive step :=
checker_b (* t_func t_atom t_form *) l b c = true ->
Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l = b.
Proof.
- unfold checker_b; intros (* t_i t_func t_atom t_form *) l b (nclauses, t, confl); case b; intros H2; case_eq (Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l); auto with smtcoq_core; intros H1; elim (checker_correct H2); auto with smtcoq_core; unfold valid; apply afold_left_andb_true; intros i Hi; rewrite get_make; auto with smtcoq_core; rewrite Lit.interp_neg, H1; auto with smtcoq_core.
+ unfold checker_b; intros (* t_i t_func t_atom t_form *) l b (nclauses, t, confl); case b; intros H2; case_eq (Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l); auto with smtcoq_core; intros H1; elim (checker_correct H2); auto with smtcoq_core; unfold valid; apply afold_left_andb_true; rewrite length_amap; intros i Hi; rewrite get_amap by assumption; rewrite get_make; auto with smtcoq_core; rewrite Lit.interp_neg, H1; auto with smtcoq_core.
Qed.
Definition checker_eq (* t_i t_func t_atom t_form *) l1 l2 l (c:certif) :=
@@ -743,9 +744,9 @@ Inductive step :=
Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l1 =
Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l2.
Proof.
- unfold checker_eq; intros (* t_i t_func t_atom t_form *) l1 l2 l (nclauses, t, confl); rewrite !andb_true_iff; case_eq (t_form .[ Lit.blit l]); [intros _ _|intros _|intros _|intros _ _ _|intros _ _|intros _ _|intros _ _|intros _ _ _|intros l1' l2' Heq|intros _ _ _ _|intros a ls Heq]; intros [[H1 H2] H3]; try discriminate; rewrite andb_true_iff in H2; rewrite !Int63Properties.eqb_spec in H2; destruct H2 as [H2 H4]; subst l1' l2'; case_eq (Lit.is_pos l); intro Heq'; rewrite Heq' in H1; try discriminate; clear H1; assert (H:PArray.default t_form = Form.Ftrue /\ Form.wf t_form).
+ unfold checker_eq; intros (* t_i t_func t_atom t_form *) l1 l2 l (nclauses, t, confl); rewrite !andb_true_iff; case_eq (t_form .[ Lit.blit l]); [intros _ _|intros _|intros _|intros _ _ _|intros _ _|intros _ _|intros _ _|intros _ _ _|intros l1' l2' Heq|intros _ _ _ _|intros a ls Heq]; intros [[H1 H2] H3]; try discriminate; rewrite andb_true_iff in H2; rewrite !Int63.eqb_spec in H2; destruct H2 as [H2 H4]; subst l1' l2'; case_eq (Lit.is_pos l); intro Heq'; rewrite Heq' in H1; try discriminate; clear H1; assert (H:PArray.default t_form = Form.Ftrue /\ Form.wf t_form).
unfold checker in H3; rewrite !andb_true_iff in H3; destruct H3 as [[[H3 _] _] _]; destruct (Form.check_form_correct (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) _ H3) as [[Ht1 Ht2] Ht3]; split; auto with smtcoq_core.
- destruct H as [H1 H2]; case_eq (Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l1); intro Heq1; case_eq (Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l2); intro Heq2; auto with smtcoq_core; elim (checker_correct H3); unfold valid; apply afold_left_andb_true; intros i Hi; rewrite get_make; unfold Lit.interp; rewrite Heq'; unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_core; rewrite Heq; simpl; rewrite Heq1, Heq2; auto with smtcoq_core.
+ destruct H as [H1 H2]; case_eq (Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l1); intro Heq1; case_eq (Lit.interp (Form.interp_state_var (Atom.interp_form_hatom t_i t_func t_atom) (Atom.interp_form_hatom_bv t_i t_func t_atom) t_form) l2); intro Heq2; auto with smtcoq_core; elim (checker_correct H3); unfold valid; apply afold_left_andb_true; rewrite length_amap; intros i Hi; rewrite get_amap by assumption; rewrite get_make; unfold Lit.interp; rewrite Heq'; unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_core; rewrite Heq; simpl; rewrite Heq1, Heq2; auto with smtcoq_core.
Qed.
diff --git a/src/_CoqProject b/src/_CoqProject
index ebecac9..41e0306 100644
--- a/src/_CoqProject
+++ b/src/_CoqProject
@@ -28,15 +28,9 @@
-I trace
-I verit
-I zchaff
--I Int63
-I Array
-I ../3rdparty/alt-ergo
-Int63/Int63.v
-Int63/Int63Native.v
-Int63/Int63Op.v
-Int63/Int63Axioms.v
-Int63/Int63Properties.v
Array/PArray.v
bva/BVList.v
diff --git a/src/array/Array_checker.v b/src/array/Array_checker.v
index de20723..f95c282 100644
--- a/src/array/Array_checker.v
+++ b/src/array/Array_checker.v
@@ -10,7 +10,7 @@
(**************************************************************************)
-Require Import Bool List Int63 PArray Psatz.
+Require Import Bool List Int63 PArray Psatz ZArith.
Require Import Misc State SMT_terms FArray SMT_classes.
Import Form.
@@ -193,7 +193,7 @@ Section certif.
rho x = Form.interp interp_form_hatom interp_form_hatom_bv t_form (t_form.[ x]).
Proof. intros x;apply wf_interp_form;trivial. Qed.
- Definition wf := PArray.forallbi lt_form t_form.
+ Definition wf := aforallbi lt_form t_form.
Hypothesis wf_t_i : wf.
Notation atom := int (only parsing).
@@ -226,14 +226,14 @@ Section certif.
apply Typ.eqb_spec in Heq6a.
apply Typ.eqb_spec in Heq6b.
apply Typ.eqb_spec in Heq6c.
- apply Int63Properties.eqb_spec in Heq6d.
- apply Int63Properties.eqb_spec in Heq6e.
+ apply Int63.eqb_spec in Heq6d.
+ apply Int63.eqb_spec in Heq6e.
pose proof (rho_interp (Lit.blit lres)) as Hrho.
rewrite Heq2 in Hrho. simpl in Hrho.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a). assert (a < PArray.length t_atom).
apply PArray.get_not_default_lt. rewrite def_t_atom. rewrite Heq3. easy.
@@ -440,7 +440,7 @@ Section certif.
subst t1 t2 t3 t4.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
assert (H15: b1 < PArray.length t_atom).
apply PArray.get_not_default_lt. rewrite def_t_atom. rewrite Heq7. discriminate.
@@ -629,9 +629,9 @@ Section certif.
apply Typ.eqb_spec in Heq11a.
apply Typ.eqb_spec in Heq11b.
- apply Int63Properties.eqb_spec in Heq11c.
+ apply Int63.eqb_spec in Heq11c.
rewrite !andb_true_iff in Heq11d.
- rewrite !Int63Properties.eqb_spec in Heq11d.
+ rewrite !Int63.eqb_spec in Heq11d.
rewrite !Atom.t_interp_wf in isif; trivial.
@@ -785,9 +785,9 @@ Section certif.
apply Typ.eqb_spec in Heq11a.
apply Typ.eqb_spec in Heq11b.
- apply Int63Properties.eqb_spec in Heq11d.
+ apply Int63.eqb_spec in Heq11d.
rewrite !andb_true_iff in Heq11c.
- rewrite !Int63Properties.eqb_spec in Heq11c.
+ rewrite !Int63.eqb_spec in Heq11c.
rewrite !Atom.t_interp_wf in isif; trivial.
@@ -955,26 +955,22 @@ Section certif.
apply Typ.eqb_spec in Heq15.
apply Typ.eqb_spec in Heq15a.
subst t3 t5 t4 t6 t7 t8.
- rewrite !Int63Properties.eqb_spec in Heq1314.
+ rewrite !Int63.eqb_spec in Heq1314.
unfold Lit.interp. rewrite Heq.
unfold Var.interp.
rewrite !wf_interp_form; trivial. rewrite Heq2. simpl.
rewrite afold_left_or.
unfold to_list.
- rewrite Int63Properties.eqb_spec in Heq3.
+ rewrite Int63.eqb_spec in Heq3.
rewrite Heq3.
- (* for native-coq compatibility *)
- assert (0 == 2 = false) as NCC.
- { auto. } rewrite NCC.
(* simpl. *)
- rewrite foldi_down_gt; auto.
+ rewrite foldi_lt_r by reflexivity.
+ rewrite foldi_lt_r by reflexivity.
+ rewrite foldi_ge by reflexivity.
+ change (2 - 1) with 1; change (2 - 1 - 1) with 0.
- (* simpl. *)
- assert (2 - 1 = 1). { auto. }
- rewrite H.
- rewrite foldi_down_eq; auto.
simpl. rewrite orb_false_r.
assert (1 - 1 = 0) as Has2. { auto. }
rewrite Has2.
@@ -985,8 +981,8 @@ Section certif.
pose proof (rho_interp (Lit.blit (a .[ 0]))).
pose proof (rho_interp (Lit.blit (a .[ 1]))).
- rewrite Heq5 in H0. rewrite Heq6 in H1.
- simpl in H0, H1.
+ rewrite Heq5 in H. rewrite Heq6 in H0.
+ simpl in H, H0.
unfold Lit.interp.
rewrite andb_true_iff in Heq4.
destruct Heq4 as (Heq4, Heq4a).
@@ -994,32 +990,32 @@ Section certif.
unfold Lit.interp in Hisa.
rewrite Heq4 in Hisa. unfold Var.interp in Hisa.
- rewrite Hisa in H0. symmetry in H0.
+ rewrite Hisa in H. symmetry in H.
rewrite Heq4a.
unfold Var.interp.
- rewrite H1.
+ rewrite H0.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
(* b *)
- pose proof (H2 b). assert (b < PArray.length t_atom).
+ pose proof (H1 b). assert (b < PArray.length t_atom).
apply PArray.get_not_default_lt. rewrite def_t_atom. rewrite Heq7. easy.
- specialize (H3 H4). simpl in H3.
- rewrite Heq7 in H3. simpl in H3.
- rewrite !andb_true_iff in H3. destruct H3. destruct H3.
- unfold get_type' in H3, H5, H6. unfold v_type in H3, H5, H6.
+ specialize (H2 H3). simpl in H2.
+ rewrite Heq7 in H2. simpl in H2.
+ rewrite !andb_true_iff in H2. destruct H2. destruct H2.
+ unfold get_type' in H2, H4, H5. unfold v_type in H2, H4, H5.
case_eq (t_interp .[ b]).
- intros v_typeb v_valb Htib. rewrite Htib in H3.
+ intros v_typeb v_valb Htib. rewrite Htib in H2.
pose proof Htib as Htib''.
- case_eq v_typeb; intros; rewrite H7 in H3; try now contradict H3.
+ case_eq v_typeb; intros; rewrite H6 in H2; try now contradict H2.
case_eq (t_interp .[ b1]).
- intros v_typeb1 v_valb1 Htib1. rewrite Htib1 in H6.
+ intros v_typeb1 v_valb1 Htib1. rewrite Htib1 in H5.
pose proof Htib1 as Htib1''.
case_eq (t_interp .[ b2]).
- intros v_typeb2 v_valb2 Htib2. rewrite Htib2 in H5.
+ intros v_typeb2 v_valb2 Htib2. rewrite Htib2 in H4.
pose proof Htib2 as Htib2''.
rewrite Atom.t_interp_wf in Htib; trivial.
rewrite Atom.t_interp_wf in Htib1; trivial.
@@ -1028,33 +1024,33 @@ Section certif.
rewrite !Atom.t_interp_wf in Htib; trivial.
rewrite Htib1, Htib2 in Htib.
unfold apply_binop in Htib.
+ apply Typ.eqb_spec in H4.
apply Typ.eqb_spec in H5.
- apply Typ.eqb_spec in H6.
generalize dependent v_valb1. generalize dependent v_valb2.
generalize dependent v_valb.
- rewrite H5, H6, H7. rewrite !Typ.cast_refl. intros.
+ rewrite H4, H5, H6. rewrite !Typ.cast_refl. intros.
specialize (Atom.Bval_inj2 t_i (Typ.Tbool) (Typ.i_eqb t_i t v_valb1 v_valb2) (v_valb)).
- intros. specialize (H8 Htib).
+ intros. specialize (H7 Htib).
(* c *)
- pose proof (H2 c). assert (c < PArray.length t_atom).
+ pose proof (H1 c). assert (c < PArray.length t_atom).
apply PArray.get_not_default_lt. rewrite def_t_atom. rewrite Heq9. easy.
- specialize (H9 H10). simpl in H9.
- rewrite Heq9 in H9. simpl in H9.
- rewrite !andb_true_iff in H9. destruct H9. destruct H9.
- unfold get_type' in H9, H11, H12. unfold v_type in H9, H11, H12.
+ specialize (H8 H9). simpl in H8.
+ rewrite Heq9 in H8. simpl in H8.
+ rewrite !andb_true_iff in H8. destruct H8. destruct H8.
+ unfold get_type' in H8, H10, H11. unfold v_type in H8, H10, H11.
case_eq (t_interp .[ c]).
- intros v_typec v_valc Htic. rewrite Htic in H9.
+ intros v_typec v_valc Htic. rewrite Htic in H8.
pose proof Htic as Htic''.
- case_eq v_typec; intros; rewrite H13 in H9; try now contradict H9.
+ case_eq v_typec; intros; rewrite H12 in H8; try now contradict H8.
case_eq (t_interp .[ c1]).
- intros v_typec1 v_valc1 Htic1. rewrite Htic1 in H12.
+ intros v_typec1 v_valc1 Htic1. rewrite Htic1 in H11.
case_eq (t_interp .[ c2]).
- intros v_typec2 v_valc2 Htic2. rewrite Htic2 in H11.
+ intros v_typec2 v_valc2 Htic2. rewrite Htic2 in H10.
rewrite Atom.t_interp_wf in Htic; trivial.
rewrite Atom.t_interp_wf in Htic1; trivial.
rewrite Atom.t_interp_wf in Htic2; trivial.
@@ -1062,33 +1058,33 @@ Section certif.
rewrite !Atom.t_interp_wf in Htic; trivial.
rewrite Htic1, Htic2 in Htic. simpl in Htic.
- apply Typ.eqb_spec in H11. apply Typ.eqb_spec in H12.
+ apply Typ.eqb_spec in H10. apply Typ.eqb_spec in H11.
generalize dependent v_valc1. generalize dependent v_valc2.
generalize dependent v_valc.
- rewrite H11, H12, H13.
+ rewrite H10, H11, H12.
rewrite !Typ.cast_refl. intros. simpl in Htic.
unfold Bval in Htic.
specialize (Atom.Bval_inj2 t_i (Typ.Tbool) (Typ.i_eqb t_i t2 v_valc1 v_valc2) (v_valc)).
- intros. specialize (H14 Htic).
+ intros. specialize (H13 Htic).
(* c1 *)
- pose proof (H2 c1). assert (c1 < PArray.length t_atom).
+ pose proof (H1 c1). assert (c1 < PArray.length t_atom).
apply PArray.get_not_default_lt. rewrite def_t_atom. rewrite Heq11. easy.
- specialize (H15 H16). simpl in H15.
- rewrite Heq11 in H15. simpl in H15.
- rewrite !andb_true_iff in H15. destruct H15. destruct H15.
- unfold get_type' in H15, H17, H18. unfold v_type in H15, H17, H18.
+ specialize (H14 H15). simpl in H14.
+ rewrite Heq11 in H14. simpl in H14.
+ rewrite !andb_true_iff in H14. destruct H14. destruct H14.
+ unfold get_type' in H14, H16, H17. unfold v_type in H14, H16, H17.
case_eq (t_interp .[ c1]).
- intros v_typec1' v_valc1' Htic1'. rewrite Htic1' in H15.
+ intros v_typec1' v_valc1' Htic1'. rewrite Htic1' in H14.
pose proof Htic1' as Htic1'''.
case_eq (t_interp .[ d1]).
- intros v_typed1 v_vald1 Htid1. rewrite Htid1 in H18.
+ intros v_typed1 v_vald1 Htid1. rewrite Htid1 in H17.
case_eq (t_interp .[ d2]).
- intros v_typed2 v_vald2 Htid2. rewrite Htid2 in H17.
+ intros v_typed2 v_vald2 Htid2. rewrite Htid2 in H16.
rewrite Atom.t_interp_wf in Htic1'; trivial.
rewrite Atom.t_interp_wf in Htid1; trivial.
rewrite Atom.t_interp_wf in Htid2; trivial.
@@ -1096,35 +1092,35 @@ Section certif.
rewrite !Atom.t_interp_wf in Htic1'; trivial.
rewrite Htid1, Htid2 in Htic1'. simpl in Htic1'.
- apply Typ.eqb_spec in H15. apply Typ.eqb_spec in H17.
- apply Typ.eqb_spec in H18.
+ apply Typ.eqb_spec in H14. apply Typ.eqb_spec in H16.
+ apply Typ.eqb_spec in H17.
generalize dependent v_vald1. generalize dependent v_vald2.
generalize dependent v_valc1'.
- rewrite H15, H17, H18.
- unfold Bval. rewrite <- H15.
+ rewrite H14, H16, H17.
+ unfold Bval. rewrite <- H14.
rewrite !Typ.cast_refl. intros.
specialize (Atom.Bval_inj2 t_i t1 (select v_vald1 v_vald2) (v_valc1')).
- intros. specialize (H19 Htic1').
+ intros. specialize (H18 Htic1').
(* c2 *)
- pose proof (H2 c2). assert (c2 < PArray.length t_atom).
+ pose proof (H1 c2). assert (c2 < PArray.length t_atom).
apply PArray.get_not_default_lt. rewrite def_t_atom. rewrite Heq12. easy.
- specialize (H20 H21). simpl in H20.
- rewrite Heq12 in H20. simpl in H20.
- rewrite !andb_true_iff in H20. destruct H20. destruct H20.
- unfold get_type' in H20, H22, H23. unfold v_type in H20, H22, H23.
+ specialize (H19 H20). simpl in H19.
+ rewrite Heq12 in H19. simpl in H19.
+ rewrite !andb_true_iff in H19. destruct H19. destruct H19.
+ unfold get_type' in H19, H21, H22. unfold v_type in H19, H21, H22.
case_eq (t_interp .[ c2]).
- intros v_typec2' v_valc2' Htic2'. rewrite Htic2' in H20.
+ intros v_typec2' v_valc2' Htic2'. rewrite Htic2' in H19.
pose proof Htic2' as Htic2'''.
case_eq (t_interp .[ e1]).
- intros v_typee1 v_vale1 Htie1. rewrite Htie1 in H23.
+ intros v_typee1 v_vale1 Htie1. rewrite Htie1 in H22.
case_eq (t_interp .[ e2]).
- intros v_typee2 v_vale2 Htie2. rewrite Htie2 in H22.
+ intros v_typee2 v_vale2 Htie2. rewrite Htie2 in H21.
pose proof Htie2 as Htie2''.
rewrite Atom.t_interp_wf in Htic2'; trivial.
rewrite Atom.t_interp_wf in Htie1; trivial.
@@ -1133,35 +1129,35 @@ Section certif.
rewrite !Atom.t_interp_wf in Htic2'; trivial.
rewrite Htie1, Htie2 in Htic2'. simpl in Htic2'.
- apply Typ.eqb_spec in H20. apply Typ.eqb_spec in H22.
- apply Typ.eqb_spec in H23.
+ apply Typ.eqb_spec in H19. apply Typ.eqb_spec in H21.
+ apply Typ.eqb_spec in H22.
generalize dependent v_valc1'. generalize dependent v_valc2'.
generalize dependent v_vale1. generalize dependent v_vale2.
- rewrite H22. rewrite H20 in *. rewrite H23.
- unfold Bval. rewrite <- H20.
+ rewrite H21. rewrite H19 in *. rewrite H22.
+ unfold Bval. rewrite <- H19.
rewrite !Typ.cast_refl. intros.
specialize (Atom.Bval_inj2 t_i t1 (select v_vale1 v_vale2) (v_valc2')).
- intros. specialize (H24 Htic2').
+ intros. specialize (H23 Htic2').
(* d2 *)
- pose proof (H2 d2). assert (d2 < PArray.length t_atom).
+ pose proof (H1 d2). assert (d2 < PArray.length t_atom).
apply PArray.get_not_default_lt. rewrite def_t_atom. rewrite Heq14. easy.
- specialize (H25 H26). simpl in H25.
- rewrite Heq14 in H25. simpl in H25.
- rewrite !andb_true_iff in H25. destruct H25. destruct H25.
- unfold get_type' in H25, H27, H28. unfold v_type in H25, H27, H28.
+ specialize (H24 H25). simpl in H24.
+ rewrite Heq14 in H24. simpl in H24.
+ rewrite !andb_true_iff in H24. destruct H24. destruct H24.
+ unfold get_type' in H24, H26, H27. unfold v_type in H24, H26, H27.
case_eq (t_interp .[ d2]).
- intros v_typed2' v_vald2' Htid2'. rewrite Htid2' in H25.
+ intros v_typed2' v_vald2' Htid2'. rewrite Htid2' in H24.
pose proof Htid2' as Htid2'''.
case_eq (t_interp .[ f1]).
- intros v_typef1 v_valf1 Htif1. rewrite Htif1 in H28.
+ intros v_typef1 v_valf1 Htif1. rewrite Htif1 in H27.
case_eq (t_interp .[ f2]).
- intros v_typef2 v_valf2 Htif2. rewrite Htif2 in H27.
+ intros v_typef2 v_valf2 Htif2. rewrite Htif2 in H26.
rewrite Atom.t_interp_wf in Htid2'; trivial.
rewrite Atom.t_interp_wf in Htif1; trivial.
rewrite Atom.t_interp_wf in Htif2; trivial.
@@ -1169,21 +1165,21 @@ Section certif.
rewrite !Atom.t_interp_wf in Htid2'; trivial.
rewrite Htif1, Htif2 in Htid2'. simpl in Htid2'.
- apply Typ.eqb_spec in H25. apply Typ.eqb_spec in H27.
- apply Typ.eqb_spec in H28.
+ apply Typ.eqb_spec in H24. apply Typ.eqb_spec in H26.
+ apply Typ.eqb_spec in H27.
generalize dependent v_valf1. generalize dependent v_valf2.
generalize dependent v_vald2'.
- rewrite H25, H27, H28.
- unfold Bval. rewrite <- H25.
+ rewrite H24, H26, H27.
+ unfold Bval. rewrite <- H24.
rewrite !Typ.cast_refl. intros.
specialize (Atom.Bval_inj2 t_i t0 (@diff _ _
(@EqbToDecType _ (@eqbtype_of_compdec _ (Typ.interp_compdec t_i t0)))
_ _ (Typ.dec_interp t_i v_typec2') _ (Typ.comp_interp t_i v_typec2') (Typ.inh_interp t_i t0) _
v_valf1 v_valf2) (v_vald2')).
- intros. specialize (H29 Htid2').
+ intros. specialize (H28 Htid2').
(* semantics *)
@@ -1278,8 +1274,8 @@ Section certif.
(interp_compdec_aux te)))
| Typ.Tindex i =>
@existT Type (fun ty : Type => CompDec ty)
- (te_carrier (@get typ_compdec t_i i))
- (te_compdec (@get typ_compdec t_i i))
+ (te_carrier (@get typ_compdec t_i (of_Z (Z.of_N i))))
+ (te_compdec (@get typ_compdec t_i (of_Z (Z.of_N i))))
| Typ.TZ =>
@existT Type (fun ty : Type => CompDec ty) BinNums.Z
SMT_classes_instances.Z_compdec
@@ -1339,8 +1335,8 @@ Section certif.
(interp_compdec_aux te)))
| Typ.Tindex i =>
@existT Type (fun ty : Type => CompDec ty)
- (te_carrier (@get typ_compdec t_i i))
- (te_compdec (@get typ_compdec t_i i))
+ (te_carrier (@get typ_compdec t_i (of_Z (Z.of_N i))))
+ (te_compdec (@get typ_compdec t_i (of_Z (Z.of_N i))))
| Typ.TZ =>
@existT Type (fun ty : Type => CompDec ty) BinNums.Z
SMT_classes_instances.Z_compdec
@@ -1358,18 +1354,18 @@ Section certif.
(Typ.ord_interp t_i v_typec1') (Typ.comp_interp t_i v_typec1')
(Typ.inh_interp t_i v_typed2') (Typ.inh_interp t_i v_typec1') v_valf1
v_valf2)).
- intros. specialize (H5 Htid2''').
- rewrite <- H5.
+ intros. specialize (H4 Htid2''').
+ rewrite <- H4.
specialize (Atom.Bval_inj2 t_i v_typed2' (v_vale2) (v_vald2)).
intros.
- unfold Atom.interp_form_hatom, interp_hatom in H0.
- rewrite !Atom.t_interp_wf in H0; trivial.
- rewrite Heq7 in H0. simpl in H0.
- rewrite !Atom.t_interp_wf in H0; trivial.
- rewrite Htib1, Htib2 in H0. simpl in H0.
- rewrite !Typ.cast_refl in H0. simpl in H0.
- apply Typ.i_eqb_spec_false in H0.
+ unfold Atom.interp_form_hatom, interp_hatom in H.
+ rewrite !Atom.t_interp_wf in H; trivial.
+ rewrite Heq7 in H. simpl in H.
+ rewrite !Atom.t_interp_wf in H; trivial.
+ rewrite Htib1, Htib2 in H. simpl in H.
+ rewrite !Typ.cast_refl in H. simpl in H.
+ apply Typ.i_eqb_spec_false in H.
destruct Heq1314 as [Heq1314 | Heq1314];
@@ -1403,7 +1399,7 @@ Section certif.
rewrite (Atom.Bval_inj2 t_i _ _ _ Htid2) in *.
apply select_at_diff.
- red in H0. red. intro. apply H0. auto.
+ red in H. red. intro. apply H. auto.
Qed.
End Correct.
diff --git a/src/bva/BVList.v b/src/bva/BVList.v
index 9e22f98..5c707f6 100644
--- a/src/bva/BVList.v
+++ b/src/bva/BVList.v
@@ -435,7 +435,7 @@ Fixpoint beq_listP (l m : list bool) {struct l} :=
Lemma bv_mk_eq l1 l2 : bv_eq l1 l2 = beq_list l1 l2.
Proof.
unfold bv_eq, size, bits.
- case_eq (Nat_eqb (length l1) (length l2)); intro Heq.
+ case_eq (Nat.eqb (length l1) (length l2)); intro Heq.
- now rewrite (EqNat.beq_nat_true _ _ Heq), N.eqb_refl.
- replace (N.of_nat (length l1) =? N.of_nat (length l2)) with false.
* revert l2 Heq. induction l1 as [ |b1 l1 IHl1]; intros [ |b2 l2]; simpl in *; auto.
@@ -696,7 +696,7 @@ Definition bvmult_bool (a b: list bool) (n: nat) : list bool :=
Definition mult_list a b := bvmult_bool a b (length a).
Definition bv_mult (a b : bitvector) : bitvector :=
- if ((@size a) =? (@size b))
+ if ((@size a) =? (@size b))%N
then mult_list a b
else zeros (@size a).
@@ -710,11 +710,11 @@ Proof. intro n.
Qed.
Definition _of_bits (a:list bool) (s: N) :=
-if (N.of_nat (length a) =? s) then a else zeros s.
+if (N.of_nat (length a) =? s)%N then a else zeros s.
Lemma _of_bits_size l s: (size (_of_bits l s)) = s.
Proof. unfold of_bits, size. unfold _of_bits.
- case_eq ( N.of_nat (length l) =? s).
+ case_eq ( N.of_nat (length l) =? s)%N.
intros. now rewrite N.eqb_eq in H.
intros. unfold zeros. rewrite length_mk_list_false.
now rewrite N2Nat.id.
@@ -800,7 +800,7 @@ Qed.
Lemma bv_eq_reflect a b : bv_eq a b = true <-> a = b.
Proof.
- unfold bv_eq. case_eq (size a =? size b); intro Heq; simpl.
+ unfold bv_eq. case_eq (size a =? size b)%N; intro Heq; simpl.
- apply List_eq.
- split; try discriminate.
intro H. rewrite H, N.eqb_refl in Heq. discriminate.
@@ -1880,11 +1880,11 @@ Qed.
Lemma bv_ult_B2P: forall x y, bv_ult x y = true <-> bv_ultP x y.
Proof. intros. split; intros; unfold bv_ult, bv_ultP in *.
- case_eq (size x =? size y); intros;
+ case_eq (size x =? size y)%N; intros;
rewrite H0 in H; unfold ult_listP. now rewrite H.
now contradict H.
unfold ult_listP in *.
- case_eq (size x =? size y); intros.
+ case_eq (size x =? size y)%N; intros.
rewrite H0 in H.
case_eq (ult_list x y); intros. easy.
rewrite H1 in H. now contradict H.
@@ -1893,11 +1893,11 @@ Qed.
Lemma bv_slt_B2P: forall x y, bv_slt x y = true <-> bv_sltP x y.
Proof. intros. split; intros; unfold bv_slt, bv_sltP in *.
- case_eq (size x =? size y); intros;
+ case_eq (size x =? size y)%N; intros;
rewrite H0 in H; unfold slt_listP. now rewrite H.
now contradict H.
unfold slt_listP in *.
- case_eq (size x =? size y); intros.
+ case_eq (size x =? size y)%N; intros.
rewrite H0 in H.
case_eq (slt_list x y); intros. easy.
rewrite H1 in H. now contradict H.
@@ -1962,28 +1962,28 @@ Qed.
Lemma bv_ult_not_eqP : forall x y, bv_ultP x y -> x <> y.
Proof. intros x y. unfold bv_ultP.
- case_eq (size x =? size y); intros.
+ case_eq (size x =? size y)%N; intros.
- now apply ult_list_not_eqP in H0.
- now contradict H0.
Qed.
Lemma bv_slt_not_eqP : forall x y, bv_sltP x y -> x <> y.
Proof. intros x y. unfold bv_sltP.
- case_eq (size x =? size y); intros.
+ case_eq (size x =? size y)%N; intros.
- now apply slt_list_not_eqP in H0.
- now contradict H0.
Qed.
Lemma bv_ult_not_eq : forall x y, bv_ult x y = true -> x <> y.
Proof. intros x y. unfold bv_ult.
- case_eq (size x =? size y); intros.
+ case_eq (size x =? size y)%N; intros.
- now apply ult_list_not_eq in H0.
- now contradict H0.
Qed.
Lemma bv_slt_not_eq : forall x y, bv_slt x y = true -> x <> y.
Proof. intros x y. unfold bv_slt.
- case_eq (size x =? size y); intros.
+ case_eq (size x =? size y)%N; intros.
- now apply slt_list_not_eq in H0.
- now contradict H0.
Qed.
@@ -2258,7 +2258,7 @@ Qed.
size a = n1 -> size (@bv_extr i n0 n1 a) = n0%N.
Proof.
intros. unfold bv_extr, size in *.
- case_eq (n1 <? n0 + i).
+ case_eq (n1 <? n0 + i)%N.
intros. now rewrite length_mk_list_false, N2Nat.id.
intros.
specialize (@length_extract a i (n0 + i)). intros.
@@ -2455,7 +2455,7 @@ Qed.
(** bit-vector extension *)
Definition bv_shl (a b : bitvector) : bitvector :=
- if ((@size a) =? (@size b))
+ if ((@size a) =? (@size b))%N
then shl_be a b
else zeros (@size a).
@@ -2510,7 +2510,7 @@ Qed.
(** bit-vector extension *)
Definition bv_shr (a b : bitvector) : bitvector :=
- if ((@size a) =? (@size b))
+ if ((@size a) =? (@size b))%N
then shr_be a b
else zeros (@size a).
diff --git a/src/bva/Bva_checker.v b/src/bva/Bva_checker.v
index 0518992..a7743b2 100644
--- a/src/bva/Bva_checker.v
+++ b/src/bva/Bva_checker.v
@@ -12,7 +12,7 @@
(** A small checker for bit-vectors bit-blasting *)
-Require Import Int63 Int63Properties PArray SMT_classes ZArith.
+Require Import Int63 PArray SMT_classes ZArith.
Require Import Misc State SMT_terms BVList Psatz.
Require Import Bool List BoolEq NZParity Nnat.
@@ -87,7 +87,7 @@ Section Checker.
Fixpoint check_bb (a: int) (bs: list _lit) (i n: nat) :=
match bs with
- | nil => Nat_eqb i n (* We go up to n *)
+ | nil => Nat.eqb i n (* We go up to n *)
| b::bs =>
if Lit.is_pos b then
match get_form (Lit.blit b) with
@@ -95,10 +95,10 @@ Section Checker.
match get_atom a' with
| Auop (UO_BVbitOf N p) a' =>
(* TODO:
- Do not need to check [Nat_eqb l (N.to_nat N)] at every iteration *)
+ Do not need to check [Nat.eqb l (N.to_nat N)] at every iteration *)
if (a == a') (* We bit blast the right bv *)
- && (Nat_eqb i p) (* We consider bits after bits *)
- && (Nat_eqb n (N.to_nat N)) (* The bv has indeed type BV n *)
+ && (Nat.eqb i p) (* We consider bits after bits *)
+ && (Nat.eqb n (N.to_nat N)) (* The bv has indeed type BV n *)
then check_bb a bs (S i) n
else false
| _ => false
@@ -262,7 +262,7 @@ Fixpoint check_symopp (bs1 bs2 bsres : list _lit) (bvop: binop) :=
if Lit.is_pos bres then
match get_form (Lit.blit bres) with
| Fand args =>
- match PArray.to_list args with
+ match to_list args with
| bres :: bsres =>
if Lit.is_pos bres then
let ires :=
@@ -530,7 +530,7 @@ Fixpoint check_symopp (bs1 bs2 bsres : list _lit) (bvop: binop) :=
end.
Definition check_mult (bs1 bs2 bsres: list _lit) : bool :=
- if Nat_eqb (length bs1) (length bs2)%nat then
+ if Nat.eqb (length bs1) (length bs2)%nat then
let bvm12 := bblast_bvmult bs1 bs2 (length bs1) in
forallb2 eq_carry_lit bvm12 bsres
else false.
@@ -1080,7 +1080,7 @@ Definition shr_lit_be (a: list _lit) (b: list bool): list _lit :=
Let rho_interp : forall x : int, rho x = Form.interp interp_form_hatom interp_form_hatom_bv t_form (t_form.[ x]).
Proof. intros x;apply wf_interp_form;trivial. Qed.
- Definition wf := PArray.forallbi lt_form t_form.
+ Definition wf := aforallbi lt_form t_form.
Hypothesis wf_t_i : wf.
Variable interp_bvatom : atom -> forall s, BITVECTOR_LIST.bitvector s.
@@ -1125,17 +1125,6 @@ Qed.
Lemma le_le_S_eq : forall (n m: nat), (n <= m)%nat -> (S n <= m)%nat \/ n = m.
Proof le_lt_or_eq.
-Lemma Nat_eqb_eq: forall n m, Nat_eqb n m = true -> n = m.
-Proof. induction n.
- intros n Hm. simpl in Hm. case_eq n. reflexivity.
- intros. rewrite H in Hm. now contradict H.
- intros m Hm.
- case_eq m. intros. rewrite H in Hm. simpl in Hm.
- now contradict Hm.
- intros. rewrite H in Hm. simpl in Hm.
- specialize (@IHn n0 Hm). now rewrite IHn.
-Qed.
-
Lemma diseq_neg_eq: forall (la lb: list bool),
List_diseqb la lb = true -> negb (RAWBITVECTOR_LIST.beq_list la lb) = true.
Proof. intro la.
@@ -1218,7 +1207,7 @@ Proof.
rewrite Typ.N_cast_refl. simpl.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
(* a *)
pose proof (H a).
@@ -1327,7 +1316,7 @@ Proof. intros a bs.
case_eq u; try (intro Heq'; rewrite Heq' in H0; now contradict H0).
intros. rewrite H2 in H0.
- case_eq ((a == i2) && Nat_eqb i n1 && Nat_eqb n (N.to_nat n0)). intros Hif.
+ case_eq ((a == i2) && Nat.eqb i n1 && Nat.eqb n (N.to_nat n0)). intros Hif.
rewrite Hif in H0.
do 2 rewrite andb_true_iff in Hif. destruct Hif as ((Hif0 & Hif1) & Hif2).
specialize (@IHys a (S i) n).
@@ -1362,7 +1351,7 @@ Proof. intros a bs.
rewrite Hif0. rewrite <- Hd.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
assert (i1 < PArray.length t_atom).
{
apply PArray.get_not_default_lt.
@@ -1412,8 +1401,8 @@ Proof. intros a bs.
apply Typ.eqb_spec in H7. inversion H7. easy.
- now apply Nat_eqb_eq in Hif2.
- now apply Nat_eqb_eq in Hif1.
+ now apply Nat.eqb_eq in Hif2.
+ now apply Nat.eqb_eq in Hif1.
lia.
destruct H1.
@@ -1525,7 +1514,7 @@ Proof. unfold Lit.interp.
unfold Var.interp.
destruct wf_rho. simpl. unfold Lit.blit.
cut (0 >> 1 = 0). intros Heq0. rewrite Heq0. exact H.
- now rewrite lsr_0_l.
+ now rewrite lsr0_l.
apply is_even_0.
Qed.
@@ -1759,7 +1748,7 @@ Proof.
+ simpl in Hcheck; now contradict Hcheck.
+ simpl in Hlen. inversion Hlen as [Hlen'].
simpl in Hcheck. rewrite andb_true_iff in Hcheck; destruct Hcheck as (Hcheck1, Hcheck2).
- apply Int63Properties.eqb_spec in Hcheck1; rewrite Hcheck1.
+ apply Int63.eqb_spec in Hcheck1; rewrite Hcheck1.
simpl. rewrite Lit.interp_neg. apply f_equal.
apply IHbs; auto.
Qed.
@@ -1804,7 +1793,7 @@ Proof.
rewrite !andb_true_iff in Hc.
destruct Hc as ((Ha, Hcheck), Hlen).
rewrite N.eqb_eq in Hlen.
- apply Int63Properties.eqb_spec in Ha.
+ apply Int63.eqb_spec in Ha.
generalize (Hs pos).
rewrite Hpos, Hnil.
unfold C.valid, C.interp; simpl; rewrite !orb_false_r.
@@ -1820,7 +1809,7 @@ Proof.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H r).
assert (r < PArray.length t_atom).
@@ -1933,12 +1922,11 @@ Qed.
Lemma to_list_two: forall {A:Type} (a: PArray.array A),
PArray.length a = 2 -> (to_list a) = a .[0] :: a .[1] :: nil.
Proof. intros A a H.
- rewrite to_list_to_list_ntr. unfold to_list_ntr.
+ unfold to_list.
rewrite H.
- cut (0 == 2 = false). intro H1.
- rewrite H1.
- unfold foldi_ntr. rewrite foldi_cont_lt; auto.
- auto.
+ rewrite 2!foldi_lt_r by reflexivity.
+ rewrite foldi_ge by reflexivity.
+ reflexivity.
Qed.
Lemma check_symopp_and: forall ibs1 ibs2 xbs1 ybs2 ibsres zbsres N,
@@ -3236,7 +3224,7 @@ Proof.
rewrite Atom.t_interp_wf; trivial.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -3569,7 +3557,7 @@ Proof.
rewrite Atom.t_interp_wf; trivial.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -3901,7 +3889,7 @@ Proof.
rewrite Atom.t_interp_wf; trivial.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -4529,7 +4517,7 @@ Lemma valid_check_bbEq pos1 pos2 lres : C.valid rho (check_bbEq pos1 pos2 lres).
rewrite wf_interp_form; trivial. rewrite Heq8. simpl.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a3).
assert (a3 < PArray.length t_atom).
@@ -5185,7 +5173,7 @@ Proof.
rewrite wf_interp_form; trivial. rewrite Heq8. simpl.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a3).
assert (a3 < PArray.length t_atom).
@@ -5410,7 +5398,7 @@ Proof.
rewrite wf_interp_form; trivial. rewrite Heq8. simpl.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a3).
assert (a3 < PArray.length t_atom).
@@ -5822,7 +5810,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -6270,7 +6258,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -6302,7 +6290,7 @@ Proof.
do 2 rewrite andb_true_iff in Heq11.
destruct Heq11 as (Heq10, Heq11).
destruct Heq10 as (Heq10a1 & Heq10a2).
- rewrite Int63Properties.eqb_spec in Heq10a1; rewrite Heq10a1 in *.
+ rewrite Int63.eqb_spec in Heq10a1; rewrite Heq10a1 in *.
(* interp_form_hatom_bv a =
interp_bv t_i (interp_atom (t_atom .[a])) *)
@@ -6556,8 +6544,8 @@ Proof. intros bs1.
+ unfold check_mult in H.
now contradict H.
- intros. unfold check_mult in H.
- case_eq (Nat_eqb (Datatypes.length (a :: bs1)) ((Datatypes.length bs2))).
- intros. now apply Nat_eqb_eq in H0.
+ case_eq (Nat.eqb (Datatypes.length (a :: bs1)) ((Datatypes.length bs2))).
+ intros. now apply Nat.eqb_eq in H0.
intros. rewrite H0 in H. now contradict H.
Qed.
@@ -6581,7 +6569,7 @@ Lemma prop_main: forall bs1 bs2 bsres,
map interp_carry (bblast_bvmult bs1 bs2 (Datatypes.length (map (Lit.interp rho) bs1))) =
map (Lit.interp rho) bsres.
Proof. intros. unfold check_mult in H.
- case_eq (Nat_eqb (Datatypes.length bs1) (Datatypes.length bs2)). intros.
+ case_eq (Nat.eqb (Datatypes.length bs1) (Datatypes.length bs2)). intros.
rewrite H0 in H. apply prop_eq_carry_lit2 in H.
rewrite map_length.
now rewrite H.
@@ -6620,7 +6608,7 @@ Proof.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -6925,7 +6913,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -7241,7 +7229,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -7501,7 +7489,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -7747,7 +7735,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -8071,7 +8059,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
@@ -8401,7 +8389,7 @@ Proof.
apply BITVECTOR_LIST.bv_eq_reflect.
generalize wt_t_atom. unfold Atom.wt. unfold is_true.
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
pose proof (H a).
assert (a < PArray.length t_atom).
diff --git a/src/classes/SMT_classes_instances.v b/src/classes/SMT_classes_instances.v
index ae8f9d6..750c506 100644
--- a/src/classes/SMT_classes_instances.v
+++ b/src/classes/SMT_classes_instances.v
@@ -474,7 +474,7 @@ Section Int63.
Local Open Scope int63_scope.
Let int_lt x y :=
- if Int63Native.ltb x y then True else False.
+ if x < y then True else False.
Global Instance int63_ord : OrdType int.
Proof.
@@ -486,13 +486,13 @@ Section Int63.
simpl; try easy.
contradict H1.
rewrite not_false_iff_true.
- rewrite Int63Axioms.ltb_spec in *.
+ rewrite ltb_spec in *.
exact (Z.lt_trans _ _ _ H H0).
- intros x y.
case_eq (x < y); intro; simpl; try easy.
intros.
- rewrite Int63Axioms.ltb_spec in *.
- rewrite <- Int63Properties.to_Z_eq.
+ rewrite ltb_spec in *.
+ rewrite <- Misc.to_Z_eq.
exact (Z.lt_neq _ _ H).
Defined.
@@ -503,19 +503,19 @@ Section Int63.
intros x y.
case_eq (x < y); intro;
case_eq (x == y); intro; unfold lt in *; simpl.
- - rewrite Int63Properties.eqb_spec in H0.
+ - rewrite Int63.eqb_spec in H0.
contradict H0.
assert (int_lt x y). unfold int_lt.
rewrite H; trivial.
remember lt_not_eq. unfold lt in *. simpl in n.
exact (n _ _ H0).
- apply LT. unfold int_lt. rewrite H; trivial.
- - apply EQ. rewrite Int63Properties.eqb_spec in H0; trivial.
+ - apply EQ. rewrite Int63.eqb_spec in H0; trivial.
- apply GT. unfold int_lt.
case_eq (y < x); intro; simpl; try easy.
- specialize (leb_ltb_eqb x y); intro.
+ specialize (Misc.leb_ltb_eqb x y); intro.
contradict H2.
- rewrite leb_negb_gtb. rewrite H1. simpl.
+ rewrite Misc.leb_negb_gtb. rewrite H1. simpl.
red. intro. symmetry in H2.
rewrite orb_true_iff in H2. destruct H2; contradict H2.
rewrite H. auto.
@@ -523,7 +523,7 @@ Section Int63.
Defined.
Global Instance int63_eqbtype : EqbType int :=
- {| eqb := Int63Native.eqb; eqb_spec := Int63Properties.eqb_spec |}.
+ {| eqb := Int63.eqb; eqb_spec := Int63.eqb_spec |}.
Global Instance int63_dec : DecType int := EqbToDecType.
diff --git a/src/cnf/Cnf.v b/src/cnf/Cnf.v
index 9817cd7..c036bc8 100644
--- a/src/cnf/Cnf.v
+++ b/src/cnf/Cnf.v
@@ -23,48 +23,49 @@ Unset Strict Implicit.
Definition or_of_imp args :=
let last := PArray.length args - 1 in
- PArray.mapi (fun i l => if i == last then l else Lit.neg l) args.
+ amapi (fun i l => if i == last then l else Lit.neg l) args.
(* Register or_of_imp as PrimInline. *)
Lemma length_or_of_imp : forall args,
PArray.length (or_of_imp args) = PArray.length args.
-Proof. intro; unfold or_of_imp; apply length_mapi. Qed.
+Proof. intro; unfold or_of_imp; apply length_amapi. Qed.
Lemma get_or_of_imp : forall args i,
i < (PArray.length args) - 1 -> (or_of_imp args).[i] = Lit.neg (args.[i]).
Proof.
unfold or_of_imp; intros args i H; case_eq (0 < PArray.length args).
- intro Heq; rewrite get_mapi.
+ intro Heq; rewrite get_amapi.
replace (i == PArray.length args - 1) with false; auto; symmetry; rewrite eqb_false_spec; intro; subst i; unfold is_true in H; rewrite ltb_spec, (to_Z_sub_1 _ _ Heq) in H; lia.
rewrite ltb_spec; unfold is_true in H; rewrite ltb_spec, (to_Z_sub_1 _ _ Heq) in H; lia.
rewrite ltb_negb_geb; case_eq (PArray.length args <= 0); try discriminate; intros Heq _; assert (H1: PArray.length args = 0).
apply to_Z_inj; rewrite leb_spec in Heq; destruct (to_Z_bounded (PArray.length args)) as [H1 _]; change [|0|] with 0%Z in *; lia.
rewrite !get_outofbound.
- rewrite default_mapi, H1; auto.
+ rewrite default_amapi, H1; auto.
rewrite H1; case_eq (i < 0); auto; intro H2; eelim ltb_0; eassumption.
- rewrite length_mapi, H1; case_eq (i < 0); auto; intro H2; eelim ltb_0; eassumption.
+ rewrite length_amapi, H1; case_eq (i < 0); auto; intro H2; eelim ltb_0; eassumption.
Qed.
Lemma get_or_of_imp2 : forall args i, 0 < PArray.length args ->
i = (PArray.length args) - 1 -> (or_of_imp args).[i] = args.[i].
Proof.
- unfold or_of_imp; intros args i Heq Hi; rewrite get_mapi; subst i.
- rewrite Int63Axioms.eqb_refl; auto.
+ unfold or_of_imp; intros args i Heq Hi; rewrite get_amapi; subst i.
+ rewrite Int63.eqb_refl; auto.
rewrite ltb_spec, (to_Z_sub_1 _ _ Heq); lia.
Qed.
Lemma afold_right_impb p a :
(forall x, p (Lit.neg x) = negb (p x)) ->
(PArray.length a == 0) = false ->
- (afold_right bool int true implb p a) =
+ (afold_right bool true implb (amap p a)) =
List.existsb p (to_list (or_of_imp a)).
Proof.
intros Hp Hl.
- case_eq (afold_right bool int true implb p a); intro Heq; symmetry.
+ case_eq (afold_right bool true implb (amap p a)); intro Heq; symmetry.
- apply afold_right_implb_true_inv in Heq.
+ rewrite length_amap in Heq.
destruct Heq as [Heq|[[i [Hi Heq]]|Heq]].
+ rewrite Heq in Hl. discriminate.
- + rewrite existsb_exists. exists (Lit.neg (a .[ i])). split.
+ + rewrite get_amap in Heq. rewrite existsb_exists. exists (Lit.neg (a .[ i])). split.
* {
apply (to_list_In_eq _ i).
- rewrite length_or_of_imp. apply (ltb_trans _ (PArray.length a - 1)); auto.
@@ -72,6 +73,8 @@ Proof.
- now rewrite get_or_of_imp.
}
* now rewrite Hp, Heq.
+ * apply (ltb_trans _ (PArray.length a - 1)); auto.
+ now apply minus_1_lt.
+ rewrite existsb_exists. exists (a.[(PArray.length a) - 1]). split.
* {
apply (to_list_In_eq _ (PArray.length a - 1)).
@@ -85,15 +88,19 @@ Proof.
clear -H H1. change [|0|] with 0%Z. lia.
}
* {
+ specialize (Heq (PArray.length a - 1)); rewrite get_amap in Heq by now apply minus_1_lt.
apply Heq. now apply minus_1_lt.
}
- apply afold_right_implb_false_inv in Heq.
destruct Heq as [H1 [H2 H3]].
- case_eq (existsb p (to_list (or_of_imp a))); auto.
+ rewrite length_amap in H1, H3.
+ case_eq (List.existsb p (to_list (or_of_imp a))); auto.
rewrite existsb_exists. intros [x [H4 H5]].
apply In_to_list in H4. destruct H4 as [i [H4 ->]].
case_eq (i < PArray.length a - 1); intro Heq.
- + assert (H6 := H2 _ Heq). now rewrite (get_or_of_imp Heq), Hp, H6 in H5.
+ + specialize (H2 i). rewrite length_amap in H2. assert (H6 := H2 Heq). rewrite get_amap in H6.
+ now rewrite (get_or_of_imp Heq), Hp, H6 in H5. apply (ltb_trans _ (PArray.length a - 1)); auto.
+ now apply minus_1_lt.
+ assert (H6:i = PArray.length a - 1).
{
clear -H4 Heq H1.
@@ -106,6 +113,7 @@ Proof.
lia.
}
rewrite get_or_of_imp2 in H5; auto.
+ rewrite get_amap in H3 by now apply minus_1_lt.
rewrite H6, H3 in H5. discriminate.
Qed.
@@ -140,17 +148,17 @@ Section CHECKER.
Definition check_BuildDef l :=
match get_hash (Lit.blit l) with
| Fand args =>
- if Lit.is_pos l then l :: List.map Lit.neg (PArray.to_list args)
+ if Lit.is_pos l then l :: List.map Lit.neg (to_list args)
else C._true
| For args =>
if Lit.is_pos l then C._true
- else l :: PArray.to_list args
+ else l :: to_list args
| Fimp args =>
if Lit.is_pos l then C._true
else if PArray.length args == 0 then C._true
else
let args := or_of_imp args in
- l :: PArray.to_list args
+ l :: to_list args
| Fxor a b =>
if Lit.is_pos l then l::a::Lit.neg b::nil
else l::a::b::nil
@@ -180,15 +188,15 @@ Section CHECKER.
match get_hash (Lit.blit l) with
| Fand args =>
if Lit.is_pos l then C._true
- else List.map Lit.neg (PArray.to_list args)
+ else List.map Lit.neg (to_list args)
| For args =>
- if Lit.is_pos l then PArray.to_list args
+ if Lit.is_pos l then to_list args
else C._true
| Fimp args =>
if PArray.length args == 0 then C._true else
if Lit.is_pos l then
let args := or_of_imp args in
- PArray.to_list args
+ to_list args
else C._true
| Fxor a b =>
if Lit.is_pos l then a::b::nil
@@ -376,12 +384,12 @@ Section CHECKER.
case_eq (i < PArray.length a);intros Hlt;auto using C.interp_true;simpl.
- rewrite Lit.interp_nlit;unfold Var.interp;rewrite rho_interp, orb_false_r, H.
simpl;rewrite afold_left_and.
- case_eq (forallb (Lit.interp rho) (to_list a));trivial.
+ case_eq (List.forallb (Lit.interp rho) (to_list a));trivial.
rewrite forallb_forall;intros Heq;rewrite Heq;trivial.
apply to_list_In; auto.
- rewrite Lit.interp_lit;unfold Var.interp;rewrite rho_interp, orb_false_r, H.
simpl;rewrite afold_left_or.
- unfold C.interp;case_eq (existsb (Lit.interp rho) (to_list a));trivial.
+ unfold C.interp;case_eq (List.existsb (Lit.interp rho) (to_list a));trivial.
rewrite <-not_true_iff_false, existsb_exists, Lit.interp_neg.
case_eq (Lit.interp rho (a .[ i]));trivial.
intros Heq Hex;elim Hex;exists (a.[i]);split;trivial.
@@ -405,7 +413,7 @@ Section CHECKER.
by (intro H; apply Hl; now apply to_Z_inj).
destruct (to_Z_bounded (PArray.length a)) as [H1 _].
lia.
- + now rewrite Int63Properties.eqb_spec in Heq.
+ + now rewrite Int63.eqb_spec in Heq.
}
* now rewrite orb_true_r.
+ rewrite orb_false_r.
@@ -487,7 +495,7 @@ Section CHECKER.
existsb_exists;case_eq (Lit.interp rho (a .[ i]));trivial;
intros Heq2 Hex;elim Hex.
exists (a.[i]);split;trivial.
- assert (H1: 0 < PArray.length a) by (apply (leb_ltb_trans _ i _); auto; apply leb_0); rewrite Int63Properties.eqb_spec in Heq'; rewrite <- (get_or_of_imp2 H1 Heq'); apply to_list_In; rewrite length_or_of_imp; auto.
+ assert (H1: 0 < PArray.length a) by (apply (leb_ltb_trans _ i _); auto; apply leb_0); rewrite Int63.eqb_spec in Heq'; rewrite <- (get_or_of_imp2 H1 Heq'); apply to_list_In; rewrite length_or_of_imp; auto.
exists (Lit.neg (a.[i]));rewrite Lit.interp_neg, Heq2;split;trivial.
assert (H1: i < PArray.length a - 1 = true) by (rewrite ltb_spec, (to_Z_sub_1 _ _ Hlt); rewrite eqb_false_spec in Heq'; assert (H1: [|i|] <> ([|PArray.length a|] - 1)%Z) by (intro H1; apply Heq', to_Z_inj; rewrite (to_Z_sub_1 _ _ Hlt); auto); rewrite ltb_spec in Hlt; lia); rewrite <- (get_or_of_imp H1); apply to_list_In; rewrite length_or_of_imp; auto.
Qed.
diff --git a/src/euf/Euf.v b/src/euf/Euf.v
index eb5ef28..ef8b505 100644
--- a/src/euf/Euf.v
+++ b/src/euf/Euf.v
@@ -214,7 +214,7 @@ Section certif.
case_eq (t_atom.[i]);trivial with smtcoq_euf;intros.
destruct b;trivial with smtcoq_euf.
generalize wt_t_atom;unfold Atom.wt;unfold is_true;
- rewrite PArray.forallbi_spec;intros.
+ rewrite Misc.aforallbi_spec;intros.
assert (i < length t_atom).
apply PArray.get_not_default_lt.
rewrite H0, def_t_atom;discriminate.
@@ -291,7 +291,7 @@ Section certif.
rewrite eqb_spec in H7. rewrite eqb_spec in H8. subst.
tunicity. subst t;rewrite interp_binop_eqb_sym in H1;rewrite H4, H1;auto with smtcoq_euf smtcoq_core.
- apply get_eq_interp;intros.
- destruct (Int63Properties.reflect_eqb t2 b);subst;tunicity; try subst t.
+ destruct (Misc.reflect_eqb t2 b);subst;tunicity; try subst t.
+ apply (IHeqs u);trivial.
simpl;unfold is_true;rewrite orb_true_iff.
rewrite Lit.interp_nlit;unfold Var.interp.
@@ -300,7 +300,7 @@ Section certif.
destruct H1;[left | auto].
apply interp_binop_eqb_trans with (4:= H1);trivial.
rewrite interp_binop_eqb_sym;trivial.
- + destruct (Int63Properties.reflect_eqb t2 a0); subst;tunicity; try subst t.
+ + destruct (Misc.reflect_eqb t2 a0); subst;tunicity; try subst t.
* apply (IHeqs u);trivial.
simpl;unfold is_true;rewrite orb_true_iff.
rewrite Lit.interp_nlit;unfold Var.interp.
@@ -308,7 +308,7 @@ Section certif.
case_eq (rho (Lit.blit a));[rewrite H4; intros | simpl;auto].
destruct H1;[left | auto].
apply interp_binop_eqb_trans with (4:= H1);trivial.
- * destruct (Int63Properties.reflect_eqb t1 b);subst;tunicity; try subst t.
+ * destruct (Misc.reflect_eqb t1 b);subst;tunicity; try subst t.
-- apply (IHeqs u);trivial.
simpl;unfold is_true;rewrite orb_true_iff.
rewrite Lit.interp_nlit;unfold Var.interp.
@@ -316,7 +316,7 @@ Section certif.
case_eq (rho (Lit.blit a));[rewrite H4; intros | simpl;auto].
destruct H1;[left | auto].
apply interp_binop_eqb_trans with (5:= H1);trivial.
- -- destruct (Int63Properties.reflect_eqb t1 a0);[subst;tunicity;try subst t|auto with smtcoq_euf smtcoq_core].
+ -- destruct (Misc.reflect_eqb t1 a0);[subst;tunicity;try subst t|auto with smtcoq_euf smtcoq_core].
apply (IHeqs u);trivial.
simpl;unfold is_true;rewrite orb_true_iff.
rewrite Lit.interp_nlit;unfold Var.interp.
@@ -333,7 +333,7 @@ Section certif.
Proof.
unfold check_trans;intros res [ | leq eqs].
- apply get_eq_interp;intros.
- destruct (Int63Properties.reflect_eqb a b).
+ destruct (Misc.reflect_eqb a b).
+ unfold C.interp; simpl; rewrite orb_false_r.
unfold Lit.interp; simpl; rewrite Lit.is_pos_lit.
unfold Var.interp; simpl; rewrite Lit.blit_lit.
@@ -384,7 +384,7 @@ Section certif.
rewrite HHa, HHb;trivial with smtcoq_euf smtcoq_core.
rewrite eqb_spec in H5. rewrite eqb_spec in H7. subst.
rewrite HHa, HHb;trivial with smtcoq_euf smtcoq_core.
- destruct (Int63Properties.reflect_eqb i i0);[subst | auto with smtcoq_euf smtcoq_core].
+ destruct (Misc.reflect_eqb i i0);[subst | auto with smtcoq_euf smtcoq_core].
apply IHlp;intros;apply H;constructor;auto with smtcoq_euf smtcoq_core.
Qed.
@@ -437,7 +437,7 @@ Section certif.
apply Typ.i_eqb_refl.
intros x y;destruct (Typ.reflect_eqb x y);auto with smtcoq_euf smtcoq_core.
(* op *)
- destruct (Int63Properties.reflect_eqb i i0);[subst | auto with smtcoq_euf smtcoq_core].
+ destruct (Misc.reflect_eqb i i0);[subst | auto with smtcoq_euf smtcoq_core].
apply build_congr_correct;intros.
simpl;rewrite Lit.interp_lit, orb_false_r;unfold Var.interp.
rewrite H1.
@@ -476,7 +476,7 @@ Section certif.
destruct (rho (Lit.blit lpa));reflexivity.
rewrite !wf_interp_form, H, H0;simpl.
generalize wt_t_atom;unfold Atom.wt;unfold is_true;
- rewrite PArray.forallbi_spec;intros.
+ rewrite Misc.aforallbi_spec;intros.
assert (i < length t_atom).
apply PArray.get_not_default_lt.
rewrite H1, def_t_atom;discriminate.
@@ -497,7 +497,7 @@ Section certif.
destruct (rho (Lit.blit lpa));reflexivity.
rewrite !wf_interp_form, H, H0;simpl.
generalize wt_t_atom;unfold Atom.wt;unfold is_true;
- rewrite PArray.forallbi_spec;intros.
+ rewrite Misc.aforallbi_spec;intros.
assert (i < length t_atom).
apply PArray.get_not_default_lt.
rewrite H1, def_t_atom. discriminate.
@@ -512,14 +512,14 @@ Section certif.
apply f_equal; apply f_equal2;trivial with smtcoq_euf smtcoq_core.
(* op *)
- destruct (Int63Properties.reflect_eqb i2 i1);[subst | auto with smtcoq_euf smtcoq_core].
+ destruct (Misc.reflect_eqb i2 i1);[subst | auto with smtcoq_euf smtcoq_core].
apply build_congr_correct;simpl;intros.
rewrite orb_false_r, Lit.interp_lit, Lit.interp_nlit;unfold Var.interp.
replace (rho (Lit.blit lpb)) with (rho (Lit.blit lpa)).
destruct (rho (Lit.blit lpa));reflexivity.
rewrite !wf_interp_form, H, H0;simpl.
generalize wt_t_atom;unfold Atom.wt;unfold is_true;
- rewrite PArray.forallbi_spec;intros.
+ rewrite Misc.aforallbi_spec;intros.
assert (i < length t_atom).
apply PArray.get_not_default_lt.
rewrite H1, def_t_atom;discriminate.
diff --git a/src/lia/Lia.v b/src/lia/Lia.v
index 145acd3..71c4020 100644
--- a/src/lia/Lia.v
+++ b/src/lia/Lia.v
@@ -10,7 +10,7 @@
(**************************************************************************)
-Require Import Bool List Int63 PArray ZArith.
+Require Import Bool List Int63 Ring63 PArray ZArith.
Require Import Misc State SMT_terms Euf.
Require Import RingMicromega ZMicromega Coq.micromega.Tauto Psatz.
@@ -44,10 +44,10 @@ Section certif.
End BuildPositive.
Definition build_positive :=
- foldi_down_cont
+ foldi
(fun i cont h =>
build_positive_atom_aux cont (get_atom h))
- (PArray.length t_atom) 0 (fun _ => None).
+ 0 (PArray.length t_atom) (fun _ => None).
Definition build_positive_atom := build_positive_atom_aux build_positive.
(* Register build_positive_atom as PrimInline. *)
@@ -120,9 +120,9 @@ Section certif.
End BuildPExpr.
Definition build_pexpr :=
- foldi_down_cont
+ foldi
(fun i cont vm h => build_pexpr_atom_aux cont vm (get_atom h))
- (PArray.length t_atom) 0 (fun vm _ => (vm,PEc 0%Z)).
+ 0 (PArray.length t_atom) (fun vm _ => (vm,PEc 0%Z)).
Definition build_pexpr_atom := build_pexpr_atom_aux build_pexpr.
@@ -157,7 +157,7 @@ Section certif.
Section Build_form.
Definition build_not2 i f :=
- fold (fun f' : BFormula (Formula Z) isProp => NOT (NOT f')) 1 i f.
+ foldi (fun _ (f' : BFormula (Formula Z) isProp) => NOT (NOT f')) 0 i f.
Variable build_var : vmap -> var -> option (vmap*(BFormula (Formula Z) isProp)).
@@ -180,23 +180,43 @@ Section certif.
| None => None
end
| Form.Fand args =>
- let n := length args in
- if n == 0 then Some (vm,TT isProp)
- else
- foldi (fun i f1 => match f1 with | Some(vm',f1') => let l := (args.[i]) in match build_var vm' (Lit.blit l) with | Some(vm2,f2) => let f2' := if Lit.is_pos l then f2 else NOT f2 in Some(vm2,AND f1' f2') | None => None end | None => None end) 1 (n-1) (let l := args.[0] in
- match build_var vm (Lit.blit l) with
- | Some (vm',f) => if Lit.is_pos l then Some (vm',f) else Some (vm',NOT f)
+ afold_left _
+ (fun vm => Some (vm, TT isProp))
+ (fun a b vm =>
+ match a vm with
+ | Some (vm1, f1) =>
+ match b vm1 with
+ | Some (vm2, f2) => Some (vm2, AND f1 f2)
+ | None => None
+ end
| None => None
end)
+ (amap
+ (fun l vm => match build_var vm (Lit.blit l) with
+ | Some (vm', f) => Some (vm', if Lit.is_pos l then f else NOT f)
+ | None => None
+ end)
+ args)
+ vm
| Form.For args =>
- let n := length args in
- if n == 0 then Some (vm,FF isProp)
- else
- foldi (fun i f1 => match f1 with | Some(vm',f1') => let l := (args.[i]) in match build_var vm' (Lit.blit l) with | Some(vm2,f2) => let f2' := if Lit.is_pos l then f2 else NOT f2 in Some(vm2,OR f1' f2') | None => None end | None => None end) 1 (n-1) (let l := args.[0] in
- match build_var vm (Lit.blit l) with
- | Some (vm',f) => if Lit.is_pos l then Some (vm',f) else Some (vm',NOT f)
+ afold_left _
+ (fun vm => Some (vm, FF isProp))
+ (fun a b vm =>
+ match a vm with
+ | Some (vm1, f1) =>
+ match b vm1 with
+ | Some (vm2, f2) => Some (vm2, OR f1 f2)
+ | None => None
+ end
| None => None
end)
+ (amap
+ (fun l vm => match build_var vm (Lit.blit l) with
+ | Some (vm', f) => Some (vm', if Lit.is_pos l then f else NOT f)
+ | None => None
+ end)
+ args)
+ vm
| Form.Fxor a b =>
match build_var vm (Lit.blit a) with
| Some (vm1, f1) =>
@@ -210,20 +230,24 @@ Section certif.
| None => None
end
| Form.Fimp args =>
- let n := length args in
- if n == 0 then Some (vm,TT isProp)
- else if n <= 1 then
- let l := args.[0] in
- match build_var vm (Lit.blit l) with
- | Some (vm',f) => if Lit.is_pos l then Some (vm',f) else Some (vm',NOT f)
- | None => None
- end
- else
- foldi_down (fun i f1 => match f1 with | Some(vm',f1') => let l := (args.[i]) in match build_var vm' (Lit.blit l) with | Some(vm2,f2) => let f2' := if Lit.is_pos l then f2 else NOT f2 in Some(vm2,IMPL f2' None f1') | None => None end | None => None end) (n-2) 0 (let l := args.[n-1] in
- match build_var vm (Lit.blit l) with
- | Some (vm',f) => if Lit.is_pos l then Some (vm',f) else Some (vm',NOT f)
+ afold_right _
+ (fun vm => Some (vm, TT isProp))
+ (fun a b vm =>
+ match b vm with
+ | Some (vm2, f2) =>
+ match a vm2 with
+ | Some (vm1, f1) => Some (vm1, IMPL f1 None f2)
+ | None => None
+ end
| None => None
end)
+ (amap
+ (fun l vm => match build_var vm (Lit.blit l) with
+ | Some (vm', f) => Some (vm', if Lit.is_pos l then f else NOT f)
+ | None => None
+ end)
+ args)
+ vm
| Form.Fiff a b =>
match build_var vm (Lit.blit a) with
| Some (vm1, f1) =>
@@ -260,9 +284,9 @@ Section certif.
Definition build_var :=
- foldi_down_cont
+ foldi
(fun i cont vm h => build_hform cont vm (get_form h))
- (PArray.length t_form) 0 (fun _ _ => None).
+ 0 (PArray.length t_form) (fun _ _ => None).
Definition build_form := build_hform build_var.
@@ -418,9 +442,10 @@ Section certif.
t_interp.[h] = Bval t_i Typ.Tpositive p.
Proof.
unfold build_positive.
- apply foldi_down_cont_ind;intros;try discriminate.
+ apply foldi_ind;intros;try discriminate.
+ apply leb_0.
rewrite t_interp_wf;trivial.
- apply build_positive_atom_aux_correct with cont;trivial.
+ apply (build_positive_atom_aux_correct a); trivial.
Qed.
Lemma build_positive_atom_correct :
@@ -870,26 +895,27 @@ Transparent build_z_atom.
t_interp.[h] = Bval t_i Typ.TZ (Zeval_expr (interp_vmap vm') pe).
Proof.
unfold build_pexpr.
- apply foldi_down_cont_ZInd.
- intros z Hz h vm vm' pe Hh.
- assert (W:=to_Z_bounded h);rewrite to_Z_0 in Hz.
+ apply foldi_ind.
+ apply leb_0.
+ intros h vm vm' pe Hh.
+ assert (W:=to_Z_bounded h);rewrite to_Z_0 in Hh.
elimtype False;lia.
intros i cont Hpos Hlen Hrec.
intros h vm vm' pe;unfold is_true;rewrite <-ltb_spec;intros.
rewrite t_interp_wf;trivial.
- apply build_pexpr_atom_aux_correct with cont h i;trivial.
+ apply build_pexpr_atom_aux_correct with cont h (i + 1);trivial.
intros;apply Hrec;auto.
- unfold is_true in H3;rewrite ltb_spec in H, H3;lia.
+ unfold is_true in H3;rewrite ltb_spec in H, H3, Hlen; rewrite to_Z_add_1_wB in H; generalize (to_Z_bounded (length t_atom)); lia.
unfold wf, is_true in wf_t_atom.
- rewrite forallbi_spec in wf_t_atom.
+ rewrite aforallbi_spec in wf_t_atom.
apply wf_t_atom.
- rewrite ltb_spec in H;rewrite leb_spec in Hlen;rewrite ltb_spec;lia.
+ rewrite ltb_spec in H, Hlen;rewrite ltb_spec; rewrite to_Z_add_1_wB in H; generalize (to_Z_bounded (length t_atom)); lia.
unfold wt, is_true in wt_t_atom.
- rewrite forallbi_spec in wt_t_atom.
+ rewrite aforallbi_spec in wt_t_atom.
change (is_true(Typ.eqb (get_type t_i t_func t_atom h) Typ.TZ)) in H0.
rewrite Typ.eqb_spec in H0;rewrite <- H0.
apply wt_t_atom.
- rewrite ltb_spec in H;rewrite leb_spec in Hlen;rewrite ltb_spec;lia.
+ rewrite ltb_spec in H, Hlen; rewrite ltb_spec; rewrite to_Z_add_1_wB in H; generalize (to_Z_bounded (length t_atom)); lia.
Qed.
Lemma build_pexpr_correct :
@@ -914,19 +940,16 @@ Transparent build_z_atom.
rewrite PArray.get_outofbound, default_t_interp.
revert H0.
unfold build_pexpr.
- case_eq (0 < length t_atom);intros Heq.
- rewrite foldi_down_cont_gt;trivial.
- rewrite PArray.get_outofbound;trivial.
+ apply foldi_ind.
+ apply leb_0.
+ discriminate.
+ intros i a _ Hi IH.
+ rewrite PArray.get_outofbound by exact H2.
Opaque build_z_atom.
- rewrite def_t_atom;simpl.
- intros HH H;revert HH H1;apply build_pexpr_atom_aux_correct_z;trivial.
- rewrite foldi_down_cont_eq;trivial.
- rewrite PArray.get_outofbound;trivial.
- rewrite def_t_atom;simpl.
- intros HH H;revert HH H1;apply build_pexpr_atom_aux_correct_z;trivial.
- rewrite <- not_true_iff_false, ltb_spec, to_Z_0 in Heq.
- assert (W:= to_Z_bounded (length t_atom)).
- apply to_Z_inj;rewrite to_Z_0;lia.
+ rewrite def_t_atom; simpl.
+ intros HH H.
+ revert HH H1.
+ apply build_pexpr_atom_aux_correct_z; trivial.
rewrite length_t_interp;trivial.
Qed.
Transparent build_z_atom.
@@ -1020,7 +1043,7 @@ Transparent build_z_atom.
rewrite t_interp_wf;trivial.
intros;apply build_formula_atom_correct with
(get_type t_i t_func t_atom h);trivial.
- unfold wt, is_true in wt_t_atom;rewrite forallbi_spec in wt_t_atom.
+ unfold wt, is_true in wt_t_atom;rewrite aforallbi_spec in wt_t_atom.
case_eq (h < length t_atom);intros Heq;unfold get_type;auto with smtcoq_core.
unfold get_type'.
rewrite !PArray.get_outofbound, default_t_interp, def_t_atom;trivial; try reflexivity.
@@ -1033,28 +1056,32 @@ Transparent build_z_atom.
Lemma build_not2_pos_correct : forall vm (f:GFormula isProp) l i,
bounded_bformula (fst vm) f -> (rho (Lit.blit l) <-> eval_f (Zeval_formula (interp_vmap vm)) f) -> Lit.is_pos l -> bounded_bformula (fst vm) (build_not2 i f) /\ (Form.interp interp_form_hatom interp_form_hatom_bv t_form (Form.Fnot2 i l) <-> eval_f (Zeval_formula (interp_vmap vm)) (build_not2 i f)).
Proof.
- simpl; intros vm f l i H1 H2 H3; split; unfold build_not2.
- apply fold_ind; auto.
- apply (fold_ind2 _ _ (fun b (f':GFormula isProp) => b = true <-> eval_f (Zeval_formula (interp_vmap vm)) f')).
+ simpl; intros vm f l i H1 H2 H3; unfold build_not2.
+ case (Z.le_gt_cases 1 [|i|]); [ intro Hle | intro Hlt ].
+ set (a := foldi _ _ _ _); set (b := foldi _ _ _ _); pattern i, a, b; subst a b; apply foldi_ind2.
+ apply leb_0.
+ unfold Lit.interp; rewrite H3; auto.
+ intros j f' b _ _; rewrite negb_involutive; simpl.
+ intros [ H H' ]; rewrite <- H'.
+ unfold is_true; rewrite not_true_iff_false, not_false_iff_true; tauto.
+ rewrite 2!foldi_ge by (rewrite leb_spec, to_Z_0; lia).
unfold Lit.interp; rewrite H3; auto.
- intros b f' H4; rewrite negb_involutive; simpl; split.
- intros Hb H5; apply H5; rewrite <- H4; auto.
- intro H5; case_eq b; auto; intro H6; elim H5; intro H7; rewrite <- H4 in H7; rewrite H7 in H6; discriminate.
Qed.
Lemma build_not2_neg_correct : forall vm (f:GFormula isProp) l i,
bounded_bformula (fst vm) f -> (rho (Lit.blit l) <-> eval_f (Zeval_formula (interp_vmap vm)) f) -> Lit.is_pos l = false -> bounded_bformula (fst vm) (NOT (build_not2 i f)) /\ (Form.interp interp_form_hatom interp_form_hatom_bv t_form (Form.Fnot2 i l) <-> eval_f (Zeval_formula (interp_vmap vm)) (NOT (build_not2 i f))).
Proof.
- simpl; intros vm f l i H1 H2 H3; split; unfold build_not2.
- apply fold_ind; auto.
- apply (fold_ind2 _ _ (fun b (f':GFormula isProp) => b = true <-> ~ eval_f (Zeval_formula (interp_vmap vm)) f')).
- unfold Lit.interp; rewrite H3; unfold Var.interp; split.
- intros H4 H5; rewrite <- H2 in H5; rewrite H5 in H4; discriminate.
- intro H4; case_eq (rho (Lit.blit l)); auto; intro H5; elim H4; rewrite <- H2; auto.
- intros b f' H4; rewrite negb_involutive; simpl; split.
- intros Hb H5; apply H5; rewrite <- H4; auto.
- intro H5; case_eq b; auto; intro H6; elim H5; intro H7; rewrite <- H4 in H7; rewrite H7 in H6; discriminate.
+ simpl; intros vm f l i H1 H2 H3; unfold build_not2.
+ case (Z.le_gt_cases 1 [|i|]); [ intro Hle | intro Hlt ].
+ set (a := foldi _ _ _ _); set (b := foldi _ _ _ _); pattern i, a, b; subst a b; apply foldi_ind2.
+ apply leb_0.
+ unfold Lit.interp; rewrite H3, <- H2; unfold is_true; rewrite negb_true_iff, not_true_iff_false; tauto.
+ intros j f' b _ _; rewrite negb_involutive; simpl.
+ intros [ H H' ]; rewrite <- H'.
+ unfold is_true; rewrite not_true_iff_false, not_false_iff_true; tauto.
+ rewrite 2!foldi_ge by (rewrite leb_spec, to_Z_0; lia).
+ unfold Lit.interp; rewrite H3, <- H2; unfold is_true; rewrite negb_true_iff, not_true_iff_false; tauto.
Qed.
@@ -1161,58 +1188,60 @@ Transparent build_z_atom.
(* Fnot2 *)
case_eq (build_var vm (Lit.blit l)); try discriminate; intros [vm0 f] Heq H H1; inversion H; subst vm0; subst bf; destruct (Hbv _ _ _ _ Heq H1) as [H2 [H3 [H4 [H5 H6]]]]; do 3 (split; auto); case_eq (Lit.is_pos l); [apply build_not2_pos_correct|apply build_not2_neg_correct]; auto.
(* Fand *)
- simpl; unfold afold_left; case (length l == 0).
+ simpl; unfold afold_left; rewrite !length_amap; case_eq (length l == 0); [ rewrite Int63.eqb_spec | rewrite eqb_false_spec, not_0_ltb ]; intro Hl.
intro H; inversion H; subst vm'; subst bf; simpl; intro H1; split; auto with smtcoq_core; split; [lia| ]; do 3 (split; auto with smtcoq_core).
- revert vm' bf; apply (foldi_ind2 _ _ (fun f1 b => forall vm' (bf:BFormula (Formula Z) isProp), f1 = Some (vm', bf) -> wf_vmap vm -> wf_vmap vm' /\ (Pos.to_nat (fst vm) <= Pos.to_nat (fst vm'))%nat /\ (forall p : positive, (Pos.to_nat p < Pos.to_nat (fst vm))%nat -> nth_error (snd vm) (Pos.to_nat (fst vm - p) - 1) = nth_error (snd vm') (Pos.to_nat (fst vm' - p) - 1)) /\ bounded_bformula (fst vm') bf /\ (b = true <-> eval_f (Zeval_formula (interp_vmap vm')) bf))).
+ revert vm' bf; rewrite !get_amap by exact Hl; set (a := foldi _ _ _ _); set (b := foldi _ _ _ _); pattern (length l), a, b; subst a b; apply foldi_ind2.
+ rewrite ltb_spec, to_Z_0 in Hl; rewrite leb_spec, to_Z_1; lia.
intros vm' bf; case_eq (build_var vm (Lit.blit (l .[ 0]))); try discriminate; intros [vm0 f] Heq; case_eq (Lit.is_pos (l .[ 0])); intros Heq2 H1 H2; inversion H1; subst vm'; subst bf; destruct (Hbv _ _ _ _ Heq H2) as [H10 [H11 [H12 [H13 H14]]]]; do 4 (split; auto); unfold Lit.interp; rewrite Heq2; auto; simpl; split.
intros H3 H4; rewrite <- H14 in H4; rewrite H4 in H3; discriminate.
intro H3; case_eq (Var.interp rho (Lit.blit (l .[ 0]))); auto; intro H4; elim H3; rewrite <- H14; auto.
- intros i a b _ H1; case a; try discriminate; intros [vm0 f0] IH vm' bf; case_eq (build_var vm0 (Lit.blit (l .[ i]))); try discriminate; intros [vm1 f1] Heq H2 H3; inversion H2; subst vm'; subst bf; destruct (IH _ _ (refl_equal (Some (vm0, f0))) H3) as [H5 [H6 [H7 [H8 H9]]]]; destruct (Hbv _ _ _ _ Heq H5) as [H10 [H11 [H12 [H13 H14]]]]; split; auto; split; [eauto with arith| ]; split.
+ intros i a b _ H1; case (a vm); try discriminate; intros [vm0 f0] IH vm' bf; rewrite get_amap by exact H1; case_eq (build_var vm0 (Lit.blit (l .[ i]))); try discriminate; intros [vm1 f1] Heq H2 H3; inversion H2; subst vm'; subst bf; destruct (IH _ _ (refl_equal (Some (vm0, f0))) H3) as [H5 [H6 [H7 [H8 H9]]]]; destruct (Hbv _ _ _ _ Heq H5) as [H10 [H11 [H12 [H13 H14]]]]; split; auto; split; [eauto with arith| ]; split.
intros p H15; rewrite H7; auto; apply H12; eauto with arith.
split.
simpl; rewrite (bounded_bformula_le _ _ H11 _ H8); case (Lit.is_pos (l .[ i])); rewrite H13; auto with smtcoq_core.
- simpl; rewrite (interp_bformula_le _ _ H12 _ H8) in H9; rewrite <- H9; case_eq (Lit.is_pos (l .[ i])); intro Heq2; simpl; rewrite <- H14; unfold Lit.interp; rewrite Heq2; split; case (Var.interp rho (Lit.blit (l .[ i]))); try rewrite andb_true_r; try rewrite andb_false_r; try (intros; split; auto with smtcoq_core); try discriminate; intros [H20 H21]; auto with smtcoq_core.
+ simpl; rewrite (interp_bformula_le _ _ H12 _ H8) in H9; rewrite <- H9; rewrite get_amap by exact H1; case_eq (Lit.is_pos (l .[ i])); intro Heq2; simpl; rewrite <- H14; unfold Lit.interp; rewrite Heq2; split; case (Var.interp rho (Lit.blit (l .[ i]))); try rewrite andb_true_r; try rewrite andb_false_r; try (intros; split; auto with smtcoq_core); try discriminate; intros [H20 H21]; auto with smtcoq_core.
(* For *)
- simpl; unfold afold_left; case (length l == 0).
+ simpl; unfold afold_left; rewrite !length_amap; case_eq (length l == 0); [ rewrite Int63.eqb_spec | rewrite eqb_false_spec, not_0_ltb ]; intro Hl.
intro H; inversion H; subst vm'; subst bf; simpl; intro H1; split; auto with smtcoq_core; split; [lia| ]; do 3 (split; auto with smtcoq_core); discriminate.
- revert vm' bf; apply (foldi_ind2 _ _ (fun f1 b => forall vm' (bf:BFormula (Formula Z) isProp), f1 = Some (vm', bf) -> wf_vmap vm -> wf_vmap vm' /\ (Pos.to_nat (fst vm) <= Pos.to_nat (fst vm'))%nat /\ (forall p : positive, (Pos.to_nat p < Pos.to_nat (fst vm))%nat -> nth_error (snd vm) (Pos.to_nat (fst vm - p) - 1) = nth_error (snd vm') (Pos.to_nat (fst vm' - p) - 1)) /\ bounded_bformula (fst vm') bf /\ (b = true <-> eval_f (Zeval_formula (interp_vmap vm')) bf))).
+ revert vm' bf; rewrite !get_amap by exact Hl; set (a := foldi _ _ _ _); set (b := foldi _ _ _ _); pattern (length l), a, b; subst a b; apply foldi_ind2.
+ rewrite ltb_spec, to_Z_0 in Hl; rewrite leb_spec, to_Z_1; lia.
intros vm' bf; case_eq (build_var vm (Lit.blit (l .[ 0]))); try discriminate; intros [vm0 f] Heq; case_eq (Lit.is_pos (l .[ 0])); intros Heq2 H1 H2; inversion H1; subst vm'; subst bf; destruct (Hbv _ _ _ _ Heq H2) as [H10 [H11 [H12 [H13 H14]]]]; do 4 (split; auto with smtcoq_core); unfold Lit.interp; rewrite Heq2; auto with smtcoq_core; simpl; split.
intros H3 H4; rewrite <- H14 in H4; rewrite H4 in H3; discriminate.
intro H3; case_eq (Var.interp rho (Lit.blit (l .[ 0]))); auto with smtcoq_core; intro H4; elim H3; rewrite <- H14; auto with smtcoq_core.
- intros i a b _ H1; case a; try discriminate; intros [vm0 f0] IH vm' bf; case_eq (build_var vm0 (Lit.blit (l .[ i]))); try discriminate; intros [vm1 f1] Heq H2 H3; inversion H2; subst vm'; subst bf; destruct (IH _ _ (refl_equal (Some (vm0, f0))) H3) as [H5 [H6 [H7 [H8 H9]]]]; destruct (Hbv _ _ _ _ Heq H5) as [H10 [H11 [H12 [H13 H14]]]]; split; auto with smtcoq_core; split; [eauto with smtcoq_core arith| ]; split.
+ intros i a b _ H1; case (a vm); try discriminate; intros [vm0 f0] IH vm' bf; rewrite get_amap by exact H1; case_eq (build_var vm0 (Lit.blit (l .[ i]))); try discriminate; intros [vm1 f1] Heq H2 H3; inversion H2; subst vm'; subst bf; destruct (IH _ _ (refl_equal (Some (vm0, f0))) H3) as [H5 [H6 [H7 [H8 H9]]]]; destruct (Hbv _ _ _ _ Heq H5) as [H10 [H11 [H12 [H13 H14]]]]; split; auto with smtcoq_core; split; [eauto with smtcoq_core arith| ]; split.
intros p H15; rewrite H7; auto with smtcoq_core; apply H12; eauto with smtcoq_core arith.
split.
simpl; rewrite (bounded_bformula_le _ _ H11 _ H8); case (Lit.is_pos (l .[ i])); rewrite H13; auto with smtcoq_core.
- simpl; rewrite (interp_bformula_le _ _ H12 _ H8) in H9; rewrite <- H9; case_eq (Lit.is_pos (l .[ i])); intro Heq2; simpl; rewrite <- H14; unfold Lit.interp; rewrite Heq2; split; case (Var.interp rho (Lit.blit (l .[ i]))); try rewrite orb_false_r; try rewrite orb_true_r; auto with smtcoq_core; try (intros [H20|H20]; auto with smtcoq_core; discriminate); right; intro H20; discriminate.
+ simpl; rewrite (interp_bformula_le _ _ H12 _ H8) in H9; rewrite <- H9; rewrite get_amap by exact H1; case_eq (Lit.is_pos (l .[ i])); intro Heq2; simpl; rewrite <- H14; unfold Lit.interp; rewrite Heq2; split; case (Var.interp rho (Lit.blit (l .[ i]))); try rewrite orb_false_r; try rewrite orb_true_r; auto with smtcoq_core; try (intros [H20|H20]; auto with smtcoq_core; discriminate); right; intro H20; discriminate.
(* Fimp *)
{
- simpl; unfold afold_right; case (length l == 0).
+ simpl; unfold afold_right; rewrite !length_amap; case_eq (length l == 0); [ rewrite Int63.eqb_spec | rewrite eqb_false_spec, not_0_ltb ]; intro Hl.
intro H; inversion H; subst vm'; subst bf; simpl; intro H1; split; auto with smtcoq_core; split; [lia| ]; do 3 (split; auto with smtcoq_core).
- case (length l <= 1).
- case_eq (build_var vm (Lit.blit (l .[ 0]))); try discriminate; intros [vm0 f] Heq; case_eq (Lit.is_pos (l .[ 0])); intros Heq2 H1 H2; inversion H1; subst vm'; subst bf; destruct (Hbv _ _ _ _ Heq H2) as [H3 [H4 [H5 [H6 H7]]]]; do 4 (split; auto with smtcoq_core); unfold Lit.interp; rewrite Heq2; auto with smtcoq_core; simpl; split.
- intros H8 H9; rewrite <- H7 in H9; rewrite H9 in H8; discriminate.
- intro H8; case_eq (Var.interp rho (Lit.blit (l .[ 0]))); auto with smtcoq_core; intro H9; rewrite H7 in H9; elim H8; auto with smtcoq_core.
- revert vm' bf; apply (foldi_down_ind2 _ _ (fun f1 b => forall vm' (bf:BFormula (Formula Z) isProp), f1 = Some (vm', bf) -> wf_vmap vm -> wf_vmap vm' /\ (Pos.to_nat (fst vm) <= Pos.to_nat (fst vm'))%nat /\ (forall p : positive, (Pos.to_nat p < Pos.to_nat (fst vm))%nat -> nth_error (snd vm) (Pos.to_nat (fst vm - p) - 1) = nth_error (snd vm') (Pos.to_nat (fst vm' - p) - 1)) /\ bounded_bformula (fst vm') bf /\ (b = true <-> eval_f (Zeval_formula (interp_vmap vm')) bf))).
+ revert vm' bf; rewrite !get_amap by (apply minus_1_lt; rewrite eqb_false_spec, not_0_ltb; exact Hl); set (a := foldi _ _ _ _); set (b := foldi _ _ _ _); pattern (length l), a, b; subst a b; apply foldi_ind2.
+ rewrite ltb_spec, to_Z_0 in Hl; rewrite leb_spec, to_Z_1; lia.
intros vm' bf; case_eq (build_var vm (Lit.blit (l .[ length l - 1]))); try discriminate; intros [vm0 f] Heq; case_eq (Lit.is_pos (l .[ length l - 1])); intros Heq2 H1 H2; inversion H1; subst vm'; subst bf; destruct (Hbv _ _ _ _ Heq H2) as [H10 [H11 [H12 [H13 H14]]]]; do 4 (split; auto with smtcoq_core); unfold Lit.interp; rewrite Heq2; auto with smtcoq_core; simpl; split.
intros H3 H4; rewrite <- H14 in H4; rewrite H4 in H3; discriminate.
intro H3; case_eq (Var.interp rho (Lit.blit (l .[ length l - 1]))); auto with smtcoq_core; intro H4; elim H3; rewrite <- H14; auto with smtcoq_core.
- intros i a b _ H1; case a; try discriminate; intros [vm0 f0] IH vm' bf; case_eq (build_var vm0 (Lit.blit (l .[ i]))); try discriminate; intros [vm1 f1] Heq H2 H3; inversion H2; subst vm'; subst bf; destruct (IH _ _ (refl_equal (Some (vm0, f0))) H3) as [H5 [H6 [H7 [H8 H9]]]]; destruct (Hbv _ _ _ _ Heq H5) as [H10 [H11 [H12 [H13 H14]]]]; split; auto with smtcoq_core; split; [eauto with smtcoq_core arith| ]; split.
+ intros i a b _ H1.
+ rewrite get_amap by (pose proof (to_Z_bounded i); pose proof (to_Z_bounded (length l)); revert H1 Hl; rewrite !ltb_spec, to_Z_0; intros; rewrite sub_spec, to_Z_sub_1_0, Z.mod_small; lia).
+ rewrite get_amap by (pose proof (to_Z_bounded i); pose proof (to_Z_bounded (length l)); revert H1 Hl; rewrite !ltb_spec, to_Z_0; intros; rewrite sub_spec, to_Z_sub_1_0, Z.mod_small; lia).
+ case a; try discriminate; intros [vm0 f0] IH vm' bf; case_eq (build_var vm0 (Lit.blit (l .[length l - 1 - i]))); try discriminate; intros [vm1 f1] Heq H2 H3; inversion H2; subst vm'; subst bf; destruct (IH _ _ (refl_equal (Some (vm0, f0))) H3) as [H5 [H6 [H7 [H8 H9]]]]; destruct (Hbv _ _ _ _ Heq H5) as [H10 [H11 [H12 [H13 H14]]]]; split; auto with smtcoq_core; split; [eauto with smtcoq_core arith| ]; split.
intros p H15; rewrite H7; auto with smtcoq_core; apply H12; eauto with smtcoq_core arith.
split.
- simpl; rewrite (bounded_bformula_le _ _ H11 _ H8); case (Lit.is_pos (l .[ i])); rewrite H13; auto with smtcoq_core.
+ simpl; rewrite (bounded_bformula_le _ _ H11 _ H8); case (Lit.is_pos (l .[length l - 1 - i])); rewrite H13; auto with smtcoq_core.
simpl; rewrite (interp_bformula_le _ _ H12 _ H8) in H9.
- case_eq (Lit.is_pos (l .[ i])); intro Heq2; simpl.
+ case_eq (Lit.is_pos (l .[length l - 1 - i])); intro Heq2; simpl.
- unfold Lit.interp. rewrite Heq2. split.
- + revert H14. case (Var.interp rho (Lit.blit (l .[ i]))); simpl.
+ + revert H14. case (Var.interp rho (Lit.blit (l .[ length l - 1 - i]))); simpl.
* intros H101 H102 H103. now rewrite <- H9.
* intros H101 H102 H103. rewrite <- H101 in H103. discriminate.
- + revert H14. case (Var.interp rho (Lit.blit (l .[ i]))); simpl; auto.
+ + revert H14. case (Var.interp rho (Lit.blit (l .[ length l - 1 - i]))); simpl; auto.
intros H101 H102. rewrite H9. apply H102. now rewrite <- H101.
- unfold Lit.interp. rewrite Heq2. split.
- + revert H14. case (Var.interp rho (Lit.blit (l .[ i]))); simpl.
+ + revert H14. case (Var.interp rho (Lit.blit (l .[ length l - 1 - i]))); simpl.
* intros H101 H102 H103. elim H103. now rewrite <- H101.
* intros H101 H102 H103. now rewrite <- H9.
- + revert H14. case (Var.interp rho (Lit.blit (l .[ i]))); simpl; auto.
+ + revert H14. case (Var.interp rho (Lit.blit (l .[ length l - 1 - i]))); simpl; auto.
intros H101 H102. rewrite H9. apply H102. now rewrite <- H101.
}
(* Fxor *)
@@ -1270,7 +1299,8 @@ Transparent build_z_atom.
bounded_bformula (fst vm') bf /\
(Var.interp rho v <-> eval_f (Zeval_formula (interp_vmap vm')) bf).
Proof.
- unfold build_var; apply foldi_down_cont_ind; try discriminate.
+ unfold build_var; apply foldi_ind; try discriminate.
+ apply leb_0.
intros i cont _ Hlen Hrec v vm vm' bf; unfold is_true; intros H1 H2; replace (Var.interp rho v) with (Form.interp interp_form_hatom interp_form_hatom_bv t_form (t_form.[v])).
apply (build_hform_correct cont); auto with smtcoq_core.
unfold Var.interp; rewrite <- wf_interp_form; auto with smtcoq_core.
@@ -1418,7 +1448,7 @@ Transparent build_z_atom.
try(case_eq (t_atom.[i]);trivial;intros); try (apply valid_C_true; trivial).
destruct b; try (apply valid_C_true; trivial).
generalize wt_t_atom;unfold Atom.wt;unfold is_true;
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
assert (i < length t_atom).
apply PArray.get_not_default_lt.
rewrite H0, def_t_atom;discriminate.
@@ -1449,7 +1479,7 @@ Transparent build_z_atom.
try(case_eq (t_atom.[i]);trivial;intros); try (apply valid_C_true; trivial).
destruct b; try (apply valid_C_true; trivial).
generalize wt_t_atom;unfold Atom.wt;unfold is_true;
- rewrite PArray.forallbi_spec;intros.
+ rewrite aforallbi_spec;intros.
assert (i < length t_atom).
apply PArray.get_not_default_lt.
rewrite H0, def_t_atom;discriminate.
@@ -1483,8 +1513,8 @@ Transparent build_z_atom.
unfold Atom.interp, Atom.interp_hatom.
rewrite HHa, HHb; simpl.
intros.
- case_eq (va <=? vb); intros; subst.
- case_eq (vb <=? va); intros; subst.
+ case_eq (va <=? vb)%Z; intros; subst.
+ case_eq (vb <=? va)%Z; intros; subst.
apply Zle_bool_imp_le in H2.
apply Zle_bool_imp_le in H3.
apply Z.eqb_neq in H.
@@ -1526,19 +1556,16 @@ Transparent build_z_atom.
case_eq ((a0 == a1) && (a0 == b1) && (b == b0) && (b == a2)); intros; subst;
try (unfold C.valid; apply valid_C_true; trivial).
repeat(apply andb_prop in H19; destruct H19).
- apply Int63Properties.eqb_spec in H19;apply Int63Properties.eqb_spec in H20;apply Int63Properties.eqb_spec in H21;apply Int63Properties.eqb_spec in H22; subst a0 b.
+ apply Int63.eqb_spec in H19;apply Int63.eqb_spec in H20;apply Int63.eqb_spec in H21;apply Int63.eqb_spec in H22; subst a0 b.
unfold C.interp; simpl; rewrite orb_false_r.
unfold Lit.interp; rewrite Lit.is_pos_lit.
unfold Var.interp; rewrite Lit.blit_lit.
rewrite wf_interp_form, H;simpl.
case_eq (Lit.interp rho (a.[0]) || Lit.interp rho (a.[1]) || Lit.interp rho (a.[2])).
intros;repeat (rewrite orb_true_iff in H19);destruct H19. destruct H19.
- apply (afold_left_orb_true int 0); subst; auto with smtcoq_core.
- apply ltb_spec;rewrite H0;compute;trivial.
- apply (afold_left_orb_true int 1); auto with smtcoq_core.
- apply ltb_spec;rewrite H0;compute;trivial.
- apply (afold_left_orb_true int 2); auto with smtcoq_core.
- apply ltb_spec;rewrite H0;compute;trivial.
+ apply (afold_left_orb_true 0); rewrite ?length_amap, ?get_amap; [ rewrite H0; reflexivity | assumption | rewrite H0; reflexivity ].
+ apply (afold_left_orb_true 1); rewrite ?length_amap, ?get_amap; [ rewrite H0; reflexivity | assumption | rewrite H0; reflexivity ].
+ apply (afold_left_orb_true 2); rewrite ?length_amap, ?get_amap; [ rewrite H0; reflexivity | assumption | rewrite H0; reflexivity ].
intros; repeat (rewrite orb_false_iff in H19);destruct H19. destruct H19.
unfold Lit.interp in H19.
rewrite H3 in H19; unfold Var.interp in H19; rewrite H4 in H19.
@@ -1574,19 +1601,16 @@ Transparent build_z_atom.
case_eq ((a0 == b0) && (a0 == a2) && (b == a1) && (b == b1)); intros; subst;
try (unfold C.valid; apply valid_C_true; trivial).
repeat(apply andb_prop in H19; destruct H19).
- apply Int63Properties.eqb_spec in H19;apply Int63Properties.eqb_spec in H20;apply Int63Properties.eqb_spec in H21;apply Int63Properties.eqb_spec in H22;subst a0 b.
+ apply Int63.eqb_spec in H19;apply Int63.eqb_spec in H20;apply Int63.eqb_spec in H21;apply Int63.eqb_spec in H22;subst a0 b.
unfold C.interp; simpl; rewrite orb_false_r.
unfold Lit.interp; rewrite Lit.is_pos_lit.
unfold Var.interp; rewrite Lit.blit_lit.
rewrite wf_interp_form, H;simpl.
case_eq (Lit.interp rho (a.[0]) || Lit.interp rho (a.[1]) || Lit.interp rho (a.[2])).
intros;repeat (rewrite orb_true_iff in H19);destruct H19. destruct H19.
- apply (afold_left_orb_true int 0); auto with smtcoq_core.
- apply ltb_spec;rewrite H0;compute;trivial.
- apply (afold_left_orb_true int 1); auto with smtcoq_core.
- apply ltb_spec;rewrite H0;compute;trivial.
- apply (afold_left_orb_true int 2); auto with smtcoq_core.
- apply ltb_spec;rewrite H0;compute;trivial.
+ apply (afold_left_orb_true 0); rewrite ?length_amap, ?get_amap; [ rewrite H0; reflexivity | assumption | rewrite H0; reflexivity ].
+ apply (afold_left_orb_true 1); rewrite ?length_amap, ?get_amap; [ rewrite H0; reflexivity | assumption | rewrite H0; reflexivity ].
+ apply (afold_left_orb_true 2); rewrite ?length_amap, ?get_amap; [ rewrite H0; reflexivity | assumption | rewrite H0; reflexivity ].
intros; repeat (rewrite orb_false_iff in H19);destruct H19. destruct H19.
unfold Lit.interp in H19.
rewrite H3 in H19; unfold Var.interp in H19; rewrite H4 in H19.
diff --git a/src/spl/Assumptions.v b/src/spl/Assumptions.v
index 32ab634..e3aa767 100644
--- a/src/spl/Assumptions.v
+++ b/src/spl/Assumptions.v
@@ -93,7 +93,7 @@ Section Checker_correct.
Proof.
induction c1 as [ |l1 c1 IHc1]; simpl; intros [ |l2 c2]; simpl; auto; try discriminate.
unfold is_true. rewrite andb_true_iff. intros [H1 H2].
- rewrite Int63Properties.eqb_spec in H1. now rewrite (IHc1 _ H2), H1.
+ rewrite Int63.eqb_spec in H1. now rewrite (IHc1 _ H2), H1.
Qed.
Lemma valid_check_clause cid c :
diff --git a/src/spl/Operators.v b/src/spl/Operators.v
index 966cbcb..7fa0a0c 100644
--- a/src/spl/Operators.v
+++ b/src/spl/Operators.v
@@ -52,7 +52,7 @@ Section Operators.
Fixpoint check_diseqs_complete_aux a dist t :=
match dist with
| nil => true
- | b::q => if PArray.existsb (fun (x:option (int*int)) =>
+ | b::q => if aexistsbi (fun _ (x:option (int*int)) =>
match x with
| Some (a',b') => ((a == a') && (b == b')) || ((a == b') && (b == a'))
| None => false
@@ -68,10 +68,10 @@ Section Operators.
Proof.
intros a dist t; induction dist as [ |b q IHq]; simpl; split; auto.
intros _ y H; inversion H.
- case_eq (PArray.existsb (fun x : option (int * int) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t); try discriminate; rewrite PArray.existsb_spec; intros [i [H1 H2]]; rewrite IHq; clear IHq; intros H3 y [H4|H4]; auto; subst y; exists i; split; auto; generalize H2; case (t .[ i]); try discriminate; intros [a' b']; rewrite orb_true_iff, !andb_true_iff, !Int63Properties.eqb_spec; intros [[H4 H5]|[H4 H5]]; subst a' b'; auto.
- intro H1; case_eq (PArray.existsb (fun x : option (int * int) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t).
+ case_eq (aexistsbi (fun _ (x : option (int * int)) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t); try discriminate; rewrite aexistsbi_spec; intros [i [H1 H2]]; rewrite IHq; clear IHq; intros H3 y [H4|H4]; auto; subst y; exists i; split; auto; generalize H2; case (t .[ i]); try discriminate; intros [a' b']; rewrite orb_true_iff, !andb_true_iff, !Int63.eqb_spec; intros [[H4 H5]|[H4 H5]]; subst a' b'; auto.
+ intro H1; case_eq (aexistsbi (fun _ (x : option (int * int)) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t).
intros _; rewrite IHq; clear IHq; intros y Hy; apply H1; auto.
- rewrite array_existsb_false_spec; destruct (H1 b (or_introl (refl_equal b))) as [i [H2 H3]]; intro H; rewrite <- (H _ H2); destruct H3 as [H3|H3]; rewrite H3; rewrite !eqb_refl; auto; rewrite orb_true_r; auto.
+ rewrite aexistsbi_false_spec; destruct (H1 b (or_introl (refl_equal b))) as [i [H2 H3]]; intro H; rewrite <- (H _ H2); destruct H3 as [H3|H3]; rewrite H3; rewrite !eqb_refl; auto; rewrite orb_true_r; auto.
Qed.
@@ -82,10 +82,10 @@ Section Operators.
Proof.
intros a dist t; induction dist as [ |b q IHq]; simpl; split; try discriminate.
intros [y [H _]]; elim H.
- case_eq (PArray.existsb (fun x : option (int * int) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t).
+ case_eq (aexistsbi (fun _ (x : option (int * int)) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t).
intros _; rewrite IHq; clear IHq; intros [y [H3 H4]]; exists y; auto.
- rewrite array_existsb_false_spec; intros H _; exists b; split; auto; intros i Hi; split; intro H1; generalize (H _ Hi); rewrite H1, !eqb_refl; try discriminate; rewrite orb_true_r; discriminate.
- intros [y [H1 H2]]; case_eq (PArray.existsb (fun x : option (int * int) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t); auto; rewrite PArray.existsb_spec; intros [i [H3 H4]]; rewrite IHq; clear IHq; exists y; destruct H1 as [H1|H1]; auto; subst y; case_eq (t.[i]); [intros [a' b'] Heq|intro Heq]; rewrite Heq in H4; try discriminate; rewrite orb_true_iff, !andb_true_iff, !eqb_spec in H4; destruct H4 as [[H4 H5]|[H4 H5]]; subst a' b'; generalize (H2 _ H3); rewrite Heq; intros [H4 H5]; [elim H4|elim H5]; auto.
+ rewrite aexistsbi_false_spec; intros H _; exists b; split; auto; intros i Hi; split; intro H1; generalize (H _ Hi); rewrite H1, !eqb_refl; try discriminate; rewrite orb_true_r; discriminate.
+ intros [y [H1 H2]]; case_eq (aexistsbi (fun _ (x : option (int * int)) => match x with | Some (a', b') => (a == a') && (b == b') || (a == b') && (b == a') | None => false end) t); auto; rewrite aexistsbi_spec; intros [i [H3 H4]]; rewrite IHq; clear IHq; exists y; destruct H1 as [H1|H1]; auto; subst y; case_eq (t.[i]); [intros [a' b'] Heq|intro Heq]; rewrite Heq in H4; try discriminate; rewrite orb_true_iff, !andb_true_iff, !eqb_spec in H4; destruct H4 as [[H4 H5]|[H4 H5]]; subst a' b'; generalize (H2 _ H3); rewrite Heq; intros [H4 H5]; [elim H4|elim H5]; auto.
Qed.
@@ -129,7 +129,7 @@ Section Operators.
Definition check_diseqs ty dist diseq :=
- let t := PArray.mapi (fun _ t =>
+ let t := amap (fun t =>
if Lit.is_pos t then None else
match get_form (Lit.blit t) with
| Fatom a =>
@@ -144,7 +144,7 @@ Section Operators.
| _ => None
end
) diseq in
- PArray.forallb (fun x => match x with | None => false | _ => true end) t &&
+ aforallbi (fun _ x => match x with | None => false | _ => true end) t &&
check_diseqs_complete dist t.
@@ -166,17 +166,17 @@ Section Operators.
get_atom a = Atom.Abop (Atom.BO_eq A) y x))).
Proof.
intros A dist diseq; unfold check_diseqs; rewrite andb_true_iff,
- PArray.forallb_spec, check_diseqs_complete_spec, length_mapi; split; intros [H1 H2]; split.
- clear H2; intros i Hi; generalize (H1 _ Hi); rewrite get_mapi;
+ aforallbi_spec, check_diseqs_complete_spec, length_amap; split; intros [H1 H2]; split.
+ clear H2; intros i Hi; generalize (H1 _ Hi); rewrite get_amap;
auto; case_eq (Lit.is_pos (diseq .[ i])); try discriminate; intro Heq1; case_eq (get_form (Lit.blit (diseq .[ i])));
try discriminate; intros a Heq2; case_eq (get_atom a); try discriminate; intros [ | | | | | | | B | | | | | | | | | | | | ]; try discriminate; intros h1 h2 Heq3; case_eq (Typ.eqb A B); try discriminate; change (Typ.eqb A B = true) with (is_true (Typ.eqb A B)); rewrite Typ.eqb_spec; intro; subst B; case_eq (h1 == h2); try discriminate; rewrite eqb_false_spec; intro H2; case_eq (check_in h1 dist); try discriminate; case_eq (check_in h2 dist); try discriminate; rewrite !check_in_spec; intros H3 H4 _; split; try discriminate; exists a; split; auto; exists h1, h2; repeat split; auto; rewrite <- In2_In; auto.
- clear H1; intros x y Hxy; destruct (H2 _ _ Hxy) as [i [H1 H4]]; clear H2; rewrite get_mapi in H4; auto; exists i; split; auto; generalize H4;
+ clear H1; intros x y Hxy; destruct (H2 _ _ Hxy) as [i [H1 H4]]; clear H2; rewrite get_amap in H4; auto; exists i; split; auto; generalize H4;
case_eq (Lit.is_pos (diseq .[ i])); intro Heq; try (intros [H|H]; discriminate); case_eq (get_form (Lit.blit (diseq .[ i]))); [intros a| | |intros a1 a2|intros a1|intros a1|intros a1|intros a1 a2|intros a1 a2| intros a1 a2 a3|intros a ls]; intro Heq2; try (intros [H|H]; discriminate); case_eq (get_atom a); [intros a1|intros a1 a2|intros [ | | | | | | | B | | | | | | | | | | | | ] h1 h2|intros a1 a2|intros a1 a2 | intros a1 a2]; intro Heq3; try (intros [H|H]; discriminate); try (case_eq (Typ.eqb A B); try (intros _ [H|H]; discriminate); change (Typ.eqb A B = true) with (is_true (Typ.eqb A B)); rewrite Typ.eqb_spec; intro; subst B; case_eq (h1 == h2); try (intros _ [H|H]; discriminate); rewrite eqb_false_spec; intro H10; case (check_in h1 dist); try (intros [H|H]; discriminate); case (check_in h2 dist); try (intros [H|H]; discriminate); simpl; intro H3; split; try discriminate; exists a; split; auto; destruct H3 as [H3|H3]; inversion H3; subst; auto).
intros. destruct H0; now contradict H0.
- clear H2; intros i Hi; rewrite get_mapi; auto; destruct (H1 _ Hi) as [H2 [a [H3 [h1 [h2 [H4 [H5 H6]]]]]]]; clear H1; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto); rewrite H3, H4, Typ.eqb_refl; simpl; replace (h1 == h2) with false by (case_eq (h1 == h2); auto; rewrite eqb_spec; intro H; elim H5; auto); simpl; rewrite <- In2_In, <- !check_in_spec in H6; auto; destruct H6 as [H6 H7]; rewrite H6, H7; auto.
- clear H1; intros x y Hxy; destruct (H2 _ _ Hxy) as [i [H1 [H3 [a [H4 [H6 H5]]]]]]; clear H2; exists i; split; auto; rewrite get_mapi; auto; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto); rewrite H4; assert (H7 := or_introl (In2 y x dist) Hxy); rewrite <- In2_In, <- !check_in_spec in H7; auto; destruct H7 as [H7 H8]; destruct H5 as [H5|H5]; rewrite H5, Typ.eqb_refl; [replace (x == y) with false by (case_eq (x == y); auto; rewrite eqb_spec; auto)|replace (y == x) with false by (case_eq (y == x); auto; rewrite eqb_spec; auto)]; simpl; rewrite H7, H8; auto.
+ clear H2; intros i Hi; rewrite get_amap; auto; destruct (H1 _ Hi) as [H2 [a [H3 [h1 [h2 [H4 [H5 H6]]]]]]]; clear H1; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto); rewrite H3, H4, Typ.eqb_refl; simpl; replace (h1 == h2) with false by (case_eq (h1 == h2); auto; rewrite eqb_spec; intro H; elim H5; auto); simpl; rewrite <- In2_In, <- !check_in_spec in H6; auto; destruct H6 as [H6 H7]; rewrite H6, H7; auto.
+ clear H1; intros x y Hxy; destruct (H2 _ _ Hxy) as [i [H1 [H3 [a [H4 [H6 H5]]]]]]; clear H2; exists i; split; auto; rewrite get_amap; auto; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto); rewrite H4; assert (H7 := or_introl (In2 y x dist) Hxy); rewrite <- In2_In, <- !check_in_spec in H7; auto; destruct H7 as [H7 H8]; destruct H5 as [H5|H5]; rewrite H5, Typ.eqb_refl; [replace (x == y) with false by (case_eq (x == y); auto; rewrite eqb_spec; auto)|replace (y == x) with false by (case_eq (y == x); auto; rewrite eqb_spec; auto)]; simpl; rewrite H7, H8; auto.
Qed.
@@ -247,7 +247,7 @@ intros. destruct H0; now contradict H0.
get_atom hb = Atom.Abop (Atom.BO_eq ty) y x).
Proof.
intros f1 f2; unfold check_distinct_two_args; split.
- case (get_form f1); try discriminate; intro ha; case (get_form f2); try discriminate; intro hb; case_eq (get_atom ha); try discriminate; intros [A] [ |x [ |y [ |l]]] Heq1; try discriminate; case_eq (get_atom hb); try discriminate; intros [ | | | | | | |B | | | | | | | | | | | | ] x' y' Heq2; try discriminate; rewrite !andb_true_iff, orb_true_iff, !andb_true_iff; change (Typ.eqb A B = true) with (is_true (Typ.eqb A B)); rewrite Typ.eqb_spec, !Int63Properties.eqb_spec; intros [H1 [[H2 H3]|[H2 H3]]]; subst B x' y'; exists ha, hb, A, x, y; auto.
+ case (get_form f1); try discriminate; intro ha; case (get_form f2); try discriminate; intro hb; case_eq (get_atom ha); try discriminate; intros [A] [ |x [ |y [ |l]]] Heq1; try discriminate; case_eq (get_atom hb); try discriminate; intros [ | | | | | | |B | | | | | | | | | | | | ] x' y' Heq2; try discriminate; rewrite !andb_true_iff, orb_true_iff, !andb_true_iff; change (Typ.eqb A B = true) with (is_true (Typ.eqb A B)); rewrite Typ.eqb_spec, !Int63.eqb_spec; intros [H1 [[H2 H3]|[H2 H3]]]; subst B x' y'; exists ha, hb, A, x, y; auto.
intros [ha [hb [A [x [y [H1 [H2 [H3 [H4|H4]]]]]]]]]; rewrite H1, H2, H3, H4, Typ.eqb_refl, !eqb_refl; auto; rewrite orb_true_r; auto.
Qed.
@@ -283,24 +283,24 @@ intros. destruct H0; now contradict H0.
Lemma interp_check_distinct : forall ha diseq,
check_distinct ha diseq = true ->
- interp_form_hatom ha = afold_left bool int true andb (Lit.interp rho) diseq.
+ interp_form_hatom ha = afold_left bool true andb (amap (Lit.interp rho) diseq).
Proof.
intros ha diseq; rewrite check_distinct_spec; intros [A [dist [H1 H2]]]; rewrite check_diseqs_spec in H2; destruct H2 as [H2 H3]; unfold Atom.interp_form_hatom, Atom.interp_bool, Atom.interp_hatom; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op; rewrite H1; simpl; generalize (Atom.compute_interp_spec_rev t_i (get (Atom.t_interp t_i t_func t_atom)) A dist); case (Atom.compute_interp t_i (get (Atom.t_interp t_i t_func t_atom)) A nil); simpl.
intros l H4; case_eq (distinct (Typ.i_eqb t_i A) (rev l)).
- rewrite distinct_spec; intro H5; symmetry; apply afold_left_andb_true; intros i Hi; destruct (H2 _ Hi) as [H9 [a [H10 [h1 [h2 [H6 [H7 H8]]]]]]]; unfold Lit.interp; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto with smtcoq_spl_op smtcoq_core); unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_spl_op smtcoq_core; rewrite H10; simpl; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H6; simpl; unfold Atom.apply_binop; unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite forallbi_spec in wt_t_atom; assert (H11: a < length t_atom).
+ rewrite distinct_spec; intro H5; symmetry; apply afold_left_andb_true; rewrite length_amap; intros i Hi; rewrite get_amap by exact Hi; destruct (H2 _ Hi) as [H9 [a [H10 [h1 [h2 [H6 [H7 H8]]]]]]]; unfold Lit.interp; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto with smtcoq_spl_op smtcoq_core); unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_spl_op smtcoq_core; rewrite H10; simpl; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H6; simpl; unfold Atom.apply_binop; unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite aforallbi_spec in wt_t_atom; assert (H11: a < length t_atom).
case_eq (a < length t_atom); auto with smtcoq_spl_op smtcoq_core; intro H11; rewrite (get_outofbound _ _ _ H11) in H6; rewrite default_t_atom in H6; inversion H6.
generalize (wt_t_atom _ H11); rewrite H6; simpl; rewrite !andb_true_iff; change (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) h1) A = true) with (is_true (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) h1) A)); change (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) h2) A = true) with (is_true (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) h2) A)); rewrite !Typ.eqb_spec; intros [[_ H13] H12]; generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom h1); rewrite H13; intros [v1 HH1]; generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom h2); rewrite H12; intros [v2 HH2]; rewrite HH1, HH2; simpl; rewrite Typ.cast_refl; simpl; destruct H8 as [H8|H8]; [ |rewrite Typ.i_eqb_sym]; rewrite H5; auto with smtcoq_spl_op smtcoq_core; rewrite H4; [exists h2; exists h1|exists h1; exists h2]; auto with smtcoq_spl_op smtcoq_core.
- rewrite distinct_false_spec; intros [v2 [v1 [H5 H6]]]; rewrite H4 in H5; destruct H5 as [a [b [H5 [H7 H8]]]]; clear H4; change (Typ.i_eqb t_i A v2 v1 = true) with (is_true (Typ.i_eqb t_i A v2 v1)) in H6; rewrite Typ.i_eqb_spec in H6; subst v2; clear H2; destruct (H3 _ _ H5) as [i [H2 [H4 [hb [H6 [H9 H10]]]]]]; clear H3; symmetry; apply (afold_left_andb_false _ i); auto with smtcoq_spl_op smtcoq_core; unfold Lit.interp; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto with smtcoq_spl_op smtcoq_core); unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_spl_op smtcoq_core; rewrite H6; simpl; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; destruct H10 as [H10|H10]; rewrite H10; simpl; rewrite H7, H8; simpl; rewrite Typ.cast_refl; simpl; replace (Typ.i_eqb t_i A v1 v1) with true; auto with smtcoq_spl_op smtcoq_core; symmetry; change (is_true (Typ.i_eqb t_i A v1 v1)); rewrite Typ.i_eqb_spec; auto with smtcoq_spl_op smtcoq_core.
+ rewrite distinct_false_spec; intros [v2 [v1 [H5 H6]]]; rewrite H4 in H5; destruct H5 as [a [b [H5 [H7 H8]]]]; clear H4; change (Typ.i_eqb t_i A v2 v1 = true) with (is_true (Typ.i_eqb t_i A v2 v1)) in H6; rewrite Typ.i_eqb_spec in H6; subst v2; clear H2; destruct (H3 _ _ H5) as [i [H2 [H4 [hb [H6 [H9 H10]]]]]]; clear H3; symmetry; apply (afold_left_andb_false i); rewrite ?length_amap; auto with smtcoq_spl_op smtcoq_core; rewrite get_amap by assumption; unfold Lit.interp; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto with smtcoq_spl_op smtcoq_core); unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_spl_op smtcoq_core; rewrite H6; simpl; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; destruct H10 as [H10|H10]; rewrite H10; simpl; rewrite H7, H8; simpl; rewrite Typ.cast_refl; simpl; replace (Typ.i_eqb t_i A v1 v1) with true; auto with smtcoq_spl_op smtcoq_core; symmetry; change (is_true (Typ.i_eqb t_i A v1 v1)); rewrite Typ.i_eqb_spec; auto with smtcoq_spl_op smtcoq_core.
intros [a [H20 H21]]; assert (H4: ha < length t_atom).
case_eq (ha < length t_atom); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H1; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_atom; discriminate.
- unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite forallbi_spec in wt_t_atom; generalize (wt_t_atom _ H4); rewrite H1; simpl; rewrite andb_true_iff, forallb_forall; intros [_ H5]; assert (H6 := H5 _ H20); generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom a); intros [va Ha]; rewrite Ha in H21; simpl in H21; elim H21; apply Typ.eqb_spec; auto with smtcoq_spl_op smtcoq_core.
+ unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite aforallbi_spec in wt_t_atom; generalize (wt_t_atom _ H4); rewrite H1; simpl; rewrite andb_true_iff, forallb_forall; intros [_ H5]; assert (H6 := H5 _ H20); generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom a); intros [va Ha]; rewrite Ha in H21; simpl in H21; elim H21; apply Typ.eqb_spec; auto with smtcoq_spl_op smtcoq_core.
Qed.
Lemma interp_check_distinct_two_args : forall f1 f2,
check_distinct_two_args f1 f2 = true ->
rho f1 = negb (rho f2).
Proof.
- intros f1 f2; rewrite check_distinct_two_args_spec; intros [ha [hb [A [x [y [H1 [H2 [H3 [H4|H4]]]]]]]]]; unfold Form.interp_state_var; assert (H5: f1 < length t_form) by (case_eq (f1 < length t_form); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H1; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_form; discriminate); assert (H6: f2 < length t_form) by (case_eq (f2 < length t_form); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H2; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_form; discriminate); rewrite !Form.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H1, H2; simpl; unfold Atom.interp_form_hatom, Atom.interp_hatom; rewrite !Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H3, H4; simpl; unfold Atom.wt,is_true in wt_t_atom; rewrite forallbi_spec in wt_t_atom; assert (H7: hb < length t_atom) by (case_eq (hb < length t_atom); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H4; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_atom; discriminate); generalize (wt_t_atom _ H7); rewrite H4; simpl; case (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) hb); try discriminate; simpl; rewrite andb_true_iff; change (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) x) A = true) with (is_true (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) x) A)); change (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) y) A = true) with (is_true (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) y) A)); rewrite !Typ.eqb_spec; intros [H8 H9]; generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom x), (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom y); rewrite H8, H9; intros [v1 HH1] [v2 HH2]; rewrite HH1, HH2; simpl; rewrite Typ.cast_refl; auto with smtcoq_spl_op smtcoq_core; rewrite Typ.i_eqb_sym; auto with smtcoq_spl_op smtcoq_core.
+ intros f1 f2; rewrite check_distinct_two_args_spec; intros [ha [hb [A [x [y [H1 [H2 [H3 [H4|H4]]]]]]]]]; unfold Form.interp_state_var; assert (H5: f1 < length t_form) by (case_eq (f1 < length t_form); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H1; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_form; discriminate); assert (H6: f2 < length t_form) by (case_eq (f2 < length t_form); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H2; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_form; discriminate); rewrite !Form.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H1, H2; simpl; unfold Atom.interp_form_hatom, Atom.interp_hatom; rewrite !Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H3, H4; simpl; unfold Atom.wt,is_true in wt_t_atom; rewrite aforallbi_spec in wt_t_atom; assert (H7: hb < length t_atom) by (case_eq (hb < length t_atom); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H4; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_atom; discriminate); generalize (wt_t_atom _ H7); rewrite H4; simpl; case (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) hb); try discriminate; simpl; rewrite andb_true_iff; change (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) x) A = true) with (is_true (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) x) A)); change (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) y) A = true) with (is_true (Typ.eqb (Atom.get_type' t_i (Atom.t_interp t_i t_func t_atom) y) A)); rewrite !Typ.eqb_spec; intros [H8 H9]; generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom x), (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom y); rewrite H8, H9; intros [v1 HH1] [v2 HH2]; rewrite HH1, HH2; simpl; rewrite Typ.cast_refl; auto with smtcoq_spl_op smtcoq_core; rewrite Typ.i_eqb_sym; auto with smtcoq_spl_op smtcoq_core.
Qed.
@@ -309,12 +309,12 @@ intros. destruct H0; now contradict H0.
(* interp_form_hatom ha -> afold_left bool int true andb (Lit.interp rho) diseq. *)
(* Proof. *)
(* intros ha diseq; rewrite check_distinct_spec; intros [A [dist [H1 H]]]; rewrite check_diseqs_spec in H; unfold Atom.interp_form_hatom, Atom.interp_bool, Atom.interp_hatom; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H1; simpl; generalize (Atom.compute_interp_spec_rev t_i (get (Atom.t_interp t_i t_func t_atom)) A dist); case (Atom.compute_interp t_i (get (Atom.t_interp t_i t_func t_atom)) A nil); simpl. *)
- (* intros l H2; unfold is_true; rewrite distinct_spec; intro H3; apply afold_left_andb_true; intros i Hi; destruct (H _ Hi) as [H4 [a [H5 [h1 [h2 [H6 [H7 H8]]]]]]]; unfold Lit.interp; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto with smtcoq_spl_op smtcoq_core); unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_spl_op smtcoq_core; rewrite H5; simpl; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H6; simpl; unfold Atom.apply_binop; unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite forallbi_spec in wt_t_atom; assert (H10: a < length t_atom). *)
+ (* intros l H2; unfold is_true; rewrite distinct_spec; intro H3; apply afold_left_andb_true; intros i Hi; destruct (H _ Hi) as [H4 [a [H5 [h1 [h2 [H6 [H7 H8]]]]]]]; unfold Lit.interp; replace (Lit.is_pos (diseq .[ i])) with false by (case_eq (Lit.is_pos (diseq .[ i])); auto with smtcoq_spl_op smtcoq_core); unfold Var.interp; rewrite Form.wf_interp_form; auto with smtcoq_spl_op smtcoq_core; rewrite H5; simpl; rewrite Atom.t_interp_wf; auto with smtcoq_spl_op smtcoq_core; rewrite H6; simpl; unfold Atom.apply_binop; unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite aforallbi_spec in wt_t_atom; assert (H10: a < length t_atom). *)
(* case_eq (a < length t_atom); auto with smtcoq_spl_op smtcoq_core; intro H10; rewrite (get_outofbound _ _ _ H10) in H6; rewrite default_t_atom in H6; inversion H6. *)
(* generalize (wt_t_atom _ H10); rewrite H6; simpl; rewrite !andb_true_iff. change (Typ.eqb (Atom.get_type t_i t_func t_atom h1) A = true) with (is_true (Typ.eqb (Atom.get_type t_i t_func t_atom h1) A)); change (Typ.eqb (Atom.get_type t_i t_func t_atom h2) A = true) with (is_true (Typ.eqb (Atom.get_type t_i t_func t_atom h2) A)); rewrite !Typ.eqb_spec; intros [[_ H11] H12]; generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom h1); rewrite H11; intros [v1 HH1]; generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom h2); rewrite H12; intros [v2 HH2]; rewrite HH1, HH2; simpl; rewrite Typ.cast_refl; simpl; destruct H8 as [H8|H8]; [ |rewrite Typ.i_eqb_sym]; rewrite H3; auto with smtcoq_spl_op smtcoq_core; rewrite H2; [exists h2; exists h1|exists h1; exists h2]; auto with smtcoq_spl_op smtcoq_core. *)
(* intros [a [H2 H3]] _; assert (H4: ha < length t_atom). *)
(* case_eq (ha < length t_atom); auto with smtcoq_spl_op smtcoq_core; intro Heq; generalize H1; rewrite get_outofbound; auto with smtcoq_spl_op smtcoq_core; rewrite default_t_atom; discriminate. *)
- (* unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite forallbi_spec in wt_t_atom; generalize (wt_t_atom _ H4); rewrite H1; simpl; rewrite andb_true_iff, forallb_forall; intros [_ H5]; assert (H6 := H5 _ H2); generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom a); intros [va Ha]; rewrite Ha in H3; simpl in H3; elim H3; apply Typ.eqb_spec; auto with smtcoq_spl_op smtcoq_core. *)
+ (* unfold Atom.wt in wt_t_atom; unfold is_true in wt_t_atom; rewrite aforallbi_spec in wt_t_atom; generalize (wt_t_atom _ H4); rewrite H1; simpl; rewrite andb_true_iff, forallb_forall; intros [_ H5]; assert (H6 := H5 _ H2); generalize (Atom.check_aux_interp_hatom _ t_func _ wf_t_atom a); intros [va Ha]; rewrite Ha in H3; simpl in H3; elim H3; apply Typ.eqb_spec; auto with smtcoq_spl_op smtcoq_core. *)
(* Qed. *)
End Valid1.
@@ -337,10 +337,10 @@ intros. destruct H0; now contradict H0.
| Ftrue, Ftrue => true
| Ffalse, Ffalse => true
| Fnot2 i1 l1, Fnot2 i2 l2 => (i1 == i2) && (check_lit l1 l2)
- | Fand a1, Fand a2 => (length a1 == length a2) && (forallbi (fun i l => check_lit l (a2.[i])) a1)
- | For a1, For a2 => (length a1 == length a2) && (forallbi (fun i l => check_lit l (a2.[i])) a1)
- | Fimp a1, Fimp a2 => (length a1 == length a2) && (forallbi (fun i l => check_lit l (a2.[i])) a1)
- (* (length a1 == length a2) && (forallbi (fun i l => if i < length a1 - 1 then check_lit (a2.[i]) l else check_lit l (a2.[i])) a1) *)
+ | Fand a1, Fand a2 => (length a1 == length a2) && (aforallbi (fun i l => check_lit l (a2.[i])) a1)
+ | For a1, For a2 => (length a1 == length a2) && (aforallbi (fun i l => check_lit l (a2.[i])) a1)
+ | Fimp a1, Fimp a2 => (length a1 == length a2) && (aforallbi (fun i l => check_lit l (a2.[i])) a1)
+ (* (length a1 == length a2) && (aforallbi (fun i l => if i < length a1 - 1 then check_lit (a2.[i]) l else check_lit l (a2.[i])) a1) *)
| Fxor l1 l2, Fxor j1 j2 => check_lit l1 j1 && check_lit l2 j2
(* check_lit l1 j1 && check_lit j1 l1 && check_lit l2 j2 && check_lit j2 l2 *)
(* (* let a := check_lit l1 j1 in *) *)
@@ -391,7 +391,7 @@ intros. destruct H0; now contradict H0.
(* check_lit l1 l2 -> Lit.interp rho l1 -> Lit.interp rho l2 = true. *)
(* Proof. *)
(* unfold check_lit; intros l1 l2; unfold is_true; rewrite !orb_true_iff, !andb_true_iff; intros [[H1|[[H1 H2] H3]]|[[H1 H2] H3]]. *)
- (* rewrite Int63Properties.eqb_spec in H1; subst l1; auto with smtcoq_core. *)
+ (* rewrite Int63.eqb_spec in H1; subst l1; auto with smtcoq_core. *)
(* unfold Lit.interp; rewrite H1, H2; apply interp_check_var; auto with smtcoq_core. *)
(* unfold Lit.interp; case_eq (Lit.is_pos l1); intro Heq; rewrite Heq in H1; try discriminate; clear Heq H1; case_eq (Lit.is_pos l2); intro Heq; rewrite Heq in H2; try discriminate; clear Heq H2; case_eq (Var.interp rho (Lit.blit l1)); try discriminate; intros H4 _; case_eq (Var.interp rho (Lit.blit l2)); auto with smtcoq_core; intro H5; rewrite (interp_check_var _ _ H3 H5) in H4; discriminate. *)
(* Qed. *)
@@ -404,17 +404,17 @@ intros. destruct H0; now contradict H0.
Proof.
intros [a| | |i1 l1|a1|a1|a1|l1 l2|l1 l2|l1 l2 l3|a l1] [b| | |j1 m1|a2|a2|a2|j1 j2|j1 j2|j1 j2 j3|b m1]; simpl; try discriminate;auto with smtcoq_core.
(* Atom *)
- unfold is_true; rewrite Int63Properties.eqb_spec; intro; subst a; auto with smtcoq_core.
+ unfold is_true; rewrite Int63.eqb_spec; intro; subst a; auto with smtcoq_core.
(* Interesting case *)
apply interp_check_distinct; auto with smtcoq_core.
(* Double negation *)
- unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H1 H2]; subst j1. rewrite (interp_check_lit _ _ H2). auto with smtcoq_core.
+ unfold is_true; rewrite andb_true_iff, Int63.eqb_spec; intros [H1 H2]; subst j1. rewrite (interp_check_lit _ _ H2). auto with smtcoq_core.
(* Conjunction *)
- unfold is_true; rewrite andb_true_iff, eqb_spec, forallbi_spec; intros [H1 H2]; apply afold_left_eq; auto with smtcoq_core; intros i Hi; apply interp_check_lit; auto with smtcoq_core.
+ unfold is_true; rewrite andb_true_iff, eqb_spec, aforallbi_spec; intros [H1 H2]; apply afold_left_eq; rewrite ?length_amap; auto with smtcoq_core; intros i Hi; rewrite 2!get_amap by (rewrite <- ?H1; assumption); apply interp_check_lit; auto with smtcoq_core.
(* Disjunction *)
- unfold is_true; rewrite andb_true_iff, eqb_spec, forallbi_spec; intros [H1 H2]; apply afold_left_eq; auto with smtcoq_core; intros i Hi; apply interp_check_lit; auto with smtcoq_core.
+ unfold is_true; rewrite andb_true_iff, eqb_spec, aforallbi_spec; intros [H1 H2]; apply afold_left_eq; rewrite ?length_amap; auto with smtcoq_core; intros i Hi; rewrite 2!get_amap by (rewrite <- ?H1; assumption); apply interp_check_lit; auto with smtcoq_core.
(* Implication *)
- unfold is_true; rewrite andb_true_iff, eqb_spec, forallbi_spec; intros [H1 H2]; apply afold_right_eq; auto with smtcoq_core; intros i Hi; apply interp_check_lit; auto with smtcoq_core.
+ unfold is_true; rewrite andb_true_iff, eqb_spec, aforallbi_spec; intros [H1 H2]; apply afold_right_eq; rewrite ?length_amap; auto with smtcoq_core; intros i Hi; rewrite 2!get_amap by (rewrite <- ?H1; assumption); apply interp_check_lit; auto with smtcoq_core.
(* Xor *)
unfold is_true; rewrite andb_true_iff; intros [H1 H2]; rewrite (interp_check_lit _ _ H1), (interp_check_lit _ _ H2); auto with smtcoq_core.
(* Iff *)
@@ -444,22 +444,22 @@ intros. destruct H0; now contradict H0.
(* apply interp_check_lit; auto with smtcoq_core. *)
(* intros a b; case a; try discriminate; intros H _; rewrite H; auto with smtcoq_core. *)
(* (* Conjunction *) *)
- (* unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H1 H2]; rewrite forallbi_spec in H2; intro H3; assert (H4 := afold_left_andb_true_inv _ _ _ H3); clear H3; apply afold_left_andb_true; rewrite <- H1; intros i Hi; eapply interp_check_lit; eauto with smtcoq_core. *)
+ (* unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H1 H2]; rewrite aforallbi_spec in H2; intro H3; assert (H4 := afold_left_andb_true_inv _ _ _ H3); clear H3; apply afold_left_andb_true; rewrite <- H1; intros i Hi; eapply interp_check_lit; eauto with smtcoq_core. *)
(* (* Disjunction *) *)
- (* unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H1 H2]; rewrite forallbi_spec in H2; intro H3; assert (H4 := afold_left_orb_true_inv _ _ _ H3); clear H3; destruct H4 as [i [H3 H4]]; eapply afold_left_orb_true. *)
+ (* unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H1 H2]; rewrite aforallbi_spec in H2; intro H3; assert (H4 := afold_left_orb_true_inv _ _ _ H3); clear H3; destruct H4 as [i [H3 H4]]; eapply afold_left_orb_true. *)
(* rewrite <- H1; eauto with smtcoq_core. *)
(* eapply interp_check_lit; eauto with smtcoq_core. *)
(* (* Implication *) *)
- (* unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H1 H2]; rewrite forallbi_spec in H2; intro H3; apply afold_right_implb_true; case_eq (length a1 == 0); intro Heq. *)
+ (* unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H1 H2]; rewrite aforallbi_spec in H2; intro H3; apply afold_right_implb_true; case_eq (length a1 == 0); intro Heq. *)
(* left; rewrite eqb_spec in Heq; rewrite <- H1; auto with smtcoq_core. *)
(* destruct (afold_right_implb_true_inv _ _ _ H3) as [H4|[[i [H4 H5]]|H4]]. *)
(* rewrite H4 in Heq; discriminate. *)
(* right; left; exists i; rewrite <- H1; split; auto with smtcoq_core; case_eq (Lit.interp rho (a2 .[ i])); auto with smtcoq_core; intro H6; assert (H7: i < length a1 = true). *)
(* rewrite ltb_spec in *; rewrite eqb_false_spec in Heq; rewrite to_Z_sub_1_diff in H4; auto with smtcoq_core; omega. *)
(* generalize (H2 _ H7); rewrite H4; intro H8; rewrite (interp_check_lit _ _ H8 H6) in H5; auto with smtcoq_core. *)
- (* right; case_eq (existsbi (fun i l => (i < length a2 - 1) && (negb (Lit.interp rho l))) a2). *)
- (* rewrite existsbi_spec; intros [i [_ H5]]; rewrite andb_true_iff in H5; destruct H5 as [H5 H6]; left; exists i; split; auto with smtcoq_core; generalize H6; case (Lit.interp rho (a2 .[ i])); auto with smtcoq_core; discriminate. *)
- (* rewrite existsbi_false_spec; intro H; right; intros i Hi; assert (Hi' := Hi); rewrite <- H1 in Hi'; generalize (H2 _ Hi') (H _ Hi); rewrite <- H1; case (i < length a1 - 1); simpl. *)
+ (* right; case_eq (aexistsbi (fun i l => (i < length a2 - 1) && (negb (Lit.interp rho l))) a2). *)
+ (* rewrite aexistsbi_spec; intros [i [_ H5]]; rewrite andb_true_iff in H5; destruct H5 as [H5 H6]; left; exists i; split; auto with smtcoq_core; generalize H6; case (Lit.interp rho (a2 .[ i])); auto with smtcoq_core; discriminate. *)
+ (* rewrite aexistsbi_false_spec; intro H; right; intros i Hi; assert (Hi' := Hi); rewrite <- H1 in Hi'; generalize (H2 _ Hi') (H _ Hi); rewrite <- H1; case (i < length a1 - 1); simpl. *)
(* intros _; case (Lit.interp rho (a2 .[ i])); auto with smtcoq_core; discriminate. *)
(* intros H5 _; apply (interp_check_lit _ _ H5); apply H4; auto with smtcoq_core. *)
(* (* Xor *) *)
@@ -473,9 +473,9 @@ intros. destruct H0; now contradict H0.
End AUX.
Definition check_hform h1 h2 :=
- foldi_down_cont
+ foldi
(fun _ cont h1 h2 => (h1 == h2) || check_form_aux cont (get_form h1) (get_form h2))
- (PArray.length t_form) 0 (fun h1 h2 => false) h1 h2.
+ 0 (PArray.length t_form) (fun h1 h2 => false) h1 h2.
Definition check_form := check_form_aux check_hform.
@@ -517,8 +517,8 @@ intros. destruct H0; now contradict H0.
Lemma interp_check_hform : forall h1 h2,
check_hform h1 h2 -> Var.interp rho h1 = Var.interp rho h2.
Proof.
- unfold check_hform; apply foldi_down_cont_ind; try discriminate. intros i cont _ _ Hrec h1 h2. unfold is_true; rewrite orb_true_iff; intros [H|H].
- rewrite Int63Properties.eqb_spec in H; rewrite H; auto with smtcoq_core.
+ unfold check_hform; apply foldi_ind; try discriminate. apply leb_0. intros i cont _ _ Hrec h1 h2. unfold is_true; rewrite orb_true_iff; intros [H|H].
+ rewrite Int63.eqb_spec in H; rewrite H; auto with smtcoq_core.
unfold Var.interp; rewrite !wf_interp_form; auto with smtcoq_core; eapply interp_check_form_aux; eauto with smtcoq_core.
Qed.
diff --git a/src/spl/Syntactic.v b/src/spl/Syntactic.v
index 054c5ea..995ea15 100644
--- a/src/spl/Syntactic.v
+++ b/src/spl/Syntactic.v
@@ -184,15 +184,15 @@ Section CheckAtom.
(* N-ary operators *)
- intros [op2|op2 i2|op2 i2 j2|op2 i2 j2|op2 li2|f2 args2]; simpl; try discriminate; destruct op1 as [t1]; destruct op2 as [t2]; unfold is_true; rewrite andb_true_iff; change (Typ.eqb t1 t2 = true) with (is_true (Typ.eqb t1 t2)); rewrite Typ.eqb_spec; intros [H1 H2]; subst t2; rewrite (list_beq_compute_interp _ _ _ H2); auto.
(* Application *)
- - intros [op2|op2 i2|op2 i2 j2|op2 i2 j2|op2 li2|f2 args2]; simpl; try discriminate; unfold is_true; rewrite andb_true_iff, Int63Properties.eqb_spec; intros [H2 H1]; subst f2; rewrite (list_beq_correct _ _ H1); auto.
+ - intros [op2|op2 i2|op2 i2 j2|op2 i2 j2|op2 li2|f2 args2]; simpl; try discriminate; unfold is_true; rewrite andb_true_iff, Int63.eqb_spec; intros [H2 H1]; subst f2; rewrite (list_beq_correct _ _ H1); auto.
Qed.
End AUX.
Definition check_hatom h1 h2 :=
- foldi_down_cont
+ foldi
(fun _ cont h1 h2 => (h1 == h2) || check_atom_aux cont (t_atom.[h1]) (t_atom.[h2]))
- (PArray.length t_atom) 0 (fun h1 h2 => false) h1 h2.
+ 0 (PArray.length t_atom) (fun h1 h2 => false) h1 h2.
Definition check_atom := check_atom_aux check_hatom.
@@ -230,10 +230,11 @@ Section CheckAtom.
interp_hatom t_i t_func t_atom h1 = interp_hatom t_i t_func t_atom h2.
Proof.
unfold check_hatom;intros Hwf Hdef.
- apply foldi_down_cont_ind;try discriminate.
+ apply foldi_ind;try discriminate.
+ apply leb_0.
intros i cont _ _ Hrec h1 h2.
unfold is_true; rewrite orb_true_iff; intros [H|H].
- rewrite Int63Properties.eqb_spec in H; rewrite H; reflexivity.
+ rewrite Int63.eqb_spec in H; rewrite H; reflexivity.
unfold interp_hatom;rewrite !t_interp_wf;trivial.
apply check_atom_aux_correct with cont;trivial.
Qed.
@@ -268,7 +269,7 @@ Section CheckAtom.
| Val _ _, Val _ _ => False
end.
Proof.
- unfold wt; unfold is_true at 1; rewrite forallbi_spec; intros Hwt Hwf Hdef h1 h2; unfold check_neg_hatom; case_eq (get_atom h1); try discriminate; intros b1 t11 t12 H1; case_eq (get_atom h2); try discriminate; intros b2 t21 t22 H2; assert (H7: h1 < length t_atom) by (apply PArray.get_not_default_lt; rewrite H1, Hdef; discriminate); generalize (Hwt _ H7); rewrite H1; simpl; generalize H1; case b1; try discriminate; clear H1 b1; simpl; intro H1; case (get_type' t_i (t_interp t_i t_func t_atom) h1); try discriminate; simpl; rewrite andb_true_iff; intros [H30 H31]; change (is_true (Typ.eqb (get_type' t_i (t_interp t_i t_func t_atom) t11) Typ.TZ)) in H30; change (is_true (Typ.eqb (get_type' t_i (t_interp t_i t_func t_atom) t12) Typ.TZ)) in H31; rewrite Typ.eqb_spec in H30, H31; generalize (check_aux_interp_hatom _ t_func _ Hwf t11), (check_aux_interp_hatom _ t_func _ Hwf t12); rewrite H30, H31; intros [v1 Hv1] [v2 Hv2]; generalize H2; case b2; try discriminate; clear H2 b2; intro H2; unfold is_true; rewrite andb_true_iff; intros [H3 H4]; generalize (check_hatom_correct Hwf Hdef _ _ H3), (check_hatom_correct Hwf Hdef _ _ H4); unfold interp_hatom; intros H5 H6; rewrite t_interp_wf; auto; rewrite H1; simpl; rewrite Hv1, Hv2; simpl; rewrite t_interp_wf; auto; rewrite H2; simpl; rewrite <- H5; rewrite <- H6, Hv1, Hv2; simpl.
+ unfold wt; unfold is_true at 1; rewrite aforallbi_spec; intros Hwt Hwf Hdef h1 h2; unfold check_neg_hatom; case_eq (get_atom h1); try discriminate; intros b1 t11 t12 H1; case_eq (get_atom h2); try discriminate; intros b2 t21 t22 H2; assert (H7: h1 < length t_atom) by (apply PArray.get_not_default_lt; rewrite H1, Hdef; discriminate); generalize (Hwt _ H7); rewrite H1; simpl; generalize H1; case b1; try discriminate; clear H1 b1; simpl; intro H1; case (get_type' t_i (t_interp t_i t_func t_atom) h1); try discriminate; simpl; rewrite andb_true_iff; intros [H30 H31]; change (is_true (Typ.eqb (get_type' t_i (t_interp t_i t_func t_atom) t11) Typ.TZ)) in H30; change (is_true (Typ.eqb (get_type' t_i (t_interp t_i t_func t_atom) t12) Typ.TZ)) in H31; rewrite Typ.eqb_spec in H30, H31; generalize (check_aux_interp_hatom _ t_func _ Hwf t11), (check_aux_interp_hatom _ t_func _ Hwf t12); rewrite H30, H31; intros [v1 Hv1] [v2 Hv2]; generalize H2; case b2; try discriminate; clear H2 b2; intro H2; unfold is_true; rewrite andb_true_iff; intros [H3 H4]; generalize (check_hatom_correct Hwf Hdef _ _ H3), (check_hatom_correct Hwf Hdef _ _ H4); unfold interp_hatom; intros H5 H6; rewrite t_interp_wf; auto; rewrite H1; simpl; rewrite Hv1, Hv2; simpl; rewrite t_interp_wf; auto; rewrite H2; simpl; rewrite <- H5; rewrite <- H6, Hv1, Hv2; simpl.
rewrite Z.ltb_antisym; auto.
rewrite Z.geb_leb, Z.ltb_antisym; auto.
rewrite Z.leb_antisym; auto.
@@ -332,20 +333,20 @@ Section FLATTEN.
(frec : list _lit -> _lit -> list _lit)
(largs:list _lit) (l:_lit) : list _lit :=
match get_op l with
- | Some a => PArray.fold_left frec largs a
+ | Some a => foldi (fun i x => frec x (a.[i])) 0 (length a) largs
| None => l::largs
end.
(* Register flatten_op_body as PrimInline. *)
Definition flatten_op_lit (get_op:_lit -> option (array _lit)) max :=
- foldi_cont (fun _ => flatten_op_body get_op) 0 max (fun largs l => l::largs).
+ foldi (fun _ => flatten_op_body get_op) 0 max (fun largs l => l::largs).
Definition flatten_and t :=
- PArray.fold_left (flatten_op_lit get_and (PArray.length t_form)) nil t.
+ foldi (fun i x => flatten_op_lit get_and (PArray.length t_form) x (t.[i])) 0 (length t) nil.
Definition flatten_or t :=
- PArray.fold_left (flatten_op_lit get_or (PArray.length t_form)) nil t.
+ foldi (fun i x => flatten_op_lit get_or (PArray.length t_form) x (t.[i])) 0 (length t) nil.
Variable check_atom check_neg_atom : atom -> atom -> bool.
@@ -371,7 +372,7 @@ Section FLATTEN.
frec l1 lf1 && frec l2 lf2
| Fimp args1, Fimp args2 =>
if PArray.length args1 == PArray.length args2 then
- PArray.forallbi (fun i l => frec l (args2.[i])) args1
+ aforallbi (fun i l => frec l (args2.[i])) args1
else false
| Fiff l1 l2, Fiff lf1 lf2 =>
frec l1 lf1 && frec l2 lf2
@@ -387,7 +388,7 @@ Section FLATTEN.
(* Register check_flatten_body as PrimInline. *)
Definition check_flatten_aux l lf :=
- foldi_cont (fun _ => check_flatten_body) 0 (PArray.length t_form) (fun _ _ => false) l lf.
+ foldi (fun _ => check_flatten_body) 0 (PArray.length t_form) (fun _ _ => false) l lf.
Definition check_flatten s cid lf :=
match S.get s cid with
@@ -412,8 +413,11 @@ Section FLATTEN.
Lemma interp_Fnot2 : forall i l, interp interp_atom interp_bvatom t_form (Fnot2 i l) = interp_lit l.
Proof.
- intros i l;simpl;apply fold_ind;trivial.
- intros a;rewrite negb_involutive;trivial.
+ intros i l;simpl.
+ apply foldi_ind.
+ apply leb_0.
+ reflexivity.
+ intros; rewrite negb_involutive; assumption.
Qed.
Lemma remove_not_correct :
@@ -452,23 +456,21 @@ Section FLATTEN.
Lemma flatten_and_correct : forall args,
List.fold_right (fun l res => andb res (interp_lit l)) true (flatten_and args) =
- afold_left _ _ true andb interp_lit args.
+ afold_left _ true andb (amap interp_lit args).
Proof.
intros;rewrite afold_left_spec;auto;unfold flatten_and.
- set (t:= true);unfold t at 2;
- change true with
- (List.fold_right (fun (l : int) (res : bool) => res && interp_lit l) true nil).
- unfold t;clear t.
- rewrite !fold_left_to_list.
+ change true with (List.fold_right (fun (l : int) (res : bool) => res && interp_lit l) true nil) at 2.
+ rewrite !foldi_to_list, to_list_amap.
generalize (@nil int);induction (to_list args);simpl;trivial.
intros l0;rewrite IHl.
clear IHl;f_equal; unfold flatten_op_lit.
- clear l;revert a l0;apply foldi_cont_ind;simpl;trivial.
- intros i cont _ Hle Hrec a l;unfold flatten_op_body.
+ clear l;revert a l0;apply foldi_ind;simpl;trivial.
+ apply leb_0.
+ intros i cont _ Hlt Hrec a l;unfold flatten_op_body.
case_eq (get_and a);intros;trivial.
rewrite get_and_correct with (1:= H);simpl.
- rewrite afold_left_spec; auto; rewrite !fold_left_to_list.
- rewrite <- !fold_left_rev_right.
+ rewrite afold_left_spec; auto; rewrite !foldi_to_list.
+ rewrite <- !fold_left_rev_right, to_list_amap, <- map_rev.
clear H a;revert l;induction (List.rev (to_list a0));simpl.
intros l;rewrite andb_true_r;trivial.
intros;rewrite Hrec, IHl, andb_assoc;trivial.
@@ -476,23 +478,21 @@ Section FLATTEN.
Lemma flatten_or_correct : forall args,
List.fold_right (fun l res => orb res (interp_lit l)) false (flatten_or args) =
- afold_left _ _ false orb interp_lit args.
+ afold_left _ false orb (amap interp_lit args).
Proof.
intros;rewrite afold_left_spec;auto;unfold flatten_or.
- set (t:= false);unfold t at 2;
- change false with
- (List.fold_right (fun (l : int) (res : bool) => res || interp_lit l) false nil).
- unfold t;clear t.
- rewrite !fold_left_to_list.
+ change false with (List.fold_right (fun (l : int) (res : bool) => res || interp_lit l) false nil) at 2.
+ rewrite !foldi_to_list, to_list_amap.
generalize (@nil int);induction (to_list args);simpl;trivial.
intros l0;rewrite IHl.
clear IHl;f_equal; unfold flatten_op_lit.
- clear l;revert a l0;apply foldi_cont_ind;simpl;trivial.
- intros i cont _ Hle Hrec a l;unfold flatten_op_body.
+ clear l;revert a l0;apply foldi_ind;simpl;trivial.
+ apply leb_0.
+ intros i cont _ Hlt Hrec a l;unfold flatten_op_body.
case_eq (get_or a);intros;trivial.
rewrite get_or_correct with (1:= H);simpl.
- rewrite afold_left_spec; auto; rewrite !fold_left_to_list.
- rewrite <- !fold_left_rev_right.
+ rewrite afold_left_spec; auto; rewrite !foldi_to_list.
+ rewrite <- !fold_left_rev_right, to_list_amap, <- map_rev.
clear H a;revert l;induction (List.rev (to_list a0));simpl.
intros l;rewrite orb_false_r;trivial.
intros;rewrite Hrec, IHl, orb_assoc;trivial.
@@ -503,7 +503,8 @@ Section FLATTEN.
interp_lit l = interp_lit lf.
Proof.
unfold check_flatten_aux.
- apply foldi_cont_ind.
+ apply foldi_ind.
+ apply leb_0.
discriminate.
intros i cont _ Hle Hrec l lf;unfold check_flatten_body.
rewrite <- (remove_not_correct l), <- (remove_not_correct lf).
@@ -513,7 +514,7 @@ Section FLATTEN.
unfold Lit.interp.
assert (Lit.is_pos l = Lit.is_pos lf).
unfold Lit.is_pos.
- rewrite <- eqb_spec, land_comm in e.
+ rewrite <- eqb_spec, landC in e.
change (is_true (is_even (l lxor lf))) in e.
rewrite is_even_xor in e.
destruct (is_even l);destruct (is_even lf);trivial;discriminate.
@@ -542,8 +543,8 @@ Section FLATTEN.
rewrite (Hrec _ _ H1), (IHl0 _ H2);trivial.
(* implb *)
revert H;destruct (reflect_eqb (length a) (length a0));[intros|discriminate].
- apply afold_right_eq;trivial.
- rewrite forallbi_spec in H;auto.
+ apply afold_right_eq;rewrite !length_amap;trivial.
+ intros;rewrite !get_amap by congruence;rewrite aforallbi_spec in H;auto.
(* xorb *)
unfold is_true in H;rewrite andb_true_iff in H;destruct H as [H H0].
rewrite (Hrec _ _ H), (Hrec _ _ H0);trivial.
@@ -556,7 +557,7 @@ Section FLATTEN.
(** opposite sign *)
assert (Lit.is_pos l = negb (Lit.is_pos lf)).
unfold Lit.is_pos.
- rewrite <- eqb_spec, land_comm in n0.
+ rewrite <- eqb_spec, landC in n0.
change (~is_true (is_even (l lxor lf))) in n0.
rewrite is_even_xor in n0.
destruct (is_even l);destruct (is_even lf);trivial;elim n0;reflexivity.
diff --git a/src/trace/coqInterface.ml b/src/trace/coqInterface.ml
index 4b0acc9..96fe6c1 100644
--- a/src/trace/coqInterface.ml
+++ b/src/trace/coqInterface.ml
@@ -101,26 +101,12 @@ let init_modules = Coqlib.init_modules
(* Int63 *)
-let int63_modules = [["SMTCoq";"Int63";"Int63Native"]]
-
-(* 31-bits integers are "called" 63 bits (this is sound) *)
-let int31_module = [["Coq";"Numbers";"Cyclic";"Int31";"Int31"]]
-let cD0 = gen_constant int31_module "D0"
-let cD1 = gen_constant int31_module "D1"
-let cI31 = gen_constant int31_module "I31"
-
-let mkInt : int -> constr = fun i ->
- let a = Array.make 31 (Lazy.force cD0) in
- let j = ref i in
- let k = ref 30 in
- while !j <> 0 do
- if !j land 1 = 1 then a.(!k) <- Lazy.force cD1;
- j := !j lsr 1;
- decr k
- done;
- mkApp (Lazy.force cI31, a)
-
-let cint = gen_constant int31_module "int31"
+let int63_module = [["Coq";"Numbers";"Cyclic";"Int63";"Int63"]]
+
+let mkInt : int -> Constr.constr =
+ fun i -> Constr.mkInt (Uint63.of_int i)
+
+let cint = gen_constant int63_module "int"
(* PArray *)
diff --git a/src/trace/coqInterface.mli b/src/trace/coqInterface.mli
index d355053..f450f6d 100644
--- a/src/trace/coqInterface.mli
+++ b/src/trace/coqInterface.mli
@@ -65,7 +65,7 @@ val init_modules : string list list
(* Int63 *)
-val int63_modules : string list list
+val int63_module : string list list
val mkInt : int -> constr
val cint : constr lazy_t
diff --git a/src/trace/coqTerms.ml b/src/trace/coqTerms.ml
index 1c4ee81..67392bb 100644
--- a/src/trace/coqTerms.ml
+++ b/src/trace/coqTerms.ml
@@ -18,7 +18,7 @@ let gen_constant = CoqInterface.gen_constant
(* Int63 *)
let cint = CoqInterface.cint
-let ceq63 = gen_constant CoqInterface.int63_modules "eqb"
+let ceq63 = gen_constant CoqInterface.int63_module "eqb"
(* PArray *)
let carray = gen_constant CoqInterface.parray_modules "array"
diff --git a/src/trace/smtBtype.ml b/src/trace/smtBtype.ml
index c9aad70..d091758 100644
--- a/src/trace/smtBtype.ml
+++ b/src/trace/smtBtype.ml
@@ -51,7 +51,7 @@ let index_tbl = Hashtbl.create 17
let index_to_coq i =
try Hashtbl.find index_tbl i
with Not_found ->
- let interp = mklApp cTindex [|mkInt i|] in
+ let interp = mklApp cTindex [|mkN i|] in
Hashtbl.add index_tbl i interp;
interp
diff --git a/unit-tests/Tests_lfsc_tactics.v b/unit-tests/Tests_lfsc_tactics.v
index cb9d7ba..4cef5f1 100644
--- a/unit-tests/Tests_lfsc_tactics.v
+++ b/unit-tests/Tests_lfsc_tactics.v
@@ -503,8 +503,8 @@ Qed.
-Goal forall x y z, ((x <=? 3) && ((y <=? 7) || (z <=? 9))) -->
- ((x + y <=? 10) || (x + z <=? 12)) = true.
+Goal forall x y z, ((x <=? 3) && ((y <=? 7) || (z <=? 9)))%Z -->
+ ((x + y <=? 10) || (x + z <=? 12))%Z = true.
Proof using.
smt.
Qed.
@@ -515,29 +515,29 @@ Proof using.
Qed.
-Goal forall x y, (x >? y) --> (y + 1 <=? x) = true.
+Goal forall x y, (x >? y)%Z --> (y + 1 <=? x)%Z = true.
Proof using.
smt.
Qed.
-Goal forall x y, Bool.eqb (x <? y) (x <=? y - 1) = true.
+Goal forall x y, Bool.eqb (x <? y)%Z (x <=? y - 1)%Z = true.
Proof using.
smt.
Qed.
Goal forall x y, ((x + y <=? - (3)) && (y >=? 0)
- || (x <=? - (3))) && (x >=? 0) = false.
+ || (x <=? - (3)))%Z && (x >=? 0) = false.
Proof using.
smt.
Qed.
-Goal forall x, (andb ((x - 3) <=? 7) (7 <=? (x - 3))) --> (x >=? 10) = true.
+Goal forall x, (andb ((x - 3) <=? 7) (7 <=? (x - 3)))%Z --> (x >=? 10) = true.
Proof using.
smt.
Qed.
-Goal forall x, (Z.eqb (x - 3) 7) --> (10 <=? x) = true.
+Goal forall x, (Z.eqb (x - 3) 7)%Z --> (10 <=? x)%Z = true.
Proof using.
smt.
Qed.