aboutsummaryrefslogtreecommitdiffstats
path: root/lib
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2015-10-20 13:32:18 +0200
committerBernhard Schommer <bernhardschommer@gmail.com>2015-10-20 13:32:18 +0200
commit4d542bc7eafadb16b845cf05d1eb4988eb55ed0f (patch)
tree1961b41815fc6e392cc0bd2beeb0fb504bc160ce /lib
parent7a6bb90048db7a254e959b1e3c308bac5fe6c418 (diff)
downloadcompcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.tar.gz
compcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.zip
Updated PR by removing whitespaces. Bug 17450.
Diffstat (limited to 'lib')
-rw-r--r--lib/Camlcoq.ml6
-rw-r--r--lib/Coqlib.v178
-rw-r--r--lib/FSetAVLplus.v88
-rw-r--r--lib/Fappli_IEEE_extra.v544
-rw-r--r--lib/Floats.v268
-rw-r--r--lib/Heaps.v76
-rw-r--r--lib/Integers.v1236
-rw-r--r--lib/Intv.v46
-rw-r--r--lib/IntvSets.v82
-rw-r--r--lib/Iteration.v46
-rw-r--r--lib/Lattice.v76
-rw-r--r--lib/Maps.v226
-rw-r--r--lib/Ordered.v18
-rw-r--r--lib/Parmov.v308
-rw-r--r--lib/Postorder.v74
-rw-r--r--lib/Readconfig.mll4
-rw-r--r--lib/UnionFind.v112
-rw-r--r--lib/Wfsimpl.v8
18 files changed, 1698 insertions, 1698 deletions
diff --git a/lib/Camlcoq.ml b/lib/Camlcoq.ml
index c50b3230..c5fb2e55 100644
--- a/lib/Camlcoq.ml
+++ b/lib/Camlcoq.ml
@@ -86,7 +86,7 @@ module P = struct
if n = 0l
then assert false
else Coq_xO (of_int32 (Int32.shift_right_logical n 1))
- else
+ else
if n = 1l
then Coq_xH
else Coq_xI (of_int32 (Int32.shift_right_logical n 1))
@@ -101,7 +101,7 @@ module P = struct
if n = 0L
then assert false
else Coq_xO (of_int64 (Int64.shift_right_logical n 1))
- else
+ else
if n = 1L
then Coq_xH
else Coq_xI (of_int64 (Int64.shift_right_logical n 1))
@@ -295,7 +295,7 @@ let intern_string s =
next_atom := Pos.succ !next_atom;
Hashtbl.add atom_of_string s a;
Hashtbl.add string_of_atom a s;
- a
+ a
let extern_atom a =
try
Hashtbl.find string_of_atom a
diff --git a/lib/Coqlib.v b/lib/Coqlib.v
index 35d53854..4ec19fa9 100644
--- a/lib/Coqlib.v
+++ b/lib/Coqlib.v
@@ -107,14 +107,14 @@ Proof.
intros. case (peq x y); intros.
elim H; auto.
auto.
-Qed.
+Qed.
Definition Plt: positive -> positive -> Prop := Pos.lt.
Lemma Plt_ne:
forall (x y: positive), Plt x y -> x <> y.
Proof.
- unfold Plt; intros. red; intro. subst y. eelim Pos.lt_irrefl; eauto.
+ unfold Plt; intros. red; intro. subst y. eelim Pos.lt_irrefl; eauto.
Qed.
Hint Resolve Plt_ne: coqlib.
@@ -125,7 +125,7 @@ Proof (Pos.lt_trans).
Lemma Plt_succ:
forall (x: positive), Plt x (Psucc x).
Proof.
- unfold Plt; intros. apply Pos.lt_succ_r. apply Pos.le_refl.
+ unfold Plt; intros. apply Pos.lt_succ_r. apply Pos.le_refl.
Qed.
Hint Resolve Plt_succ: coqlib.
@@ -139,7 +139,7 @@ Hint Resolve Plt_succ: coqlib.
Lemma Plt_succ_inv:
forall (x y: positive), Plt x (Psucc y) -> Plt x y \/ x = y.
Proof.
- unfold Plt; intros. rewrite Pos.lt_succ_r in H.
+ unfold Plt; intros. rewrite Pos.lt_succ_r in H.
apply Pos.le_lteq; auto.
Qed.
@@ -242,11 +242,11 @@ Lemma positive_Peano_ind:
Proof.
intros.
apply (well_founded_ind Plt_wf P).
- intros.
+ intros.
case (peq x0 xH); intro.
subst x0; auto.
elim (Psucc_pred x0); intro. contradiction. rewrite <- H2.
- apply H0. apply H1. apply Ppred_Plt. auto.
+ apply H0. apply H1. apply Ppred_Plt. auto.
Qed.
End POSITIVE_ITERATION.
@@ -269,14 +269,14 @@ Proof.
intros. case (zeq x y); intros.
elim H; auto.
auto.
-Qed.
+Qed.
Open Scope Z_scope.
Definition zlt: forall (x y: Z), {x < y} + {x >= y} := Z_lt_dec.
Lemma zlt_true:
- forall (A: Type) (x y: Z) (a b: A),
+ forall (A: Type) (x y: Z) (a b: A),
x < y -> (if zlt x y then a else b) = a.
Proof.
intros. case (zlt x y); intros.
@@ -285,7 +285,7 @@ Proof.
Qed.
Lemma zlt_false:
- forall (A: Type) (x y: Z) (a b: A),
+ forall (A: Type) (x y: Z) (a b: A),
x >= y -> (if zlt x y then a else b) = b.
Proof.
intros. case (zlt x y); intros.
@@ -296,7 +296,7 @@ Qed.
Definition zle: forall (x y: Z), {x <= y} + {x > y} := Z_le_gt_dec.
Lemma zle_true:
- forall (A: Type) (x y: Z) (a b: A),
+ forall (A: Type) (x y: Z) (a b: A),
x <= y -> (if zle x y then a else b) = a.
Proof.
intros. case (zle x y); intros.
@@ -305,7 +305,7 @@ Proof.
Qed.
Lemma zle_false:
- forall (A: Type) (x y: Z) (a b: A),
+ forall (A: Type) (x y: Z) (a b: A),
x > y -> (if zle x y then a else b) = b.
Proof.
intros. case (zle x y); intros.
@@ -327,7 +327,7 @@ Qed.
Lemma two_power_nat_two_p:
forall x, two_power_nat x = two_p (Z_of_nat x).
Proof.
- induction x. auto.
+ induction x. auto.
rewrite two_power_nat_S. rewrite inj_S. rewrite two_p_S. omega. omega.
Qed.
@@ -335,7 +335,7 @@ Lemma two_p_monotone:
forall x y, 0 <= x <= y -> two_p x <= two_p y.
Proof.
intros.
- replace (two_p x) with (two_p x * 1) by omega.
+ replace (two_p x) with (two_p x * 1) by omega.
replace y with (x + (y - x)) by omega.
rewrite two_p_is_exp; try omega.
apply Zmult_le_compat_l.
@@ -356,7 +356,7 @@ Lemma two_p_strict:
Proof.
intros x0 GT. pattern x0. apply natlike_ind.
simpl. omega.
- intros. rewrite two_p_S; auto. generalize (two_p_gt_ZERO x H). omega.
+ intros. rewrite two_p_S; auto. generalize (two_p_gt_ZERO x H). omega.
omega.
Qed.
@@ -366,7 +366,7 @@ Proof.
intros. assert (x = 0 \/ x - 1 >= 0) by omega. destruct H0.
subst. vm_compute. auto.
replace (two_p x) with (2 * two_p (x - 1)).
- generalize (two_p_strict _ H0). omega.
+ generalize (two_p_strict _ H0). omega.
rewrite <- two_p_S. decEq. omega. omega.
Qed.
@@ -377,7 +377,7 @@ Lemma Zmin_spec:
Proof.
intros. case (zlt x y); unfold Zlt, Zge; intro z.
unfold Zmin. rewrite z. auto.
- unfold Zmin. caseEq (x ?= y); intro.
+ unfold Zmin. caseEq (x ?= y); intro.
apply Zcompare_Eq_eq. auto.
contradiction.
reflexivity.
@@ -411,21 +411,21 @@ Qed.
Lemma Zdiv_small:
forall x y, 0 <= x < y -> x / y = 0.
Proof.
- intros. assert (y > 0). omega.
+ intros. assert (y > 0). omega.
assert (forall a b,
0 <= a < y ->
0 <= y * b + a < y ->
b = 0).
- intros.
+ intros.
assert (b = 0 \/ b > 0 \/ (-b) > 0). omega.
elim H3; intro.
auto.
elim H4; intro.
- assert (y * b >= y * 1). apply Zmult_ge_compat_l. omega. omega.
- omegaContradiction.
+ assert (y * b >= y * 1). apply Zmult_ge_compat_l. omega. omega.
+ omegaContradiction.
assert (y * (-b) >= y * 1). apply Zmult_ge_compat_l. omega. omega.
rewrite <- Zopp_mult_distr_r in H6. omegaContradiction.
- apply H1 with (x mod y).
+ apply H1 with (x mod y).
apply Z_mod_lt. auto.
rewrite <- Z_div_mod_eq. auto. auto.
Qed.
@@ -434,7 +434,7 @@ Lemma Zmod_small:
forall x y, 0 <= x < y -> x mod y = x.
Proof.
intros. assert (y > 0). omega.
- generalize (Z_div_mod_eq x y H0).
+ generalize (Z_div_mod_eq x y H0).
rewrite (Zdiv_small x y H). omega.
Qed.
@@ -442,7 +442,7 @@ Lemma Zmod_unique:
forall x y a b,
x = a * y + b -> 0 <= b < y -> x mod y = b.
Proof.
- intros. subst x. rewrite Zplus_comm.
+ intros. subst x. rewrite Zplus_comm.
rewrite Z_mod_plus. apply Zmod_small. auto. omega.
Qed.
@@ -463,12 +463,12 @@ Proof.
generalize (Z_div_mod_eq (a/b) c H0). generalize (Z_mod_lt (a/b) c H0). intros.
set (q1 := a / b) in *. set (r1 := a mod b) in *.
set (q2 := q1 / c) in *. set (r2 := q1 mod c) in *.
- symmetry. apply Zdiv_unique with (r2 * b + r1).
+ symmetry. apply Zdiv_unique with (r2 * b + r1).
rewrite H2. rewrite H4. ring.
- split.
+ split.
assert (0 <= r2 * b). apply Zmult_le_0_compat. omega. omega. omega.
assert ((r2 + 1) * b <= c * b).
- apply Zmult_le_compat_r. omega. omega.
+ apply Zmult_le_compat_r. omega. omega.
replace ((r2 + 1) * b) with (r2 * b + b) in H5 by ring.
replace (c * b) with (b * c) in H5 by ring.
omega.
@@ -490,17 +490,17 @@ Lemma Zdiv_interval_1:
lo * b <= a < hi * b ->
lo <= a/b < hi.
Proof.
- intros.
+ intros.
generalize (Z_div_mod_eq a b H1). generalize (Z_mod_lt a b H1). intros.
set (q := a/b) in *. set (r := a mod b) in *.
split.
assert (lo < (q + 1)).
- apply Zmult_lt_reg_r with b. omega.
- apply Zle_lt_trans with a. omega.
+ apply Zmult_lt_reg_r with b. omega.
+ apply Zle_lt_trans with a. omega.
replace ((q + 1) * b) with (b * q + b) by ring.
omega.
omega.
- apply Zmult_lt_reg_r with b. omega.
+ apply Zmult_lt_reg_r with b. omega.
replace (q * b) with (b * q) by ring.
omega.
Qed.
@@ -513,7 +513,7 @@ Proof.
intros.
assert (lo <= a / b < hi+1).
apply Zdiv_interval_1. omega. omega. auto.
- assert (lo * b <= lo * 1). apply Zmult_le_compat_l_neg. omega. omega.
+ assert (lo * b <= lo * 1). apply Zmult_le_compat_l_neg. omega. omega.
replace (lo * 1) with lo in H3 by ring.
assert ((hi + 1) * 1 <= (hi + 1) * b). apply Zmult_le_compat_l. omega. omega.
replace ((hi + 1) * 1) with (hi + 1) in H4 by ring.
@@ -526,19 +526,19 @@ Lemma Zmod_recombine:
a > 0 -> b > 0 ->
x mod (a * b) = ((x/b) mod a) * b + (x mod b).
Proof.
- intros.
- set (xb := x/b).
+ intros.
+ set (xb := x/b).
apply Zmod_unique with (xb/a).
generalize (Z_div_mod_eq x b H0); fold xb; intro EQ1.
generalize (Z_div_mod_eq xb a H); intro EQ2.
- rewrite EQ2 in EQ1.
+ rewrite EQ2 in EQ1.
eapply trans_eq. eexact EQ1. ring.
- generalize (Z_mod_lt x b H0). intro.
+ generalize (Z_mod_lt x b H0). intro.
generalize (Z_mod_lt xb a H). intro.
assert (0 <= xb mod a * b <= a * b - b).
split. apply Zmult_le_0_compat; omega.
replace (a * b - b) with ((a - 1) * b) by ring.
- apply Zmult_le_compat; omega.
+ apply Zmult_le_compat; omega.
omega.
Qed.
@@ -554,10 +554,10 @@ Definition Zdivide_dec:
forall (p q: Z), p > 0 -> { (p|q) } + { ~(p|q) }.
Proof.
intros. destruct (zeq (Zmod q p) 0).
- left. exists (q / p).
+ left. exists (q / p).
transitivity (p * (q / p) + (q mod p)). apply Z_div_mod_eq; auto.
transitivity (p * (q / p)). omega. ring.
- right; red; intros. elim n. apply Z_div_exact_1; auto.
+ right; red; intros. elim n. apply Z_div_exact_1; auto.
inv H0. rewrite Z_div_mult; auto. ring.
Defined.
Global Opaque Zdivide_dec.
@@ -567,7 +567,7 @@ Lemma Zdivide_interval:
0 < c -> 0 <= a < b -> (c | a) -> (c | b) -> 0 <= a <= b - c.
Proof.
intros. destruct H1 as [x EQ1]. destruct H2 as [y EQ2]. subst. destruct H0.
- split. omega. exploit Zmult_lt_reg_r; eauto. intros.
+ split. omega. exploit Zmult_lt_reg_r; eauto. intros.
replace (y * c - c) with ((y - 1) * c) by ring.
apply Zmult_le_compat_r; omega.
Qed.
@@ -585,9 +585,9 @@ Qed.
Lemma nat_of_Z_max:
forall z, Z_of_nat (nat_of_Z z) = Zmax z 0.
Proof.
- intros. unfold Zmax. destruct z; simpl; auto.
+ intros. unfold Zmax. destruct z; simpl; auto.
change (Z.of_nat (Z.to_nat (Zpos p)) = Zpos p).
- apply Z2Nat.id. compute; intuition congruence.
+ apply Z2Nat.id. compute; intuition congruence.
Qed.
Lemma nat_of_Z_eq:
@@ -607,7 +607,7 @@ Lemma nat_of_Z_plus:
p >= 0 -> q >= 0 ->
nat_of_Z (p + q) = (nat_of_Z p + nat_of_Z q)%nat.
Proof.
- unfold nat_of_Z; intros. apply Z2Nat.inj_add; omega.
+ unfold nat_of_Z; intros. apply Z2Nat.inj_add; omega.
Qed.
@@ -619,9 +619,9 @@ Definition align (n: Z) (amount: Z) :=
Lemma align_le: forall x y, y > 0 -> x <= align x y.
Proof.
- intros. unfold align.
+ intros. unfold align.
generalize (Z_div_mod_eq (x + y - 1) y H). intro.
- replace ((x + y - 1) / y * y)
+ replace ((x + y - 1) / y * y)
with ((x + y - 1) - (x + y - 1) mod y).
generalize (Z_mod_lt (x + y - 1) y H). omega.
rewrite Zmult_comm. omega.
@@ -629,7 +629,7 @@ Qed.
Lemma align_divides: forall x y, y > 0 -> (y | align x y).
Proof.
- intros. unfold align. apply Zdivide_factor_l.
+ intros. unfold align. apply Zdivide_factor_l.
Qed.
(** * Definitions and theorems on the data types [option], [sum] and [list] *)
@@ -709,14 +709,14 @@ Lemma list_length_z_cons:
list_length_z (hd :: tl) = list_length_z tl + 1.
Proof.
intros. unfold list_length_z. simpl.
- rewrite (list_length_z_aux_shift tl 1 0). omega.
+ rewrite (list_length_z_aux_shift tl 1 0). omega.
Qed.
Lemma list_length_z_pos:
forall (A: Type) (l: list A),
list_length_z l >= 0.
Proof.
- induction l; simpl. unfold list_length_z; simpl. omega.
+ induction l; simpl. unfold list_length_z; simpl. omega.
rewrite list_length_z_cons. omega.
Qed.
@@ -725,7 +725,7 @@ Lemma list_length_z_map:
list_length_z (map f l) = list_length_z l.
Proof.
induction l. reflexivity. simpl. repeat rewrite list_length_z_cons. congruence.
-Qed.
+Qed.
(** Extract the n-th element of a list, as [List.nth_error] does,
but the index [n] is of type [Z]. *)
@@ -740,7 +740,7 @@ Lemma list_nth_z_in:
forall (A: Type) (l: list A) n x,
list_nth_z l n = Some x -> In x l.
Proof.
- induction l; simpl; intros.
+ induction l; simpl; intros.
congruence.
destruct (zeq n 0). left; congruence. right; eauto.
Qed.
@@ -762,7 +762,7 @@ Proof.
discriminate.
rewrite list_length_z_cons. destruct (zeq n 0).
generalize (list_length_z_pos l); omega.
- exploit IHl; eauto. unfold Zpred. omega.
+ exploit IHl; eauto. unfold Zpred. omega.
Qed.
(** Properties of [List.incl] (list inclusion). *)
@@ -795,7 +795,7 @@ Lemma incl_same_head:
forall (A: Type) (x: A) (l1 l2: list A),
incl l1 l2 -> incl (x::l1) (x::l2).
Proof.
- intros; red; simpl; intros. intuition.
+ intros; red; simpl; intros. intuition.
Qed.
(** Properties of [List.map] (mapping a function over a list). *)
@@ -848,9 +848,9 @@ Lemma list_in_map_inv:
Proof.
induction l; simpl; intros.
contradiction.
- elim H; intro.
+ elim H; intro.
exists a; intuition auto.
- generalize (IHl y H0). intros [x [EQ IN]].
+ generalize (IHl y H0). intros [x [EQ IN]].
exists x; tauto.
Qed.
@@ -869,8 +869,8 @@ Lemma list_append_map_inv:
Proof.
induction m1; simpl; intros.
exists (@nil A); exists l; auto.
- destruct l; simpl in H; inv H.
- exploit IHm1; eauto. intros [l1 [l2 [P [Q R]]]]. subst l.
+ destruct l; simpl in H; inv H.
+ exploit IHm1; eauto. intros [l1 [l2 [P [Q R]]]]. subst l.
exists (a0 :: l1); exists l2; intuition. simpl; congruence.
Qed.
@@ -897,7 +897,7 @@ Remark list_fold_left_app:
forall l1 l2 accu,
list_fold_left accu (l1 ++ l2) = list_fold_left (list_fold_left accu l1) l2.
Proof.
- induction l1; simpl; intros.
+ induction l1; simpl; intros.
auto.
rewrite IHl1. auto.
Qed.
@@ -907,11 +907,11 @@ Lemma list_fold_right_eq:
list_fold_right l base =
match l with nil => base | x :: l' => f x (list_fold_right l' base) end.
Proof.
- unfold list_fold_right; intros.
+ unfold list_fold_right; intros.
destruct l.
auto.
- unfold rev'. rewrite <- ! rev_alt. simpl.
- rewrite list_fold_left_app. simpl. auto.
+ unfold rev'. rewrite <- ! rev_alt. simpl.
+ rewrite list_fold_left_app. simpl. auto.
Qed.
Lemma list_fold_right_spec:
@@ -943,7 +943,7 @@ Proof.
intros. apply in_or_app; simpl. elim (in_app_or _ _ _ H); intro; auto.
Qed.
-(** [list_disjoint l1 l2] holds iff [l1] and [l2] have no elements
+(** [list_disjoint l1 l2] holds iff [l1] and [l2] have no elements
in common. *)
Definition list_disjoint (A: Type) (l1 l2: list A) : Prop :=
@@ -967,21 +967,21 @@ Lemma list_disjoint_cons_left:
forall (A: Type) (a: A) (l1 l2: list A),
list_disjoint (a :: l1) l2 -> list_disjoint l1 l2.
Proof.
- unfold list_disjoint; simpl; intros. apply H; tauto.
+ unfold list_disjoint; simpl; intros. apply H; tauto.
Qed.
Lemma list_disjoint_cons_right:
forall (A: Type) (a: A) (l1 l2: list A),
list_disjoint l1 (a :: l2) -> list_disjoint l1 l2.
Proof.
- unfold list_disjoint; simpl; intros. apply H; tauto.
+ unfold list_disjoint; simpl; intros. apply H; tauto.
Qed.
Lemma list_disjoint_notin:
forall (A: Type) (l1 l2: list A) (a: A),
list_disjoint l1 l2 -> In a l1 -> ~(In a l2).
Proof.
- unfold list_disjoint; intros; red; intros.
+ unfold list_disjoint; intros; red; intros.
apply H with a a; auto.
Qed.
@@ -989,7 +989,7 @@ Lemma list_disjoint_sym:
forall (A: Type) (l1 l2: list A),
list_disjoint l1 l2 -> list_disjoint l2 l1.
Proof.
- unfold list_disjoint; intros.
+ unfold list_disjoint; intros.
apply sym_not_equal. apply H; auto.
Qed.
@@ -1000,9 +1000,9 @@ Proof.
induction l1; intros.
left; red; intros. elim H.
case (In_dec eqA_dec a l2); intro.
- right; red; intro. apply (H a a); auto with coqlib.
+ right; red; intro. apply (H a a); auto with coqlib.
case (IHl1 l2); intro.
- left; red; intros. elim H; intro.
+ left; red; intros. elim H; intro.
red; intro; subst a y. contradiction.
apply l; auto.
right; red; intros. elim n0. eapply list_disjoint_cons_left; eauto.
@@ -1029,9 +1029,9 @@ Lemma list_norepet_dec:
Proof.
induction l.
left; constructor.
- destruct IHl.
+ destruct IHl.
case (In_dec eqA_dec a l); intro.
- right. red; intro. inversion H. contradiction.
+ right. red; intro. inversion H. contradiction.
left. constructor; auto.
right. red; intro. inversion H. contradiction.
Defined.
@@ -1047,7 +1047,7 @@ Proof.
constructor.
red; intro. generalize (list_in_map_inv f _ _ H2).
intros [x [EQ IN]]. generalize EQ. change (f hd <> f x).
- apply H1. tauto. tauto.
+ apply H1. tauto. tauto.
red; intro; subst x. contradiction.
apply IHlist_norepet. intros. apply H1. tauto. tauto. auto.
Qed.
@@ -1057,20 +1057,20 @@ Remark list_norepet_append_commut:
list_norepet (a ++ b) -> list_norepet (b ++ a).
Proof.
intro A.
- assert (forall (x: A) (b: list A) (a: list A),
- list_norepet (a ++ b) -> ~(In x a) -> ~(In x b) ->
+ assert (forall (x: A) (b: list A) (a: list A),
+ list_norepet (a ++ b) -> ~(In x a) -> ~(In x b) ->
list_norepet (a ++ x :: b)).
induction a; simpl; intros.
constructor; auto.
inversion H. constructor. red; intro.
elim (in_app_or _ _ _ H6); intro.
elim H4. apply in_or_app. tauto.
- elim H7; intro. subst a. elim H0. left. auto.
+ elim H7; intro. subst a. elim H0. left. auto.
elim H4. apply in_or_app. tauto.
auto.
induction a; simpl; intros.
rewrite <- app_nil_end. auto.
- inversion H0. apply H. auto.
+ inversion H0. apply H. auto.
red; intro; elim H3. apply in_or_app. tauto.
red; intro; elim H3. apply in_or_app. tauto.
Qed.
@@ -1085,10 +1085,10 @@ Proof.
tauto.
inversion H; subst. rewrite IHl1 in H3. rewrite in_app in H2.
intuition.
- constructor; auto. red; intros. elim H2; intro. congruence. auto.
- destruct H as [B [C D]]. inversion B; subst.
- constructor. rewrite in_app. intuition. elim (D a a); auto. apply in_eq.
- rewrite IHl1. intuition. red; intros. apply D; auto. apply in_cons; auto.
+ constructor; auto. red; intros. elim H2; intro. congruence. auto.
+ destruct H as [B [C D]]. inversion B; subst.
+ constructor. rewrite in_app. intuition. elim (D a a); auto. apply in_eq.
+ rewrite IHl1. intuition. red; intros. apply D; auto. apply in_cons; auto.
Qed.
Lemma list_norepet_append:
@@ -1133,7 +1133,7 @@ Lemma is_tail_cons_left:
forall (A: Type) (i: A) c1 c2, is_tail (i :: c1) c2 -> is_tail c1 c2.
Proof.
induction c2; intros; inversion H.
- constructor. constructor. constructor. auto.
+ constructor. constructor. constructor. auto.
Qed.
Hint Resolve is_tail_refl is_tail_cons is_tail_in is_tail_cons_left: coqlib.
@@ -1171,10 +1171,10 @@ Inductive list_forall2: list A -> list B -> Prop :=
Lemma list_forall2_app:
forall a2 b2 a1 b1,
- list_forall2 a1 b1 -> list_forall2 a2 b2 ->
+ list_forall2 a1 b1 -> list_forall2 a2 b2 ->
list_forall2 (a1 ++ a2) (b1 ++ b2).
Proof.
- induction 1; intros; simpl. auto. constructor; auto.
+ induction 1; intros; simpl. auto. constructor; auto.
Qed.
Lemma list_forall2_length:
@@ -1195,7 +1195,7 @@ Lemma list_forall2_imply:
Proof.
induction 1; intros.
constructor.
- constructor. auto with coqlib. apply IHlist_forall2; auto.
+ constructor. auto with coqlib. apply IHlist_forall2; auto.
intros. auto with coqlib.
Qed.
@@ -1210,7 +1210,7 @@ Fixpoint list_drop (A: Type) (n: nat) (x: list A) {struct n} : list A :=
Lemma list_drop_incl:
forall (A: Type) (x: A) n (l: list A), In x (list_drop n l) -> In x l.
Proof.
- induction n; simpl; intros. auto.
+ induction n; simpl; intros. auto.
destruct l; auto with coqlib.
Qed.
@@ -1225,7 +1225,7 @@ Lemma list_map_drop:
forall (A B: Type) (f: A -> B) n (l: list A),
list_drop n (map f l) = map f (list_drop n l).
Proof.
- induction n; simpl; intros. auto.
+ induction n; simpl; intros. auto.
destruct l; simpl; auto.
Qed.
@@ -1267,7 +1267,7 @@ Qed.
Lemma proj_sumbool_is_true:
forall (P: Prop) (a: {P}+{~P}), P -> proj_sumbool a = true.
Proof.
- intros. unfold proj_sumbool. destruct a. auto. contradiction.
+ intros. unfold proj_sumbool. destruct a. auto. contradiction.
Qed.
Ltac InvBooleans :=
@@ -1306,7 +1306,7 @@ Lemma dec_eq_sym:
(if dec_eq x y then ifso else ifnot) =
(if dec_eq y x then ifso else ifnot).
Proof.
- intros. destruct (dec_eq x y).
+ intros. destruct (dec_eq x y).
subst y. rewrite dec_eq_true. auto.
rewrite dec_eq_false; auto.
Qed.
@@ -1352,22 +1352,22 @@ Inductive lex_ord: A*B -> A*B -> Prop :=
| lex_ord_right: forall a b1 b2,
ordB b1 b2 -> lex_ord (a,b1) (a,b2).
-Lemma wf_lex_ord:
+Lemma wf_lex_ord:
well_founded ordA -> well_founded ordB -> well_founded lex_ord.
Proof.
intros Awf Bwf.
assert (forall a, Acc ordA a -> forall b, Acc ordB b -> Acc lex_ord (a, b)).
induction 1. induction 1. constructor; intros. inv H3.
apply H0. auto. apply Bwf.
- apply H2; auto.
+ apply H2; auto.
red; intros. destruct a as [a b]. apply H; auto.
Qed.
Lemma transitive_lex_ord:
transitive _ ordA -> transitive _ ordB -> transitive _ lex_ord.
Proof.
- intros trA trB; red; intros.
- inv H; inv H0.
+ intros trA trB; red; intros.
+ inv H; inv H0.
left; eapply trA; eauto.
left; auto.
left; auto.
diff --git a/lib/FSetAVLplus.v b/lib/FSetAVLplus.v
index eab427be..f16805c6 100644
--- a/lib/FSetAVLplus.v
+++ b/lib/FSetAVLplus.v
@@ -65,7 +65,7 @@ Proof.
- discriminate.
- destruct (above_low_bound t1) eqn: LB; [destruct (below_high_bound t1) eqn: HB | idtac].
+ (* in interval *)
- exists t1; split; auto. apply Raw.IsRoot. auto.
+ exists t1; split; auto. apply Raw.IsRoot. auto.
+ (* above interval *)
exploit IHm1; auto. intros [x' [A B]]. exists x'; split; auto. apply Raw.InLeft; auto.
+ (* below interval *)
@@ -80,7 +80,7 @@ Lemma raw_mem_between_2:
Proof.
induction 1; simpl; intros.
- inv H.
-- rewrite Raw.In_node_iff in H1.
+- rewrite Raw.In_node_iff in H1.
destruct (above_low_bound x0) eqn: LB; [destruct (below_high_bound x0) eqn: HB | idtac].
+ (* in interval *)
auto.
@@ -98,7 +98,7 @@ Theorem mem_between_1:
mem_between s = true ->
exists x, In x s /\ above_low_bound x = true /\ below_high_bound x = true.
Proof.
- intros. apply raw_mem_between_1. auto.
+ intros. apply raw_mem_between_1. auto.
Qed.
Theorem mem_between_2:
@@ -138,9 +138,9 @@ Remark In_raw_elements_between_1:
Proof.
induction m; simpl; intros.
- inv H.
-- rewrite Raw.In_node_iff.
+- rewrite Raw.In_node_iff.
destruct (above_low_bound t1) eqn:LB; [destruct (below_high_bound t1) eqn: RB | idtac]; simpl in H.
- + rewrite Raw.join_spec in H. intuition.
+ + rewrite Raw.join_spec in H. intuition.
+ left; apply IHm1; auto.
+ right; right; apply IHm2; auto.
Qed.
@@ -174,7 +174,7 @@ Proof.
- inv H.
- destruct (above_low_bound t1) eqn:LB; [destruct (below_high_bound t1) eqn: RB | idtac]; simpl in H.
+ rewrite Raw.join_spec in H. intuition.
- apply above_monotone with t1; auto.
+ apply above_monotone with t1; auto.
apply below_monotone with t1; auto.
+ auto.
+ auto.
@@ -190,7 +190,7 @@ Proof.
- auto.
- rewrite Raw.In_node_iff in H1.
destruct (above_low_bound x0) eqn:LB; [destruct (below_high_bound x0) eqn: RB | idtac].
- + rewrite Raw.join_spec. intuition.
+ + rewrite Raw.join_spec. intuition.
+ assert (X.eq x x0 \/ X.lt x0 x -> False).
{ intros. exploit below_monotone; eauto. congruence. }
intuition. elim H7. apply g. auto.
@@ -204,7 +204,7 @@ Theorem elements_between_iff:
In x (elements_between s) <-> In x s /\ above_low_bound x = true /\ below_high_bound x = true.
Proof.
intros. unfold elements_between, In; simpl. split.
- intros. split. apply In_raw_elements_between_1; auto. eapply In_raw_elements_between_2; eauto.
+ intros. split. apply In_raw_elements_between_1; auto. eapply In_raw_elements_between_2; eauto.
intros [A [B C]]. apply In_raw_elements_between_3; auto. apply MSet.is_ok.
Qed.
@@ -254,24 +254,24 @@ Lemma raw_for_all_between_1:
pred x = true.
Proof.
induction 1; simpl; intros.
-- inv H0.
+- inv H0.
- destruct (above_low_bound x0) eqn: LB; [destruct (below_high_bound x0) eqn: HB | idtac].
+ (* in interval *)
destruct (andb_prop _ _ H1) as [P C]. destruct (andb_prop _ _ P) as [A B]. clear H1 P.
inv H2.
- * erewrite pred_compat; eauto.
+ * erewrite pred_compat; eauto.
* apply IHbst1; auto.
* apply IHbst2; auto.
+ (* above interval *)
inv H2.
- * assert (below_high_bound x0 = true) by (apply below_monotone with x; auto).
+ * assert (below_high_bound x0 = true) by (apply below_monotone with x; auto).
congruence.
* apply IHbst1; auto.
* assert (below_high_bound x0 = true) by (apply below_monotone with x; auto).
congruence.
+ (* below interval *)
inv H2.
- * assert (above_low_bound x0 = true) by (apply above_monotone with x; auto).
+ * assert (above_low_bound x0 = true) by (apply above_monotone with x; auto).
congruence.
* assert (above_low_bound x0 = true) by (apply above_monotone with x; auto).
congruence.
@@ -290,7 +290,7 @@ Proof.
+ (* in interval *)
rewrite IHbst1. rewrite (H1 x). rewrite IHbst2. auto.
intros. apply H1; auto. apply Raw.InRight; auto.
- apply Raw.IsRoot. reflexivity. auto. auto.
+ apply Raw.IsRoot. reflexivity. auto. auto.
intros. apply H1; auto. apply Raw.InLeft; auto.
+ (* above interval *)
apply IHbst1. intros. apply H1; auto. apply Raw.InLeft; auto.
@@ -303,7 +303,7 @@ Theorem for_all_between_iff:
for_all_between s = true <-> (forall x, In x s -> above_low_bound x = true -> below_high_bound x = true -> pred x = true).
Proof.
unfold for_all_between; intros; split; intros.
-- eapply raw_for_all_between_1; eauto. apply MSet.is_ok.
+- eapply raw_for_all_between_1; eauto. apply MSet.is_ok.
- apply raw_for_all_between_2; auto. apply MSet.is_ok.
Qed.
@@ -337,10 +337,10 @@ Remark In_raw_partition_between_1:
Proof.
induction m; simpl; intros.
- inv H.
-- destruct (raw_partition_between m1) as [l1 l2] eqn:LEQ; simpl in *.
+- destruct (raw_partition_between m1) as [l1 l2] eqn:LEQ; simpl in *.
destruct (raw_partition_between m2) as [r1 r2] eqn:REQ; simpl in *.
destruct (above_low_bound t1) eqn:LB; [destruct (below_high_bound t1) eqn: RB | idtac]; simpl in H.
- + rewrite Raw.join_spec in H. rewrite Raw.In_node_iff. intuition.
+ + rewrite Raw.join_spec in H. rewrite Raw.In_node_iff. intuition.
+ rewrite Raw.In_node_iff. intuition.
+ rewrite Raw.In_node_iff. intuition.
Qed.
@@ -351,10 +351,10 @@ Remark In_raw_partition_between_2:
Proof.
induction m; simpl; intros.
- inv H.
-- destruct (raw_partition_between m1) as [l1 l2] eqn:LEQ; simpl in *.
+- destruct (raw_partition_between m1) as [l1 l2] eqn:LEQ; simpl in *.
destruct (raw_partition_between m2) as [r1 r2] eqn:REQ; simpl in *.
destruct (above_low_bound t1) eqn:LB; [destruct (below_high_bound t1) eqn: RB | idtac]; simpl in H.
- + rewrite Raw.concat_spec in H. rewrite Raw.In_node_iff. intuition.
+ + rewrite Raw.concat_spec in H. rewrite Raw.In_node_iff. intuition.
+ rewrite Raw.join_spec in H. rewrite Raw.In_node_iff. intuition.
+ rewrite Raw.join_spec in H. rewrite Raw.In_node_iff. intuition.
Qed.
@@ -364,22 +364,22 @@ Lemma raw_partition_between_ok:
Proof.
induction 1; simpl.
- split; constructor.
-- destruct IHbst1 as [L1 L2]. destruct IHbst2 as [R1 R2].
- destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
+- destruct IHbst1 as [L1 L2]. destruct IHbst2 as [R1 R2].
+ destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
destruct (raw_partition_between r) as [r1 r2] eqn:REQ; simpl in *.
destruct (above_low_bound x) eqn:LB; [destruct (below_high_bound x) eqn: RB | idtac]; simpl.
+ split.
apply Raw.join_ok; auto.
- red; intros. apply l0. apply In_raw_partition_between_1. rewrite LEQ; auto.
+ red; intros. apply l0. apply In_raw_partition_between_1. rewrite LEQ; auto.
red; intros. apply g. apply In_raw_partition_between_1. rewrite REQ; auto.
apply Raw.concat_ok; auto.
- intros. transitivity x.
- apply l0. apply In_raw_partition_between_2. rewrite LEQ; auto.
+ intros. transitivity x.
+ apply l0. apply In_raw_partition_between_2. rewrite LEQ; auto.
apply g. apply In_raw_partition_between_2. rewrite REQ; auto.
+ split.
auto.
apply Raw.join_ok; auto.
- red; intros. apply l0. apply In_raw_partition_between_2. rewrite LEQ; auto.
+ red; intros. apply l0. apply In_raw_partition_between_2. rewrite LEQ; auto.
+ split.
auto.
apply Raw.join_ok; auto.
@@ -397,11 +397,11 @@ Remark In_raw_partition_between_3:
Proof.
induction m; simpl; intros.
- inv H.
-- destruct (raw_partition_between m1) as [l1 l2] eqn:LEQ; simpl in *.
+- destruct (raw_partition_between m1) as [l1 l2] eqn:LEQ; simpl in *.
destruct (raw_partition_between m2) as [r1 r2] eqn:REQ; simpl in *.
destruct (above_low_bound t1) eqn:LB; [destruct (below_high_bound t1) eqn: RB | idtac]; simpl in H.
+ rewrite Raw.join_spec in H. intuition.
- apply above_monotone with t1; auto.
+ apply above_monotone with t1; auto.
apply below_monotone with t1; auto.
+ auto.
+ auto.
@@ -414,17 +414,17 @@ Remark In_raw_partition_between_4:
Proof.
induction 1; simpl; intros.
- inv H.
-- destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
+- destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
destruct (raw_partition_between r) as [r1 r2] eqn:REQ; simpl in *.
destruct (above_low_bound x0) eqn:LB; [destruct (below_high_bound x0) eqn: RB | idtac]; simpl in H.
- + simpl in H1. rewrite Raw.concat_spec in H1. intuition.
+ + simpl in H1. rewrite Raw.concat_spec in H1. intuition.
+ assert (forall y, X.eq y x0 \/ X.lt x0 y -> below_high_bound y = false).
- { intros. destruct (below_high_bound y) eqn:E; auto.
+ { intros. destruct (below_high_bound y) eqn:E; auto.
assert (below_high_bound x0 = true) by (apply below_monotone with y; auto).
congruence. }
simpl in H1. rewrite Raw.join_spec in H1. intuition.
+ assert (forall y, X.eq y x0 \/ X.lt y x0 -> above_low_bound y = false).
- { intros. destruct (above_low_bound y) eqn:E; auto.
+ { intros. destruct (above_low_bound y) eqn:E; auto.
assert (above_low_bound x0 = true) by (apply above_monotone with y; auto).
congruence. }
simpl in H1. rewrite Raw.join_spec in H1. intuition.
@@ -438,23 +438,23 @@ Remark In_raw_partition_between_5:
Proof.
induction 1; simpl; intros.
- inv H.
-- destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
+- destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
destruct (raw_partition_between r) as [r1 r2] eqn:REQ; simpl in *.
destruct (above_low_bound x0) eqn:LB; [destruct (below_high_bound x0) eqn: RB | idtac]; simpl in H.
- + simpl. rewrite Raw.join_spec. inv H1.
+ + simpl. rewrite Raw.join_spec. inv H1.
auto.
- right; left; apply IHbst1; auto.
+ right; left; apply IHbst1; auto.
right; right; apply IHbst2; auto.
- + simpl. inv H1.
+ + simpl. inv H1.
assert (below_high_bound x0 = true) by (apply below_monotone with x; auto).
congruence.
- auto.
- assert (below_high_bound x0 = true) by (apply below_monotone with x; auto).
+ auto.
+ assert (below_high_bound x0 = true) by (apply below_monotone with x; auto).
congruence.
- + simpl. inv H1.
+ + simpl. inv H1.
assert (above_low_bound x0 = true) by (apply above_monotone with x; auto).
congruence.
- assert (above_low_bound x0 = true) by (apply above_monotone with x; auto).
+ assert (above_low_bound x0 = true) by (apply above_monotone with x; auto).
congruence.
eauto.
Qed.
@@ -467,7 +467,7 @@ Remark In_raw_partition_between_6:
Proof.
induction 1; simpl; intros.
- inv H.
-- destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
+- destruct (raw_partition_between l) as [l1 l2] eqn:LEQ; simpl in *.
destruct (raw_partition_between r) as [r1 r2] eqn:REQ; simpl in *.
destruct (above_low_bound x0) eqn:LB; [destruct (below_high_bound x0) eqn: RB | idtac]; simpl in H.
+ simpl. rewrite Raw.concat_spec. inv H1.
@@ -476,11 +476,11 @@ Proof.
destruct H2; congruence.
left; apply IHbst1; auto.
right; apply IHbst2; auto.
- + simpl. rewrite Raw.join_spec. inv H1.
+ + simpl. rewrite Raw.join_spec. inv H1.
auto.
- right; left; apply IHbst1; auto.
+ right; left; apply IHbst1; auto.
auto.
- + simpl. rewrite Raw.join_spec. inv H1.
+ + simpl. rewrite Raw.join_spec. inv H1.
auto.
auto.
right; right; apply IHbst2; auto.
@@ -496,7 +496,7 @@ Theorem partition_between_iff_1:
In x (fst (partition_between s)) <->
In x s /\ above_low_bound x = true /\ below_high_bound x = true.
Proof.
- intros. unfold partition_between, In; simpl. split.
+ intros. unfold partition_between, In; simpl. split.
intros. split. apply In_raw_partition_between_1; auto. eapply In_raw_partition_between_3; eauto.
intros [A [B C]]. apply In_raw_partition_between_5; auto. apply MSet.is_ok.
Qed.
@@ -506,7 +506,7 @@ Theorem partition_between_iff_2:
In x (snd (partition_between s)) <->
In x s /\ (above_low_bound x = false \/ below_high_bound x = false).
Proof.
- intros. unfold partition_between, In; simpl. split.
+ intros. unfold partition_between, In; simpl. split.
intros. split. apply In_raw_partition_between_2; auto. eapply In_raw_partition_between_4; eauto. apply MSet.is_ok.
intros [A B]. apply In_raw_partition_between_6; auto. apply MSet.is_ok.
Qed.
diff --git a/lib/Fappli_IEEE_extra.v b/lib/Fappli_IEEE_extra.v
index 3de7b103..fe7f7c6d 100644
--- a/lib/Fappli_IEEE_extra.v
+++ b/lib/Fappli_IEEE_extra.v
@@ -110,10 +110,10 @@ Proof.
subst; left; f_equal; f_equal; apply UIP_bool.
destruct (positive_eq_dec m m0); try_not_eq;
destruct (Z_eq_dec e e1); try solve [right; intro H; inversion H; congruence];
- subst; left; f_equal; apply UIP_bool.
+ subst; left; f_equal; apply UIP_bool.
destruct (positive_eq_dec m m0); try_not_eq;
destruct (Z_eq_dec e e1); try solve [right; intro H; inversion H; congruence];
- subst; left; f_equal; apply UIP_bool.
+ subst; left; f_equal; apply UIP_bool.
Defined.
(** ** Conversion from an integer to a FP number *)
@@ -134,16 +134,16 @@ Lemma integer_representable_n2p:
-2^prec < n < 2^prec -> 0 <= p -> p <= emax - prec ->
integer_representable (n * 2^p).
Proof.
- intros; split.
+ intros; split.
- red in prec_gt_0_. replace (Z.abs (n * 2^p)) with (Z.abs n * 2^p).
rewrite int_upper_bound_eq.
- apply Zmult_le_compat. zify; omega. apply (Zpower_le radix2); omega.
- zify; omega. apply (Zpower_ge_0 radix2).
- rewrite Z.abs_mul. f_equal. rewrite Z.abs_eq. auto. apply (Zpower_ge_0 radix2).
+ apply Zmult_le_compat. zify; omega. apply (Zpower_le radix2); omega.
+ zify; omega. apply (Zpower_ge_0 radix2).
+ rewrite Z.abs_mul. f_equal. rewrite Z.abs_eq. auto. apply (Zpower_ge_0 radix2).
- apply generic_format_FLT. exists (Float radix2 n p).
unfold F2R; simpl.
- split. rewrite <- Z2R_Zpower by auto. apply Z2R_mult.
- split. zify; omega.
+ split. rewrite <- Z2R_Zpower by auto. apply Z2R_mult.
+ split. zify; omega.
unfold emin; red in prec_gt_0_; omega.
Qed.
@@ -152,13 +152,13 @@ Lemma integer_representable_2p:
0 <= p <= emax - 1 ->
integer_representable (2^p).
Proof.
- intros; split.
-- red in prec_gt_0_.
- rewrite Z.abs_eq by (apply (Zpower_ge_0 radix2)).
- apply Zle_trans with (2^(emax-1)).
+ intros; split.
+- red in prec_gt_0_.
+ rewrite Z.abs_eq by (apply (Zpower_ge_0 radix2)).
+ apply Zle_trans with (2^(emax-1)).
apply (Zpower_le radix2); omega.
assert (2^emax = 2^(emax-1)*2).
- { change 2 with (2^1) at 3. rewrite <- (Zpower_plus radix2) by omega.
+ { change 2 with (2^1) at 3. rewrite <- (Zpower_plus radix2) by omega.
f_equal. omega. }
assert (2^(emax - prec) <= 2^(emax - 1)).
{ apply (Zpower_le radix2). omega. }
@@ -166,7 +166,7 @@ Proof.
- red in prec_gt_0_.
apply generic_format_FLT. exists (Float radix2 1 p).
unfold F2R; simpl.
- split. rewrite Rmult_1_l. rewrite <- Z2R_Zpower. auto. omega.
+ split. rewrite Rmult_1_l. rewrite <- Z2R_Zpower. auto. omega.
split. change 1 with (2^0). apply (Zpower_lt radix2). omega. auto.
unfold emin; omega.
Qed.
@@ -174,7 +174,7 @@ Qed.
Lemma integer_representable_opp:
forall n, integer_representable n -> integer_representable (-n).
Proof.
- intros n (A & B); split. rewrite Z.abs_opp. auto.
+ intros n (A & B); split. rewrite Z.abs_opp. auto.
rewrite Z2R_opp. apply generic_format_opp; auto.
Qed.
@@ -186,10 +186,10 @@ Proof.
intros. red in prec_gt_0_.
destruct (Z.eq_dec n (2^prec)); [idtac | destruct (Z.eq_dec n (-2^prec))].
- rewrite e. rewrite <- (Zpower_plus radix2) by omega.
- apply integer_representable_2p. omega.
-- rewrite e. rewrite <- Zopp_mult_distr_l. apply integer_representable_opp.
+ apply integer_representable_2p. omega.
+- rewrite e. rewrite <- Zopp_mult_distr_l. apply integer_representable_opp.
rewrite <- (Zpower_plus radix2) by omega.
- apply integer_representable_2p. omega.
+ apply integer_representable_2p. omega.
- apply integer_representable_n2p; omega.
Qed.
@@ -198,7 +198,7 @@ Lemma integer_representable_n:
Proof.
red in prec_gt_0_. intros.
replace n with (n * 2^0) by (change (2^0) with 1; ring).
- apply integer_representable_n2p_wide. auto. omega. omega.
+ apply integer_representable_n2p_wide. auto. omega. omega.
Qed.
Lemma round_int_no_overflow:
@@ -207,19 +207,19 @@ Lemma round_int_no_overflow:
(Rabs (round radix2 fexp (round_mode mode_NE) (Z2R n)) < bpow radix2 emax)%R.
Proof.
intros. red in prec_gt_0_.
- rewrite <- round_NE_abs.
+ rewrite <- round_NE_abs.
apply Rle_lt_trans with (Z2R (2^emax - 2^(emax-prec))).
apply round_le_generic. apply fexp_correct; auto. apply valid_rnd_N.
apply generic_format_FLT. exists (Float radix2 (2^prec-1) (emax-prec)).
rewrite int_upper_bound_eq. unfold F2R; simpl.
- split. rewrite <- Z2R_Zpower by omega. rewrite <- Z2R_mult. auto.
+ split. rewrite <- Z2R_Zpower by omega. rewrite <- Z2R_mult. auto.
split. assert (0 < 2^prec) by (apply (Zpower_gt_0 radix2); omega). zify; omega.
unfold emin; omega.
rewrite <- Z2R_abs. apply Z2R_le. auto.
- rewrite <- Z2R_Zpower by omega. apply Z2R_lt. simpl.
+ rewrite <- Z2R_Zpower by omega. apply Z2R_lt. simpl.
assert (0 < 2^(emax-prec)) by (apply (Zpower_gt_0 radix2); omega).
omega.
- apply fexp_correct. auto.
+ apply fexp_correct. auto.
Qed.
(** Conversion from an integer. Round to nearest. *)
@@ -237,17 +237,17 @@ Theorem BofZ_correct:
else
B2FF prec emax (BofZ n) = binary_overflow prec emax mode_NE (Zlt_bool n 0).
Proof.
- intros.
+ intros.
generalize (binary_normalize_correct prec emax prec_gt_0_ Hmax mode_NE n 0 false).
- fold emin; fold fexp; fold (BofZ n).
- replace (F2R {| Fnum := n; Fexp := 0 |}) with (Z2R n).
- destruct Rlt_bool.
+ fold emin; fold fexp; fold (BofZ n).
+ replace (F2R {| Fnum := n; Fexp := 0 |}) with (Z2R n).
+ destruct Rlt_bool.
- intros (A & B & C). split; [|split].
+ auto.
+ auto.
- + rewrite C. change 0%R with (Z2R 0). rewrite Rcompare_Z2R.
+ + rewrite C. change 0%R with (Z2R 0). rewrite Rcompare_Z2R.
unfold Zlt_bool. auto.
-- intros A; rewrite A. f_equal. change 0%R with (Z2R 0).
+- intros A; rewrite A. f_equal. change 0%R with (Z2R 0).
generalize (Zlt_bool_spec n 0); intros SPEC; inversion SPEC.
apply Rlt_bool_true; apply Z2R_lt; auto.
apply Rlt_bool_false; apply Z2R_le; auto.
@@ -262,7 +262,7 @@ Theorem BofZ_finite:
/\ Bsign _ _ (BofZ n) = Zlt_bool n 0%Z.
Proof.
intros.
- generalize (BofZ_correct n). rewrite Rlt_bool_true. auto.
+ generalize (BofZ_correct n). rewrite Rlt_bool_true. auto.
apply round_int_no_overflow; auto.
Qed.
@@ -274,7 +274,7 @@ Theorem BofZ_representable:
/\ Bsign _ _ (BofZ n) = (n <? 0).
Proof.
intros. destruct H as (P & Q). destruct (BofZ_finite n) as (A & B & C). auto.
- intuition. rewrite A. apply round_generic. apply valid_rnd_round_mode. auto.
+ intuition. rewrite A. apply round_generic. apply valid_rnd_round_mode. auto.
Qed.
Theorem BofZ_exact:
@@ -291,21 +291,21 @@ Lemma BofZ_finite_pos0:
forall n,
Z.abs n <= 2^emax - 2^(emax-prec) -> is_finite_pos0 (BofZ n) = true.
Proof.
- intros.
+ intros.
generalize (binary_normalize_correct prec emax prec_gt_0_ Hmax mode_NE n 0 false).
- fold emin; fold fexp; fold (BofZ n).
+ fold emin; fold fexp; fold (BofZ n).
replace (F2R {| Fnum := n; Fexp := 0 |}) with (Z2R n) by
(unfold F2R; simpl; ring).
rewrite Rlt_bool_true by (apply round_int_no_overflow; auto).
intros (A & B & C).
- destruct (BofZ n); auto; try discriminate.
- simpl in *. rewrite C. change 0%R with (Z2R 0). rewrite Rcompare_Z2R.
+ destruct (BofZ n); auto; try discriminate.
+ simpl in *. rewrite C. change 0%R with (Z2R 0). rewrite Rcompare_Z2R.
generalize (Zcompare_spec n 0); intros SPEC; inversion SPEC; auto.
assert ((round radix2 fexp ZnearestE (Z2R n) <= -1)%R).
{ change (-1)%R with (Z2R (-1)).
apply round_le_generic. apply fexp_correct. auto. apply valid_rnd_N.
- apply (integer_representable_opp 1).
- apply (integer_representable_2p 0).
+ apply (integer_representable_opp 1).
+ apply (integer_representable_2p 0).
red in prec_gt_0_; omega.
apply Z2R_le; omega.
}
@@ -325,16 +325,16 @@ Qed.
(** Commutation properties with addition, subtraction, multiplication. *)
Theorem BofZ_plus:
- forall nan p q,
+ forall nan p q,
integer_representable p -> integer_representable q ->
Bplus _ _ _ Hmax nan mode_NE (BofZ p) (BofZ q) = BofZ (p + q).
Proof.
- intros.
- destruct (BofZ_representable p) as (A & B & C); auto.
+ intros.
+ destruct (BofZ_representable p) as (A & B & C); auto.
destruct (BofZ_representable q) as (D & E & F); auto.
generalize (Bplus_correct _ _ _ Hmax nan mode_NE (BofZ p) (BofZ q) B E).
fold emin; fold fexp.
- rewrite A, D. rewrite <- Z2R_plus.
+ rewrite A, D. rewrite <- Z2R_plus.
generalize (BofZ_correct (p + q)). destruct Rlt_bool.
- intros (P & Q & R) (U & V & W).
apply B2R_Bsign_inj; auto.
@@ -342,29 +342,29 @@ Proof.
rewrite R, W, C, F.
change 0%R with (Z2R 0). rewrite Rcompare_Z2R. unfold Zlt_bool at 3.
generalize (Zcompare_spec (p + q) 0); intros SPEC; inversion SPEC; auto.
- assert (EITHER: 0 <= p \/ 0 <= q) by omega.
+ assert (EITHER: 0 <= p \/ 0 <= q) by omega.
destruct EITHER; [apply andb_false_intro1 | apply andb_false_intro2];
apply Zlt_bool_false; auto.
-- intros P (U & V).
- apply B2FF_inj.
- rewrite P, U, C. f_equal. rewrite C, F in V.
- generalize (Zlt_bool_spec p 0) (Zlt_bool_spec q 0). rewrite <- V.
+- intros P (U & V).
+ apply B2FF_inj.
+ rewrite P, U, C. f_equal. rewrite C, F in V.
+ generalize (Zlt_bool_spec p 0) (Zlt_bool_spec q 0). rewrite <- V.
intros SPEC1 SPEC2; inversion SPEC1; inversion SPEC2; try congruence; symmetry.
apply Zlt_bool_true; omega.
apply Zlt_bool_false; omega.
Qed.
Theorem BofZ_minus:
- forall nan p q,
+ forall nan p q,
integer_representable p -> integer_representable q ->
Bminus _ _ _ Hmax nan mode_NE (BofZ p) (BofZ q) = BofZ (p - q).
Proof.
- intros.
- destruct (BofZ_representable p) as (A & B & C); auto.
+ intros.
+ destruct (BofZ_representable p) as (A & B & C); auto.
destruct (BofZ_representable q) as (D & E & F); auto.
generalize (Bminus_correct _ _ _ Hmax nan mode_NE (BofZ p) (BofZ q) B E).
fold emin; fold fexp.
- rewrite A, D. rewrite <- Z2R_minus.
+ rewrite A, D. rewrite <- Z2R_minus.
generalize (BofZ_correct (p - q)). destruct Rlt_bool.
- intros (P & Q & R) (U & V & W).
apply B2R_Bsign_inj; auto.
@@ -372,14 +372,14 @@ Proof.
rewrite R, W, C, F.
change 0%R with (Z2R 0). rewrite Rcompare_Z2R. unfold Zlt_bool at 3.
generalize (Zcompare_spec (p - q) 0); intros SPEC; inversion SPEC; auto.
- assert (EITHER: 0 <= p \/ q < 0) by omega.
+ assert (EITHER: 0 <= p \/ q < 0) by omega.
destruct EITHER; [apply andb_false_intro1 | apply andb_false_intro2].
rewrite Zlt_bool_false; auto.
rewrite Zlt_bool_true; auto.
-- intros P (U & V).
- apply B2FF_inj.
- rewrite P, U, C. f_equal. rewrite C, F in V.
- generalize (Zlt_bool_spec p 0) (Zlt_bool_spec q 0). rewrite V.
+- intros P (U & V).
+ apply B2FF_inj.
+ rewrite P, U, C. f_equal. rewrite C, F in V.
+ generalize (Zlt_bool_spec p 0) (Zlt_bool_spec q 0). rewrite V.
intros SPEC1 SPEC2; inversion SPEC1; inversion SPEC2; symmetry.
rewrite <- H3 in H1; discriminate.
apply Zlt_bool_true; omega.
@@ -388,20 +388,20 @@ Proof.
Qed.
Theorem BofZ_mult:
- forall nan p q,
+ forall nan p q,
integer_representable p -> integer_representable q ->
0 < q ->
Bmult _ _ _ Hmax nan mode_NE (BofZ p) (BofZ q) = BofZ (p * q).
Proof.
- intros.
+ intros.
assert (SIGN: xorb (p <? 0) (q <? 0) = (p * q <? 0)).
{
rewrite (Zlt_bool_false q) by omega.
generalize (Zlt_bool_spec p 0); intros SPEC; inversion SPEC; simpl; symmetry.
apply Zlt_bool_true. rewrite Z.mul_comm. apply Z.mul_pos_neg; omega.
- apply Zlt_bool_false. apply Zsame_sign_imp; omega.
+ apply Zlt_bool_false. apply Zsame_sign_imp; omega.
}
- destruct (BofZ_representable p) as (A & B & C); auto.
+ destruct (BofZ_representable p) as (A & B & C); auto.
destruct (BofZ_representable q) as (D & E & F); auto.
generalize (Bmult_correct _ _ _ Hmax nan mode_NE (BofZ p) (BofZ q)).
fold emin; fold fexp.
@@ -424,27 +424,27 @@ Theorem BofZ_mult_2p:
Bmult _ _ _ Hmax nan mode_NE (BofZ x) (BofZ (2^p)) = BofZ (x * 2^p).
Proof.
intros.
- destruct (Z.eq_dec x 0).
+ destruct (Z.eq_dec x 0).
- subst x. apply BofZ_mult.
- apply integer_representable_n.
+ apply integer_representable_n.
generalize (Zpower_ge_0 radix2 prec). simpl; omega.
- apply integer_representable_2p. auto.
+ apply integer_representable_2p. auto.
apply (Zpower_gt_0 radix2).
omega.
- assert (Z2R x <> 0%R) by (apply (Z2R_neq _ _ n)).
destruct (BofZ_finite x H) as (A & B & C).
- destruct (BofZ_representable (2^p)) as (D & E & F).
+ destruct (BofZ_representable (2^p)) as (D & E & F).
apply integer_representable_2p. auto.
assert (canonic_exp radix2 fexp (Z2R (x * 2^p)) =
canonic_exp radix2 fexp (Z2R x) + p).
{
- unfold canonic_exp, fexp. rewrite Z2R_mult.
- change (2^p) with (radix2^p). rewrite Z2R_Zpower by omega.
+ unfold canonic_exp, fexp. rewrite Z2R_mult.
+ change (2^p) with (radix2^p). rewrite Z2R_Zpower by omega.
rewrite ln_beta_mult_bpow by auto.
assert (prec + 1 <= ln_beta radix2 (Z2R x)).
- { rewrite <- (ln_beta_abs radix2 (Z2R x)).
- rewrite <- (ln_beta_bpow radix2 prec).
- apply ln_beta_le.
+ { rewrite <- (ln_beta_abs radix2 (Z2R x)).
+ rewrite <- (ln_beta_bpow radix2 prec).
+ apply ln_beta_le.
apply bpow_gt_0. rewrite <- Z2R_Zpower by (red in prec_gt_0_;omega).
rewrite <- Z2R_abs. apply Z2R_le; auto. }
unfold FLT_exp.
@@ -453,25 +453,25 @@ Proof.
assert (forall m, round radix2 fexp m (Z2R x) * Z2R (2^p) =
round radix2 fexp m (Z2R (x * 2^p)))%R.
{
- intros. unfold round, scaled_mantissa. rewrite H3.
- rewrite Z2R_mult. rewrite Z.opp_add_distr. rewrite bpow_plus.
+ intros. unfold round, scaled_mantissa. rewrite H3.
+ rewrite Z2R_mult. rewrite Z.opp_add_distr. rewrite bpow_plus.
set (a := Z2R x); set (b := bpow radix2 (- canonic_exp radix2 fexp a)).
replace (a * Z2R (2^p) * (b * bpow radix2 (-p)))%R with (a * b)%R.
- unfold F2R; simpl. rewrite Rmult_assoc. f_equal.
- rewrite bpow_plus. f_equal. apply (Z2R_Zpower radix2). omega.
+ unfold F2R; simpl. rewrite Rmult_assoc. f_equal.
+ rewrite bpow_plus. f_equal. apply (Z2R_Zpower radix2). omega.
transitivity ((a * b) * (Z2R (2^p) * bpow radix2 (-p)))%R.
- rewrite (Z2R_Zpower radix2). rewrite <- bpow_plus.
+ rewrite (Z2R_Zpower radix2). rewrite <- bpow_plus.
replace (p + -p) with 0 by omega. change (bpow radix2 0) with 1%R. ring.
- omega.
+ omega.
ring.
}
assert (forall m x,
round radix2 fexp (round_mode m) (round radix2 fexp (round_mode m) x) =
round radix2 fexp (round_mode m) x).
{
- intros. apply round_generic. apply valid_rnd_round_mode.
- apply generic_format_round. apply fexp_correct; auto.
- apply valid_rnd_round_mode.
+ intros. apply round_generic. apply valid_rnd_round_mode.
+ apply generic_format_round. apply fexp_correct; auto.
+ apply valid_rnd_round_mode.
}
assert (xorb (x <? 0) (2^p <? 0) = (x * 2^p <? 0)).
{
@@ -490,7 +490,7 @@ Proof.
rewrite P, U. auto.
rewrite R, W. auto.
apply is_finite_not_is_nan; auto.
-+ intros P U.
++ intros P U.
apply B2FF_inj. rewrite P, U. f_equal; auto.
Qed.
@@ -503,8 +503,8 @@ Lemma round_odd_flt:
round radix2 fexp (Znearest choice) x.
Proof.
intros. apply round_odd_prop. auto. apply fexp_correct; auto.
- apply exists_NE_FLT. right; omega.
- apply FLT_exp_valid. red; omega.
+ apply exists_NE_FLT. right; omega.
+ apply FLT_exp_valid. red; omega.
apply exists_NE_FLT. right; omega.
unfold fexp, FLT_exp; intros. zify; omega.
Qed.
@@ -517,29 +517,29 @@ Corollary round_odd_fix:
round radix2 fexp (Znearest choice) (round radix2 (FIX_exp p) Zrnd_odd x) =
round radix2 fexp (Znearest choice) x.
Proof.
- intros. destruct (Req_EM_T x 0%R).
-- subst x. rewrite round_0. auto. apply valid_rnd_odd.
+ intros. destruct (Req_EM_T x 0%R).
+- subst x. rewrite round_0. auto. apply valid_rnd_odd.
- set (prec' := ln_beta radix2 x - p).
set (emin' := emin - 2).
assert (PREC: ln_beta radix2 (bpow radix2 (prec + p + 1)) <= ln_beta radix2 x).
{ rewrite <- (ln_beta_abs radix2 x).
apply ln_beta_le; auto. apply bpow_gt_0. }
- rewrite ln_beta_bpow in PREC.
+ rewrite ln_beta_bpow in PREC.
assert (CANON: canonic_exp radix2 (FLT_exp emin' prec') x =
canonic_exp radix2 (FIX_exp p) x).
{
unfold canonic_exp, FLT_exp, FIX_exp.
replace (ln_beta radix2 x - prec') with p by (unfold prec'; omega).
- apply Z.max_l. unfold emin', emin. red in prec_gt_0_; omega.
+ apply Z.max_l. unfold emin', emin. red in prec_gt_0_; omega.
}
assert (RND: round radix2 (FIX_exp p) Zrnd_odd x =
round radix2 (FLT_exp emin' prec') Zrnd_odd x).
{
- unfold round, scaled_mantissa. rewrite CANON. auto.
+ unfold round, scaled_mantissa. rewrite CANON. auto.
}
- rewrite RND.
- apply round_odd_flt. auto.
- unfold prec'. red in prec_gt_0_; omega.
+ rewrite RND.
+ apply round_odd_flt. auto.
+ unfold prec'. red in prec_gt_0_; omega.
unfold prec'. omega.
unfold emin'. omega.
Qed.
@@ -552,7 +552,7 @@ Lemma Zrnd_odd_int:
Zrnd_odd (Z2R n * bpow radix2 (-p)) * 2^p =
int_round_odd n p.
Proof.
- intros.
+ intros.
assert (0 < 2^p) by (apply (Zpower_gt_0 radix2); omega).
assert (n = (n / 2^p) * 2^p + n mod 2^p) by (rewrite Zmult_comm; apply Z.div_mod; omega).
assert (0 <= n mod 2^p < 2^p) by (apply Z_mod_lt; omega).
@@ -562,16 +562,16 @@ Proof.
assert (bpow radix2 p * bpow radix2 (-p) = 1)%R.
{ rewrite <- bpow_plus. replace (p + -p) with 0 by omega. auto. }
assert (Z2R n * bpow radix2 (-p) = Z2R q + Z2R r * bpow radix2 (-p))%R.
- { rewrite H1. rewrite Z2R_plus, Z2R_mult.
+ { rewrite H1. rewrite Z2R_plus, Z2R_mult.
change (Z2R (2^p)) with (Z2R (radix2^p)).
rewrite Z2R_Zpower by omega. ring_simplify.
rewrite Rmult_assoc. rewrite H4. ring. }
assert (0 <= Z2R r < bpow radix2 p)%R.
- { split. change 0%R with (Z2R 0). apply Z2R_le; omega.
+ { split. change 0%R with (Z2R 0). apply Z2R_le; omega.
rewrite <- Z2R_Zpower by omega. apply Z2R_lt; tauto. }
assert (0 <= Z2R r * bpow radix2 (-p) < 1)%R.
- { generalize (bpow_gt_0 radix2 (-p)). intros.
- split. apply Rmult_le_pos; lra.
+ { generalize (bpow_gt_0 radix2 (-p)). intros.
+ split. apply Rmult_le_pos; lra.
rewrite <- H4. apply Rmult_lt_compat_r. auto. tauto. }
assert (Zfloor (Z2R n * bpow radix2 (-p)) = q).
{ apply Zfloor_imp. rewrite H5. rewrite Z2R_plus. change (Z2R 1) with 1%R. lra. }
@@ -585,7 +585,7 @@ Proof.
destruct (Z.eqb r 0) eqn:RZ.
apply Z.eqb_eq in RZ. rewrite RZ in H9. change (Z2R 0) with 0%R in H9.
rewrite Rmult_0_l in H9. congruence.
- rewrite Zceil_floor_neq by lra. rewrite H8.
+ rewrite Zceil_floor_neq by lra. rewrite H8.
change Zeven with Z.even. rewrite Zodd_even_bool. destruct (Z.even q); auto.
Qed.
@@ -593,12 +593,12 @@ Lemma int_round_odd_le:
forall p x y, 0 <= p ->
x <= y -> int_round_odd x p <= int_round_odd y p.
Proof.
- intros.
+ intros.
assert (Zrnd_odd (Z2R x * bpow radix2 (-p)) <= Zrnd_odd (Z2R y * bpow radix2 (-p))).
- { apply Zrnd_le. apply valid_rnd_odd. apply Rmult_le_compat_r. apply bpow_ge_0.
+ { apply Zrnd_le. apply valid_rnd_odd. apply Rmult_le_compat_r. apply bpow_ge_0.
apply Z2R_le; auto. }
- rewrite <- ! Zrnd_odd_int by auto.
- apply Zmult_le_compat_r. auto. apply (Zpower_ge_0 radix2).
+ rewrite <- ! Zrnd_odd_int by auto.
+ apply Zmult_le_compat_r. auto. apply (Zpower_ge_0 radix2).
Qed.
Lemma int_round_odd_exact:
@@ -607,7 +607,7 @@ Lemma int_round_odd_exact:
Proof.
intros. unfold int_round_odd. apply Znumtheory.Zdivide_mod in H0.
rewrite H0. simpl. rewrite Zmult_comm. symmetry. apply Z_div_exact_2.
- apply Zlt_gt. apply (Zpower_gt_0 radix2). auto. auto.
+ apply Zlt_gt. apply (Zpower_gt_0 radix2). auto. auto.
Qed.
Theorem BofZ_round_odd:
@@ -621,16 +621,16 @@ Proof.
intros x p PREC XRANGE PRANGE XGE.
assert (DIV: (2^p | 2^emax - 2^(emax - prec))).
{ rewrite int_upper_bound_eq. apply Z.divide_mul_r.
- exists (2^(emax - prec - p)). red in prec_gt_0_.
+ exists (2^(emax - prec - p)). red in prec_gt_0_.
rewrite <- (Zpower_plus radix2) by omega. f_equal; omega. }
assert (YRANGE: Z.abs (int_round_odd x p) <= 2^emax - 2^(emax-prec)).
{ apply Z.abs_le. split.
replace (-(2^emax - 2^(emax-prec))) with (int_round_odd (-(2^emax - 2^(emax-prec))) p).
apply int_round_odd_le; zify; omega.
- apply int_round_odd_exact. omega. apply Z.divide_opp_r. auto.
+ apply int_round_odd_exact. omega. apply Z.divide_opp_r. auto.
replace (2^emax - 2^(emax-prec)) with (int_round_odd (2^emax - 2^(emax-prec)) p).
apply int_round_odd_le; zify; omega.
- apply int_round_odd_exact. omega. auto. }
+ apply int_round_odd_exact. omega. auto. }
destruct (BofZ_finite x XRANGE) as (X1 & X2 & X3).
destruct (BofZ_finite (int_round_odd x p) YRANGE) as (Y1 & Y2 & Y3).
apply BofZ_finite_equal; auto.
@@ -641,8 +641,8 @@ Proof.
rewrite <- Zrnd_odd_int by omega.
unfold F2R; simpl. rewrite Z2R_mult. f_equal. apply (Z2R_Zpower radix2). omega.
}
- rewrite H. symmetry. apply round_odd_fix. auto. omega.
- rewrite <- Z2R_Zpower. rewrite <- Z2R_abs. apply Z2R_le; auto.
+ rewrite H. symmetry. apply round_odd_fix. auto. omega.
+ rewrite <- Z2R_Zpower. rewrite <- Z2R_abs. apply Z2R_le; auto.
red in prec_gt_0_; omega.
Qed.
@@ -653,13 +653,13 @@ Lemma int_round_odd_shifts:
Proof.
intros.
unfold int_round_odd. rewrite Z.shiftl_mul_pow2 by auto. f_equal.
- rewrite Z.shiftr_div_pow2 by auto.
- destruct (x mod 2^p =? 0) eqn:E. auto.
+ rewrite Z.shiftr_div_pow2 by auto.
+ destruct (x mod 2^p =? 0) eqn:E. auto.
assert (forall n, (if Z.odd n then n else n + 1) = Z.lor n 1).
{ destruct n; simpl; auto.
- destruct p0; auto.
+ destruct p0; auto.
destruct p0; auto. induction p0; auto. }
- simpl. apply H0.
+ simpl. apply H0.
Qed.
Lemma int_round_odd_bits:
@@ -669,20 +669,20 @@ Lemma int_round_odd_bits:
(forall i, p < i -> Z.testbit y i = Z.testbit x i) ->
int_round_odd x p = y.
Proof.
- intros until p; intros PPOS BELOW AT ABOVE.
- rewrite int_round_odd_shifts by auto.
- apply Z.bits_inj'. intros.
+ intros until p; intros PPOS BELOW AT ABOVE.
+ rewrite int_round_odd_shifts by auto.
+ apply Z.bits_inj'. intros.
generalize (Zcompare_spec n p); intros SPEC; inversion SPEC.
-- rewrite BELOW by auto. apply Z.shiftl_spec_low; auto.
+- rewrite BELOW by auto. apply Z.shiftl_spec_low; auto.
- subst n. rewrite AT. rewrite Z.shiftl_spec_high by omega.
replace (p - p) with 0 by omega.
destruct (x mod 2^p =? 0).
- + rewrite Z.shiftr_spec by omega. f_equal; omega.
- + rewrite Z.lor_spec. apply orb_true_r.
+ + rewrite Z.shiftr_spec by omega. f_equal; omega.
+ + rewrite Z.lor_spec. apply orb_true_r.
- rewrite ABOVE by auto. rewrite Z.shiftl_spec_high by omega.
destruct (x mod 2^p =? 0).
rewrite Z.shiftr_spec by omega. f_equal; omega.
- rewrite Z.lor_spec, Z.shiftr_spec by omega.
+ rewrite Z.lor_spec, Z.shiftr_spec by omega.
change 1 with (Z.ones 1). rewrite Z.ones_spec_high by omega. rewrite orb_false_r.
f_equal; omega.
Qed.
@@ -705,18 +705,18 @@ Theorem ZofB_correct:
ZofB f = if is_finite _ _ f then Some (Ztrunc (B2R _ _ f)) else None.
Proof.
destruct f; simpl; auto.
-- f_equal. symmetry. apply (Ztrunc_Z2R 0).
-- destruct e; f_equal.
+- f_equal. symmetry. apply (Ztrunc_Z2R 0).
+- destruct e; f_equal.
+ unfold F2R; simpl. rewrite Rmult_1_r. rewrite Ztrunc_Z2R. auto.
+ unfold F2R; simpl. rewrite <- Z2R_mult. rewrite Ztrunc_Z2R. auto.
- + unfold F2R; simpl. rewrite Z2R_cond_Zopp. rewrite <- cond_Ropp_mult_l.
+ + unfold F2R; simpl. rewrite Z2R_cond_Zopp. rewrite <- cond_Ropp_mult_l.
assert (EQ: forall x, Ztrunc (cond_Ropp b x) = cond_Zopp b (Ztrunc x)).
{
intros. destruct b; simpl; auto. apply Ztrunc_opp.
}
- rewrite EQ. f_equal.
+ rewrite EQ. f_equal.
generalize (Zpower_pos_gt_0 2 p (refl_equal _)); intros.
- rewrite Ztrunc_floor. symmetry. apply Zfloor_div. omega.
+ rewrite Ztrunc_floor. symmetry. apply Zfloor_div. omega.
apply Rmult_le_pos. apply (Z2R_le 0). compute; congruence.
apply Rlt_le. apply Rinv_0_lt_compat. apply (Z2R_lt 0). auto.
Qed.
@@ -726,13 +726,13 @@ Qed.
Remark Ztrunc_range_pos:
forall x, 0 < Ztrunc x -> (Z2R (Ztrunc x) <= x < Z2R (Ztrunc x + 1)%Z)%R.
Proof.
- intros.
+ intros.
rewrite Ztrunc_floor. split. apply Zfloor_lb. rewrite Z2R_plus. apply Zfloor_ub.
generalize (Rle_bool_spec 0%R x). intros RLE; inversion RLE; subst; clear RLE.
auto.
rewrite Ztrunc_ceil in H by lra. unfold Zceil in H.
assert (-x < 0)%R.
- { apply Rlt_le_trans with (Z2R (Zfloor (-x)) + 1)%R. apply Zfloor_ub.
+ { apply Rlt_le_trans with (Z2R (Zfloor (-x)) + 1)%R. apply Zfloor_ub.
change 0%R with (Z2R 0). change 1%R with (Z2R 1). rewrite <- Z2R_plus.
apply Z2R_le. omega. }
lra.
@@ -742,14 +742,14 @@ Remark Ztrunc_range_zero:
forall x, Ztrunc x = 0 -> (-1 < x < 1)%R.
Proof.
intros; generalize (Rle_bool_spec 0%R x). intros RLE; inversion RLE; subst; clear RLE.
-- rewrite Ztrunc_floor in H by auto. split.
- + apply Rlt_le_trans with 0%R; auto. rewrite <- Ropp_0. apply Ropp_lt_contravar. apply Rlt_0_1.
- + replace 1%R with (Z2R (Zfloor x) + 1)%R. apply Zfloor_ub. rewrite H. simpl. apply Rplus_0_l.
-- rewrite Ztrunc_ceil in H by (apply Rlt_le; auto). split.
- + apply Ropp_lt_cancel. rewrite Ropp_involutive.
- replace 1%R with (Z2R (Zfloor (-x)) + 1)%R. apply Zfloor_ub.
- unfold Zceil in H. replace (Zfloor (-x)) with 0 by omega. simpl. apply Rplus_0_l.
- + apply Rlt_le_trans with 0%R; auto. apply Rle_0_1.
+- rewrite Ztrunc_floor in H by auto. split.
+ + apply Rlt_le_trans with 0%R; auto. rewrite <- Ropp_0. apply Ropp_lt_contravar. apply Rlt_0_1.
+ + replace 1%R with (Z2R (Zfloor x) + 1)%R. apply Zfloor_ub. rewrite H. simpl. apply Rplus_0_l.
+- rewrite Ztrunc_ceil in H by (apply Rlt_le; auto). split.
+ + apply Ropp_lt_cancel. rewrite Ropp_involutive.
+ replace 1%R with (Z2R (Zfloor (-x)) + 1)%R. apply Zfloor_ub.
+ unfold Zceil in H. replace (Zfloor (-x)) with 0 by omega. simpl. apply Rplus_0_l.
+ + apply Rlt_le_trans with 0%R; auto. apply Rle_0_1.
Qed.
Theorem ZofB_range_pos:
@@ -763,13 +763,13 @@ Theorem ZofB_range_neg:
forall f n, ZofB f = Some n -> n < 0 -> (Z2R (n - 1)%Z < B2R _ _ f <= Z2R n)%R.
Proof.
intros. rewrite ZofB_correct in H. destruct (is_finite prec emax f) eqn:FIN; inversion H.
- set (x := B2R prec emax f) in *. set (y := (-x)%R).
+ set (x := B2R prec emax f) in *. set (y := (-x)%R).
assert (A: (Z2R (Ztrunc y) <= y < Z2R (Ztrunc y + 1)%Z)%R).
{ apply Ztrunc_range_pos. unfold y. rewrite Ztrunc_opp. omega. }
- destruct A as [B C].
- unfold y in B, C. rewrite Ztrunc_opp in B, C.
+ destruct A as [B C].
+ unfold y in B, C. rewrite Ztrunc_opp in B, C.
replace (- Ztrunc x + 1) with (- (Ztrunc x - 1)) in C by omega.
- rewrite Z2R_opp in B, C. lra.
+ rewrite Z2R_opp in B, C. lra.
Qed.
Theorem ZofB_range_zero:
@@ -782,11 +782,11 @@ Qed.
Theorem ZofB_range_nonneg:
forall f n, ZofB f = Some n -> 0 <= n -> (-1 < B2R _ _ f < Z2R (n + 1)%Z)%R.
Proof.
- intros. destruct (Z.eq_dec n 0).
-- subst n. apply ZofB_range_zero. auto.
-- destruct (ZofB_range_pos f n) as (A & B). auto. omega.
- split; auto. apply Rlt_le_trans with (Z2R 0). simpl; lra.
- apply Rle_trans with (Z2R n); auto. apply Z2R_le; auto.
+ intros. destruct (Z.eq_dec n 0).
+- subst n. apply ZofB_range_zero. auto.
+- destruct (ZofB_range_pos f n) as (A & B). auto. omega.
+ split; auto. apply Rlt_le_trans with (Z2R 0). simpl; lra.
+ apply Rle_trans with (Z2R n); auto. apply Z2R_le; auto.
Qed.
(** For representable integers, [ZofB] is left inverse of [BofZ]. *)
@@ -795,7 +795,7 @@ Theorem ZofBofZ_exact:
forall n, integer_representable n -> ZofB (BofZ n) = Some n.
Proof.
intros. destruct (BofZ_representable n H) as (A & B & C).
- rewrite ZofB_correct. rewrite A, B. f_equal. apply Ztrunc_Z2R.
+ rewrite ZofB_correct. rewrite A, B. f_equal. apply Ztrunc_Z2R.
Qed.
(** Compatibility with subtraction *)
@@ -803,9 +803,9 @@ Qed.
Remark Zfloor_minus:
forall x n, Zfloor (x - Z2R n) = Zfloor x - n.
Proof.
- intros. apply Zfloor_imp. replace (Zfloor x - n + 1) with ((Zfloor x + 1) - n) by omega.
- rewrite ! Z2R_minus. unfold Rminus. split.
- apply Rplus_le_compat_r. apply Zfloor_lb.
+ intros. apply Zfloor_imp. replace (Zfloor x - n + 1) with ((Zfloor x + 1) - n) by omega.
+ rewrite ! Z2R_minus. unfold Rminus. split.
+ apply Rplus_le_compat_r. apply Zfloor_lb.
apply Rplus_lt_compat_r. rewrite Z2R_plus. apply Zfloor_ub.
Qed.
@@ -815,25 +815,25 @@ Theorem ZofB_minus:
ZofB (Bminus _ _ _ Hmax minus_nan m f (BofZ q)) = Some (p - q).
Proof.
intros.
- assert (Q: -2^prec <= q <= 2^prec).
+ assert (Q: -2^prec <= q <= 2^prec).
{ split; auto. generalize (Zpower_ge_0 radix2 prec); simpl; omega. }
- assert (RANGE: (-1 < B2R _ _ f < Z2R (p + 1)%Z)%R) by (apply ZofB_range_nonneg; auto; omega).
+ assert (RANGE: (-1 < B2R _ _ f < Z2R (p + 1)%Z)%R) by (apply ZofB_range_nonneg; auto; omega).
rewrite ZofB_correct in H. destruct (is_finite prec emax f) eqn:FIN; try discriminate.
- assert (PQ2: (Z2R (p + 1) <= Z2R q * 2)%R).
+ assert (PQ2: (Z2R (p + 1) <= Z2R q * 2)%R).
{ change 2%R with (Z2R 2). rewrite <- Z2R_mult. apply Z2R_le. omega. }
assert (EXACT: round radix2 fexp (round_mode m) (B2R _ _ f - Z2R q)%R = (B2R _ _ f - Z2R q)%R).
- { apply round_generic. apply valid_rnd_round_mode.
- apply sterbenz_aux. apply FLT_exp_monotone. apply generic_format_B2R.
+ { apply round_generic. apply valid_rnd_round_mode.
+ apply sterbenz_aux. apply FLT_exp_monotone. apply generic_format_B2R.
apply integer_representable_n. auto. lra. }
- destruct (BofZ_exact q Q) as (A & B & C).
+ destruct (BofZ_exact q Q) as (A & B & C).
generalize (Bminus_correct _ _ _ Hmax minus_nan m f (BofZ q) FIN B).
rewrite Rlt_bool_true.
- fold emin; fold fexp. intros (D & E & F).
- rewrite ZofB_correct. rewrite E. rewrite D. rewrite A. rewrite EXACT.
+ rewrite ZofB_correct. rewrite E. rewrite D. rewrite A. rewrite EXACT.
inversion H. f_equal. rewrite ! Ztrunc_floor. apply Zfloor_minus.
- lra. lra.
+ lra. lra.
- rewrite A. fold emin; fold fexp. rewrite EXACT.
- apply Rle_lt_trans with (bpow radix2 prec).
+ apply Rle_lt_trans with (bpow radix2 prec).
apply Rle_trans with (Z2R q). apply Rabs_le. lra.
rewrite <- Z2R_Zpower. apply Z2R_le; auto. red in prec_gt_0_; omega.
apply bpow_lt. auto.
@@ -853,8 +853,8 @@ Theorem ZofB_range_correct:
ZofB_range f min max =
if is_finite _ _ f && Zle_bool min n && Zle_bool n max then Some n else None.
Proof.
- intros. unfold ZofB_range. rewrite ZofB_correct. fold n.
- destruct (is_finite prec emax f); auto.
+ intros. unfold ZofB_range. rewrite ZofB_correct. fold n.
+ destruct (is_finite prec emax f); auto.
Qed.
Lemma ZofB_range_inversion:
@@ -862,13 +862,13 @@ Lemma ZofB_range_inversion:
ZofB_range f min max = Some n ->
min <= n /\ n <= max /\ ZofB f = Some n.
Proof.
- intros. rewrite ZofB_range_correct in H. rewrite ZofB_correct.
- destruct (is_finite prec emax f); try discriminate.
+ intros. rewrite ZofB_range_correct in H. rewrite ZofB_correct.
+ destruct (is_finite prec emax f); try discriminate.
set (n1 := Ztrunc (B2R _ _ f)) in *.
destruct (min <=? n1) eqn:MIN; try discriminate.
destruct (n1 <=? max) eqn:MAX; try discriminate.
- simpl in H. inversion H. subst n.
- split. apply Zle_bool_imp_le; auto.
+ simpl in H. inversion H. subst n.
+ split. apply Zle_bool_imp_le; auto.
split. apply Zle_bool_imp_le; auto.
auto.
Qed.
@@ -894,16 +894,16 @@ Theorem Bplus_commut:
plus_nan x y = plus_nan y x ->
Bplus _ _ _ Hmax plus_nan mode x y = Bplus _ _ _ Hmax plus_nan mode y x.
Proof.
- intros until y; intros NAN.
- pose proof (Bplus_correct _ _ _ Hmax plus_nan mode x y).
+ intros until y; intros NAN.
+ pose proof (Bplus_correct _ _ _ Hmax plus_nan mode x y).
pose proof (Bplus_correct _ _ _ Hmax plus_nan mode y x).
unfold Bplus in *; destruct x; destruct y; auto.
-- rewrite (eqb_sym b0 b). destruct (eqb b b0) eqn:EQB; auto.
+- rewrite (eqb_sym b0 b). destruct (eqb b b0) eqn:EQB; auto.
f_equal; apply eqb_prop; auto.
- rewrite NAN; auto.
-- rewrite (eqb_sym b0 b). destruct (eqb b b0) eqn:EQB.
+- rewrite (eqb_sym b0 b). destruct (eqb b b0) eqn:EQB.
f_equal; apply eqb_prop; auto.
- rewrite NAN; auto.
+ rewrite NAN; auto.
- rewrite NAN; auto.
- rewrite NAN; auto.
- rewrite NAN; auto.
@@ -912,14 +912,14 @@ Proof.
- rewrite NAN; auto.
- generalize (H (refl_equal _) (refl_equal _)); clear H.
generalize (H0 (refl_equal _) (refl_equal _)); clear H0.
- fold emin. fold fexp.
- set (x := B754_finite prec emax b0 m0 e1 e2). set (rx := B2R _ _ x).
+ fold emin. fold fexp.
+ set (x := B754_finite prec emax b0 m0 e1 e2). set (rx := B2R _ _ x).
set (y := B754_finite prec emax b m e e0). set (ry := B2R _ _ y).
rewrite (Rplus_comm ry rx). destruct Rlt_bool.
+ intros (A1 & A2 & A3) (B1 & B2 & B3).
- apply B2R_Bsign_inj; auto. rewrite <- B1 in A1. auto.
+ apply B2R_Bsign_inj; auto. rewrite <- B1 in A1. auto.
rewrite Z.add_comm. rewrite Z.min_comm. auto.
- + intros (A1 & A2) (B1 & B2). apply B2FF_inj. rewrite B2 in B1. rewrite <- B1 in A1. auto.
+ + intros (A1 & A2) (B1 & B2). apply B2FF_inj. rewrite B2 in B1. rewrite <- B1 in A1. auto.
Qed.
Theorem Bmult_commut:
@@ -927,8 +927,8 @@ Theorem Bmult_commut:
mult_nan x y = mult_nan y x ->
Bmult _ _ _ Hmax mult_nan mode x y = Bmult _ _ _ Hmax mult_nan mode y x.
Proof.
- intros until y; intros NAN.
- pose proof (Bmult_correct _ _ _ Hmax mult_nan mode x y).
+ intros until y; intros NAN.
+ pose proof (Bmult_correct _ _ _ Hmax mult_nan mode x y).
pose proof (Bmult_correct _ _ _ Hmax mult_nan mode y x).
unfold Bmult in *; destruct x; destruct y; auto.
- rewrite (xorb_comm b0 b); auto.
@@ -946,14 +946,14 @@ Proof.
- rewrite (xorb_comm b0 b); auto.
- rewrite (xorb_comm b0 b); auto.
- rewrite NAN; auto.
-- revert H H0. fold emin. fold fexp.
- set (x := B754_finite prec emax b0 m0 e1 e2). set (rx := B2R _ _ x).
+- revert H H0. fold emin. fold fexp.
+ set (x := B754_finite prec emax b0 m0 e1 e2). set (rx := B2R _ _ x).
set (y := B754_finite prec emax b m e e0). set (ry := B2R _ _ y).
rewrite (Rmult_comm ry rx). destruct Rlt_bool.
+ intros (A1 & A2 & A3) (B1 & B2 & B3).
- apply B2R_Bsign_inj; auto. rewrite <- B1 in A1. auto.
+ apply B2R_Bsign_inj; auto. rewrite <- B1 in A1. auto.
rewrite ! Bsign_FF2B. f_equal. f_equal. apply xorb_comm. apply Pos.mul_comm. apply Z.add_comm.
- + intros A B. apply B2FF_inj. etransitivity. eapply A. rewrite xorb_comm. auto.
+ + intros A B. apply B2FF_inj. etransitivity. eapply A. rewrite xorb_comm. auto.
Qed.
(** Multiplication by 2 is diagonal addition. *)
@@ -966,31 +966,31 @@ Theorem Bmult2_Bplus:
Proof.
intros until f; intros NAN.
destruct (BofZ_representable 2) as (A & B & C).
- apply (integer_representable_2p 1). red in prec_gt_0_; omega.
+ apply (integer_representable_2p 1). red in prec_gt_0_; omega.
pose proof (Bmult_correct _ _ _ Hmax mult_nan mode f (BofZ 2%Z)). fold emin in H.
- rewrite A, B, C in H. rewrite xorb_false_r in H.
+ rewrite A, B, C in H. rewrite xorb_false_r in H.
destruct (is_finite _ _ f) eqn:FIN.
-- pose proof (Bplus_correct _ _ _ Hmax plus_nan mode f f FIN FIN). fold emin in H0.
+- pose proof (Bplus_correct _ _ _ Hmax plus_nan mode f f FIN FIN). fold emin in H0.
assert (EQ: (B2R prec emax f * Z2R 2%Z = B2R prec emax f + B2R prec emax f)%R).
{ change (Z2R 2%Z) with 2%R. ring. }
- rewrite <- EQ in H0. destruct Rlt_bool.
- + destruct H0 as (P & Q & R). destruct H as (S & T & U).
+ rewrite <- EQ in H0. destruct Rlt_bool.
+ + destruct H0 as (P & Q & R). destruct H as (S & T & U).
apply B2R_Bsign_inj; auto.
rewrite P, S. auto.
- rewrite R, U.
- replace 0%R with (0 * Z2R 2%Z)%R by ring. rewrite Rcompare_mult_r.
- rewrite andb_diag, orb_diag. destruct f; try discriminate; simpl.
- rewrite Rcompare_Eq by auto. destruct mode; auto.
+ rewrite R, U.
+ replace 0%R with (0 * Z2R 2%Z)%R by ring. rewrite Rcompare_mult_r.
+ rewrite andb_diag, orb_diag. destruct f; try discriminate; simpl.
+ rewrite Rcompare_Eq by auto. destruct mode; auto.
replace 0%R with (@F2R radix2 {| Fnum := 0%Z; Fexp := e |}).
- rewrite Rcompare_F2R. destruct b; auto.
- unfold F2R. simpl. ring.
- change 0%R with (Z2R 0%Z). apply Z2R_lt. omega.
+ rewrite Rcompare_F2R. destruct b; auto.
+ unfold F2R. simpl. ring.
+ change 0%R with (Z2R 0%Z). apply Z2R_lt. omega.
destruct (Bmult prec emax prec_gt_0_ Hmax mult_nan mode f (BofZ 2)); reflexivity || discriminate.
+ destruct H0 as (P & Q). apply B2FF_inj. rewrite P, H. auto.
- destruct f; try discriminate.
+ simpl Bplus. rewrite eqb_true. destruct (BofZ 2) eqn:B2; try discriminate; simpl in *.
assert ((0 = 2)%Z) by (apply eq_Z2R; auto). discriminate.
- subst b0. rewrite xorb_false_r. auto.
+ subst b0. rewrite xorb_false_r. auto.
auto.
+ unfold Bplus, Bmult. rewrite <- NAN by auto. auto.
Qed.
@@ -1003,9 +1003,9 @@ Remark Bexact_inverse_mantissa_value:
Zpos Bexact_inverse_mantissa = 2 ^ (prec - 1).
Proof.
assert (REC: forall n, Z.pos (nat_iter n xO xH) = 2 ^ (Z.of_nat n)).
- { induction n. reflexivity.
- simpl nat_iter. transitivity (2 * Z.pos (nat_iter n xO xH)). reflexivity.
- rewrite inj_S. rewrite IHn. unfold Z.succ. rewrite Zpower_plus by omega.
+ { induction n. reflexivity.
+ simpl nat_iter. transitivity (2 * Z.pos (nat_iter n xO xH)). reflexivity.
+ rewrite inj_S. rewrite IHn. unfold Z.succ. rewrite Zpower_plus by omega.
change (2 ^ 1) with 2. ring. }
red in prec_gt_0_.
unfold Bexact_inverse_mantissa. rewrite iter_nat_of_Z by omega. rewrite REC.
@@ -1029,10 +1029,10 @@ Remark bounded_Bexact_inverse:
forall e,
emin <= e <= emax - prec <-> bounded prec emax Bexact_inverse_mantissa e = true.
Proof.
- intros. unfold bounded, canonic_mantissa. rewrite andb_true_iff.
- rewrite <- Zeq_is_eq_bool. rewrite <- Zle_is_le_bool.
- rewrite Bexact_inverse_mantissa_digits2_pos.
- split.
+ intros. unfold bounded, canonic_mantissa. rewrite andb_true_iff.
+ rewrite <- Zeq_is_eq_bool. rewrite <- Zle_is_le_bool.
+ rewrite Bexact_inverse_mantissa_digits2_pos.
+ split.
- intros; split. unfold FLT_exp. unfold emin in H. zify; omega. omega.
- intros [A B]. unfold FLT_exp in A. unfold emin. zify; omega.
Qed.
@@ -1049,7 +1049,7 @@ Program Definition Bexact_inverse (f: binary_float) : option binary_float :=
| _ => None
end.
Next Obligation.
- rewrite <- bounded_Bexact_inverse in B. rewrite <- bounded_Bexact_inverse.
+ rewrite <- bounded_Bexact_inverse in B. rewrite <- bounded_Bexact_inverse.
unfold emin in *. omega.
Qed.
@@ -1070,12 +1070,12 @@ Proof with (try discriminate).
split. auto. split. auto. split. unfold B2R. rewrite Bexact_inverse_mantissa_value.
unfold F2R; simpl. rewrite Z2R_cond_Zopp.
rewrite <- ! cond_Ropp_mult_l.
- red in prec_gt_0_.
+ red in prec_gt_0_.
replace (Z2R (2 ^ (prec - 1))) with (bpow radix2 (prec - 1))
by (symmetry; apply (Z2R_Zpower radix2); omega).
rewrite <- ! bpow_plus.
- replace (prec - 1 + e') with (- (prec - 1 + e)) by (unfold e'; omega).
- rewrite bpow_opp. unfold cond_Ropp; destruct b; auto.
+ replace (prec - 1 + e') with (- (prec - 1 + e)) by (unfold e'; omega).
+ rewrite bpow_opp. unfold cond_Ropp; destruct b; auto.
rewrite Ropp_inv_permute. auto. apply Rgt_not_eq. apply bpow_gt_0.
split. simpl. red; intros. apply F2R_eq_0_reg in H. destruct b; simpl in H; discriminate.
auto.
@@ -1091,23 +1091,23 @@ Theorem Bdiv_mult_inverse:
Proof.
intros until z; intros NAN; intros. destruct (Bexact_inverse_correct _ _ H) as (A & B & C & D & E).
pose proof (Bmult_correct _ _ _ Hmax mult_nan mode x z).
- fold emin in H0. fold fexp in H0.
+ fold emin in H0. fold fexp in H0.
pose proof (Bdiv_correct _ _ _ Hmax div_nan mode x y D).
fold emin in H1. fold fexp in H1.
- unfold Rdiv in H1. rewrite <- C in H1.
+ unfold Rdiv in H1. rewrite <- C in H1.
destruct (is_finite _ _ x) eqn:FINX.
-- destruct Rlt_bool.
- + destruct H0 as (P & Q & R). destruct H1 as (S & T & U).
+- destruct Rlt_bool.
+ + destruct H0 as (P & Q & R). destruct H1 as (S & T & U).
apply B2R_Bsign_inj; auto.
- rewrite Q. simpl. apply is_finite_strict_finite; auto.
- rewrite P, S. auto.
- rewrite R, U, E. auto.
- apply is_finite_not_is_nan; auto.
- apply is_finite_not_is_nan. rewrite Q. simpl. apply is_finite_strict_finite; auto. + apply B2FF_inj. rewrite H0, H1. rewrite E. auto.
+ rewrite Q. simpl. apply is_finite_strict_finite; auto.
+ rewrite P, S. auto.
+ rewrite R, U, E. auto.
+ apply is_finite_not_is_nan; auto.
+ apply is_finite_not_is_nan. rewrite Q. simpl. apply is_finite_strict_finite; auto. + apply B2FF_inj. rewrite H0, H1. rewrite E. auto.
- destruct y; try discriminate. destruct z; try discriminate.
destruct x; try discriminate; simpl.
+ simpl in E; congruence.
- + erewrite NAN; eauto.
+ + erewrite NAN; eauto.
Qed.
(** ** Conversion from scientific notation *)
@@ -1126,13 +1126,13 @@ Lemma pos_pow_spec:
Proof.
intros x.
assert (REC: forall y a, Pos.iter y (Pos.mul x) a = Pos.mul (pos_pow x y) a).
- { induction y; simpl; intros.
+ { induction y; simpl; intros.
- rewrite ! IHy, Pos.square_spec, ! Pos.mul_assoc. auto.
- rewrite ! IHy, Pos.square_spec, ! Pos.mul_assoc. auto.
- auto.
}
intros. simpl. rewrite <- Pos2Z.inj_pow_pos. unfold Pos.pow. rewrite REC. rewrite Pos.mul_1_r. auto.
-Qed.
+Qed.
(** Given a base [base], a mantissa [m] and an exponent [e], the following function
computes the FP number closest to [m * base ^ e], using round to odd, ties break to even.
@@ -1142,7 +1142,7 @@ Qed.
Definition Bparse (base: positive) (m: positive) (e: Z): binary_float :=
match e with
- | Z0 =>
+ | Z0 =>
BofZ (Zpos m)
| Zpos p =>
if e * Z.log2 (Zpos base) <? emax
@@ -1167,7 +1167,7 @@ Proof.
assert (B: 0 <= Z.log2 base) by apply Z.log2_nonneg.
assert (C: 0 <= Z.log2_up base) by apply Z.log2_up_nonneg.
destruct (Z.log2_spec base) as [D E]; auto.
- destruct (Z.log2_up_spec base) as [F G]. apply radix_gt_1.
+ destruct (Z.log2_up_spec base) as [F G]. apply radix_gt_1.
assert (K: 0 <= 2 ^ Z.log2 base) by (apply Z.pow_nonneg; omega).
rewrite ! (Zmult_comm n). rewrite ! Z.pow_mul_r by omega.
split; apply Z.pow_le_mono_l; omega.
@@ -1189,7 +1189,7 @@ Lemma bpow_log_neg:
(bpow base n <= bpow radix2 (n * Z.log2 base)%Z)%R.
Proof.
intros. set (m := -n). replace n with (-m) by (unfold m; omega).
- rewrite ! Z.mul_opp_l, ! bpow_opp. apply Rinv_le.
+ rewrite ! Z.mul_opp_l, ! bpow_opp. apply Rinv_le.
apply bpow_gt_0.
apply bpow_log_pos. unfold m; omega.
Qed.
@@ -1202,9 +1202,9 @@ Lemma round_integer_overflow:
emax <= e * Z.log2 base ->
(bpow radix2 emax <= round radix2 fexp (round_mode mode_NE) (Z2R (Zpos m) * bpow base e))%R.
Proof.
- intros.
+ intros.
rewrite <- (round_generic radix2 fexp (round_mode mode_NE) (bpow radix2 emax)); auto.
- apply round_le; auto. apply fexp_correct; auto. apply valid_rnd_round_mode.
+ apply round_le; auto. apply fexp_correct; auto. apply valid_rnd_round_mode.
rewrite <- (Rmult_1_l (bpow radix2 emax)). apply Rmult_le_compat.
apply Rle_0_1.
apply bpow_ge_0.
@@ -1224,24 +1224,24 @@ Proof.
intros.
set (eps := bpow radix2 (emin - 1)) in *.
assert (A: round radix2 fexp (round_mode mode_NE) eps = 0%R).
- { unfold round. simpl.
+ { unfold round. simpl.
assert (E: canonic_exp radix2 fexp eps = emin).
{ unfold canonic_exp, eps. rewrite ln_beta_bpow. unfold fexp, FLT_exp. zify; red in prec_gt_0_; omega. }
- unfold scaled_mantissa; rewrite E.
+ unfold scaled_mantissa; rewrite E.
assert (P: (eps * bpow radix2 (-emin) = / 2)%R).
{ unfold eps. rewrite <- bpow_plus. replace (emin - 1 + -emin) with (-1) by omega. auto. }
- rewrite P. unfold Znearest.
+ rewrite P. unfold Znearest.
assert (F: Zfloor (/ 2)%R = 0).
- { apply Zfloor_imp.
+ { apply Zfloor_imp.
split. apply Rlt_le. apply Rinv_0_lt_compat. apply (Z2R_lt 0 2). omega.
change (Z2R (0 + 1)) with 1%R. rewrite <- Rinv_1 at 3. apply Rinv_1_lt_contravar. apply Rle_refl. apply (Z2R_lt 1 2). omega.
}
- rewrite F. change (Z2R 0) with 0%R. rewrite Rminus_0_r. rewrite Rcompare_Eq by auto.
- simpl. unfold F2R; simpl. apply Rmult_0_l.
+ rewrite F. change (Z2R 0) with 0%R. rewrite Rminus_0_r. rewrite Rcompare_Eq by auto.
+ simpl. unfold F2R; simpl. apply Rmult_0_l.
}
apply Rle_antisym.
- rewrite <- A. apply round_le. apply fexp_correct; auto. apply valid_rnd_round_mode. tauto.
-- rewrite <- (round_0 radix2 fexp (round_mode mode_NE)).
+- rewrite <- (round_0 radix2 fexp (round_mode mode_NE)).
apply round_le. apply fexp_correct; auto. apply valid_rnd_round_mode. tauto.
Qed.
@@ -1254,16 +1254,16 @@ Proof.
intros. apply round_NE_underflows. split.
- apply Rmult_le_pos. apply (Z2R_le 0). zify; omega. apply bpow_ge_0.
- apply Rle_trans with (bpow radix2 (Z.log2_up (Z.pos m) + e * Z.log2 base)).
-+ rewrite bpow_plus. apply Rmult_le_compat.
++ rewrite bpow_plus. apply Rmult_le_compat.
apply (Z2R_le 0); zify; omega.
apply bpow_ge_0.
rewrite <- Z2R_Zpower. apply Z2R_le.
- destruct (Z.eq_dec (Z.pos m) 1).
+ destruct (Z.eq_dec (Z.pos m) 1).
rewrite e0. simpl. omega.
apply Z.log2_up_spec. zify; omega.
- apply Z.log2_up_nonneg.
+ apply Z.log2_up_nonneg.
apply bpow_log_neg. auto.
-+ apply bpow_le. omega.
++ apply bpow_le. omega.
Qed.
(** Correctness of Bparse *)
@@ -1279,29 +1279,29 @@ Theorem Bparse_correct:
else
B2FF _ _ (Bparse b m e) = F754_infinity false.
Proof.
- intros.
+ intros.
assert (A: forall x, @F2R radix2 {| Fnum := x; Fexp := 0 |} = Z2R x).
{ intros. unfold F2R, Fnum; simpl. ring. }
unfold Bparse, r. destruct e as [ | e | e].
- (* e = Z0 *)
- change (bpow base 0) with 1%R. rewrite Rmult_1_r.
+ change (bpow base 0) with 1%R. rewrite Rmult_1_r.
exact (BofZ_correct (Z.pos m)).
- (* e = Zpos e *)
destruct (Z.ltb_spec (Z.pos e * Z.log2 (Z.pos b)) emax).
+ (* no overflow *)
rewrite pos_pow_spec. rewrite <- Z2R_Zpower by (zify; omega). rewrite <- Z2R_mult.
- replace false with (Z.pos m * Z.pos b ^ Z.pos e <? 0).
+ replace false with (Z.pos m * Z.pos b ^ Z.pos e <? 0).
exact (BofZ_correct (Z.pos m * Z.pos b ^ Z.pos e)).
- rewrite Z.ltb_ge. rewrite Zmult_comm. apply Zmult_gt_0_le_0_compat. zify; omega. apply (Zpower_ge_0 base).
+ rewrite Z.ltb_ge. rewrite Zmult_comm. apply Zmult_gt_0_le_0_compat. zify; omega. apply (Zpower_ge_0 base).
+ (* overflow *)
- rewrite Rlt_bool_false. auto. eapply Rle_trans; [idtac|apply Rle_abs].
- apply (round_integer_overflow base). zify; omega. auto.
+ rewrite Rlt_bool_false. auto. eapply Rle_trans; [idtac|apply Rle_abs].
+ apply (round_integer_overflow base). zify; omega. auto.
- (* e = Zneg e *)
destruct (Z.ltb_spec (Z.neg e * Z.log2 (Z.pos b) + Z.log2_up (Z.pos m)) emin).
+ (* undeflow *)
rewrite round_integer_underflow; auto.
- rewrite Rlt_bool_true. auto.
- replace (Rabs 0)%R with 0%R. apply bpow_gt_0. apply (Z2R_abs 0).
+ rewrite Rlt_bool_true. auto.
+ replace (Rabs 0)%R with 0%R. apply bpow_gt_0. apply (Z2R_abs 0).
zify; omega.
+ (* no underflow *)
generalize (Bdiv_correct_aux prec emax prec_gt_0_ Hmax mode_NE false m 0 false (pos_pow b e) 0).
@@ -1311,19 +1311,19 @@ Proof.
binary_round_aux prec emax mode_NE (xorb false false) mz0 ez lz
| (Z.neg _, _, _) => F754_nan false 1
end).
- fold emin; fold fexp. rewrite ! A. unfold cond_Zopp. rewrite pos_pow_spec.
+ fold emin; fold fexp. rewrite ! A. unfold cond_Zopp. rewrite pos_pow_spec.
assert (B: (Z2R (Z.pos m) / Z2R (Z.pos b ^ Z.pos e) =
Z2R (Z.pos m) * bpow base (Z.neg e))%R).
{ change (Z.neg e) with (- (Z.pos e)). rewrite bpow_opp. auto. }
- rewrite B. intros [P Q].
+ rewrite B. intros [P Q].
destruct (Rlt_bool
(Rabs
(round radix2 fexp (round_mode mode_NE)
(Z2R (Z.pos m) * bpow base (Z.neg e))))
(bpow radix2 emax)).
-* destruct Q as (Q1 & Q2 & Q3).
+* destruct Q as (Q1 & Q2 & Q3).
split. rewrite B2R_FF2B, Q1. auto.
- split. rewrite is_finite_FF2B. auto.
+ split. rewrite is_finite_FF2B. auto.
rewrite Bsign_FF2B. auto.
* rewrite B2FF_FF2B. auto.
Qed.
@@ -1365,16 +1365,16 @@ Theorem Bconv_correct:
B2FF _ _ (Bconv conv_nan m f) = binary_overflow prec2 emax2 m (Bsign _ _ f).
Proof.
intros. destruct f; try discriminate.
-- simpl. rewrite round_0. rewrite Rabs_R0. rewrite Rlt_bool_true. auto.
- apply bpow_gt_0. apply valid_rnd_round_mode.
+- simpl. rewrite round_0. rewrite Rabs_R0. rewrite Rlt_bool_true. auto.
+ apply bpow_gt_0. apply valid_rnd_round_mode.
- generalize (binary_normalize_correct _ _ _ Hmax2 m (cond_Zopp b (Zpos m0)) e b).
- fold emin2; fold fexp2. simpl. destruct Rlt_bool.
- + intros (A & B & C). split. auto. split. auto. rewrite C.
- destruct b; simpl.
- rewrite Rcompare_Lt. auto. apply F2R_lt_0_compat. simpl. compute; auto.
+ fold emin2; fold fexp2. simpl. destruct Rlt_bool.
+ + intros (A & B & C). split. auto. split. auto. rewrite C.
+ destruct b; simpl.
+ rewrite Rcompare_Lt. auto. apply F2R_lt_0_compat. simpl. compute; auto.
rewrite Rcompare_Gt. auto. apply F2R_gt_0_compat. simpl. compute; auto.
+ intros A. rewrite A. f_equal. destruct b.
- apply Rlt_bool_true. apply F2R_lt_0_compat. simpl. compute; auto.
+ apply Rlt_bool_true. apply F2R_lt_0_compat. simpl. compute; auto.
apply Rlt_bool_false. apply Rlt_le. apply Rgt_lt. apply F2R_gt_0_compat. simpl. compute; auto.
Qed.
@@ -1391,8 +1391,8 @@ Proof.
intros PREC EMAX; intros. generalize (Bconv_correct conv_nan m f H).
assert (LT: (Rabs (B2R _ _ f) < bpow radix2 emax2)%R).
{
- destruct f; try discriminate; simpl.
- rewrite Rabs_R0. apply bpow_gt_0.
+ destruct f; try discriminate; simpl.
+ rewrite Rabs_R0. apply bpow_gt_0.
apply Rlt_le_trans with (bpow radix2 emax1).
rewrite F2R_cond_Zopp. rewrite abs_cond_Ropp. rewrite <- F2R_Zabs. simpl Z.abs.
eapply bounded_lt_emax; eauto.
@@ -1401,11 +1401,11 @@ Proof.
assert (EQ: round radix2 fexp2 (round_mode m) (B2R prec1 emax1 f) = B2R prec1 emax1 f).
{
apply round_generic. apply valid_rnd_round_mode. eapply generic_inclusion_le.
- 5: apply generic_format_B2R. apply fexp_correct; auto. apply fexp_correct; auto.
+ 5: apply generic_format_B2R. apply fexp_correct; auto. apply fexp_correct; auto.
instantiate (1 := emax2). intros. unfold fexp2, FLT_exp. unfold emin2. zify; omega.
apply Rlt_le; auto.
}
- rewrite EQ. rewrite Rlt_bool_true by auto. auto.
+ rewrite EQ. rewrite Rlt_bool_true by auto. auto.
Qed.
(** Conversion from integers and change of format *)
@@ -1415,18 +1415,18 @@ Theorem Bconv_BofZ:
integer_representable prec1 emax1 n ->
Bconv conv_nan mode_NE (BofZ prec1 emax1 _ Hmax1 n) = BofZ prec2 emax2 _ Hmax2 n.
Proof.
- intros.
- destruct (BofZ_representable _ _ _ Hmax1 n H) as (A & B & C).
+ intros.
+ destruct (BofZ_representable _ _ _ Hmax1 n H) as (A & B & C).
set (f := BofZ prec1 emax1 prec1_gt_0_ Hmax1 n) in *.
generalize (Bconv_correct conv_nan mode_NE f B).
- unfold BofZ.
- generalize (binary_normalize_correct _ _ _ Hmax2 mode_NE n 0 false).
- fold emin2; fold fexp2. rewrite A.
+ unfold BofZ.
+ generalize (binary_normalize_correct _ _ _ Hmax2 mode_NE n 0 false).
+ fold emin2; fold fexp2. rewrite A.
replace (F2R {| Fnum := n; Fexp := 0 |}) with (Z2R n).
- destruct Rlt_bool.
-- intros (P & Q & R) (D & E & F). apply B2R_Bsign_inj; auto.
- congruence. rewrite F, C, R. change 0%R with (Z2R 0). rewrite Rcompare_Z2R.
- unfold Zlt_bool. auto.
+ destruct Rlt_bool.
+- intros (P & Q & R) (D & E & F). apply B2R_Bsign_inj; auto.
+ congruence. rewrite F, C, R. change 0%R with (Z2R 0). rewrite Rcompare_Z2R.
+ unfold Zlt_bool. auto.
- intros P Q. apply B2FF_inj. rewrite P, Q. rewrite C. f_equal. change 0%R with (Z2R 0).
generalize (Zlt_bool_spec n 0); intros LT; inversion LT.
rewrite Rlt_bool_true; auto. apply Z2R_lt; auto.
@@ -1442,7 +1442,7 @@ Theorem ZofB_Bconv:
ZofB _ _ f = Some n -> ZofB _ _ (Bconv conv_nan m f) = Some n.
Proof.
intros. rewrite ZofB_correct in H1. destruct (is_finite _ _ f) eqn:FIN; inversion H1.
- destruct (Bconv_widen_exact H H0 conv_nan m f) as (A & B & C). auto.
+ destruct (Bconv_widen_exact H H0 conv_nan m f) as (A & B & C). auto.
rewrite ZofB_correct. rewrite B. rewrite A. auto.
Qed.
@@ -1453,9 +1453,9 @@ Theorem ZofB_range_Bconv:
ZofB_range _ _ f min1 max1 = Some n ->
ZofB_range _ _ (Bconv conv_nan m f) min2 max2 = Some n.
Proof.
- intros.
+ intros.
destruct (ZofB_range_inversion _ _ _ _ _ _ H3) as (A & B & C).
- unfold ZofB_range. erewrite ZofB_Bconv by eauto.
+ unfold ZofB_range. erewrite ZofB_Bconv by eauto.
rewrite ! Zle_bool_true by omega. auto.
Qed.
@@ -1467,7 +1467,7 @@ Theorem Bcompare_Bconv_widen:
Bcompare _ _ (Bconv conv_nan m x) (Bconv conv_nan m y) = Bcompare _ _ x y.
Proof.
intros. destruct (is_finite _ _ x && is_finite _ _ y) eqn:FIN.
-- apply andb_true_iff in FIN. destruct FIN.
+- apply andb_true_iff in FIN. destruct FIN.
destruct (Bconv_widen_exact H H0 conv_nan m x H1) as (A & B & C).
destruct (Bconv_widen_exact H H0 conv_nan m y H2) as (D & E & F).
rewrite ! Bcompare_correct by auto. rewrite A, D. auto.
@@ -1476,13 +1476,13 @@ Proof.
destruct x, y; try discriminate; simpl in P, Q; simpl;
repeat (match goal with |- context [conv_nan ?b ?pl] => destruct (conv_nan b pl) end);
auto.
- destruct Q as (D & E & F); auto.
+ destruct Q as (D & E & F); auto.
destruct (binary_normalize prec2 emax2 prec2_gt_0_ Hmax2 m (cond_Zopp b0 (Z.pos m0)) e b0);
discriminate || reflexivity.
- destruct P as (A & B & C); auto.
+ destruct P as (A & B & C); auto.
destruct (binary_normalize prec2 emax2 prec2_gt_0_ Hmax2 m (cond_Zopp b (Z.pos m0)) e b);
- try discriminate; simpl. destruct b; auto. destruct b, b1; auto.
- destruct P as (A & B & C); auto.
+ try discriminate; simpl. destruct b; auto. destruct b, b1; auto.
+ destruct P as (A & B & C); auto.
destruct (binary_normalize prec2 emax2 prec2_gt_0_ Hmax2 m (cond_Zopp b (Z.pos m0)) e b);
try discriminate; simpl. destruct b; auto.
destruct b, b2; auto.
@@ -1503,26 +1503,26 @@ Hypothesis Hmax2 : (prec2 < emax2)%Z.
Let binary_float1 := binary_float prec1 emax1.
Let binary_float2 := binary_float prec2 emax2.
-(** Converting to a higher precision then down to the original format
+(** Converting to a higher precision then down to the original format
is the identity. *)
Theorem Bconv_narrow_widen:
prec2 >= prec1 -> emax2 >= emax1 ->
- forall narrow_nan widen_nan m f,
+ forall narrow_nan widen_nan m f,
is_nan _ _ f = false ->
Bconv prec2 emax2 prec1 emax1 _ Hmax1 narrow_nan m (Bconv prec1 emax1 prec2 emax2 _ Hmax2 widen_nan m f) = f.
Proof.
- intros. destruct (is_finite _ _ f) eqn:FIN.
+ intros. destruct (is_finite _ _ f) eqn:FIN.
- assert (EQ: round radix2 fexp1 (round_mode m) (B2R prec1 emax1 f) = B2R prec1 emax1 f).
{ apply round_generic. apply valid_rnd_round_mode. apply generic_format_B2R. }
generalize (Bconv_widen_exact _ _ _ _ _ _ Hmax2 H H0 widen_nan m f FIN).
- set (f' := Bconv prec1 emax1 prec2 emax2 _ Hmax2 widen_nan m f).
+ set (f' := Bconv prec1 emax1 prec2 emax2 _ Hmax2 widen_nan m f).
intros (A & B & C).
generalize (Bconv_correct _ _ _ _ _ Hmax1 narrow_nan m f' B).
- fold emin1. fold fexp1. rewrite A, C, EQ. rewrite Rlt_bool_true.
- intros (D & E & F).
+ fold emin1. fold fexp1. rewrite A, C, EQ. rewrite Rlt_bool_true.
+ intros (D & E & F).
apply B2R_Bsign_inj; auto.
destruct f; try discriminate; simpl.
- rewrite Rabs_R0. apply bpow_gt_0.
+ rewrite Rabs_R0. apply bpow_gt_0.
rewrite F2R_cond_Zopp. rewrite abs_cond_Ropp. rewrite <- F2R_Zabs. simpl Z.abs.
eapply bounded_lt_emax; eauto.
- destruct f; try discriminate. simpl. auto.
diff --git a/lib/Floats.v b/lib/Floats.v
index e893e3e7..cf25852e 100644
--- a/lib/Floats.v
+++ b/lib/Floats.v
@@ -53,13 +53,13 @@ Lemma cmp_of_comparison_swap:
cmp_of_comparison (swap_comparison c) x =
cmp_of_comparison c (match x with None => None | Some x => Some (CompOpp x) end).
Proof.
- intros. destruct c; destruct x as [[]|]; reflexivity.
+ intros. destruct c; destruct x as [[]|]; reflexivity.
Qed.
Lemma cmp_of_comparison_ne_eq:
forall x, cmp_of_comparison Cne x = negb (cmp_of_comparison Ceq x).
Proof.
- intros. destruct x as [[]|]; reflexivity.
+ intros. destruct x as [[]|]; reflexivity.
Qed.
Lemma cmp_of_comparison_lt_eq_false:
@@ -296,23 +296,23 @@ Qed.
Theorem mul2_add:
forall f, add f f = mul f (of_int (Int.repr 2%Z)).
Proof.
- intros. apply Bmult2_Bplus.
- intros. destruct x; try discriminate. simpl.
- transitivity (b, transform_quiet_pl n).
- destruct Archi.choose_binop_pl_64; auto.
+ intros. apply Bmult2_Bplus.
+ intros. destruct x; try discriminate. simpl.
+ transitivity (b, transform_quiet_pl n).
+ destruct Archi.choose_binop_pl_64; auto.
destruct y; auto || discriminate.
Qed.
(** Divisions that can be turned into multiplication by an inverse. *)
-Definition exact_inverse : float -> option float := Bexact_inverse 53 1024 __ __.
+Definition exact_inverse : float -> option float := Bexact_inverse 53 1024 __ __.
Theorem div_mul_inverse:
forall x y z, exact_inverse y = Some z -> div x y = mul x z.
Proof.
- intros. apply Bdiv_mult_inverse; auto.
- intros. destruct x0; try discriminate. simpl.
- transitivity (b, transform_quiet_pl n).
+ intros. apply Bdiv_mult_inverse; auto.
+ intros. destruct x0; try discriminate. simpl.
+ transitivity (b, transform_quiet_pl n).
destruct y0; reflexivity || discriminate.
destruct z0; reflexivity || discriminate.
Qed.
@@ -323,13 +323,13 @@ Theorem cmp_swap:
forall c x y, cmp (swap_comparison c) x y = cmp c y x.
Proof.
unfold cmp; intros. rewrite (Bcompare_swap _ _ x y).
- apply cmp_of_comparison_swap.
+ apply cmp_of_comparison_swap.
Qed.
Theorem cmp_ne_eq:
forall f1 f2, cmp Cne f1 f2 = negb (cmp Ceq f1 f2).
Proof.
- intros; apply cmp_of_comparison_ne_eq.
+ intros; apply cmp_of_comparison_ne_eq.
Qed.
Theorem cmp_lt_eq_false:
@@ -371,7 +371,7 @@ Proof.
intros; unfold of_bits, to_bits, bits_of_b64, b64_of_bits.
rewrite Int64.unsigned_repr, binary_float_of_bits_of_binary_float; [reflexivity|].
generalize (bits_of_binary_float_range 52 11 __ __ f).
- change (2^(52+11+1)) with (Int64.max_unsigned + 1). omega.
+ change (2^(52+11+1)) with (Int64.max_unsigned + 1). omega.
Qed.
Theorem to_of_bits:
@@ -379,7 +379,7 @@ Theorem to_of_bits:
Proof.
intros; unfold of_bits, to_bits, bits_of_b64, b64_of_bits.
rewrite bits_of_binary_float_of_bits. apply Int64.repr_unsigned.
- apply Int64.unsigned_range.
+ apply Int64.unsigned_range.
Qed.
(** Conversions between floats and unsigned ints can be defined
@@ -412,7 +412,7 @@ Proof.
assert (R8: integer_representable 53 1024 (Int.unsigned ox8000_0000)).
{ apply integer_representable_2p with (p := 31);auto; smart_omega. }
rewrite BofZ_plus by auto.
- f_equal.
+ f_equal.
unfold Int.ltu in H. destruct zlt in H; try discriminate.
unfold y, Int.sub. rewrite Int.signed_repr. omega.
compute_this (Int.unsigned ox8000_0000); smart_omega.
@@ -424,10 +424,10 @@ Theorem to_intu_to_int_1:
to_intu x = Some n ->
to_int x = Some n.
Proof.
- intros. unfold to_intu in H0.
+ intros. unfold to_intu in H0.
destruct (ZofB_range 53 1024 x 0 Int.max_unsigned) as [p|] eqn:E; simpl in H0; inv H0.
exploit ZofB_range_inversion; eauto. intros (A & B & C).
- unfold to_int, ZofB_range. rewrite C.
+ unfold to_int, ZofB_range. rewrite C.
rewrite Zle_bool_true by smart_omega. rewrite Zle_bool_true; auto.
exploit (BofZ_exact 53 1024 __ __ (Int.unsigned ox8000_0000)).
vm_compute; intuition congruence.
@@ -436,16 +436,16 @@ Proof.
intros (EQy & FINy & SIGNy).
assert (FINx: is_finite _ _ x = true).
{ rewrite ZofB_correct in C. destruct (is_finite _ _ x) eqn:FINx; congruence. }
- destruct (zeq p 0).
+ destruct (zeq p 0).
subst p; smart_omega.
destruct (ZofB_range_pos 53 1024 __ __ x p C) as [P Q]. omega.
- assert (CMP: Bcompare _ _ x y = Some Lt).
+ assert (CMP: Bcompare _ _ x y = Some Lt).
{ unfold cmp, cmp_of_comparison in H. destruct (Bcompare _ _ x y) as [[]|]; auto; discriminate. }
- rewrite Bcompare_correct in CMP by auto.
+ rewrite Bcompare_correct in CMP by auto.
inv CMP. apply Rcompare_Lt_inv in H1. rewrite EQy in H1.
assert (p < Int.unsigned ox8000_0000).
{ apply lt_Z2R. eapply Rle_lt_trans; eauto. }
- change Int.max_signed with (Int.unsigned ox8000_0000 - 1). omega.
+ change Int.max_signed with (Int.unsigned ox8000_0000 - 1). omega.
Qed.
Theorem to_intu_to_int_2:
@@ -454,7 +454,7 @@ Theorem to_intu_to_int_2:
to_intu x = Some n ->
to_int (sub x (of_intu ox8000_0000)) = Some (Int.sub n ox8000_0000).
Proof.
- intros. unfold to_intu in H0.
+ intros. unfold to_intu in H0.
destruct (ZofB_range _ _ x 0 Int.max_unsigned) as [p|] eqn:E; simpl in H0; inv H0.
exploit ZofB_range_inversion; eauto. intros (A & B & C).
exploit (BofZ_exact 53 1024 __ __ (Int.unsigned ox8000_0000)).
@@ -466,19 +466,19 @@ Proof.
{ rewrite ZofB_correct in C. destruct (is_finite _ _ x) eqn:FINx; congruence. }
assert (GE: (B2R _ _ x >= Z2R (Int.unsigned ox8000_0000))%R).
{ rewrite <- EQy. unfold cmp, cmp_of_comparison in H.
- rewrite Bcompare_correct in H by auto.
+ rewrite Bcompare_correct in H by auto.
destruct (Rcompare (B2R 53 1024 x) (B2R 53 1024 y)) eqn:CMP.
apply Req_ge; apply Rcompare_Eq_inv; auto.
discriminate.
- apply Rgt_ge; apply Rcompare_Gt_inv; auto.
- }
+ apply Rgt_ge; apply Rcompare_Gt_inv; auto.
+ }
assert (EQ: ZofB_range _ _ (sub x y) Int.min_signed Int.max_signed = Some (p - Int.unsigned ox8000_0000)).
{
- apply ZofB_range_minus. exact E.
+ apply ZofB_range_minus. exact E.
compute_this (Int.unsigned ox8000_0000). smart_omega.
apply Rge_le; auto.
- }
- unfold to_int; rewrite EQ. simpl. f_equal. unfold Int.sub. f_equal. f_equal.
+ }
+ unfold to_int; rewrite EQ. simpl. f_equal. unfold Int.sub. f_equal. f_equal.
symmetry; apply Int.unsigned_repr. omega.
Qed.
@@ -522,14 +522,14 @@ Qed.
Lemma from_words_eq:
forall x, from_words ox4330_0000 x = BofZ 53 1024 __ __ (2^52 + Int.unsigned x).
Proof.
- intros.
+ intros.
pose proof (Int.unsigned_range x).
destruct (from_words_value x) as (A & B & C).
destruct (BofZ_exact 53 1024 __ __ (2^52 + Int.unsigned x)) as (D & E & F).
smart_omega.
apply B2R_Bsign_inj; auto.
- rewrite A, D. rewrite Z2R_plus. auto.
- rewrite C, F. symmetry. apply Zlt_bool_false. smart_omega.
+ rewrite A, D. rewrite Z2R_plus. auto.
+ rewrite C, F. symmetry. apply Zlt_bool_false. smart_omega.
Qed.
Theorem of_intu_from_words:
@@ -537,7 +537,7 @@ Theorem of_intu_from_words:
of_intu x = sub (from_words ox4330_0000 x) (from_words ox4330_0000 Int.zero).
Proof.
intros. pose proof (Int.unsigned_range x).
- rewrite ! from_words_eq. unfold sub. rewrite BofZ_minus.
+ rewrite ! from_words_eq. unfold sub. rewrite BofZ_minus.
unfold of_intu. f_equal. rewrite Int.unsigned_zero. omega.
apply integer_representable_n; auto; smart_omega.
apply integer_representable_n; auto; rewrite Int.unsigned_zero; smart_omega.
@@ -560,11 +560,11 @@ Theorem of_int_from_words:
of_int x = sub (from_words ox4330_0000 (Int.add x ox8000_0000))
(from_words ox4330_0000 ox8000_0000).
Proof.
- intros.
+ intros.
pose proof (Int.signed_range x).
rewrite ! from_words_eq. rewrite ox8000_0000_signed_unsigned.
change (Int.unsigned ox8000_0000) with Int.half_modulus.
- unfold sub. rewrite BofZ_minus.
+ unfold sub. rewrite BofZ_minus.
unfold of_int. f_equal. omega.
apply integer_representable_n; auto; smart_omega.
apply integer_representable_n; auto; smart_omega.
@@ -607,16 +607,16 @@ Qed.
Lemma from_words_eq':
forall x, from_words ox4530_0000 x = BofZ 53 1024 __ __ (2^84 + Int.unsigned x * 2^32).
Proof.
- intros.
+ intros.
pose proof (Int.unsigned_range x).
destruct (from_words_value' x) as (A & B & C).
destruct (BofZ_representable 53 1024 __ __ (2^84 + Int.unsigned x * 2^32)) as (D & E & F).
replace (2^84 + Int.unsigned x * 2^32)
- with ((2^52 + Int.unsigned x) * 2^32) by ring.
+ with ((2^52 + Int.unsigned x) * 2^32) by ring.
apply integer_representable_n2p; auto. smart_omega. omega. omega.
apply B2R_Bsign_inj; auto.
- rewrite A, D. rewrite <- Z2R_Zpower by omega. rewrite <- Z2R_plus. auto.
- rewrite C, F. symmetry. apply Zlt_bool_false.
+ rewrite A, D. rewrite <- Z2R_Zpower by omega. rewrite <- Z2R_plus. auto.
+ rewrite C, F. symmetry. apply Zlt_bool_false.
compute_this (2^84); compute_this (2^32); omega.
Qed.
@@ -631,7 +631,7 @@ Proof.
pose proof (Int64.unsigned_range l).
pose proof (Int.unsigned_range (Int64.hiword l)).
pose proof (Int.unsigned_range (Int64.loword l)).
- rewrite ! from_words_eq, ! from_words_eq'.
+ rewrite ! from_words_eq, ! from_words_eq'.
set (p20 := Int.unsigned (Int.repr (two_p 20))).
set (x := Int64.unsigned l) in *;
set (xl := Int.unsigned (Int64.loword l)) in *;
@@ -639,17 +639,17 @@ Proof.
unfold sub. rewrite BofZ_minus.
replace (2^84 + xh * 2^32 - (2^84 + p20 * 2^32))
with ((xh - p20) * 2^32) by ring.
- unfold add. rewrite BofZ_plus.
- unfold of_longu. f_equal.
+ unfold add. rewrite BofZ_plus.
+ unfold of_longu. f_equal.
rewrite <- (Int64.ofwords_recompose l) at 1. rewrite Int64.ofwords_add'.
fold xh; fold xl. compute_this (two_p 32); compute_this p20; ring.
apply integer_representable_n2p; auto.
compute_this p20; smart_omega. omega. omega.
- apply integer_representable_n; auto; smart_omega.
+ apply integer_representable_n; auto; smart_omega.
replace (2^84 + xh * 2^32) with ((2^52 + xh) * 2^32) by ring.
apply integer_representable_n2p; auto. smart_omega. omega. omega.
change (2^84 + p20 * 2^32) with ((2^52 + 1048576) * 2^32).
- apply integer_representable_n2p; auto. omega. omega.
+ apply integer_representable_n2p; auto. omega. omega.
Qed.
Theorem of_long_from_words:
@@ -663,29 +663,29 @@ Proof.
pose proof (Int64.signed_range l).
pose proof (Int.signed_range (Int64.hiword l)).
pose proof (Int.unsigned_range (Int64.loword l)).
- rewrite ! from_words_eq, ! from_words_eq'.
+ rewrite ! from_words_eq, ! from_words_eq'.
set (p := Int.unsigned (Int.repr (two_p 20 + two_p 31))).
set (x := Int64.signed l) in *;
set (xl := Int.unsigned (Int64.loword l)) in *;
set (xh := Int.signed (Int64.hiword l)) in *.
- rewrite ox8000_0000_signed_unsigned. fold xh.
+ rewrite ox8000_0000_signed_unsigned. fold xh.
unfold sub. rewrite BofZ_minus.
replace (2^84 + (xh + Int.half_modulus) * 2^32 - (2^84 + p * 2^32))
- with ((xh - 2^20) * 2^32)
+ with ((xh - 2^20) * 2^32)
by (compute_this p; compute_this Int.half_modulus; ring).
- unfold add. rewrite BofZ_plus.
- unfold of_long. f_equal.
+ unfold add. rewrite BofZ_plus.
+ unfold of_long. f_equal.
rewrite <- (Int64.ofwords_recompose l) at 1. rewrite Int64.ofwords_add''.
- fold xh; fold xl. compute_this (two_p 32); ring.
+ fold xh; fold xl. compute_this (two_p 32); ring.
apply integer_representable_n2p; auto.
compute_this (2^20); smart_omega. omega. omega.
apply integer_representable_n; auto; smart_omega.
replace (2^84 + (xh + Int.half_modulus) * 2^32)
- with ((2^52 + xh + Int.half_modulus) * 2^32)
+ with ((2^52 + xh + Int.half_modulus) * 2^32)
by (compute_this Int.half_modulus; ring).
apply integer_representable_n2p; auto. smart_omega. omega. omega.
change (2^84 + p * 2^32) with ((2^52 + p) * 2^32).
- apply integer_representable_n2p; auto.
+ apply integer_representable_n2p; auto.
compute_this p; smart_omega. omega.
Qed.
@@ -708,7 +708,7 @@ Proof.
assert (DECOMP: x = yh * 2^32 + yl).
{ unfold x. rewrite <- (Int64.ofwords_recompose l). apply Int64.ofwords_add'. }
rewrite BofZ_mult. rewrite BofZ_plus. rewrite DECOMP; auto.
- apply integer_representable_n2p; auto. smart_omega. omega. omega.
+ apply integer_representable_n2p; auto. smart_omega. omega. omega.
apply integer_representable_n; auto; smart_omega.
apply integer_representable_n; auto; smart_omega.
apply integer_representable_n; auto; smart_omega.
@@ -761,7 +761,7 @@ Theorem of_longu_of_long_2:
Proof.
intros. change (of_int (Int.repr 2)) with (BofZ 53 1024 __ __ (2^1)).
pose proof (Int64.unsigned_range x).
- unfold Int64.ltu in H.
+ unfold Int64.ltu in H.
change (Int64.unsigned (Int64.repr Int64.half_modulus)) with (2^63) in H.
destruct (zlt (Int64.unsigned x) (2^63)); inv H.
assert (Int64.modulus <= 2^1024 - 2^(1024-53)) by (vm_compute; intuition congruence).
@@ -771,10 +771,10 @@ Proof.
if zeq i 0 then Int64.testbit x 1 || Int64.testbit x 0
else if zeq i 63 then false else Int64.testbit x (i + 1)).
{ intros; unfold n; autorewrite with ints; auto. rewrite Int64.unsigned_one.
- rewrite Int64.bits_one. compute_this Int64.zwordsize.
+ rewrite Int64.bits_one. compute_this Int64.zwordsize.
destruct (zeq i 0); simpl proj_sumbool.
- rewrite zlt_true by omega. rewrite andb_true_r. subst i; auto.
- rewrite andb_false_r, orb_false_r.
+ rewrite zlt_true by omega. rewrite andb_true_r. subst i; auto.
+ rewrite andb_false_r, orb_false_r.
destruct (zeq i 63). subst i. apply zlt_false; omega.
apply zlt_true; omega. }
assert (NB2: forall i, 0 <= i ->
@@ -784,29 +784,29 @@ Proof.
Int64.testbit x i).
{ intros. rewrite Z.mul_pow2_bits by omega. destruct (zeq i 0).
apply Z.testbit_neg_r; omega.
- rewrite Int64.bits_signed by omega. compute_this Int64.zwordsize.
- destruct (zlt (i-1) 64).
+ rewrite Int64.bits_signed by omega. compute_this Int64.zwordsize.
+ destruct (zlt (i-1) 64).
rewrite NB by omega. destruct (zeq i 1).
subst. rewrite dec_eq_true by auto. auto.
- rewrite dec_eq_false by omega. destruct (zeq (i - 1) 63).
- symmetry. apply Int64.bits_above. compute_this Int64.zwordsize; omega.
+ rewrite dec_eq_false by omega. destruct (zeq (i - 1) 63).
+ symmetry. apply Int64.bits_above. compute_this Int64.zwordsize; omega.
f_equal; omega.
- rewrite NB by omega. rewrite dec_eq_false by omega. rewrite dec_eq_true by auto.
- rewrite dec_eq_false by omega. symmetry. apply Int64.bits_above. compute_this Int64.zwordsize; omega.
+ rewrite NB by omega. rewrite dec_eq_false by omega. rewrite dec_eq_true by auto.
+ rewrite dec_eq_false by omega. symmetry. apply Int64.bits_above. compute_this Int64.zwordsize; omega.
}
assert (EQ: Int64.signed n * 2 = int_round_odd (Int64.unsigned x) 1).
{
symmetry. apply (int_round_odd_bits 53 1024). omega.
- intros. rewrite NB2 by omega. replace i with 0 by omega. auto.
- rewrite NB2 by omega. rewrite dec_eq_false by omega. rewrite dec_eq_true.
+ intros. rewrite NB2 by omega. replace i with 0 by omega. auto.
+ rewrite NB2 by omega. rewrite dec_eq_false by omega. rewrite dec_eq_true.
rewrite orb_comm. unfold Int64.testbit. change (2^1) with 2.
destruct (Z.testbit (Int64.unsigned x) 0) eqn:B0;
[rewrite Z.testbit_true in B0 by omega|rewrite Z.testbit_false in B0 by omega];
change (2^0) with 1 in B0; rewrite Zdiv_1_r in B0; rewrite B0; auto.
intros. rewrite NB2 by omega. rewrite ! dec_eq_false by omega. auto.
}
- unfold mul, of_long, of_longu.
- rewrite BofZ_mult_2p.
+ unfold mul, of_long, of_longu.
+ rewrite BofZ_mult_2p.
- change (2^1) with 2. rewrite EQ. apply BofZ_round_odd with (p := 1).
+ omega.
+ apply Zle_trans with Int64.modulus; trivial. smart_omega.
@@ -818,7 +818,7 @@ Proof.
compute_this Int64.modulus; xomega.
- assert (2^63 <= int_round_odd (Int64.unsigned x) 1).
{ change (2^63) with (int_round_odd (2^63) 1). apply (int_round_odd_le 0 0); omega. }
- rewrite <- EQ in H1. compute_this (2^63). compute_this (2^53). xomega.
+ rewrite <- EQ in H1. compute_this (2^63). compute_this (2^53). xomega.
- omega.
Qed.
@@ -941,10 +941,10 @@ Qed.
Theorem mul2_add:
forall f, add f f = mul f (of_int (Int.repr 2%Z)).
Proof.
- intros. apply Bmult2_Bplus.
- intros. destruct x; try discriminate. simpl.
- transitivity (b, transform_quiet_pl n).
- destruct Archi.choose_binop_pl_32; auto.
+ intros. apply Bmult2_Bplus.
+ intros. destruct x; try discriminate. simpl.
+ transitivity (b, transform_quiet_pl n).
+ destruct Archi.choose_binop_pl_32; auto.
destruct y; auto || discriminate.
Qed.
@@ -955,9 +955,9 @@ Definition exact_inverse : float32 -> option float32 := Bexact_inverse 24 128 __
Theorem div_mul_inverse:
forall x y z, exact_inverse y = Some z -> div x y = mul x z.
Proof.
- intros. apply Bdiv_mult_inverse; auto.
- intros. destruct x0; try discriminate. simpl.
- transitivity (b, transform_quiet_pl n).
+ intros. apply Bdiv_mult_inverse; auto.
+ intros. destruct x0; try discriminate. simpl.
+ transitivity (b, transform_quiet_pl n).
destruct y0; reflexivity || discriminate.
destruct z0; reflexivity || discriminate.
Qed.
@@ -968,13 +968,13 @@ Theorem cmp_swap:
forall c x y, cmp (swap_comparison c) x y = cmp c y x.
Proof.
unfold cmp; intros. rewrite (Bcompare_swap _ _ x y).
- apply cmp_of_comparison_swap.
+ apply cmp_of_comparison_swap.
Qed.
Theorem cmp_ne_eq:
forall f1 f2, cmp Cne f1 f2 = negb (cmp Ceq f1 f2).
Proof.
- intros; apply cmp_of_comparison_ne_eq.
+ intros; apply cmp_of_comparison_ne_eq.
Qed.
Theorem cmp_lt_eq_false:
@@ -1031,7 +1031,7 @@ Theorem to_of_bits:
Proof.
intros; unfold of_bits, to_bits, bits_of_b32, b32_of_bits.
rewrite bits_of_binary_float_of_bits. apply Int.repr_unsigned.
- apply Int.unsigned_range.
+ apply Int.unsigned_range.
Qed.
(** Conversions from 32-bit integers to single-precision floats can
@@ -1041,15 +1041,15 @@ Qed.
Theorem of_int_double:
forall n, of_int n = of_double (Float.of_int n).
Proof.
- intros. symmetry. apply Bconv_BofZ.
- apply integer_representable_n; auto. generalize (Int.signed_range n); Float.smart_omega.
+ intros. symmetry. apply Bconv_BofZ.
+ apply integer_representable_n; auto. generalize (Int.signed_range n); Float.smart_omega.
Qed.
Theorem of_intu_double:
forall n, of_intu n = of_double (Float.of_intu n).
Proof.
intros. symmetry. apply Bconv_BofZ.
- apply integer_representable_n; auto. generalize (Int.unsigned_range n); Float.smart_omega.
+ apply integer_representable_n; auto. generalize (Int.unsigned_range n); Float.smart_omega.
Qed.
(** Conversion of single-precision floats to integers can be decomposed
@@ -1062,8 +1062,8 @@ Proof.
intros.
unfold to_int in H.
destruct (ZofB_range _ _ f Int.min_signed Int.max_signed) as [n'|] eqn:E; inv H.
- unfold Float.to_int, to_double, Float.of_single.
- erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
+ unfold Float.to_int, to_double, Float.of_single.
+ erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
Qed.
Theorem to_intu_double:
@@ -1072,8 +1072,8 @@ Proof.
intros.
unfold to_intu in H.
destruct (ZofB_range _ _ f 0 Int.max_unsigned) as [n'|] eqn:E; inv H.
- unfold Float.to_intu, to_double, Float.of_single.
- erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
+ unfold Float.to_intu, to_double, Float.of_single.
+ erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
Qed.
Theorem to_long_double:
@@ -1082,8 +1082,8 @@ Proof.
intros.
unfold to_long in H.
destruct (ZofB_range _ _ f Int64.min_signed Int64.max_signed) as [n'|] eqn:E; inv H.
- unfold Float.to_long, to_double, Float.of_single.
- erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
+ unfold Float.to_long, to_double, Float.of_single.
+ erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
Qed.
Theorem to_longu_double:
@@ -1092,8 +1092,8 @@ Proof.
intros.
unfold to_longu in H.
destruct (ZofB_range _ _ f 0 Int64.max_unsigned) as [n'|] eqn:E; inv H.
- unfold Float.to_longu, to_double, Float.of_single.
- erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
+ unfold Float.to_longu, to_double, Float.of_single.
+ erewrite ZofB_range_Bconv; eauto. auto. omega. omega. omega. omega.
Qed.
(** Conversions from 64-bit integers to single-precision floats can be expressed
@@ -1106,7 +1106,7 @@ Lemma int_round_odd_plus:
int_round_odd n p = Z.land (Z.lor n (Z.land n (2^p-1) + (2^p-1))) (-(2^p)).
Proof.
intros.
- assert (POS: 0 < 2^p) by (apply (Zpower_gt_0 radix2); auto).
+ assert (POS: 0 < 2^p) by (apply (Zpower_gt_0 radix2); auto).
assert (A: Z.land n (2^p-1) = n mod 2^p).
{ rewrite <- Z.land_ones by auto. f_equal. rewrite Z.ones_equiv. omega. }
rewrite A.
@@ -1115,29 +1115,29 @@ Proof.
set (m := n mod 2^p + (2^p-1)) in *.
assert (C: m / 2^p = if zeq (n mod 2^p) 0 then 0 else 1).
{ unfold m. destruct (zeq (n mod 2^p) 0).
- rewrite e. apply Zdiv_small. omega.
+ rewrite e. apply Zdiv_small. omega.
eapply Zdiv_unique with (n mod 2^p - 1). ring. omega. }
assert (D: Z.testbit m p = if zeq (n mod 2^p) 0 then false else true).
{ destruct (zeq (n mod 2^p) 0).
apply Z.testbit_false; auto. rewrite C; auto.
apply Z.testbit_true; auto. rewrite C; auto. }
assert (E: forall i, p < i -> Z.testbit m i = false).
- { intros. apply Z.testbit_false. omega.
- replace (m / 2^i) with 0. auto. symmetry. apply Zdiv_small.
- unfold m. split. omega. apply Zlt_le_trans with (2 * 2^p). omega.
- change 2 with (2^1) at 1. rewrite <- (Zpower_plus radix2) by omega.
+ { intros. apply Z.testbit_false. omega.
+ replace (m / 2^i) with 0. auto. symmetry. apply Zdiv_small.
+ unfold m. split. omega. apply Zlt_le_trans with (2 * 2^p). omega.
+ change 2 with (2^1) at 1. rewrite <- (Zpower_plus radix2) by omega.
apply Zpower_le. omega. }
assert (F: forall i, 0 <= i -> Z.testbit (-2^p) i = if zlt i p then false else true).
{ intros. rewrite Z.bits_opp by auto. rewrite <- Z.ones_equiv.
- destruct (zlt i p).
+ destruct (zlt i p).
rewrite Z.ones_spec_low by omega. auto.
rewrite Z.ones_spec_high by omega. auto. }
- apply int_round_odd_bits; auto.
- - intros. rewrite Z.land_spec, F, zlt_true by omega. apply andb_false_r.
+ apply int_round_odd_bits; auto.
+ - intros. rewrite Z.land_spec, F, zlt_true by omega. apply andb_false_r.
- rewrite Z.land_spec, Z.lor_spec, D, F, zlt_false, andb_true_r by omega.
- destruct (Z.eqb (n mod 2^p) 0) eqn:Z.
- rewrite Z.eqb_eq in Z. rewrite Z, zeq_true. apply orb_false_r.
- rewrite Z.eqb_neq in Z. rewrite zeq_false by auto. apply orb_true_r.
+ destruct (Z.eqb (n mod 2^p) 0) eqn:Z.
+ rewrite Z.eqb_eq in Z. rewrite Z, zeq_true. apply orb_false_r.
+ rewrite Z.eqb_neq in Z. rewrite zeq_false by auto. apply orb_true_r.
- intros. rewrite Z.land_spec, Z.lor_spec, E, F, zlt_false, andb_true_r by omega.
apply orb_false_r.
Qed.
@@ -1148,18 +1148,18 @@ Lemma of_long_round_odd:
BofZ 24 128 __ __ n = Bconv _ _ 24 128 __ __ conv_nan mode_NE (BofZ 53 1024 __ __ (Z.land (Z.lor n ((Z.land n 2047) + 2047)) (-2048))).
Proof.
intros. rewrite <- (int_round_odd_plus 11) by omega.
- assert (-2^64 <= int_round_odd n 11).
+ assert (-2^64 <= int_round_odd n 11).
{ change (-2^64) with (int_round_odd (-2^64) 11). apply (int_round_odd_le 0 0); xomega. }
- assert (int_round_odd n 11 <= 2^64).
+ assert (int_round_odd n 11 <= 2^64).
{ change (2^64) with (int_round_odd (2^64) 11). apply (int_round_odd_le 0 0); xomega. }
- rewrite Bconv_BofZ.
+ rewrite Bconv_BofZ.
apply BofZ_round_odd with (p := 11).
omega.
apply Zle_trans with (2^64). omega. compute; intuition congruence.
omega.
- exact (proj1 H).
- unfold int_round_odd. apply integer_representable_n2p_wide. auto. omega.
- unfold int_round_odd in H0, H1.
+ exact (proj1 H).
+ unfold int_round_odd. apply integer_representable_n2p_wide. auto. omega.
+ unfold int_round_odd in H0, H1.
split; (apply Zmult_le_reg_r with (2^11); [compute; auto | assumption]).
omega.
omega.
@@ -1170,46 +1170,46 @@ Theorem of_longu_double_1:
Int64.unsigned n <= 2^53 ->
of_longu n = of_double (Float.of_longu n).
Proof.
- intros. symmetry; apply Bconv_BofZ. apply integer_representable_n; auto.
+ intros. symmetry; apply Bconv_BofZ. apply integer_representable_n; auto.
pose proof (Int64.unsigned_range n); omega.
Qed.
Theorem of_longu_double_2:
forall n,
2^36 <= Int64.unsigned n ->
- of_longu n = of_double (Float.of_longu
- (Int64.and (Int64.or n
+ of_longu n = of_double (Float.of_longu
+ (Int64.and (Int64.or n
(Int64.add (Int64.and n (Int64.repr 2047))
(Int64.repr 2047)))
(Int64.repr (-2048)))).
Proof.
intros.
- pose proof (Int64.unsigned_range n).
+ pose proof (Int64.unsigned_range n).
unfold of_longu. erewrite of_long_round_odd.
- unfold of_double, Float.to_single. instantiate (1 := Float.to_single_pl).
+ unfold of_double, Float.to_single. instantiate (1 := Float.to_single_pl).
f_equal. unfold Float.of_longu. f_equal.
set (n' := Z.land (Z.lor (Int64.unsigned n) (Z.land (Int64.unsigned n) 2047 + 2047)) (-2048)).
assert (int_round_odd (Int64.unsigned n) 11 = n') by (apply int_round_odd_plus; omega).
- assert (0 <= n').
+ assert (0 <= n').
{ rewrite <- H1. change 0 with (int_round_odd 0 11). apply (int_round_odd_le 0 0); omega. }
- assert (n' < Int64.modulus).
- { apply Zle_lt_trans with (int_round_odd (Int64.modulus - 1) 11).
- rewrite <- H1. apply (int_round_odd_le 0 0); omega.
+ assert (n' < Int64.modulus).
+ { apply Zle_lt_trans with (int_round_odd (Int64.modulus - 1) 11).
+ rewrite <- H1. apply (int_round_odd_le 0 0); omega.
compute; auto. }
rewrite <- (Int64.unsigned_repr n') by (unfold Int64.max_unsigned; omega).
f_equal. Int64.bit_solve. rewrite Int64.testbit_repr by auto. unfold n'.
- rewrite Z.land_spec, Z.lor_spec. f_equal. f_equal.
+ rewrite Z.land_spec, Z.lor_spec. f_equal. f_equal.
unfold Int64.testbit. rewrite Int64.add_unsigned.
fold (Int64.testbit (Int64.repr
(Int64.unsigned (Int64.and n (Int64.repr 2047)) +
Int64.unsigned (Int64.repr 2047))) i).
rewrite Int64.testbit_repr by auto. f_equal. f_equal. unfold Int64.and.
symmetry. apply Int64.unsigned_repr. change 2047 with (Z.ones 11).
- rewrite Z.land_ones by omega.
- exploit (Z_mod_lt (Int64.unsigned n) (2^11)). compute; auto.
- assert (2^11 < Int64.max_unsigned) by (compute; auto). omega.
+ rewrite Z.land_ones by omega.
+ exploit (Z_mod_lt (Int64.unsigned n) (2^11)). compute; auto.
+ assert (2^11 < Int64.max_unsigned) by (compute; auto). omega.
apply Int64.same_bits_eqm; auto. exists (-1); auto.
- split. xomega. change (2^64) with Int64.modulus. xomega.
+ split. xomega. change (2^64) with Int64.modulus. xomega.
Qed.
Theorem of_long_double_1:
@@ -1217,50 +1217,50 @@ Theorem of_long_double_1:
Z.abs (Int64.signed n) <= 2^53 ->
of_long n = of_double (Float.of_long n).
Proof.
- intros. symmetry; apply Bconv_BofZ. apply integer_representable_n; auto. xomega.
+ intros. symmetry; apply Bconv_BofZ. apply integer_representable_n; auto. xomega.
Qed.
Theorem of_long_double_2:
forall n,
2^36 <= Z.abs (Int64.signed n) ->
of_long n = of_double (Float.of_long
- (Int64.and (Int64.or n
+ (Int64.and (Int64.or n
(Int64.add (Int64.and n (Int64.repr 2047))
(Int64.repr 2047)))
(Int64.repr (-2048)))).
Proof.
intros.
- pose proof (Int64.signed_range n).
+ pose proof (Int64.signed_range n).
unfold of_long. erewrite of_long_round_odd.
- unfold of_double, Float.to_single. instantiate (1 := Float.to_single_pl).
+ unfold of_double, Float.to_single. instantiate (1 := Float.to_single_pl).
f_equal. unfold Float.of_long. f_equal.
set (n' := Z.land (Z.lor (Int64.signed n) (Z.land (Int64.signed n) 2047 + 2047)) (-2048)).
assert (int_round_odd (Int64.signed n) 11 = n') by (apply int_round_odd_plus; omega).
- assert (Int64.min_signed <= n').
+ assert (Int64.min_signed <= n').
{ rewrite <- H1. change Int64.min_signed with (int_round_odd Int64.min_signed 11). apply (int_round_odd_le 0 0); omega. }
assert (n' <= Int64.max_signed).
- { apply Zle_trans with (int_round_odd Int64.max_signed 11).
- rewrite <- H1. apply (int_round_odd_le 0 0); omega.
+ { apply Zle_trans with (int_round_odd Int64.max_signed 11).
+ rewrite <- H1. apply (int_round_odd_le 0 0); omega.
compute; intuition congruence. }
rewrite <- (Int64.signed_repr n') by omega.
f_equal. Int64.bit_solve. rewrite Int64.testbit_repr by auto. unfold n'.
rewrite Z.land_spec, Z.lor_spec. f_equal. f_equal.
- rewrite Int64.bits_signed by omega. rewrite zlt_true by omega. auto.
+ rewrite Int64.bits_signed by omega. rewrite zlt_true by omega. auto.
unfold Int64.testbit. rewrite Int64.add_unsigned.
fold (Int64.testbit (Int64.repr
(Int64.unsigned (Int64.and n (Int64.repr 2047)) +
Int64.unsigned (Int64.repr 2047))) i).
rewrite Int64.testbit_repr by auto. f_equal. f_equal. unfold Int64.and.
- change (Int64.unsigned (Int64.repr 2047)) with 2047.
+ change (Int64.unsigned (Int64.repr 2047)) with 2047.
change 2047 with (Z.ones 11). rewrite ! Z.land_ones by omega.
- rewrite Int64.unsigned_repr. apply Int64.eqmod_mod_eq.
+ rewrite Int64.unsigned_repr. apply Int64.eqmod_mod_eq.
apply Zlt_gt. apply (Zpower_gt_0 radix2); omega.
- apply Int64.eqmod_divides with (2^64). apply Int64.eqm_signed_unsigned.
+ apply Int64.eqmod_divides with (2^64). apply Int64.eqm_signed_unsigned.
exists (2^(64-11)); auto.
- exploit (Z_mod_lt (Int64.unsigned n) (2^11)). compute; auto.
- assert (2^11 < Int64.max_unsigned) by (compute; auto). omega.
+ exploit (Z_mod_lt (Int64.unsigned n) (2^11)). compute; auto.
+ assert (2^11 < Int64.max_unsigned) by (compute; auto). omega.
apply Int64.same_bits_eqm; auto. exists (-1); auto.
- split. auto. assert (-2^64 < Int64.min_signed) by (compute; auto).
+ split. auto. assert (-2^64 < Int64.min_signed) by (compute; auto).
assert (Int64.max_signed < 2^64) by (compute; auto).
xomega.
Qed.
diff --git a/lib/Heaps.v b/lib/Heaps.v
index 0ee07a58..65334a38 100644
--- a/lib/Heaps.v
+++ b/lib/Heaps.v
@@ -152,7 +152,7 @@ Lemma le_lt_trans:
Proof.
unfold le; intros; intuition.
destruct (E.compare x1 x3).
- auto.
+ auto.
elim (@E.lt_not_eq x2 x3). auto. apply E.eq_trans with x1. apply E.eq_sym; auto. auto.
elim (@E.lt_not_eq x2 x1). eapply E.lt_trans; eauto. apply E.eq_sym; auto.
eapply E.lt_trans; eauto.
@@ -163,7 +163,7 @@ Lemma lt_le_trans:
Proof.
unfold le; intros; intuition.
destruct (E.compare x1 x3).
- auto.
+ auto.
elim (@E.lt_not_eq x1 x2). auto. apply E.eq_trans with x3. auto. apply E.eq_sym; auto.
elim (@E.lt_not_eq x3 x2). eapply E.lt_trans; eauto. apply E.eq_sym; auto.
eapply E.lt_trans; eauto.
@@ -172,7 +172,7 @@ Qed.
Lemma le_trans:
forall x1 x2 x3, le x1 x2 -> le x2 x3 -> le x1 x3.
Proof.
- intros. destruct H. destruct H0. red; left; eapply E.eq_trans; eauto.
+ intros. destruct H. destruct H0. red; left; eapply E.eq_trans; eauto.
red. right. eapply le_lt_trans; eauto. red; auto.
red. right. eapply lt_le_trans; eauto.
Qed.
@@ -181,7 +181,7 @@ Lemma lt_heap_trans:
forall x y, le x y ->
forall h, lt_heap h x -> lt_heap h y.
Proof.
- induction h; simpl; intros.
+ induction h; simpl; intros.
auto.
intuition. eapply lt_le_trans; eauto.
Qed.
@@ -190,7 +190,7 @@ Lemma gt_heap_trans:
forall x y, le y x ->
forall h, gt_heap h x -> gt_heap h y.
Proof.
- induction h; simpl; intros.
+ induction h; simpl; intros.
auto.
intuition. eapply le_lt_trans; eauto.
Qed.
@@ -205,12 +205,12 @@ Proof.
- tauto.
- tauto.
- rewrite e3 in *; simpl in *; intuition.
-- intuition. elim NEQ. eapply E.eq_trans; eauto.
+- intuition. elim NEQ. eapply E.eq_trans; eauto.
- rewrite e3 in *; simpl in *; intuition.
-- intuition. elim NEQ. eapply E.eq_trans; eauto.
-- intuition.
+- intuition. elim NEQ. eapply E.eq_trans; eauto.
+- intuition.
- rewrite e3 in *; simpl in *; intuition.
-- intuition. elim NEQ. eapply E.eq_trans; eauto.
+- intuition. elim NEQ. eapply E.eq_trans; eauto.
- rewrite e3 in *; simpl in *; intuition.
Qed.
@@ -224,7 +224,7 @@ Proof.
- rewrite e3 in *; simpl in *; tauto.
- rewrite e3 in *; simpl in *; tauto.
Qed.
-
+
Lemma partition_gt:
forall x pivot h,
gt_heap h x -> gt_heap (fst (partition pivot h)) x /\ gt_heap (snd (partition pivot h)) x.
@@ -249,18 +249,18 @@ Proof.
- intuition.
eapply lt_heap_trans; eauto. red; auto.
eapply lt_heap_trans; eauto. red; auto.
- eapply gt_heap_trans with y; eauto. red. left. apply E.eq_sym; auto.
+ eapply gt_heap_trans with y; eauto. red. left. apply E.eq_sym; auto.
- rewrite e3 in *; simpl in *; intuition.
eapply lt_heap_trans; eauto. red; auto.
eapply gt_heap_trans with y; eauto. red; auto.
-- intuition.
+- intuition.
eapply lt_heap_trans; eauto. red; auto.
eapply gt_heap_trans; eauto. red; auto.
-- intuition. eapply gt_heap_trans; eauto. red; auto.
+- intuition. eapply gt_heap_trans; eauto. red; auto.
- rewrite e3 in *; simpl in *. intuition.
eapply lt_heap_trans with y; eauto. red; auto.
eapply gt_heap_trans; eauto. red; auto.
-- intuition.
+- intuition.
eapply lt_heap_trans with y; eauto. red; auto.
eapply gt_heap_trans; eauto. red; auto.
eapply gt_heap_trans with x; eauto. red; auto.
@@ -275,16 +275,16 @@ Lemma partition_bst:
bst (fst (partition pivot h)) /\ bst (snd (partition pivot h)).
Proof.
intros pivot h0. functional induction (partition pivot h0); simpl; try tauto.
-- rewrite e3 in *; simpl in *. intuition.
+- rewrite e3 in *; simpl in *. intuition.
apply lt_heap_trans with x; auto. red; auto.
generalize (partition_gt y pivot b2 H7). rewrite e3; simpl. tauto.
-- rewrite e3 in *; simpl in *. intuition.
+- rewrite e3 in *; simpl in *. intuition.
generalize (partition_gt x pivot b1 H3). rewrite e3; simpl. tauto.
generalize (partition_lt y pivot b1 H4). rewrite e3; simpl. tauto.
- rewrite e3 in *; simpl in *. intuition.
generalize (partition_gt y pivot a2 H6). rewrite e3; simpl. tauto.
generalize (partition_lt x pivot a2 H8). rewrite e3; simpl. tauto.
-- rewrite e3 in *; simpl in *. intuition.
+- rewrite e3 in *; simpl in *. intuition.
generalize (partition_lt y pivot a1 H3). rewrite e3; simpl. tauto.
apply gt_heap_trans with x; auto. red; auto.
Qed.
@@ -294,8 +294,8 @@ Qed.
Lemma insert_bst:
forall x h, bst h -> bst (insert x h).
Proof.
- intros.
- unfold insert. case_eq (partition x h). intros a b EQ; simpl.
+ intros.
+ unfold insert. case_eq (partition x h). intros a b EQ; simpl.
generalize (partition_bst x h H).
generalize (partition_split x h H).
rewrite EQ; simpl. tauto.
@@ -305,13 +305,13 @@ Lemma In_insert:
forall x h y, bst h -> (In y (insert x h) <-> E.eq y x \/ In y h).
Proof.
intros. unfold insert.
- case_eq (partition x h). intros a b EQ; simpl.
+ case_eq (partition x h). intros a b EQ; simpl.
assert (E.eq y x \/ ~E.eq y x).
destruct (E.compare y x); auto.
right; red; intros. elim (E.lt_not_eq l). apply E.eq_sym; auto.
destruct H0.
tauto.
- generalize (In_partition y x H0 h H). rewrite EQ; simpl. tauto.
+ generalize (In_partition y x H0 h H). rewrite EQ; simpl. tauto.
Qed.
(** Properties of [findMin] and [deleteMin] *)
@@ -324,7 +324,7 @@ Opaque deleteMin.
auto.
tauto.
tauto.
- intuition. apply IHh. simpl. tauto.
+ intuition. apply IHh. simpl. tauto.
Qed.
Lemma deleteMin_bst:
@@ -336,8 +336,8 @@ Proof.
tauto.
intuition.
apply IHh. simpl; auto.
- apply deleteMin_lt; auto. simpl; auto.
- apply gt_heap_trans with y; auto. red; auto.
+ apply deleteMin_lt; auto. simpl; auto.
+ apply gt_heap_trans with y; auto. red; auto.
Qed.
Lemma In_deleteMin:
@@ -347,16 +347,16 @@ Lemma In_deleteMin:
Proof.
Transparent deleteMin.
intros y x h0. functional induction (deleteMin h0); simpl; intros.
- discriminate.
+ discriminate.
+ inv H. tauto.
inv H. tauto.
- inv H. tauto.
destruct _x. inv H. simpl. tauto. generalize (IHh H). simpl. tauto.
Qed.
Lemma gt_heap_In:
forall x y h, gt_heap h x -> In y h -> E.lt x y.
Proof.
- induction h; simpl; intros.
+ induction h; simpl; intros.
contradiction.
intuition. apply lt_le_trans with x0; auto. red. left. apply E.eq_sym; auto.
Qed.
@@ -373,7 +373,7 @@ Proof.
assert (le x x1).
apply IHh1; auto. tauto. simpl. right; left; apply E.eq_refl.
intuition.
- apply le_trans with x1. auto. apply le_trans with x0. simpl in H4. red; tauto.
+ apply le_trans with x1. auto. apply le_trans with x0. simpl in H4. red; tauto.
red; left; apply E.eq_sym; auto.
apply le_trans with x1. auto. apply le_trans with x0. simpl in H4. red; tauto.
red; right. eapply gt_heap_In; eauto.
@@ -396,8 +396,8 @@ Opaque deleteMax.
intros x h0. functional induction (deleteMax h0); simpl; intros.
auto.
tauto.
- tauto.
- intuition. apply IHh. simpl. tauto.
+ tauto.
+ intuition. apply IHh. simpl. tauto.
Qed.
Lemma deleteMax_bst:
@@ -410,7 +410,7 @@ Proof.
intuition.
apply IHh. simpl; auto.
apply lt_heap_trans with x; auto. red; auto.
- apply deleteMax_gt; auto. simpl; auto.
+ apply deleteMax_gt; auto. simpl; auto.
Qed.
Lemma In_deleteMax:
@@ -422,14 +422,14 @@ Transparent deleteMax.
intros y x h0. functional induction (deleteMax h0); simpl; intros.
congruence.
inv H. tauto.
- inv H. tauto.
+ inv H. tauto.
destruct _x1. inv H. simpl. tauto. generalize (IHh H). simpl. tauto.
Qed.
Lemma lt_heap_In:
forall x y h, lt_heap h x -> In y h -> E.lt y x.
Proof.
- induction h; simpl; intros.
+ induction h; simpl; intros.
contradiction.
intuition. apply le_lt_trans with x0; auto. red. left. apply E.eq_sym; auto.
Qed.
@@ -448,7 +448,7 @@ Proof.
intuition.
apply le_trans with x1; auto. apply le_trans with x0.
red; right. eapply lt_heap_In; eauto.
- simpl in H6. red; tauto.
+ simpl in H6. red; tauto.
apply le_trans with x1; auto. apply le_trans with x0.
red; auto.
simpl in H6. red; tauto.
@@ -511,8 +511,8 @@ Qed.
Lemma findMin_empty:
forall h y, findMin h = None -> ~In y h.
Proof.
- unfold findMin, In; intros; simpl.
- destruct (proj1_sig h).
+ unfold findMin, In; intros; simpl.
+ destruct (proj1_sig h).
simpl. tauto.
exploit R.findMin_empty; eauto. congruence.
Qed.
@@ -540,8 +540,8 @@ Qed.
Lemma findMax_empty:
forall h y, findMax h = None -> ~In y h.
Proof.
- unfold findMax, In; intros; simpl.
- destruct (proj1_sig h).
+ unfold findMax, In; intros; simpl.
+ destruct (proj1_sig h).
simpl. tauto.
exploit R.findMax_empty; eauto. congruence.
Qed.
diff --git a/lib/Integers.v b/lib/Integers.v
index a3ff5209..a0140e57 100644
--- a/lib/Integers.v
+++ b/lib/Integers.v
@@ -133,16 +133,16 @@ Proof.
induction n; simpl; intros.
- rewrite two_power_nat_O. exists (Zpos p). ring.
- rewrite two_power_nat_S. destruct p.
- + destruct (IHn p) as [y EQ]. exists y.
- change (Zpos p~1) with (2 * Zpos p + 1). rewrite EQ.
- rewrite Z.succ_double_spec. ring.
- + destruct (IHn p) as [y EQ]. exists y.
- change (Zpos p~0) with (2 * Zpos p). rewrite EQ.
+ + destruct (IHn p) as [y EQ]. exists y.
+ change (Zpos p~1) with (2 * Zpos p + 1). rewrite EQ.
+ rewrite Z.succ_double_spec. ring.
+ + destruct (IHn p) as [y EQ]. exists y.
+ change (Zpos p~0) with (2 * Zpos p). rewrite EQ.
rewrite (Z.double_spec (P_mod_two_p p n)). ring.
+ exists 0; omega.
}
- intros.
- destruct (H n p) as [y EQ].
+ intros.
+ destruct (H n p) as [y EQ].
symmetry. apply Zmod_unique with y. auto. apply P_mod_two_p_range.
Qed.
@@ -150,12 +150,12 @@ Lemma Z_mod_modulus_range:
forall x, 0 <= Z_mod_modulus x < modulus.
Proof.
intros; unfold Z_mod_modulus.
- destruct x.
+ destruct x.
- generalize modulus_pos; omega.
- apply P_mod_two_p_range.
- set (r := P_mod_two_p p wordsize).
assert (0 <= r < modulus) by apply P_mod_two_p_range.
- destruct (zeq r 0).
+ destruct (zeq r 0).
+ generalize modulus_pos; omega.
+ omega.
Qed.
@@ -171,22 +171,22 @@ Lemma Z_mod_modulus_eq:
Proof.
intros. unfold Z_mod_modulus. destruct x.
- rewrite Zmod_0_l. auto.
- - apply P_mod_two_p_eq.
+ - apply P_mod_two_p_eq.
- generalize (P_mod_two_p_range wordsize p) (P_mod_two_p_eq wordsize p).
fold modulus. intros A B.
exploit (Z_div_mod_eq (Zpos p) modulus). apply modulus_pos. intros C.
set (q := Zpos p / modulus) in *.
- set (r := P_mod_two_p p wordsize) in *.
- rewrite <- B in C.
+ set (r := P_mod_two_p p wordsize) in *.
+ rewrite <- B in C.
change (Z.neg p) with (- (Z.pos p)). destruct (zeq r 0).
- + symmetry. apply Zmod_unique with (-q). rewrite C; rewrite e. ring.
+ + symmetry. apply Zmod_unique with (-q). rewrite C; rewrite e. ring.
generalize modulus_pos; omega.
+ symmetry. apply Zmod_unique with (-q - 1). rewrite C. ring.
omega.
Qed.
(** The [unsigned] and [signed] functions return the Coq integer corresponding
- to the given machine integer, interpreted as unsigned or signed
+ to the given machine integer, interpreted as unsigned or signed
respectively. *)
Definition unsigned (n: int) : Z := intval n.
@@ -198,7 +198,7 @@ Definition signed (n: int) : Z :=
(** Conversely, [repr] takes a Coq integer and returns the corresponding
machine integer. The argument is treated modulo [modulus]. *)
-Definition repr (x: Z) : int :=
+Definition repr (x: Z) : int :=
mkint (Z_mod_modulus x) (Z_mod_modulus_range' x).
Definition zero := repr 0.
@@ -212,12 +212,12 @@ Proof.
intros. subst y.
assert (forall (n m: Z) (P1 P2: n < m), P1 = P2).
{
- unfold Zlt; intros.
- apply eq_proofs_unicity.
+ unfold Zlt; intros.
+ apply eq_proofs_unicity.
intros c1 c2. destruct c1; destruct c2; (left; reflexivity) || (right; congruence).
}
destruct Px as [Px1 Px2]. destruct Py as [Py1 Py2].
- rewrite (H _ _ Px1 Py1).
+ rewrite (H _ _ Px1 Py1).
rewrite (H _ _ Px2 Py2).
reflexivity.
Qed.
@@ -231,7 +231,7 @@ Defined.
(** * Arithmetic and logical operations over machine integers *)
-Definition eq (x y: int) : bool :=
+Definition eq (x y: int) : bool :=
if zeq (unsigned x) (unsigned y) then true else false.
Definition lt (x y: int) : bool :=
if zlt (signed x) (signed y) then true else false.
@@ -340,7 +340,7 @@ Definition Zshiftin (b: bool) (x: Z) : Z :=
*)
Definition Zzero_ext (n: Z) (x: Z) : Z :=
- Z.iter n
+ Z.iter n
(fun rec x => Zshiftin (Z.odd x) (rec (Z.div2 x)))
(fun x => 0)
x.
@@ -410,8 +410,8 @@ Definition notbool (x: int) : int := if eq x zero then one else zero.
Remark half_modulus_power:
half_modulus = two_p (zwordsize - 1).
Proof.
- unfold half_modulus. rewrite modulus_power.
- set (ws1 := zwordsize - 1).
+ unfold half_modulus. rewrite modulus_power.
+ set (ws1 := zwordsize - 1).
replace (zwordsize) with (Zsucc ws1).
rewrite two_p_S. rewrite Zmult_comm. apply Z_div_mult. omega.
unfold ws1. generalize wordsize_pos; omega.
@@ -420,8 +420,8 @@ Qed.
Remark half_modulus_modulus: modulus = 2 * half_modulus.
Proof.
- rewrite half_modulus_power. rewrite modulus_power.
- rewrite <- two_p_S. apply f_equal. omega.
+ rewrite half_modulus_power. rewrite modulus_power.
+ rewrite <- two_p_S. apply f_equal. omega.
generalize wordsize_pos; omega.
Qed.
@@ -454,8 +454,8 @@ Qed.
Remark wordsize_max_unsigned: zwordsize <= max_unsigned.
Proof.
assert (zwordsize < modulus).
- rewrite modulus_power. apply two_p_strict.
- generalize wordsize_pos. omega.
+ rewrite modulus_power. apply two_p_strict.
+ generalize wordsize_pos. omega.
unfold max_unsigned. omega.
Qed.
@@ -468,14 +468,14 @@ Qed.
Remark max_signed_unsigned: max_signed < max_unsigned.
Proof.
- unfold max_signed, max_unsigned. rewrite half_modulus_modulus.
+ unfold max_signed, max_unsigned. rewrite half_modulus_modulus.
generalize half_modulus_pos. omega.
Qed.
Lemma unsigned_repr_eq:
forall x, unsigned (repr x) = Zmod x modulus.
Proof.
- intros. simpl. apply Z_mod_modulus_eq.
+ intros. simpl. apply Z_mod_modulus_eq.
Qed.
Lemma signed_repr_eq:
@@ -528,14 +528,14 @@ Qed.
Lemma eqmod_mod_eq:
forall x y, eqmod x y -> x mod modul = y mod modul.
Proof.
- intros x y [k EQ]. subst x.
+ intros x y [k EQ]. subst x.
rewrite Zplus_comm. apply Z_mod_plus. auto.
Qed.
Lemma eqmod_mod:
forall x, eqmod x (x mod modul).
Proof.
- intros; red. exists (x / modul).
+ intros; red. exists (x / modul).
rewrite Zmult_comm. apply Z_div_mod_eq. auto.
Qed.
@@ -549,7 +549,7 @@ Qed.
Lemma eqmod_neg:
forall x y, eqmod x y -> eqmod (-x) (-y).
Proof.
- intros x y [k EQ]; red. exists (-k). rewrite EQ. ring.
+ intros x y [k EQ]; red. exists (-k). rewrite EQ. ring.
Qed.
Lemma eqmod_sub:
@@ -573,11 +573,11 @@ End EQ_MODULO.
Lemma eqmod_divides:
forall n m x y, eqmod n x y -> Zdivide m n -> eqmod m x y.
Proof.
- intros. destruct H as [k1 EQ1]. destruct H0 as [k2 EQ2].
+ intros. destruct H as [k1 EQ1]. destruct H0 as [k2 EQ2].
exists (k1*k2). rewrite <- Zmult_assoc. rewrite <- EQ2. auto.
-Qed.
+Qed.
-(** We then specialize these definitions to equality modulo
+(** We then specialize these definitions to equality modulo
$2^{wordsize}$ #2<sup>wordsize</sup>#. *)
Hint Resolve modulus_pos: ints.
@@ -630,7 +630,7 @@ Hint Resolve eqm_mult: ints.
Lemma eqm_samerepr: forall x y, eqm x y -> repr x = repr y.
Proof.
- intros. unfold repr. apply mkint_eq.
+ intros. unfold repr. apply mkint_eq.
rewrite !Z_mod_modulus_eq. apply eqmod_mod_eq. auto with ints. exact H.
Qed.
@@ -644,7 +644,7 @@ Hint Resolve eqm_unsigned_repr: ints.
Lemma eqm_unsigned_repr_l:
forall a b, eqm a b -> eqm (unsigned (repr a)) b.
Proof.
- intros. apply eqm_trans with a.
+ intros. apply eqm_trans with a.
apply eqm_sym. apply eqm_unsigned_repr. auto.
Qed.
Hint Resolve eqm_unsigned_repr_l: ints.
@@ -653,7 +653,7 @@ Lemma eqm_unsigned_repr_r:
forall a b, eqm a b -> eqm a (unsigned (repr b)).
Proof.
intros. apply eqm_trans with b. auto.
- apply eqm_unsigned_repr.
+ apply eqm_unsigned_repr.
Qed.
Hint Resolve eqm_unsigned_repr_r: ints.
@@ -662,7 +662,7 @@ Lemma eqm_signed_unsigned:
Proof.
intros; red. unfold signed. set (y := unsigned x).
case (zlt y half_modulus); intro.
- apply eqmod_refl. red; exists (-1); ring.
+ apply eqmod_refl. red; exists (-1); ring.
Qed.
Theorem unsigned_range:
@@ -675,7 +675,7 @@ Hint Resolve unsigned_range: ints.
Theorem unsigned_range_2:
forall i, 0 <= unsigned i <= max_unsigned.
Proof.
- intro; unfold max_unsigned.
+ intro; unfold max_unsigned.
generalize (unsigned_range i). omega.
Qed.
Hint Resolve unsigned_range_2: ints.
@@ -683,13 +683,13 @@ Hint Resolve unsigned_range_2: ints.
Theorem signed_range:
forall i, min_signed <= signed i <= max_signed.
Proof.
- intros. unfold signed.
+ intros. unfold signed.
generalize (unsigned_range i). set (n := unsigned i). intros.
case (zlt n half_modulus); intro.
unfold max_signed. generalize min_signed_neg. omega.
unfold min_signed, max_signed.
- rewrite half_modulus_modulus in *. omega.
-Qed.
+ rewrite half_modulus_modulus in *. omega.
+Qed.
Theorem repr_unsigned:
forall i, repr (unsigned i) = i.
@@ -702,7 +702,7 @@ Hint Resolve repr_unsigned: ints.
Lemma repr_signed:
forall i, repr (signed i) = i.
Proof.
- intros. transitivity (repr (unsigned i)).
+ intros. transitivity (repr (unsigned i)).
apply eqm_samerepr. apply eqm_signed_unsigned. auto with ints.
Qed.
Hint Resolve repr_signed: ints.
@@ -717,7 +717,7 @@ Qed.
Theorem unsigned_repr:
forall z, 0 <= z <= max_unsigned -> unsigned (repr z) = z.
Proof.
- intros. rewrite unsigned_repr_eq.
+ intros. rewrite unsigned_repr_eq.
apply Zmod_small. unfold max_unsigned in H. omega.
Qed.
Hint Resolve unsigned_repr: ints.
@@ -728,16 +728,16 @@ Proof.
intros. unfold signed. destruct (zle 0 z).
replace (unsigned (repr z)) with z.
rewrite zlt_true. auto. unfold max_signed in H. omega.
- symmetry. apply unsigned_repr. generalize max_signed_unsigned. omega.
+ symmetry. apply unsigned_repr. generalize max_signed_unsigned. omega.
pose (z' := z + modulus).
replace (repr z) with (repr z').
replace (unsigned (repr z')) with z'.
rewrite zlt_false. unfold z'. omega.
unfold z'. unfold min_signed in H.
- rewrite half_modulus_modulus. omega.
+ rewrite half_modulus_modulus. omega.
symmetry. apply unsigned_repr.
unfold z', max_unsigned. unfold min_signed, max_signed in H.
- rewrite half_modulus_modulus. omega.
+ rewrite half_modulus_modulus. omega.
apply eqm_samerepr. unfold z'; red. exists 1. omega.
Qed.
@@ -765,16 +765,16 @@ Qed.
Theorem unsigned_one: unsigned one = 1.
Proof.
- unfold one; rewrite unsigned_repr_eq. apply Zmod_small. split. omega.
- unfold modulus. replace wordsize with (S(pred wordsize)).
- rewrite two_power_nat_S. generalize (two_power_nat_pos (pred wordsize)).
+ unfold one; rewrite unsigned_repr_eq. apply Zmod_small. split. omega.
+ unfold modulus. replace wordsize with (S(pred wordsize)).
+ rewrite two_power_nat_S. generalize (two_power_nat_pos (pred wordsize)).
omega.
- generalize wordsize_pos. unfold zwordsize. omega.
+ generalize wordsize_pos. unfold zwordsize. omega.
Qed.
Theorem unsigned_mone: unsigned mone = modulus - 1.
Proof.
- unfold mone; rewrite unsigned_repr_eq.
+ unfold mone; rewrite unsigned_repr_eq.
replace (-1) with ((modulus - 1) + (-1) * modulus).
rewrite Z_mod_plus_full. apply Zmod_small.
generalize modulus_pos. omega. omega.
@@ -789,12 +789,12 @@ Theorem signed_mone: signed mone = -1.
Proof.
unfold signed. rewrite unsigned_mone.
rewrite zlt_false. omega.
- rewrite half_modulus_modulus. generalize half_modulus_pos. omega.
+ rewrite half_modulus_modulus. generalize half_modulus_pos. omega.
Qed.
Theorem one_not_zero: one <> zero.
Proof.
- assert (unsigned one <> unsigned zero).
+ assert (unsigned one <> unsigned zero).
rewrite unsigned_one; rewrite unsigned_zero; congruence.
congruence.
Qed.
@@ -802,7 +802,7 @@ Qed.
Theorem unsigned_repr_wordsize:
unsigned iwordsize = zwordsize.
Proof.
- unfold iwordsize; rewrite unsigned_repr_eq. apply Zmod_small.
+ unfold iwordsize; rewrite unsigned_repr_eq. apply Zmod_small.
generalize wordsize_pos wordsize_max_unsigned; unfold max_unsigned; omega.
Qed.
@@ -820,7 +820,7 @@ Theorem eq_spec: forall (x y: int), if eq x y then x = y else x <> y.
Proof.
intros; unfold eq. case (eq_dec x y); intro.
subst y. rewrite zeq_true. auto.
- rewrite zeq_false. auto.
+ rewrite zeq_false. auto.
destruct x; destruct y.
simpl. red; intro. elim n. apply mkint_eq. auto.
Qed.
@@ -838,11 +838,11 @@ Qed.
Theorem eq_signed:
forall x y, eq x y = if zeq (signed x) (signed y) then true else false.
Proof.
- intros. predSpec eq eq_spec x y.
- subst x. rewrite zeq_true; auto.
+ intros. predSpec eq eq_spec x y.
+ subst x. rewrite zeq_true; auto.
destruct (zeq (signed x) (signed y)); auto.
elim H. rewrite <- (repr_signed x). rewrite <- (repr_signed y). congruence.
-Qed.
+Qed.
(** ** Properties of addition *)
@@ -851,7 +851,7 @@ Proof. intros; reflexivity.
Qed.
Theorem add_signed: forall x y, add x y = repr (signed x + signed y).
-Proof.
+Proof.
intros. rewrite add_unsigned. apply eqm_samerepr.
apply eqm_add; apply eqm_sym; apply eqm_signed_unsigned.
Qed.
@@ -876,7 +876,7 @@ Proof.
set (x' := unsigned x).
set (y' := unsigned y).
set (z' := unsigned z).
- apply eqm_samerepr.
+ apply eqm_samerepr.
apply eqm_trans with ((x' + y') + z').
auto with ints.
rewrite <- Zplus_assoc. auto with ints.
@@ -884,7 +884,7 @@ Qed.
Theorem add_permut: forall x y z, add x (add y z) = add y (add x z).
Proof.
- intros. rewrite (add_commut y z). rewrite <- add_assoc. apply add_commut.
+ intros. rewrite (add_commut y z). rewrite <- add_assoc. apply add_commut.
Qed.
Theorem add_neg_zero: forall x, add x (neg x) = zero.
@@ -901,19 +901,19 @@ Proof.
intros.
unfold add, add_carry. rewrite unsigned_zero. rewrite Zplus_0_r.
rewrite unsigned_repr_eq.
- generalize (unsigned_range x) (unsigned_range y). intros.
+ generalize (unsigned_range x) (unsigned_range y). intros.
destruct (zlt (unsigned x + unsigned y) modulus).
- rewrite unsigned_zero. apply Zmod_unique with 0. omega. omega.
- rewrite unsigned_one. apply Zmod_unique with 1. omega. omega.
-Qed.
+ rewrite unsigned_zero. apply Zmod_unique with 0. omega. omega.
+ rewrite unsigned_one. apply Zmod_unique with 1. omega. omega.
+Qed.
Corollary unsigned_add_either:
forall x y,
unsigned (add x y) = unsigned x + unsigned y
\/ unsigned (add x y) = unsigned x + unsigned y - modulus.
Proof.
- intros. rewrite unsigned_add_carry. unfold add_carry.
- rewrite unsigned_zero. rewrite Zplus_0_r.
+ intros. rewrite unsigned_add_carry. unfold add_carry.
+ rewrite unsigned_zero. rewrite Zplus_0_r.
destruct (zlt (unsigned x + unsigned y) modulus).
rewrite unsigned_zero. left; omega.
rewrite unsigned_one. right; omega.
@@ -928,7 +928,7 @@ Qed.
Theorem neg_zero: neg zero = zero.
Proof.
- unfold neg. rewrite unsigned_zero. auto.
+ unfold neg. rewrite unsigned_zero. auto.
Qed.
Theorem neg_involutive: forall x, neg (neg x) = x.
@@ -936,7 +936,7 @@ Proof.
intros; unfold neg.
apply eqm_repr_eq. eapply eqm_trans. apply eqm_neg.
apply eqm_unsigned_repr_l. apply eqm_refl. apply eqm_refl2. omega.
-Qed.
+Qed.
Theorem neg_add_distr: forall x y, neg(add x y) = add (neg x) (neg y).
Proof.
@@ -952,7 +952,7 @@ Qed.
Theorem sub_zero_l: forall x, sub x zero = x.
Proof.
- intros; unfold sub. rewrite unsigned_zero.
+ intros; unfold sub. rewrite unsigned_zero.
replace (unsigned x - 0) with (unsigned x) by omega. apply repr_unsigned.
Qed.
@@ -974,7 +974,7 @@ Qed.
Theorem sub_add_l: forall x y z, sub (add x y) z = add (sub x z) y.
Proof.
- intros. repeat rewrite sub_add_opp.
+ intros. repeat rewrite sub_add_opp.
repeat rewrite add_assoc. decEq. apply add_commut.
Qed.
@@ -989,7 +989,7 @@ Theorem sub_shifted:
sub (add x z) (add y z) = sub x y.
Proof.
intros. rewrite sub_add_opp. rewrite neg_add_distr.
- rewrite add_assoc.
+ rewrite add_assoc.
rewrite (add_commut (neg y) (neg z)).
rewrite <- (add_assoc z). rewrite add_neg_zero.
rewrite (add_commut zero). rewrite add_zero.
@@ -1010,22 +1010,22 @@ Proof.
intros.
unfold sub, sub_borrow. rewrite unsigned_zero. rewrite Zminus_0_r.
rewrite unsigned_repr_eq.
- generalize (unsigned_range x) (unsigned_range y). intros.
+ generalize (unsigned_range x) (unsigned_range y). intros.
destruct (zlt (unsigned x - unsigned y) 0).
- rewrite unsigned_one. apply Zmod_unique with (-1). omega. omega.
- rewrite unsigned_zero. apply Zmod_unique with 0. omega. omega.
-Qed.
+ rewrite unsigned_one. apply Zmod_unique with (-1). omega. omega.
+ rewrite unsigned_zero. apply Zmod_unique with 0. omega. omega.
+Qed.
(** ** Properties of multiplication *)
Theorem mul_commut: forall x y, mul x y = mul y x.
Proof.
- intros; unfold mul. decEq. ring.
+ intros; unfold mul. decEq. ring.
Qed.
Theorem mul_zero: forall x, mul x zero = zero.
Proof.
- intros; unfold mul. rewrite unsigned_zero.
+ intros; unfold mul. rewrite unsigned_zero.
unfold zero. decEq. ring.
Qed.
@@ -1038,7 +1038,7 @@ Qed.
Theorem mul_mone: forall x, mul x mone = neg x.
Proof.
- intros; unfold mul, neg. rewrite unsigned_mone.
+ intros; unfold mul, neg. rewrite unsigned_mone.
apply eqm_samerepr.
replace (-unsigned x) with (0 - unsigned x) by omega.
replace (unsigned x * (modulus - 1)) with (unsigned x * modulus - unsigned x) by ring.
@@ -1074,11 +1074,11 @@ Qed.
Theorem mul_add_distr_r:
forall x y z, mul x (add y z) = add (mul x y) (mul x z).
Proof.
- intros. rewrite mul_commut. rewrite mul_add_distr_l.
+ intros. rewrite mul_commut. rewrite mul_add_distr_l.
decEq; apply mul_commut.
-Qed.
+Qed.
-Theorem neg_mul_distr_l:
+Theorem neg_mul_distr_l:
forall x y, neg(mul x y) = mul (neg x) y.
Proof.
intros. unfold mul, neg.
@@ -1093,7 +1093,7 @@ Theorem neg_mul_distr_r:
forall x y, neg(mul x y) = mul x (neg y).
Proof.
intros. rewrite (mul_commut x y). rewrite (mul_commut x (neg y)).
- apply neg_mul_distr_l.
+ apply neg_mul_distr_l.
Qed.
Theorem mul_signed:
@@ -1109,13 +1109,13 @@ Lemma modu_divu_Euclid:
forall x y, y <> zero -> x = add (mul (divu x y) y) (modu x y).
Proof.
intros. unfold add, mul, divu, modu.
- transitivity (repr (unsigned x)). auto with ints.
- apply eqm_samerepr.
+ transitivity (repr (unsigned x)). auto with ints.
+ apply eqm_samerepr.
set (x' := unsigned x). set (y' := unsigned y).
apply eqm_trans with ((x' / y') * y' + x' mod y').
apply eqm_refl2. rewrite Zmult_comm. apply Z_div_mod_eq.
generalize (unsigned_range y); intro.
- assert (unsigned y <> 0). red; intro.
+ assert (unsigned y <> 0). red; intro.
elim H. rewrite <- (repr_unsigned y). unfold zero. congruence.
unfold y'. omega.
auto with ints.
@@ -1124,7 +1124,7 @@ Qed.
Theorem modu_divu:
forall x y, y <> zero -> modu x y = sub x (mul (divu x y) y).
Proof.
- intros.
+ intros.
assert (forall a b c, a = add b c -> c = sub a b).
intros. subst a. rewrite sub_add_l. rewrite sub_idem.
rewrite add_commut. rewrite add_zero. auto.
@@ -1135,20 +1135,20 @@ Lemma mods_divs_Euclid:
forall x y, x = add (mul (divs x y) y) (mods x y).
Proof.
intros. unfold add, mul, divs, mods.
- transitivity (repr (signed x)). auto with ints.
- apply eqm_samerepr.
+ transitivity (repr (signed x)). auto with ints.
+ apply eqm_samerepr.
set (x' := signed x). set (y' := signed y).
apply eqm_trans with ((Z.quot x' y') * y' + Z.rem x' y').
apply eqm_refl2. rewrite Zmult_comm. apply Z.quot_rem'.
apply eqm_add; auto with ints.
- apply eqm_unsigned_repr_r. apply eqm_mult; auto with ints.
- unfold y'. apply eqm_signed_unsigned.
+ apply eqm_unsigned_repr_r. apply eqm_mult; auto with ints.
+ unfold y'. apply eqm_signed_unsigned.
Qed.
Theorem mods_divs:
forall x y, mods x y = sub x (mul (divs x y) y).
Proof.
- intros.
+ intros.
assert (forall a b c, a = add b c -> c = sub a b).
intros. subst a. rewrite sub_add_l. rewrite sub_idem.
rewrite add_commut. rewrite add_zero. auto.
@@ -1171,26 +1171,26 @@ Qed.
Theorem divs_mone:
forall x, divs x mone = neg x.
Proof.
- unfold divs, neg; intros.
- rewrite signed_mone.
+ unfold divs, neg; intros.
+ rewrite signed_mone.
replace (Z.quot (signed x) (-1)) with (- (signed x)).
- apply eqm_samerepr. apply eqm_neg. apply eqm_signed_unsigned.
+ apply eqm_samerepr. apply eqm_neg. apply eqm_signed_unsigned.
set (x' := signed x).
set (one := 1).
change (-1) with (- one). rewrite Zquot_opp_r.
- assert (Z.quot x' one = x').
+ assert (Z.quot x' one = x').
symmetry. apply Zquot_unique_full with 0. red.
- change (Z.abs one) with 1.
- destruct (zle 0 x'). left. omega. right. omega.
- unfold one; ring.
+ change (Z.abs one) with 1.
+ destruct (zle 0 x'). left. omega. right. omega.
+ unfold one; ring.
congruence.
Qed.
Theorem mods_mone:
forall x, mods x mone = zero.
Proof.
- intros. rewrite mods_divs. rewrite divs_mone.
- rewrite <- neg_mul_distr_l. rewrite mul_mone. rewrite neg_involutive. apply sub_idem.
+ intros. rewrite mods_divs. rewrite divs_mone.
+ rewrite <- neg_mul_distr_l. rewrite mul_mone. rewrite neg_involutive. apply sub_idem.
Qed.
(** ** Bit-level properties *)
@@ -1207,14 +1207,14 @@ Qed.
Remark Ztestbit_m1: forall n, 0 <= n -> Z.testbit (-1) n = true.
Proof.
- intros. destruct n; simpl; auto.
+ intros. destruct n; simpl; auto.
Qed.
Remark Zshiftin_spec:
forall b x, Zshiftin b x = 2 * x + (if b then 1 else 0).
Proof.
unfold Zshiftin; intros. destruct b.
- - rewrite Z.succ_double_spec. omega.
+ - rewrite Z.succ_double_spec. omega.
- rewrite Z.double_spec. omega.
Qed.
@@ -1222,7 +1222,7 @@ Remark Zshiftin_inj:
forall b1 x1 b2 x2,
Zshiftin b1 x1 = Zshiftin b2 x2 -> b1 = b2 /\ x1 = x2.
Proof.
- intros. rewrite !Zshiftin_spec in H.
+ intros. rewrite !Zshiftin_spec in H.
destruct b1; destruct b2.
split; [auto|omega].
omegaContradiction.
@@ -1235,7 +1235,7 @@ Remark Zdecomp:
Proof.
intros. destruct x; simpl.
- auto.
- - destruct p; auto.
+ - destruct p; auto.
- destruct p; auto. simpl. rewrite Pos.pred_double_succ. auto.
Qed.
@@ -1265,7 +1265,7 @@ Qed.
Remark Ztestbit_shiftin_succ:
forall b x n, 0 <= n -> Z.testbit (Zshiftin b x) (Z.succ n) = Z.testbit x n.
Proof.
- intros. rewrite Ztestbit_shiftin. rewrite zeq_false. rewrite Z.pred_succ. auto.
+ intros. rewrite Ztestbit_shiftin. rewrite zeq_false. rewrite Z.pred_succ. auto.
omega. omega.
Qed.
@@ -1273,19 +1273,19 @@ Remark Ztestbit_eq:
forall n x, 0 <= n ->
Z.testbit x n = if zeq n 0 then Z.odd x else Z.testbit (Z.div2 x) (Z.pred n).
Proof.
- intros. rewrite (Zdecomp x) at 1. apply Ztestbit_shiftin; auto.
+ intros. rewrite (Zdecomp x) at 1. apply Ztestbit_shiftin; auto.
Qed.
Remark Ztestbit_base:
forall x, Z.testbit x 0 = Z.odd x.
Proof.
- intros. rewrite Ztestbit_eq. apply zeq_true. omega.
+ intros. rewrite Ztestbit_eq. apply zeq_true. omega.
Qed.
Remark Ztestbit_succ:
forall n x, 0 <= n -> Z.testbit x (Z.succ n) = Z.testbit (Z.div2 x) n.
Proof.
- intros. rewrite Ztestbit_eq. rewrite zeq_false. rewrite Z.pred_succ. auto.
+ intros. rewrite Ztestbit_eq. rewrite zeq_false. rewrite Z.pred_succ. auto.
omega. omega.
Qed.
@@ -1296,14 +1296,14 @@ Lemma eqmod_same_bits:
Proof.
induction n; intros.
- change (two_power_nat 0) with 1. exists (x-y); ring.
- - rewrite two_power_nat_S.
+ - rewrite two_power_nat_S.
assert (eqmod (two_power_nat n) (Z.div2 x) (Z.div2 y)).
- apply IHn. intros. rewrite <- !Ztestbit_succ. apply H. rewrite inj_S; omega.
- omega. omega.
+ apply IHn. intros. rewrite <- !Ztestbit_succ. apply H. rewrite inj_S; omega.
+ omega. omega.
destruct H0 as [k EQ].
- exists k. rewrite (Zdecomp x). rewrite (Zdecomp y).
+ exists k. rewrite (Zdecomp x). rewrite (Zdecomp y).
replace (Z.odd y) with (Z.odd x).
- rewrite EQ. rewrite !Zshiftin_spec. ring.
+ rewrite EQ. rewrite !Zshiftin_spec. ring.
exploit (H 0). rewrite inj_S; omega.
rewrite !Ztestbit_base. auto.
Qed.
@@ -1321,14 +1321,14 @@ Lemma same_bits_eqmod:
Proof.
induction n; intros.
- simpl in H0. omegaContradiction.
- - rewrite inj_S in H0. rewrite two_power_nat_S in H.
- rewrite !(Ztestbit_eq i); intuition.
+ - rewrite inj_S in H0. rewrite two_power_nat_S in H.
+ rewrite !(Ztestbit_eq i); intuition.
destruct H as [k EQ].
assert (EQ': Zshiftin (Z.odd x) (Z.div2 x) =
Zshiftin (Z.odd y) (k * two_power_nat n + Z.div2 y)).
{
rewrite (Zdecomp x) in EQ. rewrite (Zdecomp y) in EQ.
- rewrite EQ. rewrite !Zshiftin_spec. ring.
+ rewrite EQ. rewrite !Zshiftin_spec. ring.
}
exploit Zshiftin_inj; eauto. intros [A B].
destruct (zeq i 0).
@@ -1348,8 +1348,8 @@ Remark two_power_nat_infinity:
Proof.
intros x0 POS0; pattern x0; apply natlike_ind; auto.
exists O. compute; auto.
- intros. destruct H0 as [n LT]. exists (S n). rewrite two_power_nat_S.
- generalize (two_power_nat_pos n). omega.
+ intros. destruct H0 as [n LT]. exists (S n). rewrite two_power_nat_S.
+ generalize (two_power_nat_pos n). omega.
Qed.
Lemma equal_same_bits:
@@ -1357,15 +1357,15 @@ Lemma equal_same_bits:
(forall i, 0 <= i -> Z.testbit x i = Z.testbit y i) ->
x = y.
Proof.
- intros.
+ intros.
set (z := if zlt x y then y - x else x - y).
assert (0 <= z).
unfold z; destruct (zlt x y); omega.
- exploit (two_power_nat_infinity z); auto. intros [n LT].
+ exploit (two_power_nat_infinity z); auto. intros [n LT].
assert (eqmod (two_power_nat n) x y).
- apply eqmod_same_bits. intros. apply H. tauto.
+ apply eqmod_same_bits. intros. apply H. tauto.
assert (eqmod (two_power_nat n) z 0).
- unfold z. destruct (zlt x y).
+ unfold z. destruct (zlt x y).
replace 0 with (y - y) by omega. apply eqmod_sub. apply eqmod_refl. auto.
replace 0 with (x - x) by omega. apply eqmod_sub. apply eqmod_refl. apply eqmod_sym; auto.
assert (z = 0).
@@ -1377,13 +1377,13 @@ Lemma Z_one_complement:
forall i, 0 <= i ->
forall x, Z.testbit (-x-1) i = negb (Z.testbit x i).
Proof.
- intros i0 POS0. pattern i0. apply Zlt_0_ind; auto.
- intros i IND POS x.
+ intros i0 POS0. pattern i0. apply Zlt_0_ind; auto.
+ intros i IND POS x.
rewrite (Zdecomp x). set (y := Z.div2 x).
replace (- Zshiftin (Z.odd x) y - 1)
with (Zshiftin (negb (Z.odd x)) (- y - 1)).
- rewrite !Ztestbit_shiftin; auto.
- destruct (zeq i 0). auto. apply IND. omega.
+ rewrite !Ztestbit_shiftin; auto.
+ destruct (zeq i 0). auto. apply IND. omega.
rewrite !Zshiftin_spec. destruct (Z.odd x); simpl negb; ring.
Qed.
@@ -1393,14 +1393,14 @@ Lemma Ztestbit_above:
i >= Z.of_nat n ->
Z.testbit x i = false.
Proof.
- induction n; intros.
- - change (two_power_nat 0) with 1 in H.
- replace x with 0 by omega.
+ induction n; intros.
+ - change (two_power_nat 0) with 1 in H.
+ replace x with 0 by omega.
apply Z.testbit_0_l.
- rewrite inj_S in H0. rewrite Ztestbit_eq. rewrite zeq_false.
- apply IHn. rewrite two_power_nat_S in H. rewrite (Zdecomp x) in H.
- rewrite Zshiftin_spec in H. destruct (Z.odd x); omega.
- omega. omega. omega.
+ apply IHn. rewrite two_power_nat_S in H. rewrite (Zdecomp x) in H.
+ rewrite Zshiftin_spec in H. destruct (Z.odd x); omega.
+ omega. omega. omega.
Qed.
Lemma Ztestbit_above_neg:
@@ -1410,11 +1410,11 @@ Lemma Ztestbit_above_neg:
Z.testbit x i = true.
Proof.
intros. set (y := -x-1).
- assert (Z.testbit y i = false).
- apply Ztestbit_above with n.
+ assert (Z.testbit y i = false).
+ apply Ztestbit_above with n.
unfold y; omega. auto.
unfold y in H1. rewrite Z_one_complement in H1.
- change true with (negb false). rewrite <- H1. rewrite negb_involutive; auto.
+ change true with (negb false). rewrite <- H1. rewrite negb_involutive; auto.
omega.
Qed.
@@ -1423,17 +1423,17 @@ Lemma Zsign_bit:
0 <= x < two_power_nat (S n) ->
Z.testbit x (Z_of_nat n) = if zlt x (two_power_nat n) then false else true.
Proof.
- induction n; intros.
- - change (two_power_nat 1) with 2 in H.
- assert (x = 0 \/ x = 1) by omega.
+ induction n; intros.
+ - change (two_power_nat 1) with 2 in H.
+ assert (x = 0 \/ x = 1) by omega.
destruct H0; subst x; reflexivity.
- - rewrite inj_S. rewrite Ztestbit_eq. rewrite zeq_false. rewrite Z.pred_succ.
- rewrite IHn. rewrite two_power_nat_S.
+ - rewrite inj_S. rewrite Ztestbit_eq. rewrite zeq_false. rewrite Z.pred_succ.
+ rewrite IHn. rewrite two_power_nat_S.
destruct (zlt (Z.div2 x) (two_power_nat n)); rewrite (Zdecomp x); rewrite Zshiftin_spec.
- rewrite zlt_true. auto. destruct (Z.odd x); omega.
- rewrite zlt_false. auto. destruct (Z.odd x); omega.
- rewrite (Zdecomp x) in H; rewrite Zshiftin_spec in H.
- rewrite two_power_nat_S in H. destruct (Z.odd x); omega.
+ rewrite zlt_true. auto. destruct (Z.odd x); omega.
+ rewrite zlt_false. auto. destruct (Z.odd x); omega.
+ rewrite (Zdecomp x) in H; rewrite Zshiftin_spec in H.
+ rewrite two_power_nat_S in H. destruct (Z.odd x); omega.
omega. omega.
Qed.
@@ -1443,7 +1443,7 @@ Lemma Zshiftin_ind:
(forall b x, 0 <= x -> P x -> P (Zshiftin b x)) ->
forall x, 0 <= x -> P x.
Proof.
- intros. destruct x.
+ intros. destruct x.
- auto.
- induction p.
+ change (P (Zshiftin true (Z.pos p))). auto.
@@ -1472,16 +1472,16 @@ Lemma Ztestbit_le:
x <= y.
Proof.
intros x y0 POS0; revert x; pattern y0; apply Zshiftin_ind; auto; intros.
- - replace x with 0. omega. apply equal_same_bits; intros.
- rewrite Ztestbit_0. destruct (Z.testbit x i) as [] eqn:E; auto.
+ - replace x with 0. omega. apply equal_same_bits; intros.
+ rewrite Ztestbit_0. destruct (Z.testbit x i) as [] eqn:E; auto.
exploit H; eauto. rewrite Ztestbit_0. auto.
- assert (Z.div2 x0 <= x).
{ apply H0. intros. exploit (H1 (Zsucc i)).
- omega. rewrite Ztestbit_succ; auto. rewrite Ztestbit_shiftin_succ; auto.
+ omega. rewrite Ztestbit_succ; auto. rewrite Ztestbit_shiftin_succ; auto.
}
- rewrite (Zdecomp x0). rewrite !Zshiftin_spec.
+ rewrite (Zdecomp x0). rewrite !Zshiftin_spec.
destruct (Z.odd x0) as [] eqn:E1; destruct b as [] eqn:E2; try omega.
- exploit (H1 0). omega. rewrite Ztestbit_base; auto.
+ exploit (H1 0). omega. rewrite Ztestbit_base; auto.
rewrite Ztestbit_shiftin_base. congruence.
Qed.
@@ -1502,25 +1502,25 @@ Lemma same_bits_eq:
(forall i, 0 <= i < zwordsize -> testbit x i = testbit y i) ->
x = y.
Proof.
- intros. rewrite <- (repr_unsigned x). rewrite <- (repr_unsigned y).
+ intros. rewrite <- (repr_unsigned x). rewrite <- (repr_unsigned y).
apply eqm_samerepr. apply eqm_same_bits. auto.
Qed.
Lemma bits_above:
forall x i, i >= zwordsize -> testbit x i = false.
Proof.
- intros. apply Ztestbit_above with wordsize; auto. apply unsigned_range.
-Qed.
+ intros. apply Ztestbit_above with wordsize; auto. apply unsigned_range.
+Qed.
Lemma bits_zero:
forall i, testbit zero i = false.
Proof.
- intros. unfold testbit. rewrite unsigned_zero. apply Ztestbit_0.
+ intros. unfold testbit. rewrite unsigned_zero. apply Ztestbit_0.
Qed.
Remark bits_one: forall n, testbit one n = zeq n 0.
Proof.
- unfold testbit; intros. rewrite unsigned_one. apply Ztestbit_1.
+ unfold testbit; intros. rewrite unsigned_one. apply Ztestbit_1.
Qed.
Lemma bits_mone:
@@ -1539,25 +1539,25 @@ Lemma sign_bit_of_unsigned:
Proof.
intros. unfold testbit.
set (ws1 := pred wordsize).
- assert (zwordsize - 1 = Z_of_nat ws1).
- unfold zwordsize, ws1, wordsize.
+ assert (zwordsize - 1 = Z_of_nat ws1).
+ unfold zwordsize, ws1, wordsize.
destruct WS.wordsize as [] eqn:E.
elim WS.wordsize_not_zero; auto.
- rewrite inj_S. simpl. omega.
+ rewrite inj_S. simpl. omega.
assert (half_modulus = two_power_nat ws1).
rewrite two_power_nat_two_p. rewrite <- H. apply half_modulus_power.
rewrite H; rewrite H0.
- apply Zsign_bit. rewrite two_power_nat_S. rewrite <- H0.
+ apply Zsign_bit. rewrite two_power_nat_S. rewrite <- H0.
rewrite <- half_modulus_modulus. apply unsigned_range.
Qed.
-
+
Lemma bits_signed:
forall x i, 0 <= i ->
Z.testbit (signed x) i = testbit x (if zlt i zwordsize then i else zwordsize - 1).
Proof.
intros.
destruct (zlt i zwordsize).
- - apply same_bits_eqm. apply eqm_signed_unsigned. omega.
+ - apply same_bits_eqm. apply eqm_signed_unsigned. omega.
- unfold signed. rewrite sign_bit_of_unsigned. destruct (zlt (unsigned x) half_modulus).
+ apply Ztestbit_above with wordsize. apply unsigned_range. auto.
+ apply Ztestbit_above_neg with wordsize.
@@ -1569,11 +1569,11 @@ Lemma bits_le:
(forall i, 0 <= i < zwordsize -> testbit x i = true -> testbit y i = true) ->
unsigned x <= unsigned y.
Proof.
- intros. apply Ztestbit_le. generalize (unsigned_range y); omega.
- intros. fold (testbit y i). destruct (zlt i zwordsize).
- apply H. omega. auto.
+ intros. apply Ztestbit_le. generalize (unsigned_range y); omega.
+ intros. fold (testbit y i). destruct (zlt i zwordsize).
+ apply H. omega. auto.
fold (testbit x i) in H1. rewrite bits_above in H1; auto. congruence.
-Qed.
+Qed.
(** ** Properties of bitwise and, or, xor *)
@@ -1654,7 +1654,7 @@ Qed.
Theorem or_zero: forall x, or x zero = x.
Proof.
- bit_solve.
+ bit_solve.
Qed.
Corollary or_zero_l: forall x, or zero x = x.
@@ -1664,7 +1664,7 @@ Qed.
Theorem or_mone: forall x, or x mone = mone.
Proof.
- bit_solve.
+ bit_solve.
Qed.
Theorem or_idem: forall x, or x x = x.
@@ -1677,7 +1677,7 @@ Theorem and_or_distrib:
and x (or y z) = or (and x y) (and x z).
Proof.
bit_solve. apply demorgan1.
-Qed.
+Qed.
Corollary and_or_distrib_l:
forall x y z,
@@ -1690,8 +1690,8 @@ Theorem or_and_distrib:
forall x y z,
or x (and y z) = and (or x y) (or x z).
Proof.
- bit_solve. apply orb_andb_distrib_r.
-Qed.
+ bit_solve. apply orb_andb_distrib_r.
+Qed.
Corollary or_and_distrib_l:
forall x y z,
@@ -1702,7 +1702,7 @@ Qed.
Theorem and_or_absorb: forall x y, and x (or x y) = x.
Proof.
- bit_solve.
+ bit_solve.
assert (forall a b, a && (a || b) = a) by destr_bool.
auto.
Qed.
@@ -1716,7 +1716,7 @@ Qed.
Theorem xor_commut: forall x y, xor x y = xor y x.
Proof.
- bit_solve. apply xorb_comm.
+ bit_solve. apply xorb_comm.
Qed.
Theorem xor_assoc: forall x y z, xor (xor x y) z = xor x (xor y z).
@@ -1726,7 +1726,7 @@ Qed.
Theorem xor_zero: forall x, xor x zero = x.
Proof.
- bit_solve. apply xorb_false.
+ bit_solve. apply xorb_false.
Qed.
Corollary xor_zero_l: forall x, xor zero x = x.
@@ -1736,7 +1736,7 @@ Qed.
Theorem xor_idem: forall x, xor x x = zero.
Proof.
- bit_solve. apply xorb_nilpotent.
+ bit_solve. apply xorb_nilpotent.
Qed.
Theorem xor_zero_one: xor zero one = one.
@@ -1747,7 +1747,7 @@ Proof. apply xor_idem. Qed.
Theorem xor_zero_equal: forall x y, xor x y = zero -> x = y.
Proof.
- intros. apply same_bits_eq; intros.
+ intros. apply same_bits_eq; intros.
assert (xorb (testbit x i) (testbit y i) = false).
rewrite <- bits_xor; auto. rewrite H. apply bits_zero.
destruct (testbit x i); destruct (testbit y i); reflexivity || discriminate.
@@ -1757,23 +1757,23 @@ Theorem and_xor_distrib:
forall x y z,
and x (xor y z) = xor (and x y) (and x z).
Proof.
- bit_solve.
+ bit_solve.
assert (forall a b c, a && (xorb b c) = xorb (a && b) (a && c)) by destr_bool.
auto.
-Qed.
+Qed.
Theorem and_le:
forall x y, unsigned (and x y) <= unsigned x.
Proof.
- intros. apply bits_le; intros.
+ intros. apply bits_le; intros.
rewrite bits_and in H0; auto. rewrite andb_true_iff in H0. tauto.
Qed.
Theorem or_le:
forall x y, unsigned x <= unsigned (or x y).
Proof.
- intros. apply bits_le; intros.
- rewrite bits_or; auto. rewrite H0; auto.
+ intros. apply bits_le; intros.
+ rewrite bits_or; auto. rewrite H0; auto.
Qed.
(** Properties of bitwise complement.*)
@@ -1781,7 +1781,7 @@ Qed.
Theorem not_involutive:
forall (x: int), not (not x) = x.
Proof.
- intros. unfold not. rewrite xor_assoc. rewrite xor_idem. apply xor_zero.
+ intros. unfold not. rewrite xor_assoc. rewrite xor_idem. apply xor_zero.
Qed.
Theorem not_zero:
@@ -1799,31 +1799,31 @@ Qed.
Theorem not_or_and_not:
forall x y, not (or x y) = and (not x) (not y).
Proof.
- bit_solve. apply negb_orb.
+ bit_solve. apply negb_orb.
Qed.
Theorem not_and_or_not:
forall x y, not (and x y) = or (not x) (not y).
Proof.
- bit_solve. apply negb_andb.
+ bit_solve. apply negb_andb.
Qed.
Theorem and_not_self:
forall x, and x (not x) = zero.
Proof.
- bit_solve.
+ bit_solve.
Qed.
Theorem or_not_self:
forall x, or x (not x) = mone.
Proof.
- bit_solve.
+ bit_solve.
Qed.
Theorem xor_not_self:
forall x, xor x (not x) = mone.
Proof.
- bit_solve. destruct (testbit x i); auto.
+ bit_solve. destruct (testbit x i); auto.
Qed.
Lemma unsigned_not:
@@ -1832,7 +1832,7 @@ Proof.
intros. transitivity (unsigned (repr(-unsigned x - 1))).
f_equal. bit_solve. rewrite testbit_repr; auto. symmetry. apply Z_one_complement. omega.
rewrite unsigned_repr_eq. apply Zmod_unique with (-1).
- unfold max_unsigned. omega.
+ unfold max_unsigned. omega.
generalize (unsigned_range x). unfold max_unsigned. omega.
Qed.
@@ -1840,30 +1840,30 @@ Theorem not_neg:
forall x, not x = add (neg x) mone.
Proof.
bit_solve.
- rewrite <- (repr_unsigned x) at 1. unfold add.
+ rewrite <- (repr_unsigned x) at 1. unfold add.
rewrite !testbit_repr; auto.
transitivity (Z.testbit (-unsigned x - 1) i).
symmetry. apply Z_one_complement. omega.
apply same_bits_eqm; auto.
replace (-unsigned x - 1) with (-unsigned x + (-1)) by omega.
- apply eqm_add.
- unfold neg. apply eqm_unsigned_repr.
- rewrite unsigned_mone. exists (-1). ring.
+ apply eqm_add.
+ unfold neg. apply eqm_unsigned_repr.
+ rewrite unsigned_mone. exists (-1). ring.
Qed.
Theorem neg_not:
forall x, neg x = add (not x) one.
Proof.
- intros. rewrite not_neg. rewrite add_assoc.
- replace (add mone one) with zero. rewrite add_zero. auto.
- apply eqm_samerepr. rewrite unsigned_mone. rewrite unsigned_one.
- exists (-1). ring.
+ intros. rewrite not_neg. rewrite add_assoc.
+ replace (add mone one) with zero. rewrite add_zero. auto.
+ apply eqm_samerepr. rewrite unsigned_mone. rewrite unsigned_one.
+ exists (-1). ring.
Qed.
Theorem sub_add_not:
forall x y, sub x y = add (add x (not y)) one.
Proof.
- intros. rewrite sub_add_opp. rewrite neg_not.
+ intros. rewrite sub_add_opp. rewrite neg_not.
rewrite ! add_assoc. auto.
Qed.
@@ -1883,13 +1883,13 @@ Theorem sub_borrow_add_carry:
b = zero \/ b = one ->
sub_borrow x y b = xor (add_carry x (not y) (xor b one)) one.
Proof.
- intros. unfold sub_borrow, add_carry. rewrite unsigned_not.
+ intros. unfold sub_borrow, add_carry. rewrite unsigned_not.
replace (unsigned (xor b one)) with (1 - unsigned b).
destruct (zlt (unsigned x - unsigned y - unsigned b)).
rewrite zlt_true. rewrite xor_zero_l; auto.
unfold max_unsigned; omega.
rewrite zlt_false. rewrite xor_idem; auto.
- unfold max_unsigned; omega.
+ unfold max_unsigned; omega.
destruct H; subst b.
rewrite xor_zero_l. rewrite unsigned_one, unsigned_zero; auto.
rewrite xor_idem. rewrite unsigned_one, unsigned_zero; auto.
@@ -1908,14 +1908,14 @@ Proof.
rewrite (Zdecomp x) in *. rewrite (Zdecomp y) in *.
transitivity (Z.testbit (Zshiftin (Z.odd x || Z.odd y) (Z.div2 x + Z.div2 y)) i).
- f_equal. rewrite !Zshiftin_spec.
- exploit (EXCL 0). omega. rewrite !Ztestbit_shiftin_base. intros.
+ exploit (EXCL 0). omega. rewrite !Ztestbit_shiftin_base. intros.
Opaque Z.mul.
destruct (Z.odd x); destruct (Z.odd y); simpl in *; discriminate || ring.
- rewrite !Ztestbit_shiftin; auto.
destruct (zeq i 0).
+ auto.
- + apply IND. omega. intros.
- exploit (EXCL (Z.succ j)). omega.
+ + apply IND. omega. intros.
+ exploit (EXCL (Z.succ j)). omega.
rewrite !Ztestbit_shiftin_succ. auto.
omega. omega.
Qed.
@@ -1926,8 +1926,8 @@ Theorem add_is_or:
add x y = or x y.
Proof.
bit_solve. unfold add. rewrite testbit_repr; auto.
- apply Z_add_is_or. omega.
- intros.
+ apply Z_add_is_or. omega.
+ intros.
assert (testbit (and x y) j = testbit zero j) by congruence.
autorewrite with ints in H2. assumption. omega.
Qed.
@@ -1935,9 +1935,9 @@ Qed.
Theorem xor_is_or:
forall x y, and x y = zero -> xor x y = or x y.
Proof.
- bit_solve.
+ bit_solve.
assert (testbit (and x y) i = testbit zero i) by congruence.
- autorewrite with ints in H1; auto.
+ autorewrite with ints in H1; auto.
destruct (testbit x i); destruct (testbit y i); simpl in *; congruence.
Qed.
@@ -1957,8 +1957,8 @@ Proof.
intros. rewrite add_is_or.
rewrite and_or_distrib; auto.
rewrite (and_commut x y).
- rewrite and_assoc.
- repeat rewrite <- (and_assoc x).
+ rewrite and_assoc.
+ repeat rewrite <- (and_assoc x).
rewrite (and_commut (and x x)).
rewrite <- and_assoc.
rewrite H. rewrite and_commut. apply and_zero.
@@ -1972,8 +1972,8 @@ Lemma bits_shl:
testbit (shl x y) i =
if zlt i (unsigned y) then false else testbit x (i - unsigned y).
Proof.
- intros. unfold shl. rewrite testbit_repr; auto.
- destruct (zlt i (unsigned y)).
+ intros. unfold shl. rewrite testbit_repr; auto.
+ destruct (zlt i (unsigned y)).
apply Z.shiftl_spec_low. auto.
apply Z.shiftl_spec_high. omega. omega.
Qed.
@@ -1984,11 +1984,11 @@ Lemma bits_shru:
testbit (shru x y) i =
if zlt (i + unsigned y) zwordsize then testbit x (i + unsigned y) else false.
Proof.
- intros. unfold shru. rewrite testbit_repr; auto.
+ intros. unfold shru. rewrite testbit_repr; auto.
rewrite Z.shiftr_spec. fold (testbit x (i + unsigned y)).
destruct (zlt (i + unsigned y) zwordsize).
auto.
- apply bits_above; auto.
+ apply bits_above; auto.
omega.
Qed.
@@ -1998,8 +1998,8 @@ Lemma bits_shr:
testbit (shr x y) i =
testbit x (if zlt (i + unsigned y) zwordsize then i + unsigned y else zwordsize - 1).
Proof.
- intros. unfold shr. rewrite testbit_repr; auto.
- rewrite Z.shiftr_spec. apply bits_signed.
+ intros. unfold shr. rewrite testbit_repr; auto.
+ rewrite Z.shiftr_spec. apply bits_signed.
generalize (unsigned_range y); omega.
omega.
Qed.
@@ -2017,10 +2017,10 @@ Lemma bitwise_binop_shl:
f' false false = false ->
f (shl x n) (shl y n) = shl (f x y) n.
Proof.
- intros. apply same_bits_eq; intros.
+ intros. apply same_bits_eq; intros.
rewrite H; auto. rewrite !bits_shl; auto.
destruct (zlt i (unsigned n)); auto.
- rewrite H; auto. generalize (unsigned_range n); omega.
+ rewrite H; auto. generalize (unsigned_range n); omega.
Qed.
Theorem and_shl:
@@ -2060,22 +2060,22 @@ Qed.
Theorem shl_shl:
forall x y z,
- ltu y iwordsize = true ->
+ ltu y iwordsize = true ->
ltu z iwordsize = true ->
ltu (add y z) iwordsize = true ->
shl (shl x y) z = shl x (add y z).
Proof.
- intros.
+ intros.
generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros.
assert (unsigned (add y z) = unsigned y + unsigned z).
- unfold add. apply unsigned_repr.
+ unfold add. apply unsigned_repr.
generalize two_wordsize_max_unsigned; omega.
- apply same_bits_eq; intros.
+ apply same_bits_eq; intros.
rewrite bits_shl; auto.
destruct (zlt i (unsigned z)).
- rewrite bits_shl; auto. rewrite zlt_true. auto. omega.
- rewrite bits_shl. destruct (zlt (i - unsigned z) (unsigned y)).
- + rewrite bits_shl; auto. rewrite zlt_true. auto. omega.
+ + rewrite bits_shl; auto. rewrite zlt_true. auto. omega.
+ rewrite bits_shl; auto. rewrite zlt_false. f_equal. omega. omega.
+ omega.
Qed.
@@ -2091,10 +2091,10 @@ Lemma bitwise_binop_shru:
f' false false = false ->
f (shru x n) (shru y n) = shru (f x y) n.
Proof.
- intros. apply same_bits_eq; intros.
+ intros. apply same_bits_eq; intros.
rewrite H; auto. rewrite !bits_shru; auto.
destruct (zlt (i + unsigned n) zwordsize); auto.
- rewrite H; auto. generalize (unsigned_range n); omega.
+ rewrite H; auto. generalize (unsigned_range n); omega.
Qed.
Theorem and_shru:
@@ -2120,7 +2120,7 @@ Qed.
Theorem shru_shru:
forall x y z,
- ltu y iwordsize = true ->
+ ltu y iwordsize = true ->
ltu z iwordsize = true ->
ltu (add y z) iwordsize = true ->
shru (shru x y) z = shru x (add y z).
@@ -2128,14 +2128,14 @@ Proof.
intros.
generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros.
assert (unsigned (add y z) = unsigned y + unsigned z).
- unfold add. apply unsigned_repr.
+ unfold add. apply unsigned_repr.
generalize two_wordsize_max_unsigned; omega.
- apply same_bits_eq; intros.
+ apply same_bits_eq; intros.
rewrite bits_shru; auto.
destruct (zlt (i + unsigned z) zwordsize).
- rewrite bits_shru. destruct (zlt (i + unsigned z + unsigned y) zwordsize).
- + rewrite bits_shru; auto. rewrite zlt_true. f_equal. omega. omega.
- + rewrite bits_shru; auto. rewrite zlt_false. auto. omega.
+ + rewrite bits_shru; auto. rewrite zlt_true. f_equal. omega. omega.
+ + rewrite bits_shru; auto. rewrite zlt_false. auto. omega.
+ omega.
- rewrite bits_shru; auto. rewrite zlt_false. auto. omega.
Qed.
@@ -2150,10 +2150,10 @@ Lemma bitwise_binop_shr:
(forall x y i, 0 <= i < zwordsize -> testbit (f x y) i = f' (testbit x i) (testbit y i)) ->
f (shr x n) (shr y n) = shr (f x y) n.
Proof.
- intros. apply same_bits_eq; intros.
+ intros. apply same_bits_eq; intros.
rewrite H; auto. rewrite !bits_shr; auto.
- rewrite H; auto.
- destruct (zlt (i + unsigned n) zwordsize).
+ rewrite H; auto.
+ destruct (zlt (i + unsigned n) zwordsize).
generalize (unsigned_range n); omega.
omega.
Qed.
@@ -2181,7 +2181,7 @@ Qed.
Theorem shr_shr:
forall x y z,
- ltu y iwordsize = true ->
+ ltu y iwordsize = true ->
ltu z iwordsize = true ->
ltu (add y z) iwordsize = true ->
shr (shr x y) z = shr x (add y z).
@@ -2189,14 +2189,14 @@ Proof.
intros.
generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros.
assert (unsigned (add y z) = unsigned y + unsigned z).
- unfold add. apply unsigned_repr.
+ unfold add. apply unsigned_repr.
generalize two_wordsize_max_unsigned; omega.
- apply same_bits_eq; intros.
+ apply same_bits_eq; intros.
rewrite !bits_shr; auto. f_equal.
destruct (zlt (i + unsigned z) zwordsize).
- rewrite H4. replace (i + (unsigned y + unsigned z)) with (i + unsigned z + unsigned y) by omega. auto.
+ rewrite H4. replace (i + (unsigned y + unsigned z)) with (i + unsigned z + unsigned y) by omega. auto.
rewrite (zlt_false _ (i + unsigned (add y z))).
- destruct (zlt (zwordsize - 1 + unsigned y) zwordsize); omega.
+ destruct (zlt (zwordsize - 1 + unsigned y) zwordsize); omega.
omega.
destruct (zlt (i + unsigned z) zwordsize); omega.
Qed.
@@ -2217,8 +2217,8 @@ Theorem shr_and_shru_and:
shru (shl z y) y = z ->
and (shr x y) z = and (shru x y) z.
Proof.
- intros.
- rewrite <- H.
+ intros.
+ rewrite <- H.
rewrite and_shru. rewrite and_shr_shru. auto.
Qed.
@@ -2228,19 +2228,19 @@ Theorem shru_lt_zero:
Proof.
intros. apply same_bits_eq; intros.
rewrite bits_shru; auto.
- rewrite unsigned_repr.
+ rewrite unsigned_repr.
destruct (zeq i 0).
subst i. rewrite Zplus_0_l. rewrite zlt_true.
rewrite sign_bit_of_unsigned.
- unfold lt. rewrite signed_zero. unfold signed.
+ unfold lt. rewrite signed_zero. unfold signed.
destruct (zlt (unsigned x) half_modulus).
- rewrite zlt_false. auto. generalize (unsigned_range x); omega.
- rewrite zlt_true. unfold one; rewrite testbit_repr; auto.
- generalize (unsigned_range x); omega.
+ rewrite zlt_false. auto. generalize (unsigned_range x); omega.
+ rewrite zlt_true. unfold one; rewrite testbit_repr; auto.
+ generalize (unsigned_range x); omega.
omega.
rewrite zlt_false.
- unfold testbit. rewrite Ztestbit_eq. rewrite zeq_false.
- destruct (lt x zero).
+ unfold testbit. rewrite Ztestbit_eq. rewrite zeq_false.
+ destruct (lt x zero).
rewrite unsigned_one. simpl Z.div2. rewrite Z.testbit_0_l; auto.
rewrite unsigned_zero. simpl Z.div2. rewrite Z.testbit_0_l; auto.
auto. omega. omega.
@@ -2256,10 +2256,10 @@ Proof.
rewrite unsigned_repr.
transitivity (testbit x (zwordsize - 1)).
f_equal. destruct (zlt (i + (zwordsize - 1)) zwordsize); omega.
- rewrite sign_bit_of_unsigned.
- unfold lt. rewrite signed_zero. unfold signed.
+ rewrite sign_bit_of_unsigned.
+ unfold lt. rewrite signed_zero. unfold signed.
destruct (zlt (unsigned x) half_modulus).
- rewrite zlt_false. rewrite bits_zero; auto. generalize (unsigned_range x); omega.
+ rewrite zlt_false. rewrite bits_zero; auto. generalize (unsigned_range x); omega.
rewrite zlt_true. rewrite bits_mone; auto. generalize (unsigned_range x); omega.
generalize wordsize_max_unsigned; omega.
Qed.
@@ -2267,18 +2267,18 @@ Qed.
(** ** Properties of rotations *)
Lemma bits_rol:
- forall x y i,
+ forall x y i,
0 <= i < zwordsize ->
testbit (rol x y) i = testbit x ((i - unsigned y) mod zwordsize).
Proof.
intros. unfold rol.
- exploit (Z_div_mod_eq (unsigned y) zwordsize). apply wordsize_pos.
- set (j := unsigned y mod zwordsize). set (k := unsigned y / zwordsize).
+ exploit (Z_div_mod_eq (unsigned y) zwordsize). apply wordsize_pos.
+ set (j := unsigned y mod zwordsize). set (k := unsigned y / zwordsize).
intros EQ.
- exploit (Z_mod_lt (unsigned y) zwordsize). apply wordsize_pos.
+ exploit (Z_mod_lt (unsigned y) zwordsize). apply wordsize_pos.
fold j. intros RANGE.
rewrite testbit_repr; auto.
- rewrite Z.lor_spec. rewrite Z.shiftr_spec. 2: omega.
+ rewrite Z.lor_spec. rewrite Z.shiftr_spec. 2: omega.
destruct (zlt i j).
- rewrite Z.shiftl_spec_low; auto. simpl.
unfold testbit. f_equal.
@@ -2289,9 +2289,9 @@ Proof.
fold (testbit x (i + (zwordsize - j))).
rewrite bits_above. rewrite orb_false_r.
fold (testbit x (i - j)).
- f_equal. symmetry. apply Zmod_unique with (-k).
+ f_equal. symmetry. apply Zmod_unique with (-k).
rewrite EQ. ring.
- omega. omega. omega. omega.
+ omega. omega. omega. omega.
Qed.
Lemma bits_ror:
@@ -2300,26 +2300,26 @@ Lemma bits_ror:
testbit (ror x y) i = testbit x ((i + unsigned y) mod zwordsize).
Proof.
intros. unfold ror.
- exploit (Z_div_mod_eq (unsigned y) zwordsize). apply wordsize_pos.
- set (j := unsigned y mod zwordsize). set (k := unsigned y / zwordsize).
+ exploit (Z_div_mod_eq (unsigned y) zwordsize). apply wordsize_pos.
+ set (j := unsigned y mod zwordsize). set (k := unsigned y / zwordsize).
intros EQ.
- exploit (Z_mod_lt (unsigned y) zwordsize). apply wordsize_pos.
+ exploit (Z_mod_lt (unsigned y) zwordsize). apply wordsize_pos.
fold j. intros RANGE.
rewrite testbit_repr; auto.
- rewrite Z.lor_spec. rewrite Z.shiftr_spec. 2: omega.
+ rewrite Z.lor_spec. rewrite Z.shiftr_spec. 2: omega.
destruct (zlt (i + j) zwordsize).
- - rewrite Z.shiftl_spec_low; auto. rewrite orb_false_r.
+ - rewrite Z.shiftl_spec_low; auto. rewrite orb_false_r.
unfold testbit. f_equal.
symmetry. apply Zmod_unique with k.
rewrite EQ. ring.
omega. omega.
- rewrite Z.shiftl_spec_high.
- fold (testbit x (i + j)).
- rewrite bits_above. simpl.
- unfold testbit. f_equal.
- symmetry. apply Zmod_unique with (k + 1).
+ fold (testbit x (i + j)).
+ rewrite bits_above. simpl.
+ unfold testbit. f_equal.
+ symmetry. apply Zmod_unique with (k + 1).
rewrite EQ. ring.
- omega. omega. omega. omega.
+ omega. omega. omega. omega.
Qed.
Hint Rewrite bits_rol bits_ror: ints.
@@ -2330,13 +2330,13 @@ Theorem shl_rolm:
shl x n = rolm x n (shl mone n).
Proof.
intros. generalize (ltu_inv _ _ H). rewrite unsigned_repr_wordsize; intros.
- unfold rolm. apply same_bits_eq; intros.
- rewrite bits_and; auto. rewrite !bits_shl; auto. rewrite bits_rol; auto.
+ unfold rolm. apply same_bits_eq; intros.
+ rewrite bits_and; auto. rewrite !bits_shl; auto. rewrite bits_rol; auto.
destruct (zlt i (unsigned n)).
- rewrite andb_false_r; auto.
- - generalize (unsigned_range n); intros.
- rewrite bits_mone. rewrite andb_true_r. f_equal.
- symmetry. apply Zmod_small. omega.
+ - generalize (unsigned_range n); intros.
+ rewrite bits_mone. rewrite andb_true_r. f_equal.
+ symmetry. apply Zmod_small. omega.
omega.
Qed.
@@ -2346,15 +2346,15 @@ Theorem shru_rolm:
shru x n = rolm x (sub iwordsize n) (shru mone n).
Proof.
intros. generalize (ltu_inv _ _ H). rewrite unsigned_repr_wordsize; intros.
- unfold rolm. apply same_bits_eq; intros.
- rewrite bits_and; auto. rewrite !bits_shru; auto. rewrite bits_rol; auto.
+ unfold rolm. apply same_bits_eq; intros.
+ rewrite bits_and; auto. rewrite !bits_shru; auto. rewrite bits_rol; auto.
destruct (zlt (i + unsigned n) zwordsize).
- - generalize (unsigned_range n); intros.
+ - generalize (unsigned_range n); intros.
rewrite bits_mone. rewrite andb_true_r. f_equal.
unfold sub. rewrite unsigned_repr. rewrite unsigned_repr_wordsize.
- symmetry. apply Zmod_unique with (-1). ring. omega.
+ symmetry. apply Zmod_unique with (-1). ring. omega.
rewrite unsigned_repr_wordsize. generalize wordsize_max_unsigned. omega.
- omega.
+ omega.
- rewrite andb_false_r; auto.
Qed.
@@ -2362,8 +2362,8 @@ Theorem rol_zero:
forall x,
rol x zero = x.
Proof.
- bit_solve. f_equal. rewrite unsigned_zero. rewrite Zminus_0_r.
- apply Zmod_small; auto.
+ bit_solve. f_equal. rewrite unsigned_zero. rewrite Zminus_0_r.
+ apply Zmod_small; auto.
Qed.
Lemma bitwise_binop_rol:
@@ -2371,8 +2371,8 @@ Lemma bitwise_binop_rol:
(forall x y i, 0 <= i < zwordsize -> testbit (f x y) i = f' (testbit x i) (testbit y i)) ->
rol (f x y) n = f (rol x n) (rol y n).
Proof.
- intros. apply same_bits_eq; intros.
- rewrite H; auto. rewrite !bits_rol; auto. rewrite H; auto.
+ intros. apply same_bits_eq; intros.
+ rewrite H; auto. rewrite !bits_rol; auto. rewrite H; auto.
apply Z_mod_lt. apply wordsize_pos.
Qed.
@@ -2402,11 +2402,11 @@ Theorem rol_rol:
Zdivide zwordsize modulus ->
rol (rol x n) m = rol x (modu (add n m) iwordsize).
Proof.
- bit_solve. f_equal. apply eqmod_mod_eq. apply wordsize_pos.
+ bit_solve. f_equal. apply eqmod_mod_eq. apply wordsize_pos.
set (M := unsigned m); set (N := unsigned n).
apply eqmod_trans with (i - M - N).
apply eqmod_sub.
- apply eqmod_sym. apply eqmod_mod. apply wordsize_pos.
+ apply eqmod_sym. apply eqmod_mod. apply wordsize_pos.
apply eqmod_refl.
replace (i - M - N) with (i - (M + N)) by omega.
apply eqmod_sub.
@@ -2416,8 +2416,8 @@ Proof.
unfold modu, add. fold M; fold N. rewrite unsigned_repr_wordsize.
assert (forall a, eqmod zwordsize a (unsigned (repr a))).
intros. eapply eqmod_divides. apply eqm_unsigned_repr. assumption.
- eapply eqmod_trans. 2: apply H1.
- apply eqmod_refl2. apply eqmod_mod_eq. apply wordsize_pos. auto.
+ eapply eqmod_trans. 2: apply H1.
+ apply eqmod_refl2. apply eqmod_mod_eq. apply wordsize_pos. auto.
apply Z_mod_lt. apply wordsize_pos.
Qed.
@@ -2436,7 +2436,7 @@ Theorem rolm_rolm:
(and (rol m1 n2) m2).
Proof.
intros.
- unfold rolm. rewrite rol_and. rewrite and_assoc.
+ unfold rolm. rewrite rol_and. rewrite and_assoc.
rewrite rol_rol. reflexivity. auto.
Qed.
@@ -2444,7 +2444,7 @@ Theorem or_rolm:
forall x n m1 m2,
or (rolm x n m1) (rolm x n m2) = rolm x n (or m1 m2).
Proof.
- intros; unfold rolm. symmetry. apply and_or_distrib.
+ intros; unfold rolm. symmetry. apply and_or_distrib.
Qed.
Theorem ror_rol:
@@ -2455,23 +2455,23 @@ Proof.
intros.
generalize (ltu_iwordsize_inv _ H); intros.
apply same_bits_eq; intros.
- rewrite bits_ror; auto. rewrite bits_rol; auto. f_equal.
+ rewrite bits_ror; auto. rewrite bits_rol; auto. f_equal.
unfold sub. rewrite unsigned_repr. rewrite unsigned_repr_wordsize.
- apply eqmod_mod_eq. apply wordsize_pos. exists 1. ring.
- rewrite unsigned_repr_wordsize.
- generalize wordsize_pos; generalize wordsize_max_unsigned; omega.
+ apply eqmod_mod_eq. apply wordsize_pos. exists 1. ring.
+ rewrite unsigned_repr_wordsize.
+ generalize wordsize_pos; generalize wordsize_max_unsigned; omega.
Qed.
Theorem ror_rol_neg:
forall x y, (zwordsize | modulus) -> ror x y = rol x (neg y).
Proof.
intros. apply same_bits_eq; intros.
- rewrite bits_ror by auto. rewrite bits_rol by auto.
- f_equal. apply eqmod_mod_eq. omega.
- apply eqmod_trans with (i - (- unsigned y)).
- apply eqmod_refl2; omega.
+ rewrite bits_ror by auto. rewrite bits_rol by auto.
+ f_equal. apply eqmod_mod_eq. omega.
+ apply eqmod_trans with (i - (- unsigned y)).
+ apply eqmod_refl2; omega.
apply eqmod_sub. apply eqmod_refl.
- apply eqmod_divides with modulus.
+ apply eqmod_divides with modulus.
apply eqm_unsigned_repr. auto.
Qed.
@@ -2484,17 +2484,17 @@ Theorem or_ror:
Proof.
intros.
generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros.
- unfold ror, or, shl, shru. apply same_bits_eq; intros.
- rewrite !testbit_repr; auto.
+ unfold ror, or, shl, shru. apply same_bits_eq; intros.
+ rewrite !testbit_repr; auto.
rewrite !Z.lor_spec. rewrite orb_comm. f_equal; apply same_bits_eqm; auto.
- apply eqm_unsigned_repr_r. apply eqm_refl2. f_equal.
- rewrite Zmod_small; auto.
- assert (unsigned (add y z) = zwordsize).
- rewrite H1. apply unsigned_repr_wordsize.
- unfold add in H5. rewrite unsigned_repr in H5.
- omega.
- generalize two_wordsize_max_unsigned; omega.
- - apply eqm_unsigned_repr_r. apply eqm_refl2. f_equal.
+ rewrite Zmod_small; auto.
+ assert (unsigned (add y z) = zwordsize).
+ rewrite H1. apply unsigned_repr_wordsize.
+ unfold add in H5. rewrite unsigned_repr in H5.
+ omega.
+ generalize two_wordsize_max_unsigned; omega.
+ - apply eqm_unsigned_repr_r. apply eqm_refl2. f_equal.
apply Zmod_small; auto.
Qed.
@@ -2509,23 +2509,23 @@ Fixpoint powerserie (l: list Z): Z :=
Lemma Z_one_bits_powerserie:
forall x, 0 <= x < modulus -> x = powerserie (Z_one_bits wordsize x 0).
Proof.
- assert (forall n x i,
+ assert (forall n x i,
0 <= i ->
0 <= x < two_power_nat n ->
x * two_p i = powerserie (Z_one_bits n x i)).
{
induction n; intros.
- simpl. rewrite two_power_nat_O in H0.
+ simpl. rewrite two_power_nat_O in H0.
assert (x = 0) by omega. subst x. omega.
rewrite two_power_nat_S in H0. simpl Z_one_bits.
rewrite (Zdecomp x) in H0. rewrite Zshiftin_spec in H0.
assert (EQ: Z.div2 x * two_p (i + 1) = powerserie (Z_one_bits n (Z.div2 x) (i + 1))).
apply IHn. omega.
- destruct (Z.odd x); omega.
+ destruct (Z.odd x); omega.
rewrite two_p_is_exp in EQ. change (two_p 1) with 2 in EQ.
- rewrite (Zdecomp x) at 1. rewrite Zshiftin_spec.
+ rewrite (Zdecomp x) at 1. rewrite Zshiftin_spec.
destruct (Z.odd x); simpl powerserie; rewrite <- EQ; ring.
- omega. omega.
+ omega. omega.
}
intros. rewrite <- H. change (two_p 0) with 1. omega.
omega. exact H0.
@@ -2543,7 +2543,7 @@ Proof.
assert (In j (Z_one_bits n (Z.div2 x) (i + 1)) -> i <= j < i + Z.succ (Z.of_nat n)).
intros. exploit IHn; eauto. omega.
destruct (Z.odd x); simpl.
- intros [A|B]. subst j. omega. auto.
+ intros [A|B]. subst j. omega. auto.
auto.
}
intros. generalize (H wordsize x 0 i H0). fold zwordsize; omega.
@@ -2572,7 +2572,7 @@ Theorem is_power2_range:
Proof.
intros. unfold ltu. rewrite unsigned_repr_wordsize.
apply zlt_true. generalize (is_power2_rng _ _ H). tauto.
-Qed.
+Qed.
Lemma is_power2_correct:
forall n logn,
@@ -2589,7 +2589,7 @@ Proof.
rewrite unsigned_repr. replace (two_p z) with (two_p z + 0).
auto. omega. elim (H z); intros.
generalize wordsize_max_unsigned; omega.
- auto with coqlib.
+ auto with coqlib.
intros; discriminate.
Qed.
@@ -2600,9 +2600,9 @@ Remark two_p_range:
Proof.
intros. split.
assert (two_p n > 0). apply two_p_gt_ZERO. omega. omega.
- generalize (two_p_monotone_strict _ _ H).
- unfold zwordsize; rewrite <- two_power_nat_two_p.
- unfold max_unsigned, modulus. omega.
+ generalize (two_p_monotone_strict _ _ H).
+ unfold zwordsize; rewrite <- two_power_nat_two_p.
+ unfold max_unsigned, modulus. omega.
Qed.
Remark Z_one_bits_zero:
@@ -2617,7 +2617,7 @@ Remark Z_one_bits_two_p:
Z_one_bits n (two_p x) i = (i + x) :: nil.
Proof.
induction n; intros; simpl. simpl in H. omegaContradiction.
- rewrite inj_S in H.
+ rewrite inj_S in H.
assert (x = 0 \/ 0 < x) by omega. destruct H0.
subst x; simpl. decEq. omega. apply Z_one_bits_zero.
assert (Z.odd (two_p x) = false /\ Z.div2 (two_p x) = two_p (x-1)).
@@ -2631,7 +2631,7 @@ Lemma is_power2_two_p:
forall n, 0 <= n < zwordsize ->
is_power2 (repr (two_p n)) = Some (repr n).
Proof.
- intros. unfold is_power2. rewrite unsigned_repr.
+ intros. unfold is_power2. rewrite unsigned_repr.
rewrite Z_one_bits_two_p. auto. auto.
apply two_p_range. auto.
Qed.
@@ -2645,7 +2645,7 @@ Lemma Zshiftl_mul_two_p:
Proof.
intros. destruct n; simpl.
- omega.
- - pattern p. apply Pos.peano_ind.
+ - pattern p. apply Pos.peano_ind.
+ change (two_power_pos 1) with 2. simpl. ring.
+ intros. rewrite Pos.iter_succ. rewrite H0.
rewrite Pplus_one_succ_l. rewrite two_power_pos_is_exp.
@@ -2658,7 +2658,7 @@ Lemma shl_mul_two_p:
shl x y = mul x (repr (two_p (unsigned y))).
Proof.
intros. unfold shl, mul. apply eqm_samerepr.
- rewrite Zshiftl_mul_two_p. auto with ints.
+ rewrite Zshiftl_mul_two_p. auto with ints.
generalize (unsigned_range y); omega.
Qed.
@@ -2666,7 +2666,7 @@ Theorem shl_mul:
forall x y,
shl x y = mul x (shl one y).
Proof.
- intros.
+ intros.
assert (shl one y = repr (two_p (unsigned y))).
{
rewrite shl_mul_two_p. rewrite mul_commut. rewrite mul_one. auto.
@@ -2690,21 +2690,21 @@ Theorem shifted_or_is_add:
unsigned y < two_p n ->
or (shl x (repr n)) y = repr(unsigned x * two_p n + unsigned y).
Proof.
- intros. rewrite <- add_is_or.
- - unfold add. apply eqm_samerepr. apply eqm_add; auto with ints.
+ intros. rewrite <- add_is_or.
+ - unfold add. apply eqm_samerepr. apply eqm_add; auto with ints.
rewrite shl_mul_two_p. unfold mul. apply eqm_unsigned_repr_l.
- apply eqm_mult; auto with ints. apply eqm_unsigned_repr_l.
- apply eqm_refl2. rewrite unsigned_repr. auto.
+ apply eqm_mult; auto with ints. apply eqm_unsigned_repr_l.
+ apply eqm_refl2. rewrite unsigned_repr. auto.
generalize wordsize_max_unsigned; omega.
- bit_solve.
- rewrite unsigned_repr.
+ rewrite unsigned_repr.
destruct (zlt i n).
+ auto.
- + replace (testbit y i) with false. apply andb_false_r.
+ + replace (testbit y i) with false. apply andb_false_r.
symmetry. unfold testbit.
assert (EQ: Z.of_nat (Z.to_nat n) = n) by (apply Z2Nat.id; omega).
apply Ztestbit_above with (Z.to_nat n).
- rewrite <- EQ in H0. rewrite <- two_power_nat_two_p in H0.
+ rewrite <- EQ in H0. rewrite <- two_power_nat_two_p in H0.
generalize (unsigned_range y); omega.
rewrite EQ; auto.
+ generalize wordsize_max_unsigned; omega.
@@ -2717,8 +2717,8 @@ Lemma Zshiftr_div_two_p:
Proof.
intros. destruct n; unfold Z.shiftr; simpl.
- rewrite Zdiv_1_r. auto.
- - pattern p. apply Pos.peano_ind.
- + change (two_power_pos 1) with 2. simpl. apply Zdiv2_div.
+ - pattern p. apply Pos.peano_ind.
+ + change (two_power_pos 1) with 2. simpl. apply Zdiv2_div.
+ intros. rewrite Pos.iter_succ. rewrite H0.
rewrite Pplus_one_succ_l. rewrite two_power_pos_is_exp.
change (two_power_pos 1) with 2.
@@ -2731,7 +2731,7 @@ Lemma shru_div_two_p:
forall x y,
shru x y = repr (unsigned x / two_p (unsigned y)).
Proof.
- intros. unfold shru.
+ intros. unfold shru.
rewrite Zshiftr_div_two_p. auto.
generalize (unsigned_range y); omega.
Qed.
@@ -2775,12 +2775,12 @@ Lemma Ztestbit_mod_two_p:
Z.testbit (x mod (two_p n)) i = if zlt i n then Z.testbit x i else false.
Proof.
intros n0 x i N0POS. revert x i; pattern n0; apply natlike_ind; auto.
- - intros. change (two_p 0) with 1. rewrite Zmod_1_r. rewrite Z.testbit_0_l.
+ - intros. change (two_p 0) with 1. rewrite Zmod_1_r. rewrite Z.testbit_0_l.
rewrite zlt_false; auto. omega.
- - intros. rewrite two_p_S; auto.
- replace (x0 mod (2 * two_p x))
+ - intros. rewrite two_p_S; auto.
+ replace (x0 mod (2 * two_p x))
with (Zshiftin (Z.odd x0) (Z.div2 x0 mod two_p x)).
- rewrite Ztestbit_shiftin; auto. rewrite (Ztestbit_eq i x0); auto. destruct (zeq i 0).
+ rewrite Ztestbit_shiftin; auto. rewrite (Ztestbit_eq i x0); auto. destruct (zeq i 0).
+ rewrite zlt_true; auto. omega.
+ rewrite H0. destruct (zlt (Z.pred i) x).
* rewrite zlt_true; auto. omega.
@@ -2800,10 +2800,10 @@ Corollary Ztestbit_two_p_m1:
forall n i, 0 <= n -> 0 <= i ->
Z.testbit (two_p n - 1) i = if zlt i n then true else false.
Proof.
- intros. replace (two_p n - 1) with ((-1) mod (two_p n)).
+ intros. replace (two_p n - 1) with ((-1) mod (two_p n)).
rewrite Ztestbit_mod_two_p; auto. destruct (zlt i n); auto. apply Ztestbit_m1; auto.
- apply Zmod_unique with (-1). ring.
- exploit (two_p_gt_ZERO n). auto. omega.
+ apply Zmod_unique with (-1). ring.
+ exploit (two_p_gt_ZERO n). auto. omega.
Qed.
Theorem modu_and:
@@ -2814,12 +2814,12 @@ Proof.
intros. generalize (is_power2_correct _ _ H); intro.
generalize (is_power2_rng _ _ H); intro.
apply same_bits_eq; intros.
- rewrite bits_and; auto.
+ rewrite bits_and; auto.
unfold sub. rewrite testbit_repr; auto.
- rewrite H0. rewrite unsigned_one.
+ rewrite H0. rewrite unsigned_one.
unfold modu. rewrite testbit_repr; auto. rewrite H0.
- rewrite Ztestbit_mod_two_p. rewrite Ztestbit_two_p_m1.
- destruct (zlt i (unsigned logn)).
+ rewrite Ztestbit_mod_two_p. rewrite Ztestbit_two_p_m1.
+ destruct (zlt i (unsigned logn)).
rewrite andb_true_r; auto.
rewrite andb_false_r; auto.
tauto. tauto. tauto. tauto.
@@ -2834,11 +2834,11 @@ Lemma Zquot_Zdiv:
Proof.
intros. destruct (zlt x 0).
- symmetry. apply Zquot_unique_full with ((x + y - 1) mod y - (y - 1)).
- + red. right; split. omega.
- exploit (Z_mod_lt (x + y - 1) y); auto.
+ + red. right; split. omega.
+ exploit (Z_mod_lt (x + y - 1) y); auto.
rewrite Z.abs_eq. omega. omega.
+ transitivity ((y * ((x + y - 1) / y) + (x + y - 1) mod y) - (y-1)).
- rewrite <- Z_div_mod_eq. ring. auto. ring.
+ rewrite <- Z_div_mod_eq. ring. auto. ring.
- apply Zquot_Zdiv_pos; omega.
Qed.
@@ -2850,20 +2850,20 @@ Proof.
intros.
set (uy := unsigned y).
assert (0 <= uy < zwordsize - 1).
- generalize (ltu_inv _ _ H). rewrite unsigned_repr. auto.
+ generalize (ltu_inv _ _ H). rewrite unsigned_repr. auto.
generalize wordsize_pos wordsize_max_unsigned; omega.
- rewrite shr_div_two_p. unfold shrx. unfold divs.
+ rewrite shr_div_two_p. unfold shrx. unfold divs.
assert (shl one y = repr (two_p uy)).
transitivity (mul one (repr (two_p uy))).
- symmetry. apply mul_pow2. replace y with (repr uy).
+ symmetry. apply mul_pow2. replace y with (repr uy).
apply is_power2_two_p. omega. apply repr_unsigned.
rewrite mul_commut. apply mul_one.
assert (two_p uy > 0). apply two_p_gt_ZERO. omega.
- assert (two_p uy < half_modulus).
- rewrite half_modulus_power.
+ assert (two_p uy < half_modulus).
+ rewrite half_modulus_power.
apply two_p_monotone_strict. auto.
assert (two_p uy < modulus).
- rewrite modulus_power. apply two_p_monotone_strict. omega.
+ rewrite modulus_power. apply two_p_monotone_strict. omega.
assert (unsigned (shl one y) = two_p uy).
rewrite H1. apply unsigned_repr. unfold max_unsigned. omega.
assert (signed (shl one y) = two_p uy).
@@ -2871,15 +2871,15 @@ Proof.
unfold max_signed. generalize min_signed_neg. omega.
rewrite H6.
rewrite Zquot_Zdiv; auto.
- unfold lt. rewrite signed_zero.
+ unfold lt. rewrite signed_zero.
destruct (zlt (signed x) 0); auto.
rewrite add_signed.
assert (signed (sub (shl one y) one) = two_p uy - 1).
- unfold sub. rewrite H5. rewrite unsigned_one.
+ unfold sub. rewrite H5. rewrite unsigned_one.
apply signed_repr.
- generalize min_signed_neg. unfold max_signed. omega.
+ generalize min_signed_neg. unfold max_signed. omega.
rewrite H7. rewrite signed_repr. f_equal. f_equal. omega.
- generalize (signed_range x). intros.
+ generalize (signed_range x). intros.
assert (two_p uy - 1 <= max_signed). unfold max_signed. omega. omega.
Qed.
@@ -2888,27 +2888,27 @@ Theorem shrx_shr_2:
ltu y (repr (zwordsize - 1)) = true ->
shrx x y = shr (add x (shru (shr x (repr (zwordsize - 1))) (sub iwordsize y))) y.
Proof.
- intros.
+ intros.
rewrite shrx_shr by auto. f_equal.
rewrite shr_lt_zero. destruct (lt x zero).
- set (uy := unsigned y).
generalize (unsigned_range y); fold uy; intros.
assert (0 <= uy < zwordsize - 1).
- generalize (ltu_inv _ _ H). rewrite unsigned_repr. auto.
+ generalize (ltu_inv _ _ H). rewrite unsigned_repr. auto.
generalize wordsize_pos wordsize_max_unsigned; omega.
assert (two_p uy < modulus).
- rewrite modulus_power. apply two_p_monotone_strict. omega.
+ rewrite modulus_power. apply two_p_monotone_strict. omega.
f_equal. rewrite shl_mul_two_p. fold uy. rewrite mul_commut. rewrite mul_one.
- unfold sub. rewrite unsigned_one. rewrite unsigned_repr.
- rewrite unsigned_repr_wordsize. fold uy.
- apply same_bits_eq; intros. rewrite bits_shru by auto.
+ unfold sub. rewrite unsigned_one. rewrite unsigned_repr.
+ rewrite unsigned_repr_wordsize. fold uy.
+ apply same_bits_eq; intros. rewrite bits_shru by auto.
rewrite testbit_repr by auto. rewrite Ztestbit_two_p_m1 by omega.
rewrite unsigned_repr by (generalize wordsize_max_unsigned; omega).
- destruct (zlt i uy).
+ destruct (zlt i uy).
rewrite zlt_true by omega. rewrite bits_mone by omega. auto.
rewrite zlt_false by omega. auto.
assert (two_p uy > 0) by (apply two_p_gt_ZERO; omega). unfold max_unsigned; omega.
-- replace (shru zero (sub iwordsize y)) with zero.
+- replace (shru zero (sub iwordsize y)) with zero.
rewrite add_zero; auto.
bit_solve. destruct (zlt (i + unsigned (sub iwordsize y)) zwordsize); auto.
Qed.
@@ -2917,9 +2917,9 @@ Lemma Zdiv_shift:
forall x y, y > 0 ->
(x + (y - 1)) / y = x / y + if zeq (Zmod x y) 0 then 0 else 1.
Proof.
- intros. generalize (Z_div_mod_eq x y H). generalize (Z_mod_lt x y H).
+ intros. generalize (Z_div_mod_eq x y H). generalize (Z_mod_lt x y H).
set (q := x / y). set (r := x mod y). intros.
- destruct (zeq r 0).
+ destruct (zeq r 0).
apply Zdiv_unique with (y - 1). rewrite H1. rewrite e. ring. omega.
apply Zdiv_unique with (r - 1). rewrite H1. ring. omega.
Qed.
@@ -2930,8 +2930,8 @@ Theorem shrx_carry:
shrx x y = add (shr x y) (shr_carry x y).
Proof.
intros. rewrite shrx_shr; auto. unfold shr_carry.
- unfold lt. set (sx := signed x). rewrite signed_zero.
- destruct (zlt sx 0); simpl.
+ unfold lt. set (sx := signed x). rewrite signed_zero.
+ destruct (zlt sx 0); simpl.
2: rewrite add_zero; auto.
set (uy := unsigned y).
assert (0 <= uy < zwordsize - 1).
@@ -2943,46 +2943,46 @@ Proof.
symmetry. rewrite H1. apply modu_and with (logn := y).
rewrite is_power2_two_p. unfold uy. rewrite repr_unsigned. auto.
omega.
- rewrite H2. rewrite H1.
+ rewrite H2. rewrite H1.
repeat rewrite shr_div_two_p. fold sx. fold uy.
assert (two_p uy > 0). apply two_p_gt_ZERO. omega.
assert (two_p uy < modulus).
rewrite modulus_power. apply two_p_monotone_strict. omega.
- assert (two_p uy < half_modulus).
- rewrite half_modulus_power.
+ assert (two_p uy < half_modulus).
+ rewrite half_modulus_power.
apply two_p_monotone_strict. auto.
assert (two_p uy < modulus).
rewrite modulus_power. apply two_p_monotone_strict. omega.
assert (sub (repr (two_p uy)) one = repr (two_p uy - 1)).
- unfold sub. apply eqm_samerepr. apply eqm_sub. apply eqm_sym; apply eqm_unsigned_repr.
+ unfold sub. apply eqm_samerepr. apply eqm_sub. apply eqm_sym; apply eqm_unsigned_repr.
rewrite unsigned_one. apply eqm_refl.
rewrite H7. rewrite add_signed. fold sx.
- rewrite (signed_repr (two_p uy - 1)). rewrite signed_repr.
- unfold modu. rewrite unsigned_repr.
- unfold eq. rewrite unsigned_zero. rewrite unsigned_repr.
+ rewrite (signed_repr (two_p uy - 1)). rewrite signed_repr.
+ unfold modu. rewrite unsigned_repr.
+ unfold eq. rewrite unsigned_zero. rewrite unsigned_repr.
assert (unsigned x mod two_p uy = sx mod two_p uy).
- apply eqmod_mod_eq; auto. apply eqmod_divides with modulus.
+ apply eqmod_mod_eq; auto. apply eqmod_divides with modulus.
fold eqm. unfold sx. apply eqm_sym. apply eqm_signed_unsigned.
- unfold modulus. rewrite two_power_nat_two_p.
+ unfold modulus. rewrite two_power_nat_two_p.
exists (two_p (zwordsize - uy)). rewrite <- two_p_is_exp.
f_equal. fold zwordsize; omega. omega. omega.
rewrite H8. rewrite Zdiv_shift; auto.
- unfold add. apply eqm_samerepr. apply eqm_add.
- apply eqm_unsigned_repr.
+ unfold add. apply eqm_samerepr. apply eqm_add.
+ apply eqm_unsigned_repr.
destruct (zeq (sx mod two_p uy) 0); simpl.
rewrite unsigned_zero. apply eqm_refl.
rewrite unsigned_one. apply eqm_refl.
generalize (Z_mod_lt (unsigned x) (two_p uy) H3). unfold max_unsigned. omega.
unfold max_unsigned; omega.
- generalize (signed_range x). fold sx. intros. split. omega. unfold max_signed. omega.
- generalize min_signed_neg. unfold max_signed. omega.
+ generalize (signed_range x). fold sx. intros. split. omega. unfold max_signed. omega.
+ generalize min_signed_neg. unfold max_signed. omega.
Qed.
(** Connections between [shr] and [shru]. *)
Lemma shr_shru_positive:
forall x y,
- signed x >= 0 ->
+ signed x >= 0 ->
shr x y = shru x y.
Proof.
intros.
@@ -2996,7 +2996,7 @@ Proof.
intros.
assert (unsigned y < half_modulus). rewrite signed_positive in H. unfold max_signed in H; omega.
generalize (sign_bit_of_unsigned y). rewrite zlt_true; auto. intros A.
- generalize (sign_bit_of_unsigned (and x y)). rewrite bits_and. rewrite A.
+ generalize (sign_bit_of_unsigned (and x y)). rewrite bits_and. rewrite A.
rewrite andb_false_r. unfold signed.
destruct (zlt (unsigned (and x y)) half_modulus).
intros. generalize (unsigned_range (and x y)); omega.
@@ -3008,7 +3008,7 @@ Theorem shr_and_is_shru_and:
forall x y z,
lt y zero = false -> shr (and x y) z = shru (and x y) z.
Proof.
- intros. apply shr_shru_positive. apply and_positive.
+ intros. apply shr_shru_positive. apply and_positive.
unfold lt in H. rewrite signed_zero in H. destruct (zlt (signed y) 0). congruence. auto.
Qed.
@@ -3017,14 +3017,14 @@ Qed.
Lemma Ziter_base:
forall (A: Type) n (f: A -> A) x, n <= 0 -> Z.iter n f x = x.
Proof.
- intros. unfold Z.iter. destruct n; auto. compute in H. elim H; auto.
+ intros. unfold Z.iter. destruct n; auto. compute in H. elim H; auto.
Qed.
Lemma Ziter_succ:
forall (A: Type) n (f: A -> A) x,
0 <= n -> Z.iter (Z.succ n) f x = f (Z.iter n f x).
Proof.
- intros. destruct n; simpl.
+ intros. destruct n; simpl.
- auto.
- rewrite Pos.add_1_r. apply Pos.iter_succ.
- compute in H. elim H; auto.
@@ -3037,7 +3037,7 @@ Lemma Znatlike_ind:
forall n, P n.
Proof.
intros. destruct (zle 0 n).
- apply natlike_ind; auto. apply H; omega.
+ apply natlike_ind; auto. apply H; omega.
apply H. omega.
Qed.
@@ -3045,12 +3045,12 @@ Lemma Zzero_ext_spec:
forall n x i, 0 <= i ->
Z.testbit (Zzero_ext n x) i = if zlt i n then Z.testbit x i else false.
Proof.
- unfold Zzero_ext. induction n using Znatlike_ind.
+ unfold Zzero_ext. induction n using Znatlike_ind.
- intros. rewrite Ziter_base; auto.
rewrite zlt_false. rewrite Ztestbit_0; auto. omega.
- - intros. rewrite Ziter_succ; auto.
- rewrite Ztestbit_shiftin; auto.
- rewrite (Ztestbit_eq i x); auto.
+ - intros. rewrite Ziter_succ; auto.
+ rewrite Ztestbit_shiftin; auto.
+ rewrite (Ztestbit_eq i x); auto.
destruct (zeq i 0).
+ subst i. rewrite zlt_true; auto. omega.
+ rewrite IHn. destruct (zlt (Z.pred i) n).
@@ -3059,12 +3059,12 @@ Proof.
omega.
Qed.
-Lemma bits_zero_ext:
+Lemma bits_zero_ext:
forall n x i, 0 <= i ->
testbit (zero_ext n x) i = if zlt i n then testbit x i else false.
Proof.
intros. unfold zero_ext. destruct (zlt i zwordsize).
- rewrite testbit_repr; auto. rewrite Zzero_ext_spec. auto. auto.
+ rewrite testbit_repr; auto. rewrite Zzero_ext_spec. auto. auto.
rewrite !bits_above; auto. destruct (zlt i n); auto.
Qed.
@@ -3072,20 +3072,20 @@ Lemma Zsign_ext_spec:
forall n x i, 0 <= i -> 0 < n ->
Z.testbit (Zsign_ext n x) i = Z.testbit x (if zlt i n then i else n - 1).
Proof.
- intros n0 x i I0 N0.
+ intros n0 x i I0 N0.
revert x i I0. pattern n0. apply Zlt_lower_bound_ind with (z := 1).
- unfold Zsign_ext. intros.
destruct (zeq x 1).
- + subst x; simpl.
+ + subst x; simpl.
replace (if zlt i 1 then i else 0) with 0.
rewrite Ztestbit_base.
- destruct (Z.odd x0).
- apply Ztestbit_m1; auto.
+ destruct (Z.odd x0).
+ apply Ztestbit_m1; auto.
apply Ztestbit_0.
destruct (zlt i 1); omega.
+ set (x1 := Z.pred x). replace x1 with (Z.succ (Z.pred x1)).
- rewrite Ziter_succ. rewrite Ztestbit_shiftin.
- destruct (zeq i 0).
+ rewrite Ziter_succ. rewrite Ztestbit_shiftin.
+ destruct (zeq i 0).
* subst i. rewrite zlt_true. rewrite Ztestbit_base; auto. omega.
* rewrite H. unfold x1. destruct (zlt (Z.pred i) (Z.pred x)).
rewrite zlt_true. rewrite (Ztestbit_eq i x0); auto. rewrite zeq_false; auto. omega.
@@ -3097,13 +3097,13 @@ Proof.
- omega.
Qed.
-Lemma bits_sign_ext:
+Lemma bits_sign_ext:
forall n x i, 0 <= i < zwordsize -> 0 < n ->
testbit (sign_ext n x) i = testbit x (if zlt i n then i else n - 1).
Proof.
intros. unfold sign_ext.
rewrite testbit_repr; auto. rewrite Zsign_ext_spec. destruct (zlt i n); auto.
- omega. auto.
+ omega. auto.
Qed.
Hint Rewrite bits_zero_ext bits_sign_ext: ints.
@@ -3119,7 +3119,7 @@ Theorem sign_ext_above:
forall n x, n >= zwordsize -> sign_ext n x = x.
Proof.
intros. apply same_bits_eq; intros.
- unfold sign_ext; rewrite testbit_repr; auto.
+ unfold sign_ext; rewrite testbit_repr; auto.
rewrite Zsign_ext_spec. rewrite zlt_true. auto. omega. omega. omega.
Qed.
@@ -3127,22 +3127,22 @@ Theorem zero_ext_and:
forall n x, 0 <= n -> zero_ext n x = and x (repr (two_p n - 1)).
Proof.
bit_solve. rewrite testbit_repr; auto. rewrite Ztestbit_two_p_m1; intuition.
- destruct (zlt i n).
- rewrite andb_true_r; auto.
+ destruct (zlt i n).
+ rewrite andb_true_r; auto.
rewrite andb_false_r; auto.
tauto.
Qed.
Theorem zero_ext_mod:
- forall n x, 0 <= n < zwordsize ->
+ forall n x, 0 <= n < zwordsize ->
unsigned (zero_ext n x) = Zmod (unsigned x) (two_p n).
Proof.
intros. apply equal_same_bits. intros.
rewrite Ztestbit_mod_two_p; auto.
- fold (testbit (zero_ext n x) i).
+ fold (testbit (zero_ext n x) i).
destruct (zlt i zwordsize).
rewrite bits_zero_ext; auto.
- rewrite bits_above. rewrite zlt_false; auto. omega. omega.
+ rewrite bits_above. rewrite zlt_false; auto. omega. omega.
omega.
Qed.
@@ -3150,8 +3150,8 @@ Theorem zero_ext_widen:
forall x n n', 0 <= n <= n' ->
zero_ext n' (zero_ext n x) = zero_ext n x.
Proof.
- bit_solve. destruct (zlt i n).
- apply zlt_true. omega.
+ bit_solve. destruct (zlt i n).
+ apply zlt_true. omega.
destruct (zlt i n'); auto.
tauto. tauto.
Qed.
@@ -3163,11 +3163,11 @@ Proof.
intros. destruct (zlt n' zwordsize).
bit_solve. destruct (zlt i n').
auto.
- rewrite (zlt_false _ i n).
+ rewrite (zlt_false _ i n).
destruct (zlt (n' - 1) n); f_equal; omega.
omega. omega.
destruct (zlt i n'); omega.
- omega. omega.
+ omega. omega.
apply sign_ext_above; auto.
Qed.
@@ -3176,10 +3176,10 @@ Theorem sign_zero_ext_widen:
sign_ext n' (zero_ext n x) = zero_ext n x.
Proof.
intros. destruct (zlt n' zwordsize).
- bit_solve.
+ bit_solve.
destruct (zlt i n').
auto.
- rewrite !zlt_false. auto. omega. omega. omega.
+ rewrite !zlt_false. auto. omega. omega. omega.
destruct (zlt i n'); omega.
omega.
apply sign_ext_above; auto.
@@ -3189,8 +3189,8 @@ Theorem zero_ext_narrow:
forall x n n', 0 <= n <= n' ->
zero_ext n (zero_ext n' x) = zero_ext n x.
Proof.
- bit_solve. destruct (zlt i n).
- apply zlt_true. omega.
+ bit_solve. destruct (zlt i n).
+ apply zlt_true. omega.
auto.
omega. omega. omega.
Qed.
@@ -3201,10 +3201,10 @@ Theorem sign_ext_narrow:
Proof.
intros. destruct (zlt n zwordsize).
bit_solve. destruct (zlt i n); f_equal; apply zlt_true; omega.
- omega.
+ omega.
destruct (zlt i n); omega.
omega. omega.
- rewrite (sign_ext_above n'). auto. omega.
+ rewrite (sign_ext_above n'). auto. omega.
Qed.
Theorem zero_sign_ext_narrow:
@@ -3212,7 +3212,7 @@ Theorem zero_sign_ext_narrow:
zero_ext n (sign_ext n' x) = zero_ext n x.
Proof.
intros. destruct (zlt n' zwordsize).
- bit_solve.
+ bit_solve.
destruct (zlt i n); auto.
rewrite zlt_true; auto. omega.
omega. omega. omega.
@@ -3235,10 +3235,10 @@ Theorem sign_ext_zero_ext:
forall n x, 0 < n -> sign_ext n (zero_ext n x) = sign_ext n x.
Proof.
intros. destruct (zlt n zwordsize).
- bit_solve.
- destruct (zlt i n).
+ bit_solve.
+ destruct (zlt i n).
rewrite zlt_true; auto.
- rewrite zlt_true; auto. omega.
+ rewrite zlt_true; auto. omega.
destruct (zlt i n); omega.
rewrite zero_ext_above; auto.
Qed.
@@ -3268,12 +3268,12 @@ Proof.
assert (unsigned y = zwordsize - n).
unfold y. apply unsigned_repr. generalize wordsize_max_unsigned. omega.
apply same_bits_eq; intros.
- rewrite bits_zero_ext.
+ rewrite bits_zero_ext.
rewrite bits_shru; auto.
destruct (zlt i n).
- rewrite zlt_true. rewrite bits_shl. rewrite zlt_false. f_equal. omega.
- omega. omega. omega.
- rewrite zlt_false. auto. omega.
+ rewrite zlt_true. rewrite bits_shl. rewrite zlt_false. f_equal. omega.
+ omega. omega. omega.
+ rewrite zlt_false. auto. omega.
omega.
Qed.
@@ -3287,13 +3287,13 @@ Proof.
assert (unsigned y = zwordsize - n).
unfold y. apply unsigned_repr. generalize wordsize_max_unsigned. omega.
apply same_bits_eq; intros.
- rewrite bits_sign_ext.
+ rewrite bits_sign_ext.
rewrite bits_shr; auto.
destruct (zlt i n).
- rewrite zlt_true. rewrite bits_shl. rewrite zlt_false. f_equal. omega.
- omega. omega. omega.
- rewrite zlt_false. rewrite bits_shl. rewrite zlt_false. f_equal. omega.
- omega. omega. omega. omega. omega.
+ rewrite zlt_true. rewrite bits_shl. rewrite zlt_false. f_equal. omega.
+ omega. omega. omega.
+ rewrite zlt_false. rewrite bits_shl. rewrite zlt_false. f_equal. omega.
+ omega. omega. omega. omega. omega.
Qed.
(** [zero_ext n x] is the unique integer congruent to [x] modulo [2^n]
@@ -3302,13 +3302,13 @@ Qed.
Lemma zero_ext_range:
forall n x, 0 <= n < zwordsize -> 0 <= unsigned (zero_ext n x) < two_p n.
Proof.
- intros. rewrite zero_ext_mod; auto. apply Z_mod_lt. apply two_p_gt_ZERO. omega.
+ intros. rewrite zero_ext_mod; auto. apply Z_mod_lt. apply two_p_gt_ZERO. omega.
Qed.
Lemma eqmod_zero_ext:
forall n x, 0 <= n < zwordsize -> eqmod (two_p n) (unsigned (zero_ext n x)) (unsigned x).
Proof.
- intros. rewrite zero_ext_mod; auto. apply eqmod_sym. apply eqmod_mod.
+ intros. rewrite zero_ext_mod; auto. apply eqmod_sym. apply eqmod_mod.
apply two_p_gt_ZERO. omega.
Qed.
@@ -3318,25 +3318,25 @@ Qed.
Lemma sign_ext_range:
forall n x, 0 < n < zwordsize -> -two_p (n-1) <= signed (sign_ext n x) < two_p (n-1).
Proof.
- intros. rewrite sign_ext_shr_shl; auto.
+ intros. rewrite sign_ext_shr_shl; auto.
set (X := shl x (repr (zwordsize - n))).
assert (two_p (n - 1) > 0) by (apply two_p_gt_ZERO; omega).
assert (unsigned (repr (zwordsize - n)) = zwordsize - n).
- apply unsigned_repr.
+ apply unsigned_repr.
split. omega. generalize wordsize_max_unsigned; omega.
rewrite shr_div_two_p.
rewrite signed_repr.
rewrite H1.
- apply Zdiv_interval_1.
+ apply Zdiv_interval_1.
omega. omega. apply two_p_gt_ZERO; omega.
replace (- two_p (n - 1) * two_p (zwordsize - n))
with (- (two_p (n - 1) * two_p (zwordsize - n))) by ring.
rewrite <- two_p_is_exp.
replace (n - 1 + (zwordsize - n)) with (zwordsize - 1) by omega.
rewrite <- half_modulus_power.
- generalize (signed_range X). unfold min_signed, max_signed. omega.
+ generalize (signed_range X). unfold min_signed, max_signed. omega.
omega. omega.
- apply Zdiv_interval_2. apply signed_range.
+ apply Zdiv_interval_2. apply signed_range.
generalize min_signed_neg; omega.
generalize max_signed_pos; omega.
rewrite H1. apply two_p_gt_ZERO. omega.
@@ -3346,13 +3346,13 @@ Lemma eqmod_sign_ext':
forall n x, 0 < n < zwordsize ->
eqmod (two_p n) (unsigned (sign_ext n x)) (unsigned x).
Proof.
- intros.
+ intros.
set (N := Z.to_nat n).
assert (Z.of_nat N = n) by (apply Z2Nat.id; omega).
- rewrite <- H0. rewrite <- two_power_nat_two_p.
- apply eqmod_same_bits; intros.
- rewrite H0 in H1. rewrite H0.
- fold (testbit (sign_ext n x) i). rewrite bits_sign_ext.
+ rewrite <- H0. rewrite <- two_power_nat_two_p.
+ apply eqmod_same_bits; intros.
+ rewrite H0 in H1. rewrite H0.
+ fold (testbit (sign_ext n x) i). rewrite bits_sign_ext.
rewrite zlt_true. auto. omega. omega. omega.
Qed.
@@ -3360,11 +3360,11 @@ Lemma eqmod_sign_ext:
forall n x, 0 < n < zwordsize ->
eqmod (two_p n) (signed (sign_ext n x)) (unsigned x).
Proof.
- intros. apply eqmod_trans with (unsigned (sign_ext n x)).
- apply eqmod_divides with modulus. apply eqm_signed_unsigned.
- exists (two_p (zwordsize - n)).
+ intros. apply eqmod_trans with (unsigned (sign_ext n x)).
+ apply eqmod_divides with modulus. apply eqm_signed_unsigned.
+ exists (two_p (zwordsize - n)).
unfold modulus. rewrite two_power_nat_two_p. fold zwordsize.
- rewrite <- two_p_is_exp. f_equal. omega. omega. omega.
+ rewrite <- two_p_is_exp. f_equal. omega. omega. omega.
apply eqmod_sign_ext'; auto.
Qed.
@@ -3374,8 +3374,8 @@ Theorem one_bits_range:
forall x i, In i (one_bits x) -> ltu i iwordsize = true.
Proof.
assert (A: forall p, 0 <= p < zwordsize -> ltu (repr p) iwordsize = true).
- intros. unfold ltu, iwordsize. apply zlt_true.
- repeat rewrite unsigned_repr. tauto.
+ intros. unfold ltu, iwordsize. apply zlt_true.
+ repeat rewrite unsigned_repr. tauto.
generalize wordsize_max_unsigned; omega.
generalize wordsize_max_unsigned; omega.
intros. unfold one_bits in H.
@@ -3392,21 +3392,21 @@ Fixpoint int_of_one_bits (l: list int) : int :=
Theorem one_bits_decomp:
forall x, x = int_of_one_bits (one_bits x).
Proof.
- intros.
+ intros.
transitivity (repr (powerserie (Z_one_bits wordsize (unsigned x) 0))).
transitivity (repr (unsigned x)).
auto with ints. decEq. apply Z_one_bits_powerserie.
auto with ints.
- unfold one_bits.
+ unfold one_bits.
generalize (Z_one_bits_range (unsigned x)).
generalize (Z_one_bits wordsize (unsigned x) 0).
induction l.
intros; reflexivity.
intros; simpl. rewrite <- IHl. unfold add. apply eqm_samerepr.
- apply eqm_add. rewrite shl_mul_two_p. rewrite mul_commut.
- rewrite mul_one. apply eqm_unsigned_repr_r.
+ apply eqm_add. rewrite shl_mul_two_p. rewrite mul_commut.
+ rewrite mul_one. apply eqm_unsigned_repr_r.
rewrite unsigned_repr. auto with ints.
- generalize (H a (in_eq _ _)). generalize wordsize_max_unsigned. omega.
+ generalize (H a (in_eq _ _)). generalize wordsize_max_unsigned. omega.
auto with ints.
intros; apply H; auto with coqlib.
Qed.
@@ -3443,7 +3443,7 @@ Lemma translate_eq:
Proof.
intros. unfold eq. case (zeq (unsigned x) (unsigned y)); intro.
unfold add. rewrite e. apply zeq_true.
- apply zeq_false. unfold add. red; intro. apply n.
+ apply zeq_false. unfold add. red; intro. apply n.
apply eqm_small_eq; auto with ints.
replace (unsigned x) with ((unsigned x + unsigned d) - unsigned d).
replace (unsigned y) with ((unsigned y + unsigned d) - unsigned d).
@@ -3473,7 +3473,7 @@ Theorem translate_cmpu:
Proof.
intros. unfold cmpu.
rewrite translate_eq. repeat rewrite translate_ltu; auto.
-Qed.
+Qed.
Lemma translate_lt:
forall x y d,
@@ -3495,13 +3495,13 @@ Theorem translate_cmp:
Proof.
intros. unfold cmp.
rewrite translate_eq. repeat rewrite translate_lt; auto.
-Qed.
+Qed.
Theorem notbool_isfalse_istrue:
forall x, is_false x -> is_true (notbool x).
Proof.
- unfold is_false, is_true, notbool; intros; subst x.
- rewrite eq_true. apply one_not_zero.
+ unfold is_false, is_true, notbool; intros; subst x.
+ rewrite eq_true. apply one_not_zero.
Qed.
Theorem notbool_istrue_isfalse:
@@ -3527,7 +3527,7 @@ Theorem lt_sub_overflow:
forall x y,
xor (sub_overflow x y zero) (negative (sub x y)) = if lt x y then one else zero.
Proof.
- intros. unfold negative, sub_overflow, lt. rewrite sub_signed.
+ intros. unfold negative, sub_overflow, lt. rewrite sub_signed.
rewrite signed_zero. rewrite Zminus_0_r.
generalize (signed_range x) (signed_range y).
set (X := signed x); set (Y := signed y). intros RX RY.
@@ -3540,19 +3540,19 @@ Proof.
+ unfold proj_sumbool; rewrite zle_true by omega.
rewrite signed_repr. rewrite zlt_false by omega. apply xor_idem.
unfold min_signed, max_signed; omega.
- + unfold proj_sumbool; rewrite zle_false by omega.
+ + unfold proj_sumbool; rewrite zle_false by omega.
replace (signed (repr (X - Y))) with (X - Y - modulus).
- rewrite zlt_true by omega. apply xor_idem.
- rewrite signed_repr_eq. replace ((X - Y) mod modulus) with (X - Y).
- rewrite zlt_false; auto.
+ rewrite zlt_true by omega. apply xor_idem.
+ rewrite signed_repr_eq. replace ((X - Y) mod modulus) with (X - Y).
+ rewrite zlt_false; auto.
symmetry. apply Zmod_unique with 0; omega.
- unfold proj_sumbool at 2. rewrite zle_true at 1 by omega. rewrite andb_true_r.
rewrite (zlt_true _ X) by omega.
destruct (zlt (X - Y) (-half_modulus)).
- + unfold proj_sumbool; rewrite zle_false by omega.
+ + unfold proj_sumbool; rewrite zle_false by omega.
replace (signed (repr (X - Y))) with (X - Y + modulus).
rewrite zlt_false by omega. apply xor_zero.
- rewrite signed_repr_eq. replace ((X - Y) mod modulus) with (X - Y + modulus).
+ rewrite signed_repr_eq. replace ((X - Y) mod modulus) with (X - Y + modulus).
rewrite zlt_true by omega; auto.
symmetry. apply Zmod_unique with (-1); omega.
+ unfold proj_sumbool; rewrite zle_true by omega.
@@ -3573,8 +3573,8 @@ Lemma no_overlap_sound:
unsigned (add base ofs1) + sz1 <= unsigned (add base ofs2)
\/ unsigned (add base ofs2) + sz2 <= unsigned (add base ofs1).
Proof.
- intros.
- destruct (andb_prop _ _ H1). clear H1.
+ intros.
+ destruct (andb_prop _ _ H1). clear H1.
destruct (andb_prop _ _ H2). clear H2.
exploit proj_sumbool_true. eexact H1. intro A; clear H1.
exploit proj_sumbool_true. eexact H4. intro B; clear H4.
@@ -3610,18 +3610,18 @@ Lemma Zsize_shiftin:
forall b x, 0 < x -> Zsize (Zshiftin b x) = Zsucc (Zsize x).
Proof.
intros. destruct x; compute in H; try discriminate.
- destruct b.
+ destruct b.
change (Zshiftin true (Zpos p)) with (Zpos (p~1)).
- simpl. f_equal. rewrite Pos.add_1_r; auto.
+ simpl. f_equal. rewrite Pos.add_1_r; auto.
change (Zshiftin false (Zpos p)) with (Zpos (p~0)).
- simpl. f_equal. rewrite Pos.add_1_r; auto.
+ simpl. f_equal. rewrite Pos.add_1_r; auto.
Qed.
Lemma Ztestbit_size_1:
forall x, 0 < x -> Z.testbit x (Zpred (Zsize x)) = true.
Proof.
intros x0 POS0; pattern x0; apply Zshiftin_pos_ind; auto.
- intros. rewrite Zsize_shiftin; auto.
+ intros. rewrite Zsize_shiftin; auto.
replace (Z.pred (Z.succ (Zsize x))) with (Z.succ (Z.pred (Zsize x))) by omega.
rewrite Ztestbit_shiftin_succ. auto. generalize (Zsize_pos' x H); omega.
Qed.
@@ -3629,14 +3629,14 @@ Qed.
Lemma Ztestbit_size_2:
forall x, 0 <= x -> forall i, i >= Zsize x -> Z.testbit x i = false.
Proof.
- intros x0 POS0. destruct (zeq x0 0).
- - subst x0; intros. apply Ztestbit_0.
+ intros x0 POS0. destruct (zeq x0 0).
+ - subst x0; intros. apply Ztestbit_0.
- pattern x0; apply Zshiftin_pos_ind.
- + simpl. intros. change 1 with (Zshiftin true 0). rewrite Ztestbit_shiftin.
+ + simpl. intros. change 1 with (Zshiftin true 0). rewrite Ztestbit_shiftin.
rewrite zeq_false. apply Ztestbit_0. omega. omega.
+ intros. rewrite Zsize_shiftin in H1; auto.
generalize (Zsize_pos' _ H); intros.
- rewrite Ztestbit_shiftin. rewrite zeq_false. apply H0. omega.
+ rewrite Ztestbit_shiftin. rewrite zeq_false. apply H0. omega.
omega. omega.
+ omega.
Qed.
@@ -3644,25 +3644,25 @@ Qed.
Lemma Zsize_interval_1:
forall x, 0 <= x -> 0 <= x < two_p (Zsize x).
Proof.
- intros.
+ intros.
assert (x = x mod (two_p (Zsize x))).
apply equal_same_bits; intros.
- rewrite Ztestbit_mod_two_p; auto.
- destruct (zlt i (Zsize x)). auto. apply Ztestbit_size_2; auto.
- apply Zsize_pos; auto.
- rewrite H0 at 1. rewrite H0 at 3. apply Z_mod_lt. apply two_p_gt_ZERO. apply Zsize_pos; auto.
+ rewrite Ztestbit_mod_two_p; auto.
+ destruct (zlt i (Zsize x)). auto. apply Ztestbit_size_2; auto.
+ apply Zsize_pos; auto.
+ rewrite H0 at 1. rewrite H0 at 3. apply Z_mod_lt. apply two_p_gt_ZERO. apply Zsize_pos; auto.
Qed.
Lemma Zsize_interval_2:
forall x n, 0 <= n -> 0 <= x < two_p n -> n >= Zsize x.
Proof.
- intros. set (N := Z.to_nat n).
+ intros. set (N := Z.to_nat n).
assert (Z.of_nat N = n) by (apply Z2Nat.id; auto).
- rewrite <- H1 in H0. rewrite <- two_power_nat_two_p in H0.
- destruct (zeq x 0).
- subst x; simpl; omega.
+ rewrite <- H1 in H0. rewrite <- two_power_nat_two_p in H0.
+ destruct (zeq x 0).
+ subst x; simpl; omega.
destruct (zlt n (Zsize x)); auto.
- exploit (Ztestbit_above N x (Zpred (Zsize x))). auto. omega.
+ exploit (Ztestbit_above N x (Zpred (Zsize x))). auto. omega.
rewrite Ztestbit_size_1. congruence. omega.
Qed.
@@ -3670,15 +3670,15 @@ Lemma Zsize_monotone:
forall x y, 0 <= x <= y -> Zsize x <= Zsize y.
Proof.
intros. apply Zge_le. apply Zsize_interval_2. apply Zsize_pos.
- exploit (Zsize_interval_1 y). omega.
- omega.
+ exploit (Zsize_interval_1 y). omega.
+ omega.
Qed.
Theorem size_zero: size zero = 0.
Proof.
unfold size; rewrite unsigned_zero; auto.
Qed.
-
+
Theorem bits_size_1:
forall x, x = zero \/ testbit x (Zpred (size x)) = true.
Proof.
@@ -3690,8 +3690,8 @@ Qed.
Theorem bits_size_2:
forall x i, size x <= i -> testbit x i = false.
Proof.
- intros. apply Ztestbit_size_2. generalize (unsigned_range x); omega.
- fold (size x); omega.
+ intros. apply Ztestbit_size_2. generalize (unsigned_range x); omega.
+ fold (size x); omega.
Qed.
Theorem size_range:
@@ -3700,9 +3700,9 @@ Proof.
intros; split. apply Zsize_pos.
destruct (bits_size_1 x).
subst x; unfold size; rewrite unsigned_zero; simpl. generalize wordsize_pos; omega.
- destruct (zle (size x) zwordsize); auto.
+ destruct (zle (size x) zwordsize); auto.
rewrite bits_above in H. congruence. omega.
-Qed.
+Qed.
Theorem bits_size_3:
forall x n,
@@ -3710,10 +3710,10 @@ Theorem bits_size_3:
(forall i, n <= i < zwordsize -> testbit x i = false) ->
size x <= n.
Proof.
- intros. destruct (zle (size x) n). auto.
- destruct (bits_size_1 x).
+ intros. destruct (zle (size x) n). auto.
+ destruct (bits_size_1 x).
subst x. unfold size; rewrite unsigned_zero; assumption.
- rewrite (H0 (Z.pred (size x))) in H1. congruence.
+ rewrite (H0 (Z.pred (size x))) in H1. congruence.
generalize (size_range x); omega.
Qed.
@@ -3728,7 +3728,7 @@ Proof.
assert (size x <= n).
apply bits_size_3; auto.
destruct (zlt (size x) n).
- rewrite bits_size_2 in H0. congruence. omega.
+ rewrite bits_size_2 in H0. congruence. omega.
omega.
Qed.
@@ -3747,24 +3747,24 @@ Qed.
Theorem size_and:
forall a b, size (and a b) <= Z.min (size a) (size b).
Proof.
- intros.
+ intros.
assert (0 <= Z.min (size a) (size b)).
- generalize (size_range a) (size_range b). zify; omega.
+ generalize (size_range a) (size_range b). zify; omega.
apply bits_size_3. auto. intros.
- rewrite bits_and. zify. subst z z0. destruct H1.
- rewrite (bits_size_2 a). auto. omega.
- rewrite (bits_size_2 b). apply andb_false_r. omega.
+ rewrite bits_and. zify. subst z z0. destruct H1.
+ rewrite (bits_size_2 a). auto. omega.
+ rewrite (bits_size_2 b). apply andb_false_r. omega.
omega.
Qed.
Corollary and_interval:
forall a b, 0 <= unsigned (and a b) < two_p (Z.min (size a) (size b)).
Proof.
- intros.
- generalize (size_interval_1 (and a b)); intros.
+ intros.
+ generalize (size_interval_1 (and a b)); intros.
assert (two_p (size (and a b)) <= two_p (Z.min (size a) (size b))).
- apply two_p_monotone. split. generalize (size_range (and a b)); omega.
- apply size_and.
+ apply two_p_monotone. split. generalize (size_range (and a b)); omega.
+ apply size_and.
omega.
Qed.
@@ -3776,42 +3776,42 @@ Proof.
subst a. rewrite size_zero. rewrite or_zero_l. zify; omega.
destruct (bits_size_1 b).
subst b. rewrite size_zero. rewrite or_zero. zify; omega.
- zify. destruct H3 as [[P Q] | [P Q]]; subst.
- apply bits_size_4. tauto. rewrite bits_or. rewrite H2. apply orb_true_r.
- omega.
- intros. rewrite bits_or. rewrite !bits_size_2. auto. omega. omega. omega.
- apply bits_size_4. tauto. rewrite bits_or. rewrite H1. apply orb_true_l.
+ zify. destruct H3 as [[P Q] | [P Q]]; subst.
+ apply bits_size_4. tauto. rewrite bits_or. rewrite H2. apply orb_true_r.
+ omega.
+ intros. rewrite bits_or. rewrite !bits_size_2. auto. omega. omega. omega.
+ apply bits_size_4. tauto. rewrite bits_or. rewrite H1. apply orb_true_l.
destruct (zeq (size a) 0). unfold testbit in H1. rewrite Z.testbit_neg_r in H1.
- congruence. omega. omega.
- intros. rewrite bits_or. rewrite !bits_size_2. auto. omega. omega. omega.
+ congruence. omega. omega.
+ intros. rewrite bits_or. rewrite !bits_size_2. auto. omega. omega. omega.
Qed.
Corollary or_interval:
forall a b, 0 <= unsigned (or a b) < two_p (Z.max (size a) (size b)).
Proof.
- intros. rewrite <- size_or. apply size_interval_1.
+ intros. rewrite <- size_or. apply size_interval_1.
Qed.
Theorem size_xor:
forall a b, size (xor a b) <= Z.max (size a) (size b).
Proof.
- intros.
+ intros.
assert (0 <= Z.max (size a) (size b)).
- generalize (size_range a) (size_range b). zify; omega.
+ generalize (size_range a) (size_range b). zify; omega.
apply bits_size_3. auto. intros.
- rewrite bits_xor. rewrite !bits_size_2. auto.
+ rewrite bits_xor. rewrite !bits_size_2. auto.
+ zify; omega.
zify; omega.
- zify; omega.
- omega.
+ omega.
Qed.
Corollary xor_interval:
forall a b, 0 <= unsigned (xor a b) < two_p (Z.max (size a) (size b)).
Proof.
- intros.
- generalize (size_interval_1 (xor a b)); intros.
+ intros.
+ generalize (size_interval_1 (xor a b)); intros.
assert (two_p (size (xor a b)) <= two_p (Z.max (size a) (size b))).
- apply two_p_monotone. split. generalize (size_range (xor a b)); omega.
+ apply two_p_monotone. split. generalize (size_range (xor a b)); omega.
apply size_xor.
omega.
Qed.
@@ -3837,7 +3837,7 @@ Notation int := Int.int.
Remark int_wordsize_divides_modulus:
Zdivide (Z_of_nat Int.wordsize) Int.modulus.
Proof.
- exists (two_p (32-5)); reflexivity.
+ exists (two_p (32-5)); reflexivity.
Qed.
Module Wordsize_8.
@@ -3883,8 +3883,8 @@ Lemma bits_shl':
testbit (shl' x y) i =
if zlt i (Int.unsigned y) then false else testbit x (i - Int.unsigned y).
Proof.
- intros. unfold shl'. rewrite testbit_repr; auto.
- destruct (zlt i (Int.unsigned y)).
+ intros. unfold shl'. rewrite testbit_repr; auto.
+ destruct (zlt i (Int.unsigned y)).
apply Z.shiftl_spec_low. auto.
apply Z.shiftl_spec_high. omega. omega.
Qed.
@@ -3895,11 +3895,11 @@ Lemma bits_shru':
testbit (shru' x y) i =
if zlt (i + Int.unsigned y) zwordsize then testbit x (i + Int.unsigned y) else false.
Proof.
- intros. unfold shru'. rewrite testbit_repr; auto.
+ intros. unfold shru'. rewrite testbit_repr; auto.
rewrite Z.shiftr_spec. fold (testbit x (i + Int.unsigned y)).
destruct (zlt (i + Int.unsigned y) zwordsize).
auto.
- apply bits_above; auto.
+ apply bits_above; auto.
omega.
Qed.
@@ -3909,8 +3909,8 @@ Lemma bits_shr':
testbit (shr' x y) i =
testbit x (if zlt (i + Int.unsigned y) zwordsize then i + Int.unsigned y else zwordsize - 1).
Proof.
- intros. unfold shr'. rewrite testbit_repr; auto.
- rewrite Z.shiftr_spec. apply bits_signed.
+ intros. unfold shr'. rewrite testbit_repr; auto.
+ rewrite Z.shiftr_spec. apply bits_signed.
generalize (Int.unsigned_range y); omega.
omega.
Qed.
@@ -3927,7 +3927,7 @@ Definition ofwords (hi lo: Int.int) : int :=
Lemma bits_loword:
forall n i, 0 <= i < Int.zwordsize -> Int.testbit (loword n) i = testbit n i.
Proof.
- intros. unfold loword. rewrite Int.testbit_repr; auto.
+ intros. unfold loword. rewrite Int.testbit_repr; auto.
Qed.
Lemma bits_hiword:
@@ -3937,7 +3937,7 @@ Proof.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
fold (testbit (shru n (repr Int.zwordsize)) i). rewrite bits_shru.
change (unsigned (repr Int.zwordsize)) with Int.zwordsize.
- apply zlt_true. omega. omega.
+ apply zlt_true. omega. omega.
Qed.
Lemma bits_ofwords:
@@ -3945,53 +3945,53 @@ Lemma bits_ofwords:
testbit (ofwords hi lo) i =
if zlt i Int.zwordsize then Int.testbit lo i else Int.testbit hi (i - Int.zwordsize).
Proof.
- intros. unfold ofwords. rewrite bits_or; auto. rewrite bits_shl; auto.
- change (unsigned (repr Int.zwordsize)) with Int.zwordsize.
+ intros. unfold ofwords. rewrite bits_or; auto. rewrite bits_shl; auto.
+ change (unsigned (repr Int.zwordsize)) with Int.zwordsize.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
destruct (zlt i Int.zwordsize).
- rewrite testbit_repr; auto.
+ rewrite testbit_repr; auto.
rewrite !testbit_repr; auto.
- fold (Int.testbit lo i). rewrite Int.bits_above. apply orb_false_r. auto.
+ fold (Int.testbit lo i). rewrite Int.bits_above. apply orb_false_r. auto.
omega.
Qed.
Lemma lo_ofwords:
forall hi lo, loword (ofwords hi lo) = lo.
Proof.
- intros. apply Int.same_bits_eq; intros.
- rewrite bits_loword; auto. rewrite bits_ofwords. apply zlt_true. omega.
+ intros. apply Int.same_bits_eq; intros.
+ rewrite bits_loword; auto. rewrite bits_ofwords. apply zlt_true. omega.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity. omega.
Qed.
Lemma hi_ofwords:
forall hi lo, hiword (ofwords hi lo) = hi.
Proof.
- intros. apply Int.same_bits_eq; intros.
+ intros. apply Int.same_bits_eq; intros.
rewrite bits_hiword; auto. rewrite bits_ofwords.
- rewrite zlt_false. f_equal. omega. omega.
+ rewrite zlt_false. f_equal. omega. omega.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity. omega.
Qed.
Lemma ofwords_recompose:
forall n, ofwords (hiword n) (loword n) = n.
Proof.
- intros. apply same_bits_eq; intros. rewrite bits_ofwords; auto.
- destruct (zlt i Int.zwordsize).
- apply bits_loword. omega.
- rewrite bits_hiword. f_equal. omega.
+ intros. apply same_bits_eq; intros. rewrite bits_ofwords; auto.
+ destruct (zlt i Int.zwordsize).
+ apply bits_loword. omega.
+ rewrite bits_hiword. f_equal. omega.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity. omega.
Qed.
Lemma ofwords_add:
forall lo hi, ofwords hi lo = repr (Int.unsigned hi * two_p 32 + Int.unsigned lo).
Proof.
- intros. unfold ofwords. rewrite shifted_or_is_add.
- apply eqm_samerepr. apply eqm_add. apply eqm_mult.
+ intros. unfold ofwords. rewrite shifted_or_is_add.
+ apply eqm_samerepr. apply eqm_add. apply eqm_mult.
apply eqm_sym; apply eqm_unsigned_repr.
- apply eqm_refl.
+ apply eqm_refl.
apply eqm_sym; apply eqm_unsigned_repr.
change Int.zwordsize with 32; change zwordsize with 64; omega.
- rewrite unsigned_repr. generalize (Int.unsigned_range lo). intros [A B]. exact B.
+ rewrite unsigned_repr. generalize (Int.unsigned_range lo). intros [A B]. exact B.
assert (Int.max_unsigned < max_unsigned) by (compute; auto).
generalize (Int.unsigned_range_2 lo); omega.
Qed.
@@ -4000,7 +4000,7 @@ Lemma ofwords_add':
forall lo hi, unsigned (ofwords hi lo) = Int.unsigned hi * two_p 32 + Int.unsigned lo.
Proof.
intros. rewrite ofwords_add. apply unsigned_repr.
- generalize (Int.unsigned_range hi) (Int.unsigned_range lo).
+ generalize (Int.unsigned_range hi) (Int.unsigned_range lo).
change (two_p 32) with Int.modulus.
change Int.modulus with 4294967296.
change max_unsigned with 18446744073709551615.
@@ -4011,7 +4011,7 @@ Remark eqm_mul_2p32:
forall x y, Int.eqm x y -> eqm (x * two_p 32) (y * two_p 32).
Proof.
intros. destruct H as [k EQ]. exists k. rewrite EQ.
- change Int.modulus with (two_p 32).
+ change Int.modulus with (two_p 32).
change modulus with (two_p 32 * two_p 32).
ring.
Qed.
@@ -4023,7 +4023,7 @@ Proof.
replace (repr (Int.unsigned hi * two_p 32 + Int.unsigned lo))
with (repr (Int.signed hi * two_p 32 + Int.unsigned lo)).
apply signed_repr.
- generalize (Int.signed_range hi) (Int.unsigned_range lo).
+ generalize (Int.signed_range hi) (Int.unsigned_range lo).
change (two_p 32) with Int.modulus.
change min_signed with (Int.min_signed * Int.modulus).
change max_signed with (Int.max_signed * Int.modulus + Int.modulus - 1).
@@ -4074,7 +4074,7 @@ Lemma decompose_not:
forall xh xl,
not (ofwords xh xl) = ofwords (Int.not xh) (Int.not xl).
Proof.
- intros. unfold not, Int.not. rewrite <- decompose_xor. f_equal.
+ intros. unfold not, Int.not. rewrite <- decompose_xor. f_equal.
apply (Int64.eq_spec mone (ofwords Int.mone Int.mone)).
Qed.
@@ -4087,21 +4087,21 @@ Lemma decompose_shl_1:
Proof.
intros.
assert (Int.unsigned (Int.sub Int.iwordsize y) = Int.zwordsize - Int.unsigned y).
- { unfold Int.sub. rewrite Int.unsigned_repr. auto.
+ { unfold Int.sub. rewrite Int.unsigned_repr. auto.
rewrite Int.unsigned_repr_wordsize. generalize Int.wordsize_max_unsigned; omega. }
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
apply Int64.same_bits_eq; intros.
rewrite bits_shl' by auto. symmetry. rewrite bits_ofwords by auto.
- destruct (zlt i Int.zwordsize). rewrite Int.bits_shl by omega.
- destruct (zlt i (Int.unsigned y)). auto.
+ destruct (zlt i Int.zwordsize). rewrite Int.bits_shl by omega.
+ destruct (zlt i (Int.unsigned y)). auto.
rewrite bits_ofwords by omega. rewrite zlt_true by omega. auto.
- rewrite zlt_false by omega. rewrite bits_ofwords by omega.
- rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega.
+ rewrite zlt_false by omega. rewrite bits_ofwords by omega.
+ rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega.
rewrite Int.bits_shru by omega. rewrite H0.
destruct (zlt (i - Int.unsigned y) (Int.zwordsize)).
rewrite zlt_true by omega. rewrite zlt_true by omega.
- rewrite orb_false_l. f_equal. omega.
- rewrite zlt_false by omega. rewrite zlt_false by omega.
+ rewrite orb_false_l. f_equal. omega.
+ rewrite zlt_false by omega. rewrite zlt_false by omega.
rewrite orb_false_r. f_equal. omega.
Qed.
@@ -4114,16 +4114,16 @@ Proof.
intros.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
assert (Int.unsigned (Int.sub y Int.iwordsize) = Int.unsigned y - Int.zwordsize).
- { unfold Int.sub. rewrite Int.unsigned_repr. auto.
+ { unfold Int.sub. rewrite Int.unsigned_repr. auto.
rewrite Int.unsigned_repr_wordsize. generalize (Int.unsigned_range_2 y). omega. }
apply Int64.same_bits_eq; intros.
rewrite bits_shl' by auto. symmetry. rewrite bits_ofwords by auto.
destruct (zlt i Int.zwordsize). rewrite zlt_true by omega. apply Int.bits_zero.
- rewrite Int.bits_shl by omega.
+ rewrite Int.bits_shl by omega.
destruct (zlt i (Int.unsigned y)).
- rewrite zlt_true by omega. auto.
- rewrite zlt_false by omega.
- rewrite bits_ofwords by omega. rewrite zlt_true by omega. f_equal. omega.
+ rewrite zlt_true by omega. auto.
+ rewrite zlt_false by omega.
+ rewrite bits_ofwords by omega. rewrite zlt_true by omega. f_equal. omega.
Qed.
Lemma decompose_shru_1:
@@ -4135,25 +4135,25 @@ Lemma decompose_shru_1:
Proof.
intros.
assert (Int.unsigned (Int.sub Int.iwordsize y) = Int.zwordsize - Int.unsigned y).
- { unfold Int.sub. rewrite Int.unsigned_repr. auto.
+ { unfold Int.sub. rewrite Int.unsigned_repr. auto.
rewrite Int.unsigned_repr_wordsize. generalize Int.wordsize_max_unsigned; omega. }
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
apply Int64.same_bits_eq; intros.
rewrite bits_shru' by auto. symmetry. rewrite bits_ofwords by auto.
destruct (zlt i Int.zwordsize).
- rewrite zlt_true by omega.
+ rewrite zlt_true by omega.
rewrite bits_ofwords by omega.
- rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega.
+ rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega.
rewrite Int.bits_shru by omega. rewrite H0.
destruct (zlt (i + Int.unsigned y) (Int.zwordsize)).
rewrite zlt_true by omega.
rewrite orb_false_r. auto.
- rewrite zlt_false by omega.
+ rewrite zlt_false by omega.
rewrite orb_false_l. f_equal. omega.
- rewrite Int.bits_shru by omega.
+ rewrite Int.bits_shru by omega.
destruct (zlt (i + Int.unsigned y) zwordsize).
- rewrite bits_ofwords by omega.
- rewrite zlt_true by omega. rewrite zlt_false by omega. f_equal. omega.
+ rewrite bits_ofwords by omega.
+ rewrite zlt_true by omega. rewrite zlt_false by omega. f_equal. omega.
rewrite zlt_false by omega. auto.
Qed.
@@ -4166,15 +4166,15 @@ Proof.
intros.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
assert (Int.unsigned (Int.sub y Int.iwordsize) = Int.unsigned y - Int.zwordsize).
- { unfold Int.sub. rewrite Int.unsigned_repr. auto.
+ { unfold Int.sub. rewrite Int.unsigned_repr. auto.
rewrite Int.unsigned_repr_wordsize. generalize (Int.unsigned_range_2 y). omega. }
apply Int64.same_bits_eq; intros.
rewrite bits_shru' by auto. symmetry. rewrite bits_ofwords by auto.
destruct (zlt i Int.zwordsize).
rewrite Int.bits_shru by omega. rewrite H1.
destruct (zlt (i + Int.unsigned y) zwordsize).
- rewrite zlt_true by omega. rewrite bits_ofwords by omega.
- rewrite zlt_false by omega. f_equal; omega.
+ rewrite zlt_true by omega. rewrite bits_ofwords by omega.
+ rewrite zlt_false by omega. f_equal; omega.
rewrite zlt_false by omega. auto.
rewrite zlt_false by omega. apply Int.bits_zero.
Qed.
@@ -4188,26 +4188,26 @@ Lemma decompose_shr_1:
Proof.
intros.
assert (Int.unsigned (Int.sub Int.iwordsize y) = Int.zwordsize - Int.unsigned y).
- { unfold Int.sub. rewrite Int.unsigned_repr. auto.
+ { unfold Int.sub. rewrite Int.unsigned_repr. auto.
rewrite Int.unsigned_repr_wordsize. generalize Int.wordsize_max_unsigned; omega. }
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
apply Int64.same_bits_eq; intros.
rewrite bits_shr' by auto. symmetry. rewrite bits_ofwords by auto.
destruct (zlt i Int.zwordsize).
- rewrite zlt_true by omega.
+ rewrite zlt_true by omega.
rewrite bits_ofwords by omega.
- rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega.
+ rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega.
rewrite Int.bits_shru by omega. rewrite H0.
destruct (zlt (i + Int.unsigned y) (Int.zwordsize)).
rewrite zlt_true by omega.
rewrite orb_false_r. auto.
- rewrite zlt_false by omega.
+ rewrite zlt_false by omega.
rewrite orb_false_l. f_equal. omega.
- rewrite Int.bits_shr by omega.
+ rewrite Int.bits_shr by omega.
destruct (zlt (i + Int.unsigned y) zwordsize).
- rewrite bits_ofwords by omega.
- rewrite zlt_true by omega. rewrite zlt_false by omega. f_equal. omega.
- rewrite zlt_false by omega. rewrite bits_ofwords by omega.
+ rewrite bits_ofwords by omega.
+ rewrite zlt_true by omega. rewrite zlt_false by omega. f_equal. omega.
+ rewrite zlt_false by omega. rewrite bits_ofwords by omega.
rewrite zlt_false by omega. f_equal.
Qed.
@@ -4221,24 +4221,24 @@ Proof.
intros.
assert (zwordsize = 2 * Int.zwordsize) by reflexivity.
assert (Int.unsigned (Int.sub y Int.iwordsize) = Int.unsigned y - Int.zwordsize).
- { unfold Int.sub. rewrite Int.unsigned_repr. auto.
+ { unfold Int.sub. rewrite Int.unsigned_repr. auto.
rewrite Int.unsigned_repr_wordsize. generalize (Int.unsigned_range_2 y). omega. }
apply Int64.same_bits_eq; intros.
rewrite bits_shr' by auto. symmetry. rewrite bits_ofwords by auto.
destruct (zlt i Int.zwordsize).
rewrite Int.bits_shr by omega. rewrite H1.
destruct (zlt (i + Int.unsigned y) zwordsize).
- rewrite zlt_true by omega. rewrite bits_ofwords by omega.
- rewrite zlt_false by omega. f_equal; omega.
- rewrite zlt_false by omega. rewrite bits_ofwords by omega.
+ rewrite zlt_true by omega. rewrite bits_ofwords by omega.
+ rewrite zlt_false by omega. f_equal; omega.
+ rewrite zlt_false by omega. rewrite bits_ofwords by omega.
rewrite zlt_false by omega. auto.
- rewrite Int.bits_shr by omega.
+ rewrite Int.bits_shr by omega.
change (Int.unsigned (Int.sub Int.iwordsize Int.one)) with (Int.zwordsize - 1).
destruct (zlt (i + Int.unsigned y) zwordsize);
rewrite bits_ofwords by omega.
- symmetry. rewrite zlt_false by omega. f_equal.
+ symmetry. rewrite zlt_false by omega. f_equal.
destruct (zlt (i - Int.zwordsize + (Int.zwordsize - 1)) Int.zwordsize); omega.
- symmetry. rewrite zlt_false by omega. f_equal.
+ symmetry. rewrite zlt_false by omega. f_equal.
destruct (zlt (i - Int.zwordsize + (Int.zwordsize - 1)) Int.zwordsize); omega.
Qed.
@@ -4249,8 +4249,8 @@ Lemma decompose_add:
(Int.add xl yl).
Proof.
intros. symmetry. rewrite ofwords_add. rewrite add_unsigned.
- apply eqm_samerepr.
- rewrite ! ofwords_add'. rewrite (Int.unsigned_add_carry xl yl).
+ apply eqm_samerepr.
+ rewrite ! ofwords_add'. rewrite (Int.unsigned_add_carry xl yl).
set (cc := Int.add_carry xl yl Int.zero).
set (Xl := Int.unsigned xl); set (Xh := Int.unsigned xh);
set (Yl := Int.unsigned yl); set (Yh := Int.unsigned yh).
@@ -4264,8 +4264,8 @@ Proof.
apply eqm_add. 2: apply eqm_refl. apply eqm_mul_2p32.
replace (Xh + Yh) with ((Xh + Yh + Int.unsigned cc) - Int.unsigned cc) by ring.
apply Int.eqm_sub. 2: apply Int.eqm_refl.
- apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. 2: apply Int.eqm_refl.
- apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl.
+ apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. 2: apply Int.eqm_refl.
+ apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl.
Qed.
Lemma decompose_sub:
@@ -4275,8 +4275,8 @@ Lemma decompose_sub:
(Int.sub xl yl).
Proof.
intros. symmetry. rewrite ofwords_add.
- apply eqm_samerepr.
- rewrite ! ofwords_add'. rewrite (Int.unsigned_sub_borrow xl yl).
+ apply eqm_samerepr.
+ rewrite ! ofwords_add'. rewrite (Int.unsigned_sub_borrow xl yl).
set (bb := Int.sub_borrow xl yl Int.zero).
set (Xl := Int.unsigned xl); set (Xh := Int.unsigned xh);
set (Yl := Int.unsigned yl); set (Yh := Int.unsigned yh).
@@ -4290,8 +4290,8 @@ Proof.
apply eqm_add. 2: apply eqm_refl. apply eqm_mul_2p32.
replace (Xh - Yh) with ((Xh - Yh - Int.unsigned bb) + Int.unsigned bb) by ring.
apply Int.eqm_add. 2: apply Int.eqm_refl.
- apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. 2: apply Int.eqm_refl.
- apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl.
+ apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. 2: apply Int.eqm_refl.
+ apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl.
Qed.
Lemma decompose_sub':
@@ -4300,9 +4300,9 @@ Lemma decompose_sub':
ofwords (Int.add (Int.add xh (Int.not yh)) (Int.add_carry xl (Int.not yl) Int.one))
(Int.sub xl yl).
Proof.
- intros. rewrite decompose_sub. f_equal.
+ intros. rewrite decompose_sub. f_equal.
rewrite Int.sub_borrow_add_carry by auto.
- rewrite Int.sub_add_not_3. rewrite Int.xor_assoc. rewrite Int.xor_idem.
+ rewrite Int.sub_add_not_3. rewrite Int.xor_assoc. rewrite Int.xor_idem.
rewrite Int.xor_zero. auto.
rewrite Int.xor_zero_l. unfold Int.add_carry.
destruct (zlt (Int.unsigned xl + Int.unsigned (Int.not yl) + Int.unsigned Int.one) Int.modulus);
@@ -4314,12 +4314,12 @@ Definition mul' (x y: Int.int) : int := repr (Int.unsigned x * Int.unsigned y).
Lemma mul'_mulhu:
forall x y, mul' x y = ofwords (Int.mulhu x y) (Int.mul x y).
Proof.
- intros.
- rewrite ofwords_add. unfold mul', Int.mulhu, Int.mul.
+ intros.
+ rewrite ofwords_add. unfold mul', Int.mulhu, Int.mul.
set (p := Int.unsigned x * Int.unsigned y).
- set (ph := p / Int.modulus). set (pl := p mod Int.modulus).
+ set (ph := p / Int.modulus). set (pl := p mod Int.modulus).
transitivity (repr (ph * Int.modulus + pl)).
-- f_equal. rewrite Zmult_comm. apply Z_div_mod_eq. apply Int.modulus_pos.
+- f_equal. rewrite Zmult_comm. apply Z_div_mod_eq. apply Int.modulus_pos.
- apply eqm_samerepr. apply eqm_add. apply eqm_mul_2p32. auto with ints.
rewrite Int.unsigned_repr_eq. apply eqm_refl.
Qed.
@@ -4330,7 +4330,7 @@ Lemma decompose_mul:
ofwords (Int.add (Int.add (hiword (mul' xl yl)) (Int.mul xl yh)) (Int.mul xh yl))
(loword (mul' xl yl)).
Proof.
- intros.
+ intros.
set (pl := loword (mul' xl yl)); set (ph := hiword (mul' xl yl)).
assert (EQ0: unsigned (mul' xl yl) = Int.unsigned ph * two_p 32 + Int.unsigned pl).
{ rewrite <- (ofwords_recompose (mul' xl yl)). apply ofwords_add'. }
@@ -4339,7 +4339,7 @@ Proof.
set (YL := Int.unsigned yl); set (YH := Int.unsigned yh).
set (PH := Int.unsigned ph) in *. set (PL := Int.unsigned pl) in *.
transitivity (repr (((PH + XL * YH) + XH * YL) * two_p 32 + PL)).
- apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl.
+ apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl.
apply eqm_mul_2p32.
rewrite Int.add_unsigned. apply Int.eqm_unsigned_repr_l. apply Int.eqm_add.
rewrite Int.add_unsigned. apply Int.eqm_unsigned_repr_l. apply Int.eqm_add.
@@ -4349,14 +4349,14 @@ Proof.
transitivity (repr (unsigned (mul' xl yl) + (XL * YH + XH * YL) * two_p 32)).
rewrite EQ0. f_equal. ring.
transitivity (repr ((XL * YL + (XL * YH + XH * YL) * two_p 32))).
- apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl.
+ apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl.
unfold mul'. apply eqm_unsigned_repr_l. apply eqm_refl.
transitivity (repr (0 + (XL * YL + (XL * YH + XH * YL) * two_p 32))).
- rewrite Zplus_0_l; auto.
+ rewrite Zplus_0_l; auto.
transitivity (repr (XH * YH * (two_p 32 * two_p 32) + (XL * YL + (XL * YH + XH * YL) * two_p 32))).
- apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl.
- change (two_p 32 * two_p 32) with modulus. exists (- XH * YH). ring.
- f_equal. ring.
+ apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl.
+ change (two_p 32 * two_p 32) with modulus. exists (- XH * YH). ring.
+ f_equal. ring.
Qed.
Lemma decompose_mul_2:
@@ -4365,7 +4365,7 @@ Lemma decompose_mul_2:
ofwords (Int.add (Int.add (Int.mulhu xl yl) (Int.mul xl yh)) (Int.mul xh yl))
(Int.mul xl yl).
Proof.
- intros. rewrite decompose_mul. rewrite mul'_mulhu.
+ intros. rewrite decompose_mul. rewrite mul'_mulhu.
rewrite hi_ofwords, lo_ofwords. auto.
Qed.
@@ -4375,11 +4375,11 @@ Lemma decompose_ltu:
Proof.
intros. unfold ltu. rewrite ! ofwords_add'. unfold Int.ltu, Int.eq.
destruct (zeq (Int.unsigned xh) (Int.unsigned yh)).
- rewrite e. destruct (zlt (Int.unsigned xl) (Int.unsigned yl)).
+ rewrite e. destruct (zlt (Int.unsigned xl) (Int.unsigned yl)).
apply zlt_true; omega.
apply zlt_false; omega.
- change (two_p 32) with Int.modulus.
- generalize (Int.unsigned_range xl) (Int.unsigned_range yl).
+ change (two_p 32) with Int.modulus.
+ generalize (Int.unsigned_range xl) (Int.unsigned_range yl).
change Int.modulus with 4294967296. intros.
destruct (zlt (Int.unsigned xh) (Int.unsigned yh)).
apply zlt_true; omega.
@@ -4392,9 +4392,9 @@ Lemma decompose_leu:
if Int.eq xh yh then negb (Int.ltu yl xl) else Int.ltu xh yh.
Proof.
intros. rewrite decompose_ltu. rewrite Int.eq_sym.
- unfold Int.eq. destruct (zeq (Int.unsigned xh) (Int.unsigned yh)).
+ unfold Int.eq. destruct (zeq (Int.unsigned xh) (Int.unsigned yh)).
auto.
- unfold Int.ltu. destruct (zlt (Int.unsigned xh) (Int.unsigned yh)).
+ unfold Int.ltu. destruct (zlt (Int.unsigned xh) (Int.unsigned yh)).
rewrite zlt_false by omega; auto.
rewrite zlt_true by omega; auto.
Qed.
@@ -4403,13 +4403,13 @@ Lemma decompose_lt:
forall xh xl yh yl,
lt (ofwords xh xl) (ofwords yh yl) = if Int.eq xh yh then Int.ltu xl yl else Int.lt xh yh.
Proof.
- intros. unfold lt. rewrite ! ofwords_add''. rewrite Int.eq_signed.
+ intros. unfold lt. rewrite ! ofwords_add''. rewrite Int.eq_signed.
destruct (zeq (Int.signed xh) (Int.signed yh)).
- rewrite e. unfold Int.ltu. destruct (zlt (Int.unsigned xl) (Int.unsigned yl)).
+ rewrite e. unfold Int.ltu. destruct (zlt (Int.unsigned xl) (Int.unsigned yl)).
apply zlt_true; omega.
apply zlt_false; omega.
- change (two_p 32) with Int.modulus.
- generalize (Int.unsigned_range xl) (Int.unsigned_range yl).
+ change (two_p 32) with Int.modulus.
+ generalize (Int.unsigned_range xl) (Int.unsigned_range yl).
change Int.modulus with 4294967296. intros.
unfold Int.lt. destruct (zlt (Int.signed xh) (Int.signed yh)).
apply zlt_true; omega.
@@ -4422,9 +4422,9 @@ Lemma decompose_le:
if Int.eq xh yh then negb (Int.ltu yl xl) else Int.lt xh yh.
Proof.
intros. rewrite decompose_lt. rewrite Int.eq_sym.
- rewrite Int.eq_signed. destruct (zeq (Int.signed xh) (Int.signed yh)).
+ rewrite Int.eq_signed. destruct (zeq (Int.signed xh) (Int.signed yh)).
auto.
- unfold Int.lt. destruct (zlt (Int.signed xh) (Int.signed yh)).
+ unfold Int.lt. destruct (zlt (Int.signed xh) (Int.signed yh)).
rewrite zlt_false by omega; auto.
rewrite zlt_true by omega; auto.
Qed.
diff --git a/lib/Intv.v b/lib/Intv.v
index a8fbd714..090ff408 100644
--- a/lib/Intv.v
+++ b/lib/Intv.v
@@ -30,18 +30,18 @@ Lemma In_dec:
forall x i, {In x i} + {~In x i}.
Proof.
unfold In; intros.
- case (zle (fst i) x); intros.
+ case (zle (fst i) x); intros.
case (zlt x (snd i)); intros.
left; auto.
- right; intuition.
+ right; intuition.
right; intuition.
Qed.
-Lemma notin_range:
+Lemma notin_range:
forall x i,
x < fst i \/ x >= snd i -> ~In x i.
Proof.
- unfold In; intros; omega.
+ unfold In; intros; omega.
Qed.
Lemma range_notin:
@@ -58,7 +58,7 @@ Definition empty (i: interv) : Prop := fst i >= snd i.
Lemma empty_dec:
forall i, {empty i} + {~empty i}.
Proof.
- unfold empty; intros.
+ unfold empty; intros.
case (zle (snd i) (fst i)); intros.
left; omega.
right; omega.
@@ -90,7 +90,7 @@ Definition disjoint (i j: interv) : Prop :=
Lemma disjoint_sym:
forall i j, disjoint i j -> disjoint j i.
Proof.
- unfold disjoint; intros; red; intros. elim (H x); auto.
+ unfold disjoint; intros; red; intros. elim (H x); auto.
Qed.
Lemma empty_disjoint_r:
@@ -102,7 +102,7 @@ Qed.
Lemma empty_disjoint_l:
forall i j, empty i -> disjoint i j.
Proof.
- intros. apply disjoint_sym. apply empty_disjoint_r; auto.
+ intros. apply disjoint_sym. apply empty_disjoint_r; auto.
Qed.
Lemma disjoint_range:
@@ -147,12 +147,12 @@ Qed.
Lemma disjoint_dec:
forall i j, {disjoint i j} + {~disjoint i j}.
Proof.
- intros.
+ intros.
destruct (empty_dec i). left; apply empty_disjoint_l; auto.
destruct (empty_dec j). left; apply empty_disjoint_r; auto.
destruct (zle (snd i) (fst j)). left; apply disjoint_range; auto.
destruct (zle (snd j) (fst i)). left; apply disjoint_range; auto.
- right; red; intro. exploit range_disjoint; eauto. intuition.
+ right; red; intro. exploit range_disjoint; eauto. intuition.
Qed.
(** * Shifting an interval by some amount *)
@@ -170,7 +170,7 @@ Lemma in_shift_inv:
forall x i delta,
In x (shift i delta) -> In (x - delta) i.
Proof.
- unfold shift, In; simpl; intros. omega.
+ unfold shift, In; simpl; intros. omega.
Qed.
(** * Enumerating the elements of an interval *)
@@ -182,7 +182,7 @@ Variable lo: Z.
Function elements_rec (hi: Z) {wf (Zwf lo) hi} : list Z :=
if zlt lo hi then (hi-1) :: elements_rec (hi-1) else nil.
Proof.
- intros. red. omega.
+ intros. red. omega.
apply Zwf_well_founded.
Qed.
@@ -190,11 +190,11 @@ Lemma In_elements_rec:
forall hi x,
List.In x (elements_rec hi) <-> lo <= x < hi.
Proof.
- intros. functional induction (elements_rec hi).
+ intros. functional induction (elements_rec hi).
simpl; split; intros.
destruct H. clear IHl. omega. rewrite IHl in H. clear IHl. omega.
destruct (zeq (hi - 1) x); auto. right. rewrite IHl. clear IHl. omega.
- simpl; intuition.
+ simpl; intuition.
Qed.
End ELEMENTS.
@@ -213,8 +213,8 @@ Lemma elements_in:
forall x i,
List.In x (elements i) -> In x i.
Proof.
- unfold elements; intros.
- rewrite In_elements_rec in H. auto.
+ unfold elements; intros.
+ rewrite In_elements_rec in H. auto.
Qed.
(** * Checking properties on all elements of an interval *)
@@ -241,11 +241,11 @@ Program Fixpoint forall_rec (hi: Z) {wf (Zwf lo) hi}:
left _ _
.
Next Obligation.
- red. omega.
+ red. omega.
Qed.
Next Obligation.
assert (x = hi - 1 \/ x < hi - 1) by omega.
- destruct H2. congruence. auto.
+ destruct H2. congruence. auto.
Qed.
Next Obligation.
exists wildcard'0; split; auto. omega.
@@ -276,7 +276,7 @@ Variable a: A.
Function fold_rec (hi: Z) {wf (Zwf lo) hi} : A :=
if zlt lo hi then f (hi - 1) (fold_rec (hi - 1)) else a.
Proof.
- intros. red. omega.
+ intros. red. omega.
apply Zwf_well_founded.
Qed.
@@ -284,9 +284,9 @@ Lemma fold_rec_elements:
forall hi, fold_rec hi = List.fold_right f a (elements_rec lo hi).
Proof.
intros. functional induction (fold_rec hi).
- rewrite elements_rec_equation. rewrite zlt_true; auto.
- simpl. congruence.
- rewrite elements_rec_equation. rewrite zlt_false; auto.
+ rewrite elements_rec_equation. rewrite zlt_true; auto.
+ simpl. congruence.
+ rewrite elements_rec_equation. rewrite zlt_false; auto.
Qed.
End FOLD.
@@ -298,7 +298,7 @@ Lemma fold_elements:
forall (A: Type) (f: Z -> A -> A) a i,
fold f a i = List.fold_right f a (elements i).
Proof.
- intros. unfold fold, elements. apply fold_rec_elements.
+ intros. unfold fold, elements. apply fold_rec_elements.
Qed.
(** Hints *)
@@ -313,4 +313,4 @@ Hint Resolve
-
+
diff --git a/lib/IntvSets.v b/lib/IntvSets.v
index 9f1a895f..78c20cc5 100644
--- a/lib/IntvSets.v
+++ b/lib/IntvSets.v
@@ -55,12 +55,12 @@ Lemma mem_In:
Proof.
induction 1; simpl.
- intuition congruence.
-- destruct (zlt x h).
+- destruct (zlt x h).
+ destruct (zle l x); simpl.
* tauto.
* split; intros. congruence.
- exfalso. destruct H0. omega. exploit BELOW; eauto. omega.
-+ rewrite IHok. intuition.
+ exfalso. destruct H0. omega. exploit BELOW; eauto. omega.
++ rewrite IHok. intuition.
Qed.
Fixpoint contains (L H: Z) (s: t) : bool :=
@@ -78,9 +78,9 @@ Proof.
- destruct (zle h0 h); simpl.
destruct (zle l l0); simpl.
intuition.
- rewrite IHok. intuition. destruct (H3 x); auto. exfalso.
+ rewrite IHok. intuition. destruct (H3 x); auto. exfalso.
destruct (H3 l0). omega. omega. exploit BELOW; eauto. omega.
- rewrite IHok. intuition. destruct (H3 x); auto. exfalso.
+ rewrite IHok. intuition. destruct (H3 x); auto. exfalso.
destruct (H3 h). omega. omega. exploit BELOW; eauto. omega.
Qed.
@@ -102,7 +102,7 @@ Proof.
simpl. rewrite IHok. tauto.
destruct (zlt h0 l).
simpl. tauto.
- rewrite IHok. intuition.
+ rewrite IHok. intuition.
assert (l0 <= x < h0 \/ l <= x < h) by xomega. tauto.
left; xomega.
left; xomega.
@@ -115,10 +115,10 @@ Proof.
constructor. auto. intros. inv H0. constructor.
destruct (zlt h l0).
constructor; auto. intros. rewrite In_add in H1; auto.
- destruct H1. omega. auto.
+ destruct H1. omega. auto.
destruct (zlt h0 l).
- constructor. auto. simpl; intros. destruct H1. omega. exploit BELOW; eauto. omega.
- constructor. omega. auto. auto.
+ constructor. auto. simpl; intros. destruct H1. omega. exploit BELOW; eauto. omega.
+ constructor. omega. auto. auto.
apply IHok. xomega.
Qed.
@@ -130,7 +130,7 @@ Fixpoint remove (L H: Z) (s: t) {struct s} : t :=
else if zlt H l then s
else if zlt l L then
if zlt H h then Cons l L (Cons H h s') else Cons l L (remove L H s')
- else
+ else
if zlt H h then Cons H h s' else remove L H s'
end.
@@ -141,22 +141,22 @@ Proof.
induction 1; simpl.
tauto.
destruct (zlt h l0).
- simpl. rewrite IHok. intuition omega.
+ simpl. rewrite IHok. intuition omega.
destruct (zlt h0 l).
- simpl. intuition. exploit BELOW; eauto. omega.
+ simpl. intuition. exploit BELOW; eauto. omega.
destruct (zlt l l0).
destruct (zlt h0 h); simpl. clear IHok. split.
- intros [A | [A | A]].
- split. omega. left; omega.
- split. omega. left; omega.
+ intros [A | [A | A]].
+ split. omega. left; omega.
+ split. omega. left; omega.
split. exploit BELOW; eauto. omega. auto.
intros [A [B | B]].
- destruct (zlt x l0). left; omega. right; left; omega.
+ destruct (zlt x l0). left; omega. right; left; omega.
auto.
- intuition omega.
+ intuition omega.
destruct (zlt h0 h); simpl.
intuition. exploit BELOW; eauto. omega.
- rewrite IHok. intuition. omegaContradiction.
+ rewrite IHok. intuition. omegaContradiction.
Qed.
Lemma remove_ok:
@@ -165,14 +165,14 @@ Proof.
induction 2; simpl.
constructor.
destruct (zlt h l0).
- constructor; auto. intros; apply BELOW. rewrite In_remove in H1; tauto.
+ constructor; auto. intros; apply BELOW. rewrite In_remove in H1; tauto.
destruct (zlt h0 l).
- constructor; auto.
+ constructor; auto.
destruct (zlt l l0).
destruct (zlt h0 h).
- constructor. omega. intros. inv H1. omega. exploit BELOW; eauto. omega.
+ constructor. omega. intros. inv H1. omega. exploit BELOW; eauto. omega.
constructor. omega. auto. auto.
- constructor; auto. intros. rewrite In_remove in H1 by auto. destruct H1. exploit BELOW; eauto. omega.
+ constructor; auto. intros. rewrite In_remove in H1 by auto. destruct H1. exploit BELOW; eauto. omega.
destruct (zlt h0 h).
constructor; auto.
auto.
@@ -204,7 +204,7 @@ Proof.
tauto.
assert (ok (Cons l0 h0 s0)) by (constructor; auto).
destruct (zle h l0).
- rewrite IHok; auto. simpl. intuition. omegaContradiction.
+ rewrite IHok; auto. simpl. intuition. omegaContradiction.
exploit BELOW0; eauto. intros. omegaContradiction.
destruct (zle h0 l).
simpl in IHok0; rewrite IHok0. intuition. omegaContradiction.
@@ -212,10 +212,10 @@ Proof.
destruct (zle l l0).
destruct (zle h0 h).
simpl. simpl in IHok0; rewrite IHok0. intuition.
- simpl. rewrite IHok; auto. simpl. intuition. exploit BELOW0; eauto. intros; omegaContradiction.
+ simpl. rewrite IHok; auto. simpl. intuition. exploit BELOW0; eauto. intros; omegaContradiction.
destruct (zle h h0).
simpl. rewrite IHok; auto. simpl. intuition.
- simpl. simpl in IHok0; rewrite IHok0. intuition.
+ simpl. simpl in IHok0; rewrite IHok0. intuition.
exploit BELOW; eauto. intros; omegaContradiction.
Qed.
@@ -237,8 +237,8 @@ Proof.
constructor; auto. intros.
assert (In x (inter (Cons l h s) s0)) by exact H3.
rewrite In_inter in H4; auto. apply BELOW0. tauto.
- constructor. omega. intros. rewrite In_inter in H3; auto. apply BELOW. tauto.
- auto.
+ constructor. omega. intros. rewrite In_inter in H3; auto. apply BELOW. tauto.
+ auto.
destruct (zle h h0).
constructor. omega. intros. rewrite In_inter in H3; auto. apply BELOW. tauto.
auto.
@@ -265,7 +265,7 @@ Proof.
split. constructor; auto. tauto.
destruct (IHok s0) as [A B]; auto.
split. apply add_ok; auto. apply add_ok; auto.
- intros. rewrite In_add. rewrite In_add. rewrite <- B. tauto. auto. apply add_ok; auto.
+ intros. rewrite In_add. rewrite In_add. rewrite <- B. tauto. auto. apply add_ok; auto.
Qed.
Fixpoint beq (s1 s2: t) : bool :=
@@ -281,13 +281,13 @@ Lemma beq_spec:
Proof.
induction 1; destruct 1; simpl.
- tauto.
-- split; intros. discriminate. exfalso. apply (H0 l). left; omega.
- split; intros. discriminate. exfalso. apply (H0 l). left; omega.
-- split; intros.
-+ InvBooleans. subst. rewrite IHok in H3 by auto. rewrite H3. tauto.
-+ destruct (zeq l l0). destruct (zeq h h0). simpl. subst.
+- split; intros. discriminate. exfalso. apply (H0 l). left; omega.
+- split; intros.
++ InvBooleans. subst. rewrite IHok in H3 by auto. rewrite H3. tauto.
++ destruct (zeq l l0). destruct (zeq h h0). simpl. subst.
apply IHok. auto. intros; split; intros.
- destruct (proj1 (H1 x)); auto. exfalso. exploit BELOW; eauto. omega.
+ destruct (proj1 (H1 x)); auto. exfalso. exploit BELOW; eauto. omega.
destruct (proj2 (H1 x)); auto. exfalso. exploit BELOW0; eauto. omega.
exfalso. subst l0. destruct (zlt h h0).
destruct (proj2 (H1 h)). left; omega. omega. exploit BELOW; eauto. omega.
@@ -310,7 +310,7 @@ Next Obligation. constructor. Qed.
Theorem In_empty: forall x, ~(In x empty).
Proof.
- unfold In; intros; simpl. tauto.
+ unfold In; intros; simpl. tauto.
Qed.
Program Definition interval (l h: Z) : t :=
@@ -337,16 +337,16 @@ Qed.
Theorem In_add: forall x l h s, In x (add l h s) <-> l <= x < h \/ In x s.
Proof.
- unfold add, In; intros.
+ unfold add, In; intros.
destruct (zlt l h).
simpl. apply R.In_add. apply proj2_sig.
- intuition. omegaContradiction.
+ intuition. omegaContradiction.
Qed.
Program Definition remove (l h: Z) (s: t) : t :=
if zlt l h then R.remove l h s else s.
Next Obligation.
- apply R.remove_ok. auto. apply proj2_sig.
+ apply R.remove_ok. auto. apply proj2_sig.
Qed.
Theorem In_remove: forall x l h s, In x (remove l h s) <-> ~(l <= x < h) /\ In x s.
@@ -362,11 +362,11 @@ Next Obligation. apply R.inter_ok; apply proj2_sig. Qed.
Theorem In_inter: forall x s1 s2, In x (inter s1 s2) <-> In x s1 /\ In x s2.
Proof.
- unfold inter, In; intros; simpl. apply R.In_inter; apply proj2_sig.
+ unfold inter, In; intros; simpl. apply R.In_inter; apply proj2_sig.
Qed.
Program Definition union (s1 s2: t) : t := R.union s1 s2.
-Next Obligation.
+Next Obligation.
destruct (R.In_ok_union _ (proj2_sig s1) _ (proj2_sig s2)). auto.
Qed.
@@ -381,7 +381,7 @@ Program Definition mem (x: Z) (s: t) := R.mem x s.
Theorem mem_spec: forall x s, mem x s = true <-> In x s.
Proof.
- unfold mem, In; intros. apply R.mem_In. apply proj2_sig.
+ unfold mem, In; intros. apply R.mem_In. apply proj2_sig.
Qed.
Program Definition contains (l h: Z) (s: t) :=
@@ -392,7 +392,7 @@ Theorem contains_spec:
Proof.
unfold contains, In; intros. destruct (zlt l h).
apply R.contains_In. auto. apply proj2_sig.
- split; intros. omegaContradiction. auto.
+ split; intros. omegaContradiction. auto.
Qed.
Program Definition beq (s1 s2: t) : bool := R.beq s1 s2.
diff --git a/lib/Iteration.v b/lib/Iteration.v
index f3507fe6..4398f96d 100644
--- a/lib/Iteration.v
+++ b/lib/Iteration.v
@@ -50,7 +50,7 @@ Hypothesis step_decr: forall a a', step a = inr _ a' -> ord a' a.
Definition step_info (a: A) : {b | step a = inl _ b} + {a' | step a = inr _ a' & ord a' a}.
Proof.
- caseEq (step a); intros. left; exists b; auto. right; exists a0; auto.
+ caseEq (step a); intros. left; exists b; auto. right; exists a0; auto.
Defined.
Definition iterate_F (a: A) (rec: forall a', ord a' a -> B) : B :=
@@ -75,10 +75,10 @@ Lemma iterate_prop:
forall a, P a -> Q (iterate a).
Proof.
intros a0. pattern a0. apply well_founded_ind with (R := ord). auto.
- intros. unfold iterate; rewrite unroll_Fix. unfold iterate_F.
- destruct (step_info x) as [[b U] | [a' U V]].
+ intros. unfold iterate; rewrite unroll_Fix. unfold iterate_F.
+ destruct (step_info x) as [[b U] | [a' U V]].
exploit step_prop; eauto. rewrite U; auto.
- apply H. auto. exploit step_prop; eauto. rewrite U; auto.
+ apply H. auto. exploit step_prop; eauto. rewrite U; auto.
Qed.
End ITERATION.
@@ -105,7 +105,7 @@ End WfIter.
Since we know (informally) that our computations terminate, we can
take a very large constant as the maximal number of iterations.
Failure will therefore never happen in practice, but of
- course our proofs also cover the failure case and show that
+ course our proofs also cover the failure case and show that
nothing bad happens in this hypothetical case either. *)
Module PrimIter.
@@ -169,11 +169,11 @@ Hypothesis step_prop:
Lemma iter_prop:
forall n a b, P a -> iter n a = Some b -> Q b.
Proof.
- apply (well_founded_ind Plt_wf
+ apply (well_founded_ind Plt_wf
(fun p => forall a b, P a -> iter p a = Some b -> Q b)).
intros. unfold iter in H1. rewrite unroll_Fix in H1. unfold iter_step in H1.
destruct (peq x 1). discriminate.
- specialize (step_prop a H0).
+ specialize (step_prop a H0).
destruct (step a) as [b'|a'] eqn:?.
inv H1. auto.
apply H with (Ppred x) a'. apply Ppred_Plt; auto. auto. auto.
@@ -222,8 +222,8 @@ Definition F_iter (next: A -> option B) (a: A) : option B :=
Lemma F_iter_monot:
forall f g, F_le f g -> F_le (F_iter f) (F_iter g).
Proof.
- intros; red; intros. unfold F_iter.
- destruct (step a) as [b | a']. red; auto. apply H.
+ intros; red; intros. unfold F_iter.
+ destruct (step a) as [b | a']. red; auto. apply H.
Qed.
Fixpoint iter (n: nat) : A -> option B :=
@@ -235,9 +235,9 @@ Fixpoint iter (n: nat) : A -> option B :=
Lemma iter_monot:
forall p q, (p <= q)%nat -> F_le (iter p) (iter q).
Proof.
- induction p; intros.
+ induction p; intros.
simpl. red; intros; red; auto.
- destruct q. elimtype False; omega.
+ destruct q. elimtype False; omega.
simpl. apply F_iter_monot. apply IHp. omega.
Qed.
@@ -249,7 +249,7 @@ Proof.
intro a. elim (classic (forall n, iter n a = None)); intro.
right; assumption.
left. generalize (not_all_ex_not nat (fun n => iter n a = None) H).
- intros [n D]. exists n. generalize D.
+ intros [n D]. exists n. generalize D.
case (iter n a); intros. exists b; auto. congruence.
Qed.
@@ -259,23 +259,23 @@ Definition converges_to (a: A) (b: option B) : Prop :=
Lemma converges_to_Some:
forall a n b, iter n a = Some b -> converges_to a (Some b).
Proof.
- intros. exists n. intros.
+ intros. exists n. intros.
assert (B_le (iter n a) (iter m a)). apply iter_monot. auto.
- elim H1; intro; congruence.
+ elim H1; intro; congruence.
Qed.
Lemma converges_to_exists:
forall a, exists b, converges_to a b.
Proof.
- intros. elim (iter_either a).
+ intros. elim (iter_either a).
intros [n [b EQ]]. exists (Some b). apply converges_to_Some with n. assumption.
- intro. exists (@None B). exists O. intros. auto.
+ intro. exists (@None B). exists O. intros. auto.
Qed.
Lemma converges_to_unique:
forall a b, converges_to a b -> forall b', converges_to a b' -> b = b'.
Proof.
- intros a b [n C] b' [n' C'].
+ intros a b [n C] b' [n' C'].
rewrite <- (C (max n n')). rewrite <- (C' (max n n')). auto.
apply le_max_r. apply le_max_l.
Qed.
@@ -283,7 +283,7 @@ Qed.
Lemma converges_to_exists_uniquely:
forall a, exists! b, converges_to a b .
Proof.
- intro. destruct (converges_to_exists a) as [b CT].
+ intro. destruct (converges_to_exists a) as [b CT].
exists b. split. assumption. exact (converges_to_unique _ _ CT).
Qed.
@@ -293,7 +293,7 @@ Definition iterate (a: A) : option B :=
Lemma converges_to_iterate:
forall a b, converges_to a b -> iterate a = b.
Proof.
- intros. unfold iterate.
+ intros. unfold iterate.
destruct (constructive_definite_description (converges_to a) (converges_to_exists_uniquely a)) as [b' P].
simpl. apply converges_to_unique with a; auto.
Qed.
@@ -301,7 +301,7 @@ Qed.
Lemma iterate_converges_to:
forall a, converges_to a (iterate a).
Proof.
- intros. unfold iterate.
+ intros. unfold iterate.
destruct (constructive_definite_description (converges_to a) (converges_to_exists_uniquely a)) as [b' P].
simpl; auto.
Qed.
@@ -320,15 +320,15 @@ Lemma iter_prop:
Proof.
induction n; intros until b; intro H; simpl.
congruence.
- unfold F_iter. generalize (step_prop a H).
- case (step a); intros. congruence.
+ unfold F_iter. generalize (step_prop a H).
+ case (step a); intros. congruence.
apply IHn with a0; auto.
Qed.
Lemma iterate_prop:
forall a b, iterate a = Some b -> P a -> Q b.
Proof.
- intros. destruct (iterate_converges_to a) as [n IT].
+ intros. destruct (iterate_converges_to a) as [n IT].
rewrite H in IT. apply iter_prop with n a. auto. apply IT. auto.
Qed.
diff --git a/lib/Lattice.v b/lib/Lattice.v
index 5a941a13..352b4479 100644
--- a/lib/Lattice.v
+++ b/lib/Lattice.v
@@ -27,7 +27,7 @@ Local Unset Case Analysis Schemes.
(** * Signatures of semi-lattices *)
(** A semi-lattice is a type [t] equipped with an equivalence relation [eq],
- a boolean equivalence test [beq], a partial order [ge], a smallest element
+ a boolean equivalence test [beq], a partial order [ge], a smallest element
[bot], and an upper bound operation [lub].
Note that we do not demand that [lub] computes the least upper bound. *)
@@ -86,9 +86,9 @@ Lemma gsspec:
forall p v x q,
L.eq (get q (set p v x)) (if peq q p then v else get q x).
Proof.
- intros. unfold set, get.
+ intros. unfold set, get.
destruct (L.beq v L.bot) eqn:EBOT.
- rewrite PTree.grspec. unfold PTree.elt_eq. destruct (peq q p).
+ rewrite PTree.grspec. unfold PTree.elt_eq. destruct (peq q p).
apply L.eq_sym. apply L.beq_correct; auto.
apply L.eq_refl.
rewrite PTree.gsspec. destruct (peq q p); apply L.eq_refl.
@@ -117,7 +117,7 @@ Definition beq (x y: t) : bool := PTree.beq L.beq x y.
Lemma beq_correct: forall x y, beq x y = true -> eq x y.
Proof.
unfold beq; intros; red; intros. unfold get.
- rewrite PTree.beq_correct in H. specialize (H p).
+ rewrite PTree.beq_correct in H. specialize (H p).
destruct (x!p); destruct (y!p); intuition.
apply L.beq_correct; auto.
apply L.eq_refl.
@@ -165,17 +165,17 @@ Definition opt_eq (ox oy: option L.t) : Prop :=
Lemma opt_eq_refl: forall ox, opt_eq ox ox.
Proof.
- intros. unfold opt_eq. destruct ox. apply L.eq_refl. auto.
+ intros. unfold opt_eq. destruct ox. apply L.eq_refl. auto.
Qed.
Lemma opt_eq_sym: forall ox oy, opt_eq ox oy -> opt_eq oy ox.
Proof.
- unfold opt_eq. destruct ox; destruct oy; auto. apply L.eq_sym.
+ unfold opt_eq. destruct ox; destruct oy; auto. apply L.eq_sym.
Qed.
Lemma opt_eq_trans: forall ox oy oz, opt_eq ox oy -> opt_eq oy oz -> opt_eq ox oz.
Proof.
- unfold opt_eq. destruct ox; destruct oy; destruct oz; intuition.
+ unfold opt_eq. destruct ox; destruct oy; destruct oz; intuition.
eapply L.eq_trans; eauto.
Qed.
@@ -189,8 +189,8 @@ Definition opt_beq (ox oy: option L.t) : bool :=
Lemma opt_beq_correct:
forall ox oy, opt_beq ox oy = true -> opt_eq ox oy.
Proof.
- unfold opt_beq, opt_eq. destruct ox; destruct oy; try congruence.
- intros. apply L.beq_correct; auto.
+ unfold opt_beq, opt_eq. destruct ox; destruct oy; try congruence.
+ intros. apply L.beq_correct; auto.
auto.
Qed.
@@ -206,7 +206,7 @@ Proof. intros; red; intros; apply opt_eq_sym; auto. Qed.
Lemma tree_eq_trans: forall m1 m2 m3, tree_eq m1 m2 -> tree_eq m2 m3 -> tree_eq m1 m3.
Proof. intros; red; intros; apply opt_eq_trans with (PTree.get i m2); auto. Qed.
-Lemma tree_eq_node:
+Lemma tree_eq_node:
forall l1 o1 r1 l2 o2 r2,
tree_eq l1 l2 -> tree_eq r1 r2 -> opt_eq o1 o2 ->
tree_eq (PTree.Node l1 o1 r1) (PTree.Node l2 o2 r2).
@@ -214,23 +214,23 @@ Proof.
intros; red; intros. destruct i; simpl; auto.
Qed.
-Lemma tree_eq_node':
+Lemma tree_eq_node':
forall l1 o1 r1 l2 o2 r2,
tree_eq l1 l2 -> tree_eq r1 r2 -> opt_eq o1 o2 ->
tree_eq (PTree.Node l1 o1 r1) (PTree.Node' l2 o2 r2).
Proof.
- intros; red; intros. rewrite PTree.gnode'. apply tree_eq_node; auto.
+ intros; red; intros. rewrite PTree.gnode'. apply tree_eq_node; auto.
Qed.
-Lemma tree_eq_node'':
+Lemma tree_eq_node'':
forall l1 o1 r1 l2 o2 r2,
tree_eq l1 l2 -> tree_eq r1 r2 -> opt_eq o1 o2 ->
tree_eq (PTree.Node' l1 o1 r1) (PTree.Node' l2 o2 r2).
Proof.
- intros; red; intros. repeat rewrite PTree.gnode'. apply tree_eq_node; auto.
+ intros; red; intros. repeat rewrite PTree.gnode'. apply tree_eq_node; auto.
Qed.
-Hint Resolve opt_beq_correct opt_eq_refl opt_eq_sym
+Hint Resolve opt_beq_correct opt_eq_refl opt_eq_sym
tree_eq_refl tree_eq_sym
tree_eq_node tree_eq_node' tree_eq_node'' : combine.
@@ -296,7 +296,7 @@ Inductive changed2 : Type :=
Fixpoint xcombine (m1 m2 : PTree.t L.t) {struct m1} : changed2 :=
match m1, m2 with
- | PTree.Leaf, PTree.Leaf =>
+ | PTree.Leaf, PTree.Leaf =>
Same
| PTree.Leaf, _ =>
match combine_r m2 with
@@ -333,7 +333,7 @@ Fixpoint xcombine (m1 m2 : PTree.t L.t) {struct m1} : changed2 :=
end.
Lemma xcombine_eq:
- forall m1 m2,
+ forall m1 m2,
match xcombine m1 m2 with
| Same => tree_eq m1 (PTree.combine f m1 m2) /\ tree_eq m2 (PTree.combine f m1 m2)
| Same1 => tree_eq m1 (PTree.combine f m1 m2)
@@ -348,7 +348,7 @@ Opaque combine_l combine_r PTree.xcombine_l PTree.xcombine_r.
destruct (combine_r (PTree.Node m2_1 o m2_2)); auto.
generalize (combine_l_eq (PTree.Node m1_1 o m1_2)).
destruct (combine_l (PTree.Node m1_1 o m1_2)); auto.
- generalize (IHm1_1 m2_1) (IHm1_2 m2_2).
+ generalize (IHm1_1 m2_1) (IHm1_2 m2_2).
destruct (xcombine m1_1 m2_1);
destruct (xcombine m1_2 m2_2); auto with combine;
intuition; case_eq (opt_beq (f o o0) o); case_eq (opt_beq (f o o0) o0); auto with combine.
@@ -390,7 +390,7 @@ Lemma gcombine_bot:
L.eq (get p (combine f t1 t2))
(match f t1!p t2!p with Some x => x | None => L.bot end).
Proof.
- intros. unfold get. generalize (gcombine f H t1 t2 p). unfold opt_eq.
+ intros. unfold get. generalize (gcombine f H t1 t2 p). unfold opt_eq.
destruct ((combine f t1 t2)!p); destruct (f t1!p t2!p).
auto. contradiction. contradiction. intros; apply L.eq_refl.
Qed.
@@ -398,9 +398,9 @@ Qed.
Lemma ge_lub_left:
forall x y, ge (lub x y) x.
Proof.
- unfold ge, lub; intros.
+ unfold ge, lub; intros.
eapply L.ge_trans. apply L.ge_refl. apply gcombine_bot; auto.
- unfold get. destruct x!p. destruct y!p.
+ unfold get. destruct x!p. destruct y!p.
apply L.ge_lub_left.
apply L.ge_refl. apply L.eq_refl.
apply L.ge_bot.
@@ -409,9 +409,9 @@ Qed.
Lemma ge_lub_right:
forall x y, ge (lub x y) y.
Proof.
- unfold ge, lub; intros.
+ unfold ge, lub; intros.
eapply L.ge_trans. apply L.ge_refl. apply gcombine_bot; auto.
- unfold get. destruct y!p. destruct x!p.
+ unfold get. destruct y!p. destruct x!p.
apply L.ge_lub_right.
apply L.ge_refl. apply L.eq_refl.
apply L.ge_bot.
@@ -451,11 +451,11 @@ Lemma gsspec:
x <> Bot -> ~L.eq v L.bot ->
L.eq (get q (set p v x)) (if peq q p then v else get q x).
Proof.
- intros. unfold set. destruct x. congruence.
- destruct (L.beq v L.bot) eqn:EBOT.
+ intros. unfold set. destruct x. congruence.
+ destruct (L.beq v L.bot) eqn:EBOT.
elim H0. apply L.beq_correct; auto.
destruct (L.beq v L.top) eqn:ETOP; simpl.
- rewrite PTree.grspec. unfold PTree.elt_eq. destruct (peq q p).
+ rewrite PTree.grspec. unfold PTree.elt_eq. destruct (peq q p).
apply L.eq_sym. apply L.beq_correct; auto.
apply L.eq_refl.
rewrite PTree.gsspec. destruct (peq q p); apply L.eq_refl.
@@ -561,7 +561,7 @@ Lemma gcombine_top:
L.eq (get p (Top_except (LM.combine f t1 t2)))
(match f t1!p t2!p with Some x => x | None => L.top end).
Proof.
- intros. simpl. generalize (LM.gcombine f H t1 t2 p). unfold LM.opt_eq.
+ intros. simpl. generalize (LM.gcombine f H t1 t2 p). unfold LM.opt_eq.
destruct ((LM.combine f t1 t2)!p); destruct (f t1!p t2!p).
auto. contradiction. contradiction. intros; apply L.eq_refl.
Qed.
@@ -574,10 +574,10 @@ Proof.
rewrite get_bot. apply L.ge_bot.
apply L.ge_refl. apply L.eq_refl.
eapply L.ge_trans. apply L.ge_refl. apply gcombine_top; auto.
- unfold get. destruct t0!p. destruct t1!p.
+ unfold get. destruct t0!p. destruct t1!p.
unfold opt_lub. destruct (L.beq (L.lub t2 t3) L.top) eqn:E.
- apply L.ge_top. apply L.ge_lub_left.
- apply L.ge_top.
+ apply L.ge_top. apply L.ge_lub_left.
+ apply L.ge_top.
apply L.ge_top.
Qed.
@@ -589,10 +589,10 @@ Proof.
apply L.ge_refl. apply L.eq_refl.
rewrite get_bot. apply L.ge_bot.
eapply L.ge_trans. apply L.ge_refl. apply gcombine_top; auto.
- unfold get. destruct t0!p; destruct t1!p.
+ unfold get. destruct t0!p; destruct t1!p.
unfold opt_lub. destruct (L.beq (L.lub t2 t3) L.top) eqn:E.
- apply L.ge_top. apply L.ge_lub_right.
- apply L.ge_top.
+ apply L.ge_top. apply L.ge_lub_right.
+ apply L.ge_top.
apply L.ge_top.
apply L.ge_top.
Qed.
@@ -618,7 +618,7 @@ Module LFSet (S: FSetInterface.WS) <: SEMILATTICE.
Definition ge (x y: t) := S.Subset y x.
Lemma ge_refl: forall x y, eq x y -> ge x y.
Proof.
- unfold eq, ge, S.Equal, S.Subset; intros. firstorder.
+ unfold eq, ge, S.Equal, S.Subset; intros. firstorder.
Qed.
Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z.
Proof.
@@ -635,12 +635,12 @@ Module LFSet (S: FSetInterface.WS) <: SEMILATTICE.
Lemma ge_lub_left: forall x y, ge (lub x y) x.
Proof.
- unfold lub, ge, S.Subset; intros. apply S.union_2; auto.
+ unfold lub, ge, S.Subset; intros. apply S.union_2; auto.
Qed.
Lemma ge_lub_right: forall x y, ge (lub x y) y.
Proof.
- unfold lub, ge, S.Subset; intros. apply S.union_3; auto.
+ unfold lub, ge, S.Subset; intros. apply S.union_3; auto.
Qed.
End LFSet.
@@ -650,7 +650,7 @@ End LFSet.
(** Given a type with decidable equality [X], the following functor
returns a semi-lattice structure over [X.t] complemented with
a top and a bottom element. The ordering is the flat ordering
- [Bot < Inj x < Top]. *)
+ [Bot < Inj x < Top]. *)
Module LFlat(X: EQUALITY_TYPE) <: SEMILATTICE_WITH_TOP.
@@ -735,7 +735,7 @@ Proof.
Qed.
End LFlat.
-
+
(** * Boolean semi-lattice *)
(** This semi-lattice has only two elements, [bot] and [top], trivially
diff --git a/lib/Maps.v b/lib/Maps.v
index 63ac0c09..39fec9fd 100644
--- a/lib/Maps.v
+++ b/lib/Maps.v
@@ -126,7 +126,7 @@ Module Type TREE.
forall (A: Type) (m: t A) (i: elt) (v: A),
In (i, v) (elements m) -> get i m = Some v.
Hypothesis elements_keys_norepet:
- forall (A: Type) (m: t A),
+ forall (A: Type) (m: t A),
list_norepet (List.map (@fst elt A) (elements m)).
Hypothesis elements_extensional:
forall (A: Type) (m n: t A),
@@ -396,7 +396,7 @@ Module PTree <: TREE.
generalize (H xH); simpl; congruence.
destruct (andb_prop _ _ H). rewrite IHm1 in H0. rewrite IHm2 in H1.
destruct x; simpl; auto.
- apply andb_true_intro; split.
+ apply andb_true_intro; split.
apply IHm1. intros; apply (H (xO x)).
apply IHm2. intros; apply (H (xI x)).
Qed.
@@ -414,18 +414,18 @@ Module PTree <: TREE.
induction m1; intros.
- simpl. rewrite bempty_correct. split; intros.
rewrite gleaf. rewrite H. auto.
- generalize (H x). rewrite gleaf. destruct (get x m2); tauto.
+ generalize (H x). rewrite gleaf. destruct (get x m2); tauto.
- destruct m2.
+ unfold beq. rewrite bempty_correct. split; intros.
rewrite H. rewrite gleaf. auto.
generalize (H x). rewrite gleaf. destruct (get x (Node m1_1 o m1_2)); tauto.
+ simpl. split; intros.
* destruct (andb_prop _ _ H). destruct (andb_prop _ _ H0).
- rewrite IHm1_1 in H3. rewrite IHm1_2 in H1.
+ rewrite IHm1_1 in H3. rewrite IHm1_2 in H1.
destruct x; simpl. apply H1. apply H3.
destruct o; destruct o0; auto || congruence.
* apply andb_true_intro. split. apply andb_true_intro. split.
- generalize (H xH); simpl. destruct o; destruct o0; tauto.
+ generalize (H xH); simpl. destruct o; destruct o0; tauto.
apply IHm1_1. intros; apply (H (xO x)).
apply IHm1_2. intros; apply (H (xI x)).
Qed.
@@ -450,7 +450,7 @@ Module PTree <: TREE.
intros j. simpl. rewrite IH. reflexivity.
intros j. simpl. rewrite IH. reflexivity.
Qed.
-
+
Lemma prev_involutive i :
prev (prev i) = i.
Proof (prev_append_prev i xH).
@@ -513,7 +513,7 @@ Module PTree <: TREE.
forall (A: Type) (l r: t A) (x: option A) (i: positive),
get i (Node' l x r) = get i (Node l x r).
Proof.
- intros. unfold Node'.
+ intros. unfold Node'.
destruct l; destruct x; destruct r; auto.
destruct i; simpl; auto; rewrite gleaf; auto.
Qed.
@@ -531,9 +531,9 @@ Module PTree <: TREE.
get i (filter1 pred m) =
match get i m with None => None | Some x => if pred x then Some x else None end.
Proof.
- intros until m. revert m i. induction m; simpl; intros.
+ intros until m. revert m i. induction m; simpl; intros.
rewrite gleaf; auto.
- rewrite gnode'. destruct i; simpl; auto. destruct o; auto.
+ rewrite gnode'. destruct i; simpl; auto. destruct o; auto.
Qed.
Section COMBINE.
@@ -589,7 +589,7 @@ Module PTree <: TREE.
induction m1; intros; simpl.
rewrite gleaf. apply xgcombine_r.
destruct m2; simpl.
- rewrite gleaf. rewrite <- xgcombine_l. auto.
+ rewrite gleaf. rewrite <- xgcombine_l. auto.
repeat rewrite gnode'. destruct i; simpl; auto.
Qed.
@@ -622,10 +622,10 @@ Module PTree <: TREE.
auto.
rewrite IHm1_1.
rewrite IHm1_2.
- auto.
+ auto.
Qed.
- Fixpoint xelements (A : Type) (m : t A) (i : positive)
+ Fixpoint xelements (A : Type) (m : t A) (i : positive)
(k: list (positive * A)) {struct m}
: list (positive * A) :=
match m with
@@ -651,7 +651,7 @@ Module PTree <: TREE.
Remark xelements_leaf:
forall A i, xelements (@Leaf A) i nil = nil.
Proof.
- intros; reflexivity.
+ intros; reflexivity.
Qed.
Remark xelements_node:
@@ -685,8 +685,8 @@ Module PTree <: TREE.
apply xelements_incl. right. auto.
auto.
inv H. apply xelements_incl. left. reflexivity.
- apply xelements_incl. auto.
- auto.
+ apply xelements_incl. auto.
+ auto.
inv H.
Qed.
@@ -694,7 +694,7 @@ Module PTree <: TREE.
forall (A: Type) (m: t A) (i: positive) (v: A),
get i m = Some v -> In (i, v) (elements m).
Proof.
- intros A m i v H.
+ intros A m i v H.
generalize (xelements_correct m i xH nil H). rewrite prev_append_prev. exact id.
Qed.
@@ -703,11 +703,11 @@ Module PTree <: TREE.
In (k, v) (xelements m i nil) ->
exists j, k = prev (prev_append j i) /\ get j m = Some v.
Proof.
- induction m; intros.
+ induction m; intros.
- rewrite xelements_leaf in H. contradiction.
- rewrite xelements_node in H. rewrite ! in_app_iff in H. destruct H as [P | [P | P]].
+ exploit IHm1; eauto. intros (j & Q & R). exists (xO j); auto.
- + destruct o; simpl in P; intuition auto. inv H. exists xH; auto.
+ + destruct o; simpl in P; intuition auto. inv H. exists xH; auto.
+ exploit IHm2; eauto. intros (j & Q & R). exists (xI j); auto.
Qed.
@@ -715,8 +715,8 @@ Module PTree <: TREE.
forall (A: Type) (m: t A) (i: positive) (v: A),
In (i, v) (elements m) -> get i m = Some v.
Proof.
- unfold elements. intros A m i v H. exploit in_xelements; eauto. intros (j & P & Q).
- rewrite prev_append_prev in P. change i with (prev_append 1 i) in P.
+ unfold elements. intros A m i v H. exploit in_xelements; eauto. intros (j & P & Q).
+ rewrite prev_append_prev in P. change i with (prev_append 1 i) in P.
exploit prev_append_inj; eauto. intros; congruence.
Qed.
@@ -726,7 +726,7 @@ Module PTree <: TREE.
Remark xkeys_leaf:
forall A i, xkeys (@Leaf A) i = nil.
Proof.
- intros; reflexivity.
+ intros; reflexivity.
Qed.
Remark xkeys_node:
@@ -736,7 +736,7 @@ Module PTree <: TREE.
++ match o with None => nil | Some v => prev i :: nil end
++ xkeys m2 (xI i).
Proof.
- intros. unfold xkeys. rewrite xelements_node. rewrite ! map_app. destruct o; auto.
+ intros. unfold xkeys. rewrite xelements_node. rewrite ! map_app. destruct o; auto.
Qed.
Lemma in_xkeys:
@@ -746,7 +746,7 @@ Module PTree <: TREE.
Proof.
unfold xkeys; intros.
apply (list_in_map_inv) in H. destruct H as ((j, v) & -> & H).
- exploit in_xelements; eauto. intros (k & P & Q). exists k; auto.
+ exploit in_xelements; eauto. intros (k & P & Q). exists k; auto.
Qed.
Lemma xelements_keys_norepet:
@@ -756,26 +756,26 @@ Module PTree <: TREE.
induction m; intros.
- rewrite xkeys_leaf; constructor.
- assert (NOTIN1: ~ In (prev i) (xkeys m1 (xO i))).
- { red; intros. exploit in_xkeys; eauto. intros (j & EQ).
+ { red; intros. exploit in_xkeys; eauto. intros (j & EQ).
rewrite prev_append_prev in EQ. simpl in EQ. apply prev_append_inj in EQ. discriminate. }
assert (NOTIN2: ~ In (prev i) (xkeys m2 (xI i))).
- { red; intros. exploit in_xkeys; eauto. intros (j & EQ).
+ { red; intros. exploit in_xkeys; eauto. intros (j & EQ).
rewrite prev_append_prev in EQ. simpl in EQ. apply prev_append_inj in EQ. discriminate. }
assert (DISJ: forall x, In x (xkeys m1 (xO i)) -> In x (xkeys m2 (xI i)) -> False).
- { intros. exploit in_xkeys. eexact H. intros (j1 & EQ1).
- exploit in_xkeys. eexact H0. intros (j2 & EQ2).
+ { intros. exploit in_xkeys. eexact H. intros (j1 & EQ1).
+ exploit in_xkeys. eexact H0. intros (j2 & EQ2).
rewrite prev_append_prev in *. simpl in *. rewrite EQ2 in EQ1. apply prev_append_inj in EQ1. discriminate. }
- rewrite xkeys_node. apply list_norepet_append. auto.
+ rewrite xkeys_node. apply list_norepet_append. auto.
destruct o; simpl; auto. constructor; auto.
- red; intros. red; intros; subst y. destruct o; simpl in H0.
- destruct H0. subst x. tauto. eauto. eauto.
+ red; intros. red; intros; subst y. destruct o; simpl in H0.
+ destruct H0. subst x. tauto. eauto. eauto.
Qed.
Theorem elements_keys_norepet:
- forall (A: Type) (m: t A),
+ forall (A: Type) (m: t A),
list_norepet (List.map (@fst elt A) (elements m)).
Proof.
- intros. apply (xelements_keys_norepet m xH).
+ intros. apply (xelements_keys_norepet m xH).
Qed.
Remark xelements_empty:
@@ -783,7 +783,7 @@ Module PTree <: TREE.
Proof.
induction m; intros.
auto.
- rewrite xelements_node. rewrite IHm1, IHm2. destruct o; auto.
+ rewrite xelements_node. rewrite IHm1, IHm2. destruct o; auto.
generalize (H xH); simpl; congruence.
intros. apply (H (xI i0)).
intros. apply (H (xO i0)).
@@ -806,29 +806,29 @@ Module PTree <: TREE.
(xelements m j nil) (xelements n j nil)).
{
induction m; intros.
- - rewrite xelements_leaf. rewrite xelements_empty. constructor.
- intros. destruct (get i n) eqn:E; auto. exploit H0; eauto.
+ - rewrite xelements_leaf. rewrite xelements_empty. constructor.
+ intros. destruct (get i n) eqn:E; auto. exploit H0; eauto.
intros [x [P Q]]. rewrite gleaf in P; congruence.
- - destruct n as [ | n1 o' n2 ].
+ - destruct n as [ | n1 o' n2 ].
+ rewrite xelements_leaf, xelements_empty. constructor.
- intros. destruct (get i (Node m1 o m2)) eqn:E; auto. exploit H; eauto.
+ intros. destruct (get i (Node m1 o m2)) eqn:E; auto. exploit H; eauto.
intros [x [P Q]]. rewrite gleaf in P; congruence.
+ rewrite ! xelements_node. apply list_forall2_app.
- apply IHm1.
- intros. apply (H (xO i) x); auto.
+ apply IHm1.
+ intros. apply (H (xO i) x); auto.
intros. apply (H0 (xO i) y); auto.
- apply list_forall2_app.
+ apply list_forall2_app.
destruct o, o'.
- destruct (H xH a) as [x [P Q]]. auto. simpl in P. inv P.
+ destruct (H xH a) as [x [P Q]]. auto. simpl in P. inv P.
constructor. auto. constructor.
destruct (H xH a) as [x [P Q]]. auto. simpl in P. inv P.
destruct (H0 xH b) as [x [P Q]]. auto. simpl in P. inv P.
constructor.
- apply IHm2.
- intros. apply (H (xI i) x); auto.
+ apply IHm2.
+ intros. apply (H (xI i) x); auto.
intros. apply (H0 (xI i) y); auto.
}
- intros. apply H with (j := xH); auto.
+ intros. apply H with (j := xH); auto.
Qed.
Theorem elements_extensional:
@@ -836,8 +836,8 @@ Module PTree <: TREE.
(forall i, get i m = get i n) ->
elements m = elements n.
Proof.
- intros.
- exploit (elements_canonical_order (fun (x y: A) => x = y) m n).
+ intros.
+ exploit (elements_canonical_order (fun (x y: A) => x = y) m n).
intros. rewrite H in H0. exists x; auto.
intros. rewrite <- H in H0. exists y; auto.
induction 1. auto. destruct a1 as [a2 a3]; destruct b1 as [b2 b3]; simpl in *.
@@ -862,8 +862,8 @@ Module PTree <: TREE.
{
destruct i; simpl remove.
destruct m1; auto. destruct o; auto. destruct (remove i m2); auto.
- destruct o; auto. destruct m2; auto. destruct (remove i m1); auto.
- destruct m1; auto. destruct m2; auto.
+ destruct o; auto. destruct m2; auto. destruct (remove i m1); auto.
+ destruct m1; auto. destruct m2; auto.
}
rewrite REMOVE. destruct i; simpl in H.
+ destruct (IHm2 i (xI j) H) as (l1 & l2 & EQ & EQ').
@@ -883,7 +883,7 @@ Module PTree <: TREE.
rewrite xelements_node, EQ', ! app_ass. auto.
+ subst o. exists (xelements m1 (xO j) nil); exists (xelements m2 (xI j) nil); split.
rewrite xelements_node. rewrite prev_append_prev. auto.
- rewrite xelements_node; auto.
+ rewrite xelements_node; auto.
Qed.
Theorem elements_remove:
@@ -891,8 +891,8 @@ Module PTree <: TREE.
get i m = Some v ->
exists l1 l2, elements m = l1 ++ (i,v) :: l2 /\ elements (remove i m) = l1 ++ l2.
Proof.
- intros. exploit xelements_remove. eauto. instantiate (1 := xH).
- rewrite prev_append_prev. auto.
+ intros. exploit xelements_remove. eauto. instantiate (1 := xH).
+ rewrite prev_append_prev. auto.
Qed.
Fixpoint xfold (A B: Type) (f: B -> positive -> A -> B)
@@ -920,7 +920,7 @@ Module PTree <: TREE.
simpl. auto.
destruct o; simpl.
rewrite <- IHm1. simpl. rewrite <- IHm2. auto.
- rewrite <- IHm1. rewrite <- IHm2. auto.
+ rewrite <- IHm1. rewrite <- IHm2. auto.
Qed.
Theorem fold_spec:
@@ -928,7 +928,7 @@ Module PTree <: TREE.
fold f m v =
List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v.
Proof.
- intros. unfold fold, elements. rewrite <- xfold_xelements. auto.
+ intros. unfold fold, elements. rewrite <- xfold_xelements. auto.
Qed.
Fixpoint fold1 (A B: Type) (f: B -> A -> B) (m: t A) (v: B) {struct m} : B :=
@@ -952,7 +952,7 @@ Module PTree <: TREE.
simpl. auto.
destruct o; simpl.
rewrite <- IHm1. simpl. rewrite <- IHm2. auto.
- rewrite <- IHm1. rewrite <- IHm2. auto.
+ rewrite <- IHm1. rewrite <- IHm2. auto.
Qed.
Theorem fold1_spec:
@@ -960,7 +960,7 @@ Module PTree <: TREE.
fold1 f m v =
List.fold_left (fun a p => f a (snd p)) (elements m) v.
Proof.
- intros. apply fold1_xelements with (l := @nil (positive * A)).
+ intros. apply fold1_xelements with (l := @nil (positive * A)).
Qed.
End PTree.
@@ -1064,7 +1064,7 @@ Module IMap(X: INDEXED_TYPE).
Lemma gi:
forall (A: Type) (x: A) (i: X.t), get i (init x) = x.
Proof.
- intros. unfold get, init. apply PMap.gi.
+ intros. unfold get, init. apply PMap.gi.
Qed.
Lemma gss:
@@ -1077,19 +1077,19 @@ Module IMap(X: INDEXED_TYPE).
forall (A: Type) (i j: X.t) (x: A) (m: t A),
i <> j -> get i (set j x m) = get i m.
Proof.
- intros. unfold get, set. apply PMap.gso.
- red. intro. apply H. apply X.index_inj; auto.
+ intros. unfold get, set. apply PMap.gso.
+ red. intro. apply H. apply X.index_inj; auto.
Qed.
Lemma gsspec:
forall (A: Type) (i j: X.t) (x: A) (m: t A),
get i (set j x m) = if X.eq i j then x else get i m.
Proof.
- intros. unfold get, set.
+ intros. unfold get, set.
rewrite PMap.gsspec.
case (X.eq i j); intro.
subst j. rewrite peq_true. reflexivity.
- rewrite peq_false. reflexivity.
+ rewrite peq_false. reflexivity.
red; intro. elim n. apply X.index_inj; auto.
Qed.
@@ -1097,7 +1097,7 @@ Module IMap(X: INDEXED_TYPE).
forall (A B: Type) (f: A -> B) (i: X.t) (m: t A),
get i (map f m) = f(get i m).
Proof.
- intros. unfold map, get. apply PMap.gmap.
+ intros. unfold map, get. apply PMap.gmap.
Qed.
Lemma set2:
@@ -1225,7 +1225,7 @@ Hypothesis P_compat:
(forall x, T.get x m = T.get x m') ->
P m a -> P m' a.
-Hypothesis H_base:
+Hypothesis H_base:
P (T.empty _) init.
Hypothesis H_rec:
@@ -1253,23 +1253,23 @@ Remark H_rec':
P' l a ->
P' (l ++ (k, v) :: nil) (f a k v).
Proof.
- unfold P'; intros.
- set (m0 := T.remove k m).
+ unfold P'; intros.
+ set (m0 := T.remove k m).
apply P_compat with (T.set k v m0).
intros. unfold m0. rewrite T.gsspec. destruct (T.elt_eq x k).
symmetry. apply T.elements_complete. rewrite <- (H2 (x, v)).
apply in_or_app. simpl. intuition congruence.
apply T.gro. auto.
- apply H_rec. unfold m0. apply T.grs. apply T.elements_complete. auto.
- apply H1. red. intros [k' v'].
- split; intros.
- apply T.elements_correct. unfold m0. rewrite T.gro. apply T.elements_complete.
- rewrite <- (H2 (k', v')). apply in_or_app. auto.
+ apply H_rec. unfold m0. apply T.grs. apply T.elements_complete. auto.
+ apply H1. red. intros [k' v'].
+ split; intros.
+ apply T.elements_correct. unfold m0. rewrite T.gro. apply T.elements_complete.
+ rewrite <- (H2 (k', v')). apply in_or_app. auto.
red; intro; subst k'. elim H. change k with (fst (k, v')). apply in_map. auto.
assert (T.get k' m0 = Some v'). apply T.elements_complete. auto.
unfold m0 in H4. rewrite T.grspec in H4. destruct (T.elt_eq k' k). congruence.
assert (In (k', v') (T.elements m)). apply T.elements_correct; auto.
- rewrite <- (H2 (k', v')) in H5. destruct (in_app_or _ _ _ H5). auto.
+ rewrite <- (H2 (k', v')) in H5. destruct (in_app_or _ _ _ H5). auto.
simpl in H6. intuition congruence.
Qed.
@@ -1282,10 +1282,10 @@ Lemma fold_rec_aux:
Proof.
induction l1; intros; simpl.
rewrite <- List.app_nil_end. auto.
- destruct a as [k v]; simpl in *. inv H1.
+ destruct a as [k v]; simpl in *. inv H1.
change ((k, v) :: l1) with (((k, v) :: nil) ++ l1). rewrite <- List.app_ass. apply IHl1.
rewrite app_ass. auto.
- red; intros. rewrite map_app in H3. destruct (in_app_or _ _ _ H3). apply H0; auto with coqlib.
+ red; intros. rewrite map_app in H3. destruct (in_app_or _ _ _ H3). apply H0; auto with coqlib.
simpl in H4. intuition congruence.
auto.
unfold f'. simpl. apply H_rec'; auto. eapply list_disjoint_notin; eauto with coqlib.
@@ -1300,8 +1300,8 @@ Proof.
apply fold_rec_aux.
simpl. red; intros; tauto.
simpl. red; intros. elim H0.
- apply T.elements_keys_norepet.
- apply H_base'.
+ apply T.elements_keys_norepet.
+ apply H_base'.
simpl in H. red in H. apply H. red; intros. tauto.
Qed.
@@ -1319,18 +1319,18 @@ Theorem cardinal_remove:
forall x m y, T.get x m = Some y -> (cardinal (T.remove x m) < cardinal m)%nat.
Proof.
unfold cardinal; intros.
- exploit T.elements_remove; eauto. intros (l1 & l2 & P & Q).
+ exploit T.elements_remove; eauto. intros (l1 & l2 & P & Q).
rewrite P, Q. rewrite ! app_length. simpl. omega.
Qed.
Theorem cardinal_set:
forall x m y, T.get x m = None -> (cardinal m < cardinal (T.set x y m))%nat.
Proof.
- intros. set (m' := T.set x y m).
- replace (cardinal m) with (cardinal (T.remove x m')).
- apply cardinal_remove with y. unfold m'; apply T.gss.
- unfold cardinal. f_equal. apply T.elements_extensional.
- intros. unfold m'. rewrite T.grspec, T.gsspec.
+ intros. set (m' := T.set x y m).
+ replace (cardinal m) with (cardinal (T.remove x m')).
+ apply cardinal_remove with y. unfold m'; apply T.gss.
+ unfold cardinal. f_equal. apply T.elements_extensional.
+ intros. unfold m'. rewrite T.grspec, T.gsspec.
destruct (T.elt_eq i x); auto. congruence.
Qed.
@@ -1352,16 +1352,16 @@ Proof.
intros m0 f.
unfold for_all. apply fold_rec; intros.
- (* Extensionality *)
- rewrite H0. split; intros. rewrite <- H in H2; auto. rewrite H in H2; auto.
+ rewrite H0. split; intros. rewrite <- H in H2; auto. rewrite H in H2; auto.
- (* Base case *)
split; intros. rewrite T.gempty in H0; congruence. auto.
- (* Inductive case *)
split; intros.
- destruct (andb_prop _ _ H2). rewrite T.gsspec in H3. destruct (T.elt_eq x k).
+ destruct (andb_prop _ _ H2). rewrite T.gsspec in H3. destruct (T.elt_eq x k).
inv H3. auto.
apply H1; auto.
- apply andb_true_intro. split.
- rewrite H1. intros. apply H2. rewrite T.gso; auto. congruence.
+ apply andb_true_intro. split.
+ rewrite H1. intros. apply H2. rewrite T.gso; auto. congruence.
apply H2. apply T.gss.
Qed.
@@ -1375,7 +1375,7 @@ Proof.
intros m0 f.
unfold exists_. apply fold_rec; intros.
- (* Extensionality *)
- rewrite H0. split; intros (x0 & a0 & P & Q); exists x0; exists a0; split; auto; congruence.
+ rewrite H0. split; intros (x0 & a0 & P & Q); exists x0; exists a0; split; auto; congruence.
- (* Base case *)
split; intros. congruence. destruct H as (x & a & P & Q). rewrite T.gempty in P; congruence.
- (* Inductive case *)
@@ -1383,7 +1383,7 @@ Proof.
destruct (orb_true_elim _ _ H2).
rewrite H1 in e. destruct e as (x1 & a1 & P & Q).
exists x1; exists a1; split; auto. rewrite T.gso; auto. congruence.
- exists k; exists v; split; auto. apply T.gss.
+ exists k; exists v; split; auto. apply T.gss.
destruct H2 as (x1 & a1 & P & Q). apply orb_true_intro.
rewrite T.gsspec in P. destruct (T.elt_eq x1 k).
inv P. right; auto.
@@ -1394,11 +1394,11 @@ Remark exists_for_all:
forall m f,
exists_ m f = negb (for_all m (fun x a => negb (f x a))).
Proof.
- intros. unfold exists_, for_all. rewrite ! T.fold_spec.
- change false with (negb true). generalize (T.elements m) true.
+ intros. unfold exists_, for_all. rewrite ! T.fold_spec.
+ change false with (negb true). generalize (T.elements m) true.
induction l; simpl; intros.
auto.
- rewrite <- IHl. f_equal.
+ rewrite <- IHl. f_equal.
destruct b; destruct (f (fst a) (snd a)); reflexivity.
Qed.
@@ -1406,11 +1406,11 @@ Remark for_all_exists:
forall m f,
for_all m f = negb (exists_ m (fun x a => negb (f x a))).
Proof.
- intros. unfold exists_, for_all. rewrite ! T.fold_spec.
- change true with (negb false). generalize (T.elements m) false.
+ intros. unfold exists_, for_all. rewrite ! T.fold_spec.
+ change true with (negb false). generalize (T.elements m) false.
induction l; simpl; intros.
auto.
- rewrite <- IHl. f_equal.
+ rewrite <- IHl. f_equal.
destruct b; destruct (f (fst a) (snd a)); reflexivity.
Qed.
@@ -1418,20 +1418,20 @@ Lemma for_all_false:
forall m f,
for_all m f = false <-> (exists x a, T.get x m = Some a /\ f x a = false).
Proof.
- intros. rewrite for_all_exists.
- rewrite negb_false_iff. rewrite exists_correct.
- split; intros (x & a & P & Q); exists x; exists a; split; auto.
+ intros. rewrite for_all_exists.
+ rewrite negb_false_iff. rewrite exists_correct.
+ split; intros (x & a & P & Q); exists x; exists a; split; auto.
rewrite negb_true_iff in Q. auto.
- rewrite Q; auto.
+ rewrite Q; auto.
Qed.
Lemma exists_false:
forall m f,
exists_ m f = false <-> (forall x a, T.get x m = Some a -> f x a = false).
Proof.
- intros. rewrite exists_for_all.
+ intros. rewrite exists_for_all.
rewrite negb_false_iff. rewrite for_all_correct.
- split; intros. apply H in H0. rewrite negb_true_iff in H0. auto. rewrite H; auto.
+ split; intros. apply H in H0. rewrite negb_true_iff in H0. auto. rewrite H; auto.
Qed.
End FORALL_EXISTS.
@@ -1457,20 +1457,20 @@ Proof.
set (p1 := fun x a1 => match T.get x m2 with None => false | Some a2 => beqA a1 a2 end).
set (p2 := fun x a2 => match T.get x m1 with None => false | Some a1 => beqA a1 a2 end).
destruct (for_all m1 p1) eqn:F1; [destruct (for_all m2 p2) eqn:F2 | idtac].
- + cut (T.beq beqA m1 m2 = true). congruence.
- rewrite for_all_correct in *. rewrite T.beq_correct; intros.
- destruct (T.get x m1) as [a1|] eqn:X1.
+ + cut (T.beq beqA m1 m2 = true). congruence.
+ rewrite for_all_correct in *. rewrite T.beq_correct; intros.
+ destruct (T.get x m1) as [a1|] eqn:X1.
generalize (F1 _ _ X1). unfold p1. destruct (T.get x m2); congruence.
destruct (T.get x m2) as [a2|] eqn:X2; auto.
- generalize (F2 _ _ X2). unfold p2. rewrite X1. congruence.
+ generalize (F2 _ _ X2). unfold p2. rewrite X1. congruence.
+ rewrite for_all_false in F2. destruct F2 as (x & a & P & Q).
- exists x. rewrite P. unfold p2 in Q. destruct (T.get x m1); auto.
+ exists x. rewrite P. unfold p2 in Q. destruct (T.get x m1); auto.
+ rewrite for_all_false in F1. destruct F1 as (x & a & P & Q).
exists x. rewrite P. unfold p1 in Q. destruct (T.get x m2); auto.
- (* existence -> beq = false *)
destruct H as [x P].
- destruct (T.beq beqA m1 m2) eqn:E; auto.
- rewrite T.beq_correct in E.
+ destruct (T.beq beqA m1 m2) eqn:E; auto.
+ rewrite T.beq_correct in E.
generalize (E x). destruct (T.get x m1); destruct (T.get x m2); tauto || congruence.
Qed.
@@ -1493,7 +1493,7 @@ Definition Equal (m1 m2: T.t A) : Prop :=
Lemma Equal_refl: forall m, Equal m m.
Proof.
- intros; red; intros. destruct (T.get x m); auto. reflexivity.
+ intros; red; intros. destruct (T.get x m); auto. reflexivity.
Qed.
Lemma Equal_sym: forall m1 m2, Equal m1 m2 -> Equal m2 m1.
@@ -1505,7 +1505,7 @@ Lemma Equal_trans: forall m1 m2 m3, Equal m1 m2 -> Equal m2 m3 -> Equal m1 m3.
Proof.
intros; red; intros. generalize (H x) (H0 x).
destruct (T.get x m1); destruct (T.get x m2); try tauto;
- destruct (T.get x m3); try tauto.
+ destruct (T.get x m3); try tauto.
intros. transitivity a0; auto.
Qed.
@@ -1525,15 +1525,15 @@ Program Definition Equal_dec (m1 m2: T.t A) : { m1 === m2 } + { m1 =/= m2 } :=
Next Obligation.
rename Heq_anonymous into B.
symmetry in B. rewrite T.beq_correct in B.
- red; intros. generalize (B x).
- destruct (T.get x m1); destruct (T.get x m2); auto.
- intros. eapply proj_sumbool_true; eauto.
+ red; intros. generalize (B x).
+ destruct (T.get x m1); destruct (T.get x m2); auto.
+ intros. eapply proj_sumbool_true; eauto.
Qed.
Next Obligation.
assert (T.beq (fun a1 a2 => proj_sumbool (a1 == a2)) m1 m2 = true).
- apply T.beq_correct; intros.
- generalize (H x).
- destruct (T.get x m1); destruct (T.get x m2); try tauto.
+ apply T.beq_correct; intros.
+ generalize (H x).
+ destruct (T.get x m1); destruct (T.get x m2); try tauto.
intros. apply proj_sumbool_is_true; auto.
unfold equiv, complement in H0. congruence.
Qed.
diff --git a/lib/Ordered.v b/lib/Ordered.v
index 5d02586d..a2c36673 100644
--- a/lib/Ordered.v
+++ b/lib/Ordered.v
@@ -31,7 +31,7 @@ Definition eq (x y: t) := x = y.
Definition lt := Plt.
Lemma eq_refl : forall x : t, eq x x.
-Proof (@refl_equal t).
+Proof (@refl_equal t).
Lemma eq_sym : forall x y : t, eq x y -> eq y x.
Proof (@sym_equal t).
Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
@@ -61,7 +61,7 @@ Definition eq (x y: t) := x = y.
Definition lt := Zlt.
Lemma eq_refl : forall x : t, eq x x.
-Proof (@refl_equal t).
+Proof (@refl_equal t).
Lemma eq_sym : forall x y : t, eq x y -> eq y x.
Proof (@sym_equal t).
Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
@@ -69,7 +69,7 @@ Proof (@trans_equal t).
Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
Proof Zlt_trans.
Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
-Proof. unfold lt, eq, t; intros. omega. Qed.
+Proof. unfold lt, eq, t; intros. omega. Qed.
Lemma compare : forall x y : t, Compare lt eq x y.
Proof.
intros. destruct (Z.compare x y) as [] eqn:E.
@@ -91,7 +91,7 @@ Definition eq (x y: t) := x = y.
Definition lt (x y: t) := Int.unsigned x < Int.unsigned y.
Lemma eq_refl : forall x : t, eq x x.
-Proof (@refl_equal t).
+Proof (@refl_equal t).
Lemma eq_sym : forall x y : t, eq x y -> eq y x.
Proof (@sym_equal t).
Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
@@ -129,7 +129,7 @@ Definition eq (x y: t) := x = y.
Definition lt (x y: t) := Plt (A.index x) (A.index y).
Lemma eq_refl : forall x : t, eq x x.
-Proof (@refl_equal t).
+Proof (@refl_equal t).
Lemma eq_sym : forall x y : t, eq x y -> eq y x.
Proof (@sym_equal t).
Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
@@ -149,7 +149,7 @@ Qed.
Lemma compare : forall x y : t, Compare lt eq x y.
Proof.
intros. case (OrderedPositive.compare (A.index x) (A.index y)); intro.
- apply LT. exact l.
+ apply LT. exact l.
apply EQ. red; red in e. apply A.index_inj; auto.
apply GT. exact l.
Defined.
@@ -158,7 +158,7 @@ Lemma eq_dec : forall x y, { eq x y } + { ~ eq x y }.
Proof.
intros. case (peq (A.index x) (A.index y)); intros.
left. apply A.index_inj; auto.
- right; red; unfold eq; intros; subst. congruence.
+ right; red; unfold eq; intros; subst. congruence.
Defined.
End OrderedIndexed.
@@ -211,7 +211,7 @@ Proof.
case (A.compare (fst x) (fst z)); intro.
assumption.
generalize (A.lt_not_eq H1); intro. elim H5.
- apply A.eq_trans with (fst x).
+ apply A.eq_trans with (fst x).
apply A.eq_sym. auto. auto.
generalize (@A.lt_not_eq (fst y) (fst x)); intro.
elim H5. apply A.lt_trans with (fst z); auto.
@@ -231,7 +231,7 @@ Proof.
elim H3; intros.
apply (@B.lt_not_eq _ _ H5 H2).
Qed.
-
+
Lemma compare : forall x y : t, Compare lt eq x y.
Proof.
intros.
diff --git a/lib/Parmov.v b/lib/Parmov.v
index f96a692e..92bba559 100644
--- a/lib/Parmov.v
+++ b/lib/Parmov.v
@@ -61,7 +61,7 @@ Section PARMOV.
(** * Registers, moves, and their semantics *)
-(** The development is parameterized by the type of registers,
+(** The development is parameterized by the type of registers,
equipped with a decidable equality. *)
Variable reg: Type.
@@ -102,7 +102,7 @@ Lemma env_ext:
(forall r, e1 r = e2 r) -> e1 = e2.
Proof (@extensionality reg val).
-(** The main operation over environments is update: it assigns
+(** The main operation over environments is update: it assigns
a value [v] to a register [r] and preserves the values of other
registers. *)
@@ -132,7 +132,7 @@ Lemma update_commut:
r1 <> r2 ->
update r1 v1 (update r2 v2 e) = update r2 v2 (update r1 v1 e).
Proof.
- intros. apply env_ext; intro; unfold update.
+ intros. apply env_ext; intro; unfold update.
destruct (reg_eq r r1); destruct (reg_eq r r2); auto.
congruence.
Qed.
@@ -174,7 +174,7 @@ Fixpoint exec_seq (m: moves) (e: env) {struct m}: env :=
Fixpoint exec_seq_rev (m: moves) (e: env) {struct m}: env :=
match m with
| nil => e
- | (s, d) :: m' =>
+ | (s, d) :: m' =>
let e' := exec_seq_rev m' e in
update d (e' s) e'
end.
@@ -274,7 +274,7 @@ Inductive state_wf: state -> Prop :=
Lemma dests_append:
forall m1 m2, dests (m1 ++ m2) = dests m1 ++ dests m2.
Proof.
- intros. unfold dests. apply map_app.
+ intros. unfold dests. apply map_app.
Qed.
Lemma dests_decomp:
@@ -286,7 +286,7 @@ Qed.
Lemma srcs_append:
forall m1 m2, srcs (m1 ++ m2) = srcs m1 ++ srcs m2.
Proof.
- intros. unfold srcs. apply map_app.
+ intros. unfold srcs. apply map_app.
Qed.
Lemma srcs_decomp:
@@ -317,7 +317,7 @@ Lemma dests_disjoint_sym:
forall m1 m2,
dests_disjoint m1 m2 <-> dests_disjoint m2 m1.
Proof.
- unfold dests_disjoint; intros.
+ unfold dests_disjoint; intros.
split; intros; apply list_disjoint_sym; auto.
Qed.
@@ -326,9 +326,9 @@ Lemma dests_disjoint_cons_left:
dests_disjoint ((s, d) :: m1) m2 <->
dests_disjoint m1 m2 /\ ~In d (dests m2).
Proof.
- unfold dests_disjoint, list_disjoint.
+ unfold dests_disjoint, list_disjoint.
simpl; intros; split; intros.
- split. auto. firstorder.
+ split. auto. firstorder.
destruct H. elim H0; intro.
red; intro; subst. contradiction.
apply H; auto.
@@ -339,7 +339,7 @@ Lemma dests_disjoint_cons_right:
dests_disjoint m1 ((s, d) :: m2) <->
dests_disjoint m1 m2 /\ ~In d (dests m1).
Proof.
- intros. rewrite dests_disjoint_sym. rewrite dests_disjoint_cons_left.
+ intros. rewrite dests_disjoint_sym. rewrite dests_disjoint_cons_left.
rewrite dests_disjoint_sym. tauto.
Qed.
@@ -348,11 +348,11 @@ Lemma dests_disjoint_append_left:
dests_disjoint (m1 ++ m2) m3 <->
dests_disjoint m1 m3 /\ dests_disjoint m2 m3.
Proof.
- unfold dests_disjoint, list_disjoint.
+ unfold dests_disjoint, list_disjoint.
intros; split; intros. split; intros.
apply H; eauto. rewrite dests_append. apply in_or_app. auto.
apply H; eauto. rewrite dests_append. apply in_or_app. auto.
- destruct H.
+ destruct H.
rewrite dests_append in H0. elim (in_app_or _ _ _ H0); auto.
Qed.
@@ -361,7 +361,7 @@ Lemma dests_disjoint_append_right:
dests_disjoint m1 (m2 ++ m3) <->
dests_disjoint m1 m2 /\ dests_disjoint m1 m3.
Proof.
- intros. rewrite dests_disjoint_sym. rewrite dests_disjoint_append_left.
+ intros. rewrite dests_disjoint_sym. rewrite dests_disjoint_append_left.
intuition; rewrite dests_disjoint_sym; assumption.
Qed.
@@ -371,7 +371,7 @@ Lemma is_mill_cons:
is_mill m /\ ~In d (dests m).
Proof.
unfold is_mill, dests_disjoint; intros. simpl.
- split; intros.
+ split; intros.
inversion H; tauto.
constructor; tauto.
Qed.
@@ -381,7 +381,7 @@ Lemma is_mill_append:
is_mill (m1 ++ m2) <->
is_mill m1 /\ is_mill m2 /\ dests_disjoint m1 m2.
Proof.
- unfold is_mill, dests_disjoint; intros. rewrite dests_append.
+ unfold is_mill, dests_disjoint; intros. rewrite dests_append.
apply list_norepet_app.
Qed.
@@ -391,7 +391,7 @@ Lemma move_no_temp_append:
forall m1 m2,
move_no_temp m1 -> move_no_temp m2 -> move_no_temp (m1 ++ m2).
Proof.
- intros; red; intros. elim (in_app_or _ _ _ H1); intro.
+ intros; red; intros. elim (in_app_or _ _ _ H1); intro.
apply H; auto. apply H0; auto.
Qed.
@@ -418,12 +418,12 @@ Lemma temp_last_push:
is_not_temp s1 -> is_not_temp d1 ->
temp_last ((s1, d1) :: (s2, d2) :: sigma).
Proof.
- unfold temp_last; intros. simpl. simpl in H.
+ unfold temp_last; intros. simpl. simpl in H.
destruct (rev sigma); simpl in *.
- intuition. red; simpl; intros.
- elim H; intros. inversion H4. subst; tauto. tauto.
- destruct p as [sN dN]. intuition.
- red; intros. elim (in_app_or _ _ _ H); intros.
+ intuition. red; simpl; intros.
+ elim H; intros. inversion H4. subst; tauto. tauto.
+ destruct p as [sN dN]. intuition.
+ red; intros. elim (in_app_or _ _ _ H); intros.
apply H3; auto.
elim H4; intros. inversion H5; subst; tauto. elim H5.
Qed.
@@ -433,12 +433,12 @@ Lemma temp_last_pop:
temp_last ((s1, d1) :: sigma ++ (s2, d2) :: nil) ->
temp_last (sigma ++ (s2, d2) :: nil).
Proof.
- intros until d2.
+ intros until d2.
change ((s1, d1) :: sigma ++ (s2, d2) :: nil)
with ((((s1, d1) :: nil) ++ sigma) ++ ((s2, d2) :: nil)).
- unfold temp_last. repeat rewrite rev_unit.
- intuition. simpl in H1. red; intros. apply H1.
- apply in_or_app. auto.
+ unfold temp_last. repeat rewrite rev_unit.
+ intuition. simpl in H1. red; intros. apply H1.
+ apply in_or_app. auto.
Qed.
(** Some properties of [is_path]. *)
@@ -458,7 +458,7 @@ Proof.
induction sigma; simpl; intros.
constructor. red; auto. constructor.
inversion H; subst; clear H.
- constructor.
+ constructor.
destruct sigma as [ | [s1 d1] sigma']; simpl; simpl in H2; auto.
auto.
Qed.
@@ -471,12 +471,12 @@ Lemma path_sources_dests:
Proof.
induction sigma; simpl; intros.
red; simpl; tauto.
- inversion H; subst; clear H. simpl.
+ inversion H; subst; clear H. simpl.
assert (In s (dests (sigma ++ (s0, d0) :: nil))).
- destruct sigma as [ | [s1 d1] sigma']; simpl; simpl in H2; intuition.
- apply incl_cons. simpl; tauto.
+ destruct sigma as [ | [s1 d1] sigma']; simpl; simpl in H2; intuition.
+ apply incl_cons. simpl; tauto.
apply incl_tran with (s0 :: dests (sigma ++ (s0, d0) :: nil)).
- eapply IHsigma; eauto.
+ eapply IHsigma; eauto.
red; simpl; tauto.
Qed.
@@ -484,11 +484,11 @@ Lemma no_read_path:
forall d1 sigma s0 d0,
d1 <> s0 ->
is_path (sigma ++ (s0, d0) :: nil) ->
- ~In d1 (dests (sigma ++ (s0, d0) :: nil)) ->
+ ~In d1 (dests (sigma ++ (s0, d0) :: nil)) ->
no_read (sigma ++ (s0, d0) :: nil) d1.
Proof.
intros.
- generalize (path_sources_dests _ _ _ H0). intro.
+ generalize (path_sources_dests _ _ _ H0). intro.
intro. elim H1. elim (H2 _ H3); intro. congruence. auto.
Qed.
@@ -499,7 +499,7 @@ Lemma notin_dests_cons:
forall x s d m,
~In x (dests ((s, d) :: m)) <-> x <> d /\ ~In x (dests m).
Proof.
- intros. simpl. intuition auto.
+ intros. simpl. intuition auto.
Qed.
Lemma notin_dests_append:
@@ -509,7 +509,7 @@ Proof.
intros. rewrite dests_append. rewrite in_app. tauto.
Qed.
-Hint Rewrite is_mill_cons is_mill_append
+Hint Rewrite is_mill_cons is_mill_append
dests_disjoint_cons_left dests_disjoint_cons_right
dests_disjoint_append_left dests_disjoint_append_right
notin_dests_cons notin_dests_append: pmov.
@@ -525,29 +525,29 @@ Proof.
autorewrite with pmov in A; constructor; autorewrite with pmov.
(* Nop *)
- tauto.
+ tauto.
red; intros. apply B. apply list_in_insert; auto.
auto. auto.
(* Start *)
tauto.
red; intros. apply B. apply list_in_insert; auto.
- red. simpl. split. elim (B s d). auto.
- apply in_or_app. right. apply in_eq.
+ red. simpl. split. elim (B s d). auto.
+ apply in_or_app. right. apply in_eq.
red; simpl; tauto.
- constructor. exact I. constructor.
+ constructor. exact I. constructor.
(* Push *)
intuition.
red; intros. apply B. apply list_in_insert; auto.
- elim (B d r). apply temp_last_push; auto.
+ elim (B d r). apply temp_last_push; auto.
apply in_or_app; right; apply in_eq.
constructor. simpl. auto. auto.
(* Loop *)
- tauto.
+ tauto.
auto.
- eapply temp_last_change_last_source; eauto.
+ eapply temp_last_change_last_source; eauto.
eapply is_path_change_last_source; eauto.
(* Pop *)
@@ -557,7 +557,7 @@ Proof.
eapply is_path_pop; eauto.
(* Last *)
- intuition.
+ intuition.
auto.
unfold temp_last. simpl. auto.
constructor.
@@ -577,7 +577,7 @@ Qed.
reverse sequential order, then the moves [mu ++ sigma] in parallel. *)
Definition statemove (st: state) (e: env) :=
- match st with
+ match st with
| State mu sigma tau => exec_par (mu ++ sigma) (exec_seq_rev tau e)
end.
@@ -589,7 +589,7 @@ Lemma exec_par_outside:
forall m e r, ~In r (dests m) -> exec_par m e r = e r.
Proof.
induction m; simpl; intros. auto.
- destruct a as [s d]. rewrite update_o. apply IHm. tauto.
+ destruct a as [s d]. rewrite update_o. apply IHm. tauto.
simpl in H. intuition.
Qed.
@@ -600,8 +600,8 @@ Lemma exec_par_lift:
Proof.
induction m1; simpl; intros.
auto.
- destruct a as [s0 d0]. simpl in H. rewrite IHm1. simpl.
- apply update_commut. tauto. tauto.
+ destruct a as [s0 d0]. simpl in H. rewrite IHm1. simpl.
+ apply update_commut. tauto. tauto.
Qed.
Lemma exec_par_ident:
@@ -609,9 +609,9 @@ Lemma exec_par_ident:
is_mill (m1 ++ (r, r) :: m2) ->
exec_par (m1 ++ (r, r) :: m2) e = exec_par (m1 ++ m2) e.
Proof.
- intros. autorewrite with pmov in H.
- rewrite exec_par_lift. simpl.
- replace (e r) with (exec_par (m1 ++ m2) e r). apply update_ident.
+ intros. autorewrite with pmov in H.
+ rewrite exec_par_lift. simpl.
+ replace (e r) with (exec_par (m1 ++ m2) e r). apply update_ident.
apply exec_par_outside. autorewrite with pmov. tauto. tauto.
Qed.
@@ -619,7 +619,7 @@ Lemma exec_seq_app:
forall m1 m2 e,
exec_seq (m1 ++ m2) e = exec_seq m2 (exec_seq m1 e).
Proof.
- induction m1; simpl; intros. auto.
+ induction m1; simpl; intros. auto.
destruct a as [s d]. rewrite IHm1. auto.
Qed.
@@ -627,7 +627,7 @@ Lemma exec_seq_rev_app:
forall m1 m2 e,
exec_seq_rev (m1 ++ m2) e = exec_seq_rev m1 (exec_seq_rev m2 e).
Proof.
- induction m1; simpl; intros. auto.
+ induction m1; simpl; intros. auto.
destruct a as [s d]. rewrite IHm1. auto.
Qed.
@@ -657,8 +657,8 @@ Lemma exec_par_update_no_read:
Proof.
unfold no_read; induction m; simpl; intros.
auto.
- destruct a as [s0 d0]; simpl in *. rewrite IHm.
- rewrite update_commut. f_equal. f_equal.
+ destruct a as [s0 d0]; simpl in *. rewrite IHm.
+ rewrite update_commut. f_equal. f_equal.
apply update_o. tauto. tauto. tauto. tauto.
Qed.
@@ -682,14 +682,14 @@ Lemma exec_par_combine:
Proof.
induction sl; destruct dl; simpl; intros; try discriminate.
split; auto.
- inversion H0; subst; clear H0.
+ inversion H0; subst; clear H0.
injection H; intro; clear H.
destruct (IHsl dl H0 H4) as [A B].
set (e' := exec_par (combine sl dl) e) in *.
split.
- decEq. apply update_s.
+ decEq. apply update_s.
rewrite <- A. apply list_map_exten; intros.
- rewrite update_o. auto. congruence.
+ rewrite update_o. auto. congruence.
intros. rewrite update_o. apply B. tauto. intuition.
Qed.
@@ -733,10 +733,10 @@ Lemma exec_par_env_equiv:
Proof.
unfold move_no_temp; induction m; simpl; intros.
auto.
- destruct a as [s d].
+ destruct a as [s d].
red; intros. unfold update. destruct (reg_eq r d).
- apply H0. elim (H s d); tauto.
- apply IHm; auto.
+ apply H0. elim (H s d); tauto.
+ apply IHm; auto.
Qed.
(** The proof that transitions preserve semantics (up to the values of
@@ -750,57 +750,57 @@ Proof.
induction 1; intro WF; inversion WF as [mu0 sigma0 tau0 A B C D]; subst; simpl.
(* nop *)
- apply env_equiv_refl'. unfold statemove.
- repeat rewrite app_ass. simpl. symmetry. apply exec_par_ident.
- rewrite app_ass in A. exact A.
+ apply env_equiv_refl'. unfold statemove.
+ repeat rewrite app_ass. simpl. symmetry. apply exec_par_ident.
+ rewrite app_ass in A. exact A.
(* start *)
- apply env_equiv_refl'. unfold statemove.
- autorewrite with pmov in A.
+ apply env_equiv_refl'. unfold statemove.
+ autorewrite with pmov in A.
rewrite exec_par_lift. repeat rewrite app_ass. simpl. rewrite exec_par_lift. reflexivity.
- tauto. autorewrite with pmov. tauto.
+ tauto. autorewrite with pmov. tauto.
(* push *)
- apply env_equiv_refl'. unfold statemove.
+ apply env_equiv_refl'. unfold statemove.
autorewrite with pmov in A.
rewrite exec_par_lift. rewrite exec_par_lift. simpl.
- rewrite exec_par_lift. repeat rewrite app_ass. simpl. rewrite exec_par_lift.
- simpl. apply update_commut. intuition.
- tauto. autorewrite with pmov; tauto.
- autorewrite with pmov; intuition.
+ rewrite exec_par_lift. repeat rewrite app_ass. simpl. rewrite exec_par_lift.
+ simpl. apply update_commut. intuition.
+ tauto. autorewrite with pmov; tauto.
+ autorewrite with pmov; intuition.
autorewrite with pmov; intuition.
(* loop *)
- unfold statemove. simpl exec_seq_rev.
- set (e1 := exec_seq_rev tau e).
+ unfold statemove. simpl exec_seq_rev.
+ set (e1 := exec_seq_rev tau e).
autorewrite with pmov in A.
- repeat rewrite <- app_ass.
- assert (~In d (dests (mu ++ sigma))). autorewrite with pmov. tauto.
- repeat rewrite exec_par_lift; auto. simpl.
+ repeat rewrite <- app_ass.
+ assert (~In d (dests (mu ++ sigma))). autorewrite with pmov. tauto.
+ repeat rewrite exec_par_lift; auto. simpl.
repeat rewrite <- app_nil_end.
assert (move_no_temp (mu ++ sigma)).
- red in C. rewrite rev_unit in C. destruct C.
+ red in C. rewrite rev_unit in C. destruct C.
apply move_no_temp_append; auto. apply move_no_temp_rev; auto.
set (e2 := update (temp s) (e1 s) e1).
set (e3 := exec_par (mu ++ sigma) e1).
set (e4 := exec_par (mu ++ sigma) e2).
assert (env_equiv e2 e1).
- unfold e2; red; intros. apply update_o. apply H1.
+ unfold e2; red; intros. apply update_o. apply H1.
assert (env_equiv e4 e3).
unfold e4, e3. apply exec_par_env_equiv; auto.
- red; intros. unfold update. destruct (reg_eq r d).
+ red; intros. unfold update. destruct (reg_eq r d).
unfold e2. apply update_s. apply H2. auto.
(* pop *)
- apply env_equiv_refl'. unfold statemove. simpl exec_seq_rev.
+ apply env_equiv_refl'. unfold statemove. simpl exec_seq_rev.
set (e1 := exec_seq_rev tau e).
autorewrite with pmov in A.
- apply exec_par_append_eq. simpl.
- apply exec_par_update_no_read.
- apply no_read_path; auto. eapply is_path_pop; eauto.
+ apply exec_par_append_eq. simpl.
+ apply exec_par_update_no_read.
+ apply no_read_path; auto. eapply is_path_pop; eauto.
autorewrite with pmov; tauto.
autorewrite with pmov; tauto.
- intros. apply update_o. red; intro; subst r. elim (H H1).
+ intros. apply update_o. red; intro; subst r. elim (H H1).
(* last *)
apply env_equiv_refl'. unfold statemove. simpl exec_seq_rev.
@@ -814,9 +814,9 @@ Lemma transitions_preserve_semantics:
transitions st st' -> state_wf st ->
env_equiv (statemove st' e) (statemove st e).
Proof.
- induction 1; intros.
+ induction 1; intros.
eapply transition_preserves_semantics; eauto.
- apply env_equiv_refl.
+ apply env_equiv_refl.
apply env_equiv_trans with (statemove y e); auto.
apply IHclos_refl_trans2. eapply transitions_preserve_wf; eauto.
Qed.
@@ -828,10 +828,10 @@ Lemma state_wf_start:
state_wf (State mu nil nil).
Proof.
intros. constructor. rewrite <- app_nil_end. auto.
- auto.
+ auto.
red. simpl. auto.
constructor.
-Qed.
+Qed.
(** The main correctness result in this section is the following:
if we can transition repeatedly from an initial state [mu, nil, nil]
@@ -846,10 +846,10 @@ Theorem transitions_correctness:
transitions (State mu nil nil) (State nil nil tau) ->
forall e, env_equiv (exec_seq (List.rev tau) e) (exec_par mu e).
Proof.
- intros.
+ intros.
generalize (transitions_preserve_semantics _ _ e H1
(state_wf_start _ H H0)).
- unfold statemove. simpl. rewrite <- app_nil_end.
+ unfold statemove. simpl. rewrite <- app_nil_end.
rewrite exec_seq_exec_seq_rev. auto.
Qed.
@@ -897,23 +897,23 @@ Lemma transition_determ:
transitions st st'.
Proof.
induction 1; intro; unfold transitions.
- apply rt_step. exact (tr_nop nil r mu nil tau).
- apply rt_step. exact (tr_start nil s d mu tau).
- apply rt_step. apply tr_push.
+ apply rt_step. exact (tr_nop nil r mu nil tau).
+ apply rt_step. exact (tr_start nil s d mu tau).
+ apply rt_step. apply tr_push.
eapply rt_trans.
- apply rt_step.
+ apply rt_step.
change ((s, r0) :: sigma ++ (r0, d) :: nil)
with (((s, r0) :: sigma) ++ (r0, d) :: nil).
- apply tr_loop.
- apply rt_step. simpl. apply tr_pop. auto.
- inv H0. generalize H6.
+ apply tr_loop.
+ apply rt_step. simpl. apply tr_pop. auto.
+ inv H0. generalize H6.
change ((s, r0) :: sigma ++ (r0, d) :: nil)
with (((s, r0) :: sigma) ++ (r0, d) :: nil).
unfold temp_last; rewrite List.rev_unit. intros [E F].
- elim (F s r0). unfold is_not_temp. auto.
+ elim (F s r0). unfold is_not_temp. auto.
rewrite <- List.In_rev. apply in_eq.
- apply rt_step. apply tr_pop. auto. auto.
- apply rt_step. apply tr_last. auto.
+ apply rt_step. apply tr_pop. auto. auto.
+ apply rt_step. apply tr_last. auto.
Qed.
Lemma transitions_determ:
@@ -939,9 +939,9 @@ Theorem dtransitions_correctness:
dtransitions (State mu nil nil) (State nil nil tau) ->
forall e, env_equiv (exec_seq (List.rev tau) e) (exec_par mu e).
Proof.
- intros.
+ intros.
eapply transitions_correctness; eauto.
- apply transitions_determ. auto. apply state_wf_start; auto.
+ apply transitions_determ. auto. apply state_wf_start; auto.
Qed.
(** * The compilation function *)
@@ -1017,10 +1017,10 @@ Lemma split_move_charact:
Proof.
unfold no_read. induction m; simpl; intros.
- tauto.
-- destruct a as [s d]. destruct (reg_eq s r).
+- destruct a as [s d]. destruct (reg_eq s r).
+ subst s. auto.
+ specialize (IHm r). destruct (split_move m r) as [[[before d'] after] | ].
- * destruct IHm. subst m. simpl. intuition.
+ * destruct IHm. subst m. simpl. intuition.
* simpl; intuition.
Qed.
@@ -1030,11 +1030,11 @@ Lemma is_last_source_charact:
then s = r
else s <> r.
Proof.
- induction m; simpl.
+ induction m; simpl.
destruct (reg_eq s r); congruence.
destruct a as [s0 d0]. case_eq (m ++ (s, d) :: nil); intros.
generalize (app_cons_not_nil m nil (s, d)). congruence.
- rewrite <- H. auto.
+ rewrite <- H. auto.
Qed.
Lemma replace_last_source_charact:
@@ -1055,24 +1055,24 @@ Lemma parmove_step_compatible:
dtransition st (parmove_step st).
Proof.
intros st NOTFINAL. destruct st as [mu sigma tau]. unfold parmove_step.
- case_eq mu; [intros MEQ | intros [ms md] mtl MEQ].
- case_eq sigma; [intros SEQ | intros [ss sd] stl SEQ].
- subst mu sigma. simpl in NOTFINAL. discriminate.
- simpl.
+ case_eq mu; [intros MEQ | intros [ms md] mtl MEQ].
+ case_eq sigma; [intros SEQ | intros [ss sd] stl SEQ].
+ subst mu sigma. simpl in NOTFINAL. discriminate.
+ simpl.
case_eq stl; [intros STLEQ | intros xx1 xx2 STLEQ].
apply dtr_last. red; simpl; auto.
- elim (@exists_last _ stl). 2:congruence. intros sigma1 [[ss1 sd1] SEQ2].
- rewrite <- STLEQ. clear STLEQ xx1 xx2.
- generalize (is_last_source_charact sd ss1 sd1 sigma1).
+ elim (@exists_last _ stl). 2:congruence. intros sigma1 [[ss1 sd1] SEQ2].
+ rewrite <- STLEQ. clear STLEQ xx1 xx2.
+ generalize (is_last_source_charact sd ss1 sd1 sigma1).
rewrite SEQ2. destruct (is_last_source sd (sigma1 ++ (ss1, sd1) :: nil)).
- intro. subst ss1.
- rewrite replace_last_source_charact. apply dtr_loop_pop.
+ intro. subst ss1.
+ rewrite replace_last_source_charact. apply dtr_loop_pop.
red; simpl; auto.
- intro. apply dtr_pop. red; simpl; auto. auto.
+ intro. apply dtr_pop. red; simpl; auto. auto.
- case_eq sigma; [intros SEQ | intros [ss sd] stl SEQ].
- destruct (reg_eq ms md).
- subst. apply dtr_nop.
+ case_eq sigma; [intros SEQ | intros [ss sd] stl SEQ].
+ destruct (reg_eq ms md).
+ subst. apply dtr_nop.
apply dtr_start. auto.
generalize (split_move_charact ((ms, md) :: mtl) sd).
@@ -1082,9 +1082,9 @@ Proof.
intro NOREAD.
case_eq stl; [intros STLEQ | intros xx1 xx2 STLEQ].
apply dtr_last. auto.
- elim (@exists_last _ stl). 2:congruence. intros sigma1 [[ss1 sd1] SEQ2].
- rewrite <- STLEQ. clear STLEQ xx1 xx2.
- generalize (is_last_source_charact sd ss1 sd1 sigma1).
+ elim (@exists_last _ stl). 2:congruence. intros sigma1 [[ss1 sd1] SEQ2].
+ rewrite <- STLEQ. clear STLEQ xx1 xx2.
+ generalize (is_last_source_charact sd ss1 sd1 sigma1).
rewrite SEQ2. destruct (is_last_source sd (sigma1 ++ (ss1, sd1) :: nil)).
intro. subst ss1.
rewrite replace_last_source_charact. apply dtr_loop_pop. auto.
@@ -1120,7 +1120,7 @@ Qed.
(** Compilation function for parallel moves. *)
Function parmove_aux (st: state) {measure measure st} : moves :=
- if final_state st
+ if final_state st
then match st with State _ _ tau => tau end
else parmove_aux (parmove_step st).
Proof.
@@ -1134,7 +1134,7 @@ Proof.
unfold dtransitions. intro st. functional induction (parmove_aux st).
destruct _x; destruct _x0; simpl in e; discriminate || apply rt_refl.
eapply rt_trans. apply rt_step. apply parmove_step_compatible; eauto.
- auto.
+ auto.
Qed.
Definition parmove (mu: moves) : moves :=
@@ -1151,7 +1151,7 @@ Theorem parmove_correctness:
env_equiv (exec_seq (parmove mu) e) (exec_par mu e).
Proof.
intros. unfold parmove. apply dtransitions_correctness; auto.
- apply parmove_aux_transitions.
+ apply parmove_aux_transitions.
Qed.
(** Here is an alternate formulation of [parmove], where the
@@ -1172,7 +1172,7 @@ Theorem parmove2_correctness:
List.map e' dl = List.map e sl /\
forall r, ~In r dl -> is_not_temp r -> e' r = e r.
Proof.
- intros.
+ intros.
destruct (srcs_dests_combine sl dl H) as [A B].
assert (env_equiv e' (exec_par (List.combine sl dl) e)).
unfold e', parmove2. apply parmove_correctness.
@@ -1182,8 +1182,8 @@ Proof.
red. rewrite B. auto.
destruct (exec_par_combine e sl dl H H0) as [C D].
set (e1 := exec_par (combine sl dl) e) in *.
- split. rewrite <- C. apply list_map_exten; intros.
- symmetry. apply H3. auto.
+ split. rewrite <- C. apply list_map_exten; intros.
+ symmetry. apply H3. auto.
intros. transitivity (e1 r); auto.
Qed.
@@ -1213,7 +1213,7 @@ Definition wf_moves (m: moves) : Prop :=
Lemma wf_moves_cons: forall s d m,
wf_moves ((s, d) :: m) <-> wf_move s d /\ wf_moves m.
Proof.
- unfold wf_moves; intros; simpl. firstorder.
+ unfold wf_moves; intros; simpl. firstorder.
inversion H0; subst s0 d0. auto.
Qed.
@@ -1237,7 +1237,7 @@ Lemma dtransition_preserves_wf_state:
dtransition st st' -> wf_state st -> wf_state st'.
Proof.
induction 1; intro WF; inv WF; constructor; autorewrite with pmov in *; intuition.
- apply wf_move_temp_left; auto.
+ apply wf_move_temp_left; auto.
eapply wf_move_temp_right; eauto.
Qed.
@@ -1245,31 +1245,31 @@ Lemma dtransitions_preserve_wf_state:
forall st st',
dtransitions st st' -> wf_state st -> wf_state st'.
Proof.
- induction 1; intros; eauto.
+ induction 1; intros; eauto.
eapply dtransition_preserves_wf_state; eauto.
-Qed.
+Qed.
End PROPERTIES.
Lemma parmove_wf_moves:
forall mu, wf_moves mu (parmove mu).
Proof.
- intros.
+ intros.
assert (wf_state mu (State mu nil nil)).
constructor. red; intros. apply wf_move_same. auto.
red; simpl; tauto. red; simpl; tauto.
generalize (dtransitions_preserve_wf_state mu
_ _
(parmove_aux_transitions (State mu nil nil)) H).
- intro WFS. inv WFS.
- unfold parmove. red; intros. apply H5.
+ intro WFS. inv WFS.
+ unfold parmove. red; intros. apply H5.
rewrite List.In_rev. auto.
Qed.
Lemma parmove2_wf_moves:
forall sl dl, wf_moves (List.combine sl dl) (parmove2 sl dl).
Proof.
- intros. unfold parmove2. apply parmove_wf_moves.
+ intros. unfold parmove2. apply parmove_wf_moves.
Qed.
(** As a corollary, we show that all sources of [parmove mu]
@@ -1278,12 +1278,12 @@ Qed.
or temporaries. *)
Remark wf_move_initial_reg_or_temp:
- forall mu s d,
+ forall mu s d,
wf_move mu s d ->
(In s (srcs mu) \/ is_temp s) /\ (In d (dests mu) \/ is_temp d).
Proof.
- induction 1.
- split; left.
+ induction 1.
+ split; left.
change s with (fst (s, d)). unfold srcs. apply List.in_map; auto.
change d with (snd (s, d)). unfold dests. apply List.in_map; auto.
split. right. exists s; auto. tauto.
@@ -1316,7 +1316,7 @@ Lemma parmove_srcs_initial_reg_or_temp:
forall mu s,
In s (srcs (parmove mu)) -> In s (srcs mu) \/ is_temp s.
Proof.
- intros. destruct (in_srcs _ _ H) as [d A].
+ intros. destruct (in_srcs _ _ H) as [d A].
destruct (parmove_initial_reg_or_temp _ _ _ A). auto.
Qed.
@@ -1324,7 +1324,7 @@ Lemma parmove_dests_initial_reg_or_temp:
forall mu d,
In d (dests (parmove mu)) -> In d (dests mu) \/ is_temp d.
Proof.
- intros. destruct (in_dests _ _ H) as [s A].
+ intros. destruct (in_dests _ _ H) as [s A].
destruct (parmove_initial_reg_or_temp _ _ _ A). auto.
Qed.
@@ -1455,11 +1455,11 @@ Hypothesis temps_no_overlap:
Lemma disjoint_list_notin:
forall r l, disjoint_list r l -> ~In r l.
Proof.
- intros. red; intro.
- assert (r <> r). apply disjoint_not_equal. apply H; auto.
+ intros. red; intro.
+ assert (r <> r). apply disjoint_not_equal. apply H; auto.
congruence.
Qed.
-
+
Lemma pairwise_disjoint_norepet:
forall l, pairwise_disjoint l -> list_norepet l.
Proof.
@@ -1513,7 +1513,7 @@ Proof.
subst. right. apply H. auto.
subst. right. apply disjoint_sym. apply H. auto.
auto.
-Qed.
+Qed.
Lemma no_adherence_dst:
forall d, In d (dests mu) -> no_adherence d.
@@ -1547,7 +1547,7 @@ Qed.
Definition env_match (e1 e2: env) : Prop :=
forall r, no_adherence r -> e1 r = e2 r.
-(** The following lemmas relate the effect of executing moves
+(** The following lemmas relate the effect of executing moves
using normal, overlap-unaware update and weak, overlap-aware update. *)
Lemma weak_update_match:
@@ -1558,9 +1558,9 @@ Lemma weak_update_match:
env_match (update d (e1 s) e1)
(weak_update d (e2 s) e2).
Proof.
- intros. red; intros.
+ intros. red; intros.
assert (no_overlap d r). apply H2. auto.
- destruct H3.
+ destruct H3.
subst. rewrite update_s. rewrite weak_update_s. apply H1.
destruct H. apply no_adherence_src; auto. apply no_adherence_tmp; auto.
rewrite update_o. rewrite weak_update_d. apply H1. auto.
@@ -1576,8 +1576,8 @@ Lemma weak_exec_seq_match:
Proof.
induction m; intros; simpl.
auto.
- destruct a as [s d]. simpl in H. simpl in H0.
- apply IHm; auto.
+ destruct a as [s d]. simpl in H. simpl in H0.
+ apply IHm; auto.
apply weak_update_match; auto.
Qed.
@@ -1600,7 +1600,7 @@ Theorem parmove2_correctness_with_overlap:
forall r, disjoint_list r dl -> disjoint_temps r -> e' r = e r.
Proof.
intros.
- assert (list_norepet dl).
+ assert (list_norepet dl).
apply pairwise_disjoint_norepet; auto.
assert (forall r : reg, In r sl -> is_not_temp r).
intros. apply disjoint_temps_not_temp; auto.
@@ -1608,7 +1608,7 @@ Proof.
intros. apply disjoint_temps_not_temp; auto.
generalize (parmove2_correctness sl dl H H5 H6 H7 e).
set (e1 := exec_seq (parmove2 sl dl) e). intros [A B].
- destruct (srcs_dests_combine sl dl H) as [C D].
+ destruct (srcs_dests_combine sl dl H) as [C D].
assert (env_match (combine sl dl) e1 e').
unfold parmove2. unfold e1, e'.
apply weak_exec_seq_match; try (rewrite C); try (rewrite D); auto.
@@ -1616,11 +1616,11 @@ Proof.
intros. rewrite <- D. apply parmove_dests_initial_reg_or_temp. auto.
red; auto.
split.
- rewrite <- A.
+ rewrite <- A.
apply list_map_exten; intros. apply H8.
- apply no_adherence_dst. rewrite D; auto. rewrite D; auto. rewrite D; auto.
+ apply no_adherence_dst. rewrite D; auto. rewrite D; auto. rewrite D; auto.
intros. transitivity (e1 r).
- symmetry. apply H8. red. rewrite D. intros. destruct H11.
+ symmetry. apply H8. red. rewrite D. intros. destruct H11.
right. apply disjoint_sym. apply H9. auto.
right. apply disjoint_sym. apply H10. auto.
apply B. apply disjoint_list_notin; auto. apply disjoint_temps_not_temp; auto.
diff --git a/lib/Postorder.v b/lib/Postorder.v
index 4a83ea50..0215a829 100644
--- a/lib/Postorder.v
+++ b/lib/Postorder.v
@@ -41,8 +41,8 @@ Module PositiveOrd.
Theorem leb_total : forall x y, is_true (leb x y) \/ is_true (leb y x).
Proof.
unfold leb, is_true; intros.
- destruct (plt x y); auto. destruct (plt y x); auto.
- elim (Plt_strict x). eapply Plt_trans; eauto.
+ destruct (plt x y); auto. destruct (plt y x); auto.
+ elim (Plt_strict x). eapply Plt_trans; eauto.
Qed.
End PositiveOrd.
@@ -114,7 +114,7 @@ Inductive invariant (s: state) : Prop :=
(REM: forall x y, s.(gr)!x = Some y -> s.(map)!x = None)
(* black nodes have no white son *)
(COLOR: forall x succs n y,
- ginit!x = Some succs -> s.(map)!x = Some n ->
+ ginit!x = Some succs -> s.(map)!x = Some n ->
In y succs -> s.(gr)!y = None)
(* worklist is well-formed *)
(WKLIST: forall x l, In (x, l) s.(wrk) ->
@@ -140,15 +140,15 @@ Proof.
Qed.
Lemma transition_spec:
- forall s, invariant s ->
+ forall s, invariant s ->
match transition s with inr s' => invariant s' | inl m => postcondition m end.
Proof.
- intros. inv H. unfold transition. destruct (wrk s) as [ | [x succ_x] l].
+ intros. inv H. unfold transition. destruct (wrk s) as [ | [x succ_x] l].
(* finished *)
constructor; intros.
eauto.
caseEq (s.(map)!root); intros. congruence. exploit GREY; eauto. intros [? ?]; contradiction.
- destruct (s.(map)!x) eqn:?; try congruence.
+ destruct (s.(map)!x) eqn:?; try congruence.
destruct (s.(map)!y) eqn:?; try congruence.
exploit COLOR; eauto. intros. exploit GREY; eauto. intros [? ?]; contradiction.
(* not finished *)
@@ -160,30 +160,30 @@ Proof.
(* root *)
eauto.
(* below *)
- rewrite PTree.gsspec in H. destruct (peq x0 x). inv H.
+ rewrite PTree.gsspec in H. destruct (peq x0 x). inv H.
apply Plt_succ.
apply Plt_trans_succ. eauto.
(* inj *)
- rewrite PTree.gsspec in H. rewrite PTree.gsspec in H0.
+ rewrite PTree.gsspec in H. rewrite PTree.gsspec in H0.
destruct (peq x1 x); destruct (peq x2 x); subst.
auto.
- inv H. exploit BELOW; eauto. intros. eelim Plt_strict; eauto.
+ inv H. exploit BELOW; eauto. intros. eelim Plt_strict; eauto.
inv H0. exploit BELOW; eauto. intros. eelim Plt_strict; eauto.
eauto.
(* rem *)
- intros. rewrite PTree.gso; eauto. red; intros; subst x0.
+ intros. rewrite PTree.gso; eauto. red; intros; subst x0.
exploit (WKLIST x nil); auto with coqlib. intros [A B]. congruence.
(* color *)
- rewrite PTree.gsspec in H0. destruct (peq x0 x).
- inv H0. exploit (WKLIST x nil); auto with coqlib.
+ rewrite PTree.gsspec in H0. destruct (peq x0 x).
+ inv H0. exploit (WKLIST x nil); auto with coqlib.
intros [A [l' [B C]]].
assert (l' = succs) by congruence. subst l'.
- apply C; auto.
+ apply C; auto.
eauto.
(* wklist *)
apply WKLIST. auto with coqlib.
(* grey *)
- rewrite PTree.gsspec in H1. destruct (peq x0 x). inv H1.
+ rewrite PTree.gsspec in H1. destruct (peq x0 x). inv H1.
exploit GREY; eauto. intros [l' A]. simpl in A; destruct A.
congruence.
exists l'; auto.
@@ -191,11 +191,11 @@ Proof.
(* children y needs traversing *)
destruct ((gr s)!y) as [ succs_y | ] eqn:?.
(* y has children *)
- constructor; simpl; intros.
+ constructor; simpl; intros.
(* sub *)
rewrite PTree.grspec in H. destruct (PTree.elt_eq x0 y); eauto. inv H.
(* root *)
- rewrite PTree.gro. auto. congruence.
+ rewrite PTree.gro. auto. congruence.
(* below *)
eauto.
(* inj *)
@@ -203,33 +203,33 @@ Proof.
(* rem *)
rewrite PTree.grspec in H. destruct (PTree.elt_eq x0 y); eauto. inv H.
(* color *)
- rewrite PTree.grspec. destruct (PTree.elt_eq y0 y); eauto.
+ rewrite PTree.grspec. destruct (PTree.elt_eq y0 y); eauto.
(* wklist *)
- destruct H.
- inv H. split. apply PTree.grs. exists succs_y; split. eauto.
+ destruct H.
+ inv H. split. apply PTree.grs. exists succs_y; split. eauto.
intros. rewrite In_sort in H. tauto.
destruct H.
inv H. exploit WKLIST; eauto with coqlib. intros [A [l' [B C]]].
- split. rewrite PTree.grspec. destruct (PTree.elt_eq x0 y); auto.
+ split. rewrite PTree.grspec. destruct (PTree.elt_eq x0 y); auto.
exists l'; split. auto. intros. rewrite PTree.grspec. destruct (PTree.elt_eq y0 y); auto.
- apply C; auto. simpl. intuition congruence.
+ apply C; auto. simpl. intuition congruence.
exploit (WKLIST x0 l0); eauto with coqlib. intros [A [l' [B C]]].
- split. rewrite PTree.grspec. destruct (PTree.elt_eq x0 y); auto.
- exists l'; split; auto. intros.
+ split. rewrite PTree.grspec. destruct (PTree.elt_eq x0 y); auto.
+ exists l'; split; auto. intros.
rewrite PTree.grspec. destruct (PTree.elt_eq y0 y); auto.
(* grey *)
- rewrite PTree.grspec in H0. destruct (PTree.elt_eq x0 y) in H0.
+ rewrite PTree.grspec in H0. destruct (PTree.elt_eq x0 y) in H0.
subst. exists (Sort.sort succs_y); auto with coqlib.
- exploit GREY; eauto. simpl. intros [l1 A]. destruct A.
- inv H2. exists succ_x; auto.
+ exploit GREY; eauto. simpl. intros [l1 A]. destruct A.
+ inv H2. exists succ_x; auto.
exists l1; auto.
(* y has no children *)
constructor; simpl; intros; eauto.
(* wklist *)
- destruct H. inv H.
+ destruct H. inv H.
exploit (WKLIST x0); eauto with coqlib. intros [A [l' [B C]]].
- split. auto. exists l'; split. auto.
+ split. auto. exists l'; split. auto.
intros. destruct (peq y y0). congruence. apply C; auto. simpl. intuition congruence.
eapply WKLIST; eauto with coqlib.
(* grey *)
@@ -257,18 +257,18 @@ Proof.
(* color *)
rewrite PTree.gempty in H0; inv H0.
(* wklist *)
- destruct H; inv H.
- split. apply PTree.grs. exists succs; split; auto.
+ destruct H; inv H.
+ split. apply PTree.grs. exists succs; split; auto.
intros. rewrite In_sort in H. intuition.
(* grey *)
rewrite PTree.grspec in H0. destruct (PTree.elt_eq x root).
- subst. exists (Sort.sort succs); auto.
+ subst. exists (Sort.sort succs); auto.
contradiction.
(* root has no succs *)
constructor; simpl; intros.
(* sub *)
- auto.
+ auto.
(* root *)
auto.
(* below *)
@@ -302,12 +302,12 @@ Lemma lt_state_wf: well_founded lt_state.
Proof.
set (f := fun s => (PTree_Properties.cardinal s.(gr), size_worklist s.(wrk))).
change (well_founded (fun s1 s2 => lex_ord lt lt (f s1) (f s2))).
- apply wf_inverse_image.
- apply wf_lex_ord.
+ apply wf_inverse_image.
+ apply wf_lex_ord.
apply lt_wf. apply lt_wf.
Qed.
-Lemma transition_decreases:
+Lemma transition_decreases:
forall s s', transition s = inr _ s' -> lt_state s' s.
Proof.
unfold transition, lt_state; intros.
@@ -338,11 +338,11 @@ Theorem postorder_correct:
(forall x1 x2 y, m!x1 = Some y -> m!x2 = Some y -> x1 = x2)
/\ (forall x, reachable g root x -> g!x <> None -> m!x <> None).
Proof.
- intros.
+ intros.
assert (postcondition g root m).
unfold m. unfold postorder.
apply WfIter.iterate_prop with (P := invariant g root).
- apply transition_spec.
+ apply transition_spec.
apply initial_state_spec.
inv H.
split. auto.
diff --git a/lib/Readconfig.mll b/lib/Readconfig.mll
index 27ef32cf..ec2f6bb0 100644
--- a/lib/Readconfig.mll
+++ b/lib/Readconfig.mll
@@ -28,7 +28,7 @@ let error msg lexbuf =
lexbuf.lex_curr_p.pos_lnum,
msg)))
-let ill_formed_line lexbuf = error "Ill-formed line" lexbuf
+let ill_formed_line lexbuf = error "Ill-formed line" lexbuf
let unterminated_quote lexbuf = error "Unterminated quote" lexbuf
let lone_backslash lexbuf = error "Lone \\ (backslash) at end of file" lexbuf
@@ -41,7 +41,7 @@ let ident = ['A'-'Z' 'a'-'z' '_'] ['A'-'Z' 'a'-'z' '0'-'9' '_' '.']*
rule begline = parse
| '#' [^ '\n']* ('\n' | eof)
{ Lexing.new_line lexbuf; begline lexbuf }
- | whitespace* (ident as key) whitespace* '='
+ | whitespace* (ident as key) whitespace* '='
{ let words = unquoted false [] lexbuf in
Hashtbl.add key_val_tbl key (List.rev words);
begline lexbuf }
diff --git a/lib/UnionFind.v b/lib/UnionFind.v
index 46a886ea..76dd6b31 100644
--- a/lib/UnionFind.v
+++ b/lib/UnionFind.v
@@ -90,7 +90,7 @@ Module Type UNIONFIND.
Hypothesis sameclass_union_3:
forall uf a b x y,
sameclass (union uf a b) x y ->
- sameclass uf x y
+ sameclass uf x y
\/ sameclass uf x a /\ sameclass uf y b
\/ sameclass uf x b /\ sameclass uf y a.
@@ -120,7 +120,7 @@ Module Type UNIONFIND.
pathlen (merge uf a b) x =
if elt_eq (repr uf a) (repr uf b) then
pathlen uf x
- else if elt_eq (repr uf x) (repr uf a) then
+ else if elt_eq (repr uf x) (repr uf a) then
pathlen uf x + pathlen uf b + 1
else
pathlen uf x.
@@ -155,7 +155,7 @@ Definition t := unionfind.
Definition getlink (m: M.t elt) (a: elt) : {a' | M.get a m = Some a'} + {M.get a m = None}.
Proof.
- destruct (M.get a m). left. exists e; auto. right; auto.
+ destruct (M.get a m). left. exists e; auto. right; auto.
Defined.
(* The canonical representative of an element *)
@@ -175,11 +175,11 @@ Definition repr (a: elt) : elt := Fix uf.(mwf) (fun _ => elt) F_repr a.
Lemma repr_unroll:
forall a, repr a = match M.get a uf.(m) with Some a' => repr a' | None => a end.
Proof.
- intros. unfold repr at 1. rewrite Fix_eq.
- unfold F_repr. destruct (getlink uf.(m) a) as [[a' P] | Q].
+ intros. unfold repr at 1. rewrite Fix_eq.
+ unfold F_repr. destruct (getlink uf.(m) a) as [[a' P] | Q].
rewrite P; auto.
rewrite Q; auto.
- intros. unfold F_repr. destruct (getlink (m uf) x) as [[a' P] | Q]; auto.
+ intros. unfold F_repr. destruct (getlink (m uf) x) as [[a' P] | Q]; auto.
Qed.
Lemma repr_none:
@@ -187,36 +187,36 @@ Lemma repr_none:
M.get a uf.(m) = None ->
repr a = a.
Proof.
- intros. rewrite repr_unroll. rewrite H; auto.
+ intros. rewrite repr_unroll. rewrite H; auto.
Qed.
Lemma repr_some:
- forall a a',
+ forall a a',
M.get a uf.(m) = Some a' ->
repr a = repr a'.
Proof.
- intros. rewrite repr_unroll. rewrite H; auto.
+ intros. rewrite repr_unroll. rewrite H; auto.
Qed.
Lemma repr_res_none:
forall (a: elt), M.get (repr a) uf.(m) = None.
Proof.
- apply (well_founded_ind (mwf uf)). intros.
+ apply (well_founded_ind (mwf uf)). intros.
rewrite repr_unroll. destruct (M.get x (m uf)) as [y|] eqn:X; auto.
Qed.
Lemma repr_canonical:
forall (a: elt), repr (repr a) = repr a.
Proof.
- intros. apply repr_none. apply repr_res_none.
+ intros. apply repr_none. apply repr_res_none.
Qed.
Lemma repr_some_diff:
forall a a', M.get a uf.(m) = Some a' -> a <> repr a'.
Proof.
- intros; red; intros.
- assert (repr a = a). rewrite (repr_some a a'); auto.
- assert (M.get a uf.(m) = None). rewrite <- H1. apply repr_res_none.
+ intros; red; intros.
+ assert (repr a = a). rewrite (repr_some a a'); auto.
+ assert (M.get a uf.(m) = None). rewrite <- H1. apply repr_res_none.
congruence.
Qed.
@@ -297,9 +297,9 @@ Remark identify_Acc_b:
Proof.
induction 1; intros. constructor; intros.
rewrite identify_order in H2. destruct H2 as [A | [A B]].
- apply H0; auto. rewrite <- (repr_some uf _ _ A). auto.
+ apply H0; auto. rewrite <- (repr_some uf _ _ A). auto.
subst. elim H1. apply repr_none. auto.
-Qed.
+Qed.
Remark identify_Acc:
forall x,
@@ -308,13 +308,13 @@ Proof.
induction 1. constructor; intros.
rewrite identify_order in H1. destruct H1 as [A | [A B]].
auto.
- subst. apply identify_Acc_b; auto. apply uf.(mwf).
+ subst. apply identify_Acc_b; auto. apply uf.(mwf).
Qed.
Lemma identify_wf:
well_founded (order (M.set a b uf.(m))).
Proof.
- red; intros. apply identify_Acc. apply uf.(mwf).
+ red; intros. apply identify_Acc. apply uf.(mwf).
Qed.
Definition identify := mk (M.set a b uf.(m)) identify_wf.
@@ -323,11 +323,11 @@ Lemma repr_identify_1:
forall x, repr uf x <> a -> repr identify x = repr uf x.
Proof.
intros x0; pattern x0. apply (well_founded_ind (mwf uf)); intros.
- rewrite (repr_unroll uf) in *.
+ rewrite (repr_unroll uf) in *.
destruct (M.get x (m uf)) as [a'|] eqn:X.
rewrite <- H; auto.
apply repr_some. simpl. rewrite M.gsspec. rewrite dec_eq_false; auto. congruence.
- apply repr_none. simpl. rewrite M.gsspec. rewrite dec_eq_false; auto.
+ apply repr_none. simpl. rewrite M.gsspec. rewrite dec_eq_false; auto.
Qed.
Lemma repr_identify_2:
@@ -335,9 +335,9 @@ Lemma repr_identify_2:
Proof.
intros x0; pattern x0. apply (well_founded_ind (mwf uf)); intros.
rewrite (repr_unroll uf) in H0. destruct (M.get x (m uf)) as [a'|] eqn:X.
- rewrite <- (H a'); auto.
+ rewrite <- (H a'); auto.
apply repr_some. simpl. rewrite M.gsspec. rewrite dec_eq_false; auto. congruence.
- subst x. rewrite (repr_unroll identify). simpl. rewrite M.gsspec.
+ subst x. rewrite (repr_unroll identify). simpl. rewrite M.gsspec.
rewrite dec_eq_true. apply repr_identify_1. auto.
Qed.
@@ -348,7 +348,7 @@ End IDENTIFY.
Remark union_not_same_class:
forall uf a b, repr uf a <> repr uf b -> repr uf (repr uf b) <> repr uf a.
Proof.
- intros. rewrite repr_canonical. auto.
+ intros. rewrite repr_canonical. auto.
Qed.
Definition union (uf: t) (a b: elt) : t :=
@@ -402,7 +402,7 @@ Qed.
Lemma sameclass_union_3:
forall uf a b x y,
sameclass (union uf a b) x y ->
- sameclass uf x y
+ sameclass uf x y
\/ sameclass uf x a /\ sameclass uf y b
\/ sameclass uf x b /\ sameclass uf y a.
Proof.
@@ -448,14 +448,14 @@ Definition path_ord (uf: t) : elt -> elt -> Prop := order uf.(m).
Lemma path_ord_wellfounded:
forall uf, well_founded (path_ord uf).
Proof.
- intros. apply mwf.
+ intros. apply mwf.
Qed.
Lemma path_ord_canonical:
forall uf x y, repr uf x = x -> ~path_ord uf y x.
Proof.
intros; red; intros. hnf in H0.
- assert (M.get x (m uf) = None). rewrite <- H. apply repr_res_none.
+ assert (M.get x (m uf) = None). rewrite <- H. apply repr_res_none.
congruence.
Qed.
@@ -463,10 +463,10 @@ Lemma path_ord_merge_1:
forall uf a b x y,
path_ord uf x y -> path_ord (merge uf a b) x y.
Proof.
- intros. unfold merge.
+ intros. unfold merge.
destruct (M.elt_eq (repr uf a) (repr uf b)).
auto.
- red. simpl. red. rewrite M.gsspec. rewrite dec_eq_false. apply H.
+ red. simpl. red. rewrite M.gsspec. rewrite dec_eq_false. apply H.
red; intros. hnf in H. generalize (repr_res_none uf a). congruence.
Qed.
@@ -497,11 +497,11 @@ Definition pathlen (a: elt) : nat := Fix uf.(mwf) (fun _ => nat) F_pathlen a.
Lemma pathlen_unroll:
forall a, pathlen a = match M.get a uf.(m) with Some a' => S(pathlen a') | None => O end.
Proof.
- intros. unfold pathlen at 1. rewrite Fix_eq.
- unfold F_pathlen. destruct (getlink uf.(m) a) as [[a' P] | Q].
+ intros. unfold pathlen at 1. rewrite Fix_eq.
+ unfold F_pathlen. destruct (getlink uf.(m) a) as [[a' P] | Q].
rewrite P; auto.
rewrite Q; auto.
- intros. unfold F_pathlen. destruct (getlink (m uf) x) as [[a' P] | Q]; auto.
+ intros. unfold F_pathlen. destruct (getlink (m uf) x) as [[a' P] | Q]; auto.
Qed.
Lemma pathlen_none:
@@ -509,11 +509,11 @@ Lemma pathlen_none:
M.get a uf.(m) = None ->
pathlen a = 0.
Proof.
- intros. rewrite pathlen_unroll. rewrite H; auto.
+ intros. rewrite pathlen_unroll. rewrite H; auto.
Qed.
Lemma pathlen_some:
- forall a a',
+ forall a a',
M.get a uf.(m) = Some a' ->
pathlen a = S (pathlen a').
Proof.
@@ -524,8 +524,8 @@ Lemma pathlen_zero:
forall a, repr uf a = a <-> pathlen a = O.
Proof.
intros; split; intros.
- apply pathlen_none. rewrite <- H. apply repr_res_none.
- apply repr_none. rewrite pathlen_unroll in H.
+ apply pathlen_none. rewrite <- H. apply repr_res_none.
+ apply repr_none. rewrite pathlen_unroll in H.
destruct (M.get a (m uf)); congruence.
Qed.
@@ -538,27 +538,27 @@ Lemma pathlen_merge:
pathlen (merge uf a b) x =
if M.elt_eq (repr uf a) (repr uf b) then
pathlen uf x
- else if M.elt_eq (repr uf x) (repr uf a) then
+ else if M.elt_eq (repr uf x) (repr uf a) then
pathlen uf x + pathlen uf b + 1
else
pathlen uf x.
Proof.
- intros. unfold merge.
+ intros. unfold merge.
destruct (M.elt_eq (repr uf a) (repr uf b)).
auto.
set (uf' := identify uf (repr uf a) b (repr_res_none uf a) (not_eq_sym n)).
pattern x. apply (well_founded_ind (mwf uf')); intros.
rewrite (pathlen_unroll uf'). destruct (M.get x0 (m uf')) as [x'|] eqn:G.
- rewrite H; auto. simpl in G. rewrite M.gsspec in G.
+ rewrite H; auto. simpl in G. rewrite M.gsspec in G.
destruct (M.elt_eq x0 (repr uf a)). rewrite e. rewrite repr_canonical. rewrite dec_eq_true.
- inversion G. subst x'. rewrite dec_eq_false; auto.
- replace (pathlen uf (repr uf a)) with 0. omega.
- symmetry. apply pathlen_none. apply repr_res_none.
+ inversion G. subst x'. rewrite dec_eq_false; auto.
+ replace (pathlen uf (repr uf a)) with 0. omega.
+ symmetry. apply pathlen_none. apply repr_res_none.
rewrite (repr_unroll uf x0), (pathlen_unroll uf x0); rewrite G.
- destruct (M.elt_eq (repr uf x') (repr uf a)); omega.
+ destruct (M.elt_eq (repr uf x') (repr uf a)); omega.
simpl in G. rewrite M.gsspec in G. destruct (M.elt_eq x0 (repr uf a)); try discriminate.
- rewrite (repr_none uf x0) by auto. rewrite dec_eq_false; auto.
- symmetry. apply pathlen_zero; auto. apply repr_none; auto.
+ rewrite (repr_none uf x0) by auto. rewrite dec_eq_false; auto.
+ symmetry. apply pathlen_zero; auto. apply repr_none; auto.
Qed.
Lemma pathlen_gt_merge:
@@ -567,7 +567,7 @@ Lemma pathlen_gt_merge:
pathlen uf x > pathlen uf y ->
pathlen (merge uf a b) x > pathlen (merge uf a b) y.
Proof.
- intros. repeat rewrite pathlen_merge.
+ intros. repeat rewrite pathlen_merge.
destruct (M.elt_eq (repr uf a) (repr uf b)). auto.
rewrite H. destruct (M.elt_eq (repr uf y) (repr uf a)).
omega. auto.
@@ -600,7 +600,7 @@ Proof.
induction 1. constructor; intros.
destruct (compress_order _ _ H1) as [A | [A B]].
auto.
- subst x y. constructor; intros.
+ subst x y. constructor; intros.
destruct (compress_order _ _ H2) as [A | [A B]].
red in A. generalize (repr_res_none uf a). congruence.
congruence.
@@ -609,7 +609,7 @@ Qed.
Lemma compress_wf:
well_founded (order (M.set a b uf.(m))).
Proof.
- red; intros. apply compress_Acc. apply uf.(mwf).
+ red; intros. apply compress_Acc. apply uf.(mwf).
Qed.
Definition compress := mk (M.set a b uf.(m)) compress_wf.
@@ -620,11 +620,11 @@ Proof.
apply (well_founded_ind (mwf compress)); intros.
rewrite (repr_unroll compress).
destruct (M.get x (m compress)) as [y|] eqn:G.
- rewrite H; auto.
- simpl in G. rewrite M.gsspec in G. destruct (M.elt_eq x a).
+ rewrite H; auto.
+ simpl in G. rewrite M.gsspec in G. destruct (M.elt_eq x a).
inversion G. subst x y. rewrite <- a_repr_b. apply repr_canonical.
symmetry; apply repr_some; auto.
- simpl in G. rewrite M.gsspec in G. destruct (M.elt_eq x a).
+ simpl in G. rewrite M.gsspec in G. destruct (M.elt_eq x a).
congruence.
symmetry; apply repr_none; auto.
Qed.
@@ -637,7 +637,7 @@ Section FIND.
Variable uf: t.
-Program Fixpoint find_x (a: elt) {wf (order uf.(m)) a} :
+Program Fixpoint find_x (a: elt) {wf (order uf.(m)) a} :
{ r: elt * t | fst r = repr uf a /\ forall x, repr (snd r) x = repr uf x } :=
match M.get a uf.(m) with
| Some a' =>
@@ -664,7 +664,7 @@ Next Obligation.
destruct (find_x a')
as [[b' uf''] [A B]]. simpl in *. inv Heq_anonymous0.
symmetry. apply repr_some. auto.
- intros. rewrite repr_compress.
+ intros. rewrite repr_compress.
destruct (find_x a')
as [[b' uf''] [A B]]. simpl in *. inv Heq_anonymous0. auto.
Qed.
@@ -672,7 +672,7 @@ Next Obligation.
split; auto. symmetry. apply repr_none. auto.
Qed.
Next Obligation.
- apply mwf.
+ apply mwf.
Defined.
Definition find (a: elt) : elt * t := proj1_sig (find_x a).
@@ -680,15 +680,15 @@ Definition find (a: elt) : elt * t := proj1_sig (find_x a).
Lemma find_repr:
forall a, fst (find a) = repr uf a.
Proof.
- unfold find; intros. destruct (find_x a) as [[b uf'] [A B]]. simpl. auto.
+ unfold find; intros. destruct (find_x a) as [[b uf'] [A B]]. simpl. auto.
Qed.
Lemma find_unchanged:
forall a x, repr (snd (find a)) x = repr uf x.
Proof.
- unfold find; intros. destruct (find_x a) as [[b uf'] [A B]]. simpl. auto.
+ unfold find; intros. destruct (find_x a) as [[b uf'] [A B]]. simpl. auto.
Qed.
-
+
Lemma sameclass_find_1:
forall a x y, sameclass (snd (find a)) x y <-> sameclass uf x y.
Proof.
diff --git a/lib/Wfsimpl.v b/lib/Wfsimpl.v
index 1ed6326a..4f80822e 100644
--- a/lib/Wfsimpl.v
+++ b/lib/Wfsimpl.v
@@ -35,9 +35,9 @@ Definition Fix (x: A) : B := Wf.Fix Rwf (fun (x: A) => B) F x.
Theorem unroll_Fix:
forall x, Fix x = F (fun (y: A) (P: R y x) => Fix y).
Proof.
- unfold Fix; intros. apply Wf.Fix_eq with (P := fun (x: A) => B).
+ unfold Fix; intros. apply Wf.Fix_eq with (P := fun (x: A) => B).
intros. assert (f = g). apply functional_extensionality_dep; intros.
- apply functional_extensionality; intros. auto.
+ apply functional_extensionality; intros. auto.
subst g; auto.
Qed.
@@ -56,9 +56,9 @@ Definition Fixm (x: A) : B := Wf.Fix (well_founded_ltof A measure) (fun (x: A) =
Theorem unroll_Fixm:
forall x, Fixm x = F (fun (y: A) (P: measure y < measure x) => Fixm y).
Proof.
- unfold Fixm; intros. apply Wf.Fix_eq with (P := fun (x: A) => B).
+ unfold Fixm; intros. apply Wf.Fix_eq with (P := fun (x: A) => B).
intros. assert (f = g). apply functional_extensionality_dep; intros.
- apply functional_extensionality; intros. auto.
+ apply functional_extensionality; intros. auto.
subst g; auto.
Qed.