From 6b62e98a0809a9ac22e64f82943473994e79bfaf Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 6 Dec 2021 17:11:43 +0100 Subject: ExtZ --- kvx/ExtZ.v | 12 ++++++++++++ 1 file changed, 12 insertions(+) create mode 100644 kvx/ExtZ.v diff --git a/kvx/ExtZ.v b/kvx/ExtZ.v new file mode 100644 index 00000000..095fd0cc --- /dev/null +++ b/kvx/ExtZ.v @@ -0,0 +1,12 @@ +Require Import Coq.ZArith.Zdiv. + +Open Scope Z. + +Definition Zdiv_ne (a b : Z) := + let q := Z.div a b in + let q1 := Z.succ q in + match Z.compare (a-b*q) (b*q1-a) with + | Lt => q + | Gt => q1 + | Eq => (if Z.even q then q else q1) + end. -- cgit From 9e62bc5ba6af27e47b7d3f22cef8c2dd24a8fc5e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 6 Dec 2021 19:28:40 +0100 Subject: Zdiv_ne --- kvx/ExtIEEE754.v | 12 ++++++++++++ 1 file changed, 12 insertions(+) create mode 100644 kvx/ExtIEEE754.v diff --git a/kvx/ExtIEEE754.v b/kvx/ExtIEEE754.v new file mode 100644 index 00000000..095fd0cc --- /dev/null +++ b/kvx/ExtIEEE754.v @@ -0,0 +1,12 @@ +Require Import Coq.ZArith.Zdiv. + +Open Scope Z. + +Definition Zdiv_ne (a b : Z) := + let q := Z.div a b in + let q1 := Z.succ q in + match Z.compare (a-b*q) (b*q1-a) with + | Lt => q + | Gt => q1 + | Eq => (if Z.even q then q else q1) + end. -- cgit From 54e63022f92795ab9ffa5c67023a98b7d30ebc69 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 6 Dec 2021 19:34:23 +0100 Subject: Zdiv_ne --- lib/IEEE754_extra.v | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index b0d1944e..d2372f17 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -879,6 +879,17 @@ Proof. unfold ZofB_range. rewrite D. rewrite Zle_bool_true by lia. rewrite Zle_bool_true by lia. auto. Qed. +(** ZofB_ne : convert float to integer, round to nearest *) + +Definition Zdiv_ne (a b : Z) := + let q := Z.div a b in + let q1 := Z.succ q in + match Z.compare (a-b*q) (b*q1-a) with + | Lt => q + | Gt => q1 + | Eq => (if Z.even q then q else q1) + end. + (** ** Algebraic identities *) (** Commutativity of addition and multiplication *) -- cgit From c5e74c17f2ab7f40d314417f2b20b18d4d2d057b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 6 Dec 2021 21:14:46 +0100 Subject: progress on ZofB_ne --- lib/IEEE754_extra.v | 51 +++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 51 insertions(+) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index d2372f17..1ee9c2e0 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -890,6 +890,57 @@ Definition Zdiv_ne (a b : Z) := | Eq => (if Z.even q then q else q1) end. +Definition ZofB_ne (f: binary_float): option Z := + match f with + | B754_finite _ _ s m (Zpos e) _ => Some (cond_Zopp s (Zpos m) * Z.pow_pos radix2 e)%Z + | B754_finite _ _ s m 0 _ => Some (cond_Zopp s (Zpos m)) + | B754_finite _ _ s m (Zneg e) _ => Some (cond_Zopp s (Zdiv_ne (Zpos m) (Z.pow_pos radix2 e)))%Z + | B754_zero _ _ _ => Some 0%Z + | _ => None + end. + +Lemma ZnearestE_IZR: + forall n, (ZnearestE (IZR n)) = n. +Proof. + intro n. + unfold ZnearestE. + case Rcompare_spec ; intro ORDER. + - apply Zfloor_IZR. + - destruct negb. + + apply Zceil_IZR. + + apply Zfloor_IZR. + - apply Zceil_IZR. +Qed. + +Lemma Zfloor_opp : + forall x : R, (Zfloor (- x)) = - (Zceil x). +Proof. + unfold Zceil, Zfloor. + intro x. + rewrite Z.opp_involutive. + reflexivity. +Qed. + +Lemma Zceil_opp : + forall x : R, (Zceil (- x)) = - (Zfloor x). +Proof. + unfold Zceil, Zfloor. + intro x. + rewrite Ropp_involutive. + reflexivity. +Qed. + +Lemma ZnearestE_opp + : forall x : R, ZnearestE (- x) = - ZnearestE x. +Admitted. + + +Theorem ZofB_ne_correct: + forall f, + ZofB_ne f = if is_finite _ _ f then Some (ZnearestE (B2R _ _ f)) else None. +Admitted. + + (** ** Algebraic identities *) (** Commutativity of addition and multiplication *) -- cgit From fbead38d40df6904dc0e446f673747b008c40022 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 7 Dec 2021 12:00:54 +0100 Subject: ZnearestE_opp --- lib/IEEE754_extra.v | 87 ++++++++++++++++++++++++++++++++++++++++++++++++++--- 1 file changed, 83 insertions(+), 4 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 1ee9c2e0..db4fbec7 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -929,12 +929,91 @@ Proof. rewrite Ropp_involutive. reflexivity. Qed. - -Lemma ZnearestE_opp - : forall x : R, ZnearestE (- x) = - ZnearestE x. -Admitted. +Lemma Zceil_non_floor: + forall x : R, (x > IZR(Zfloor x))%R -> Zceil x = Z.succ(Zfloor x). +Proof. + intros x BETWEEN. + unfold Z.succ. + apply Zceil_imp. + split. + { rewrite minus_IZR. + rewrite plus_IZR. + lra. + } + rewrite plus_IZR. + pose proof (Zfloor_ub x). + lra. +Qed. + +Lemma Zceil_non_ceil: + forall x : R, (x < IZR(Zceil x))%R -> Zceil x = Z.succ(Zfloor x). +Proof. + intros x BETWEEN. + unfold Z.succ. + cut (Zfloor x = (Zceil x) - 1). { intros; lia. } + apply Zfloor_imp. + split. + { rewrite minus_IZR. + pose proof (Zceil_lb x). + lra. + } + rewrite plus_IZR. + rewrite minus_IZR. + lra. +Qed. +Lemma ZnearestE_opp + : forall x : R, ZnearestE (- x) = - ZnearestE x. +Proof. + intro x. + unfold ZnearestE. + case (Rcompare_spec (x - IZR (Zfloor x)) (/ 2)); intro CMP. + - pose proof (Zfloor_lb x) as LB. + destruct (Rcompare_spec x (IZR (Zfloor x))) as [ ABSURD | EXACT | INEXACT]. + lra. + { set (n := Zfloor x) in *. + rewrite EXACT. + rewrite <- opp_IZR. + rewrite Zfloor_IZR. + rewrite opp_IZR. + rewrite Rcompare_Lt by lra. + reflexivity. + } + rewrite Rcompare_Gt. + { apply Zceil_opp. } + rewrite Zfloor_opp. + rewrite opp_IZR. + rewrite Zceil_non_floor by assumption. + unfold Z.succ. + rewrite plus_IZR. + lra. + - rewrite Rcompare_Eq. + { rewrite Zceil_opp. + rewrite Zfloor_opp. + rewrite Z.even_opp. + rewrite Zceil_non_floor by lra. + rewrite Z.even_succ. + rewrite Z.negb_odd. + destruct (Z.even (Zfloor x)); reflexivity. + } + rewrite Zfloor_opp. + rewrite opp_IZR. + ring_simplify. + rewrite Zceil_non_floor by lra. + unfold Z.succ. + rewrite plus_IZR. + lra. + - rewrite Rcompare_Lt. + { apply Zfloor_opp. } + rewrite Zfloor_opp. + rewrite opp_IZR. + rewrite Zceil_non_floor by lra. + unfold Z.succ. + rewrite plus_IZR. + lra. +Qed. + Theorem ZofB_ne_correct: forall f, ZofB_ne f = if is_finite _ _ f then Some (ZnearestE (B2R _ _ f)) else None. -- cgit From 0825b433dd2ebd947d5d4566ce8a3dfd8a98db88 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 7 Dec 2021 12:00:54 +0100 Subject: ZnearestE_opp --- lib/IEEE754_extra.v | 124 ++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 121 insertions(+), 3 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 1ee9c2e0..83cacb2b 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -929,15 +929,133 @@ Proof. rewrite Ropp_involutive. reflexivity. Qed. - + +Lemma Zceil_non_floor: + forall x : R, (x > IZR(Zfloor x))%R -> Zceil x = Z.succ(Zfloor x). +Proof. + intros x BETWEEN. + unfold Z.succ. + apply Zceil_imp. + split. + { rewrite minus_IZR. + rewrite plus_IZR. + lra. + } + rewrite plus_IZR. + pose proof (Zfloor_ub x). + lra. +Qed. + +Lemma Zceil_non_ceil: + forall x : R, (x < IZR(Zceil x))%R -> Zceil x = Z.succ(Zfloor x). +Proof. + intros x BETWEEN. + unfold Z.succ. + cut (Zfloor x = (Zceil x) - 1). { intros; lia. } + apply Zfloor_imp. + split. + { rewrite minus_IZR. + pose proof (Zceil_lb x). + lra. + } + rewrite plus_IZR. + rewrite minus_IZR. + lra. +Qed. + +(* Can be also done with Znearest_opp and functional extensionality *) Lemma ZnearestE_opp : forall x : R, ZnearestE (- x) = - ZnearestE x. -Admitted. - +Proof. + intro x. + unfold ZnearestE. + case (Rcompare_spec (x - IZR (Zfloor x)) (/ 2)); intro CMP. + - pose proof (Zfloor_lb x) as LB. + destruct (Rcompare_spec x (IZR (Zfloor x))) as [ ABSURD | EXACT | INEXACT]. + lra. + { set (n := Zfloor x) in *. + rewrite EXACT. + rewrite <- opp_IZR. + rewrite Zfloor_IZR. + rewrite opp_IZR. + rewrite Rcompare_Lt by lra. + reflexivity. + } + rewrite Rcompare_Gt. + { apply Zceil_opp. } + rewrite Zfloor_opp. + rewrite opp_IZR. + rewrite Zceil_non_floor by assumption. + unfold Z.succ. + rewrite plus_IZR. + lra. + - rewrite Rcompare_Eq. + { rewrite Zceil_opp. + rewrite Zfloor_opp. + rewrite Z.even_opp. + rewrite Zceil_non_floor by lra. + rewrite Z.even_succ. + rewrite Z.negb_odd. + destruct (Z.even (Zfloor x)); reflexivity. + } + rewrite Zfloor_opp. + rewrite opp_IZR. + ring_simplify. + rewrite Zceil_non_floor by lra. + unfold Z.succ. + rewrite plus_IZR. + lra. + - rewrite Rcompare_Lt. + { apply Zfloor_opp. } + rewrite Zfloor_opp. + rewrite opp_IZR. + rewrite Zceil_non_floor by lra. + unfold Z.succ. + rewrite plus_IZR. + lra. +Qed. Theorem ZofB_ne_correct: forall f, ZofB_ne f = if is_finite _ _ f then Some (ZnearestE (B2R _ _ f)) else None. +Proof. + destruct f as [s|s|s p H|s m e H]; simpl; auto. +- f_equal. symmetry. apply (ZnearestE_IZR 0). +- destruct e; f_equal. + + unfold F2R; cbn. rewrite Rmult_1_r. rewrite ZnearestE_IZR. auto. + + unfold F2R; cbn. rewrite <- mult_IZR. rewrite ZnearestE_IZR. auto. + + unfold F2R; cbn. rewrite IZR_cond_Zopp. rewrite <- cond_Ropp_mult_l. + assert (EQ: forall x, ZnearestE (cond_Ropp s x) = cond_Zopp s (ZnearestE x)). + { intros. destruct s; cbn; auto. apply ZnearestE_opp. } + rewrite EQ. f_equal. + generalize (Zpower_pos_gt_0 2 p (eq_refl _)); intros. + set (p2p := (Z.pow_pos 2 p)) in *. + set (zm := Z.pos m) in *. + unfold Zdiv_ne, Z.succ in *. + case Z.compare_spec; intro CMP. + * admit. + * symmetry. + apply Znearest_imp with (n := zm / p2p). + apply Rabs_lt. + split. + { admit. } + assert (0 < 2*(IZR p2p * IZR (zm / p2p) - IZR zm) + (IZR p2p))%R as LT. + { rewrite <- mult_IZR. + rewrite <- minus_IZR. + rewrite <- mult_IZR. + rewrite <- plus_IZR. + apply IZR_lt. + lia. } + assert (0 < IZR p2p)%R. + { apply IZR_lt. assumption. } + assert (0 < -(IZR zm * / IZR p2p - IZR (zm / p2p) - / 2))%R as GT. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + lra. + * admit. Admitted. -- cgit From b08a818cbc7872b1a062a80ea478b2e32e3a3746 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 8 Dec 2021 22:32:41 +0100 Subject: ZofB_ne_correct unfinished --- lib/IEEE754_extra.v | 19 ++++++++++++++++++- 1 file changed, 18 insertions(+), 1 deletion(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 56b90727..a0ae3089 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1038,7 +1038,24 @@ Proof. apply Znearest_imp with (n := zm / p2p). apply Rabs_lt. split. - { admit. } + { assert (p2p > 0) as POS by lia. + pose proof (Z_mult_div_ge zm p2p POS). + assert (0 <= IZR zm * / IZR p2p - IZR (zm / p2p))%R. + 2: lra. + assert (0 < IZR p2p)%R as POS2. + { apply IZR_lt. lia. } + field_simplify. + 2: lra. + apply Rmult_le_pos. + { rewrite <- mult_IZR. + rewrite <- minus_IZR. + apply IZR_le. + lia. + } + assert (0 < / IZR p2p)%R. + 2: lra. + apply Rinv_0_lt_compat. assumption. + } assert (0 < 2*(IZR p2p * IZR (zm / p2p) - IZR zm) + (IZR p2p))%R as LT. { rewrite <- mult_IZR. rewrite <- minus_IZR. -- cgit From bfbcc770319a0d3b8935314b16f1e0c205436685 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 15:19:01 +0100 Subject: progress on ZnearestE --- lib/IEEE754_extra.v | 78 ++++++++++++++++++++++++++++++++++++----------------- 1 file changed, 54 insertions(+), 24 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index a0ae3089..edb3f105 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1031,19 +1031,18 @@ Proof. generalize (Zpower_pos_gt_0 2 p (eq_refl _)); intros. set (p2p := (Z.pow_pos 2 p)) in *. set (zm := Z.pos m) in *. + assert (p2p > 0) as POS by lia. + assert (0 < IZR p2p)%R as POS2. + { apply IZR_lt. assumption. } unfold Zdiv_ne, Z.succ in *. case Z.compare_spec; intro CMP. * admit. * symmetry. apply Znearest_imp with (n := zm / p2p). - apply Rabs_lt. - split. - { assert (p2p > 0) as POS by lia. - pose proof (Z_mult_div_ge zm p2p POS). + apply Rabs_lt. split. + ** pose proof (Z_mult_div_ge zm p2p POS). assert (0 <= IZR zm * / IZR p2p - IZR (zm / p2p))%R. 2: lra. - assert (0 < IZR p2p)%R as POS2. - { apply IZR_lt. lia. } field_simplify. 2: lra. apply Rmult_le_pos. @@ -1055,24 +1054,55 @@ Proof. assert (0 < / IZR p2p)%R. 2: lra. apply Rinv_0_lt_compat. assumption. - } - assert (0 < 2*(IZR p2p * IZR (zm / p2p) - IZR zm) + (IZR p2p))%R as LT. - { rewrite <- mult_IZR. - rewrite <- minus_IZR. - rewrite <- mult_IZR. - rewrite <- plus_IZR. - apply IZR_lt. - lia. } - assert (0 < IZR p2p)%R. - { apply IZR_lt. assumption. } - assert (0 < -(IZR zm * / IZR p2p - IZR (zm / p2p) - / 2))%R as GT. - 2: lra. - field_simplify. - 2: lra. - apply Rdiv_lt_0_compat. - 2: lra. - lra. - * admit. + ** assert (0 < 2*(IZR p2p * IZR (zm / p2p) - IZR zm) + (IZR p2p))%R as LT. + { rewrite <- mult_IZR. + rewrite <- minus_IZR. + rewrite <- mult_IZR. + rewrite <- plus_IZR. + apply IZR_lt. + lia. } + assert (0 < -(IZR zm * / IZR p2p - IZR (zm / p2p) - / 2))%R as GT. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + lra. + * symmetry. + apply Znearest_imp. + apply Rabs_lt. split. + ** assert (0 < (IZR zm - IZR p2p * IZR (zm / p2p)) - (IZR p2p * (IZR (zm / p2p) + 1) - IZR zm))%R. + { ring_simplify. + rewrite <- mult_IZR. + rewrite <- mult_IZR. + rewrite <- mult_IZR. + rewrite <- minus_IZR. + rewrite <- minus_IZR. + apply IZR_lt. + lia. + } + assert (0 < (/ 2) + IZR zm * / IZR p2p - IZR (zm / p2p + 1))%R. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + rewrite plus_IZR. + lra. + ** assert (0 < IZR (zm / p2p + 1) - (IZR zm * / IZR p2p))%R. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + rewrite <- mult_IZR. + rewrite <- minus_IZR. + apply IZR_lt. + pose proof (Z_div_mod_eq_full zm p2p) as DECOMPOSE. + ring_simplify. + set (q := (zm / p2p)) in *. + pose proof (Z_mod_lt zm p2p POS) as MOD. + lia. Admitted. -- cgit From 30d3a679e09e5aef65a34b2d12e7d8ccfa9e4ae2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 19:02:35 +0100 Subject: ZofB_ne_correct --- lib/IEEE754_extra.v | 33 +++++++++++++++++++++++++++++++-- 1 file changed, 31 insertions(+), 2 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index edb3f105..d9329960 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1036,7 +1036,36 @@ Proof. { apply IZR_lt. assumption. } unfold Zdiv_ne, Z.succ in *. case Z.compare_spec; intro CMP. - * admit. + * pose proof (Z_div_mod_eq_full zm p2p) as DECOMPOSE. + destruct (Z_mod_lt zm p2p POS) as [MOD1 MOD2]. + set (q := zm / p2p) in *. + set (r := zm mod p2p) in *. + rewrite inbetween_int_NE with (m := q) (l := loc_Inexact Eq). + { cbn. unfold cond_incr. + destruct Z.even; reflexivity. + } + constructor. + split. + ** assert (0 < IZR zm / IZR p2p - IZR q)%R. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + rewrite <- mult_IZR. + rewrite <- minus_IZR. + apply IZR_lt. + lia. + ** assert (0 < IZR (q + 1) - (IZR zm * / IZR p2p))%R. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + rewrite <- mult_IZR. + rewrite <- minus_IZR. + apply IZR_lt. + lia. * symmetry. apply Znearest_imp with (n := zm / p2p). apply Rabs_lt. split. @@ -1103,7 +1132,7 @@ Proof. set (q := (zm / p2p)) in *. pose proof (Z_mod_lt zm p2p POS) as MOD. lia. -Admitted. +Qed. (** ** Algebraic identities *) -- cgit From d3f9053b783714151aab27eb153f1b727e843ad1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 19:02:35 +0100 Subject: ZofB_ne_correct --- lib/IEEE754_extra.v | 33 +++++++++++++++++++++++++++++++-- 1 file changed, 31 insertions(+), 2 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index edb3f105..d9329960 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1036,7 +1036,36 @@ Proof. { apply IZR_lt. assumption. } unfold Zdiv_ne, Z.succ in *. case Z.compare_spec; intro CMP. - * admit. + * pose proof (Z_div_mod_eq_full zm p2p) as DECOMPOSE. + destruct (Z_mod_lt zm p2p POS) as [MOD1 MOD2]. + set (q := zm / p2p) in *. + set (r := zm mod p2p) in *. + rewrite inbetween_int_NE with (m := q) (l := loc_Inexact Eq). + { cbn. unfold cond_incr. + destruct Z.even; reflexivity. + } + constructor. + split. + ** assert (0 < IZR zm / IZR p2p - IZR q)%R. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + rewrite <- mult_IZR. + rewrite <- minus_IZR. + apply IZR_lt. + lia. + ** assert (0 < IZR (q + 1) - (IZR zm * / IZR p2p))%R. + 2: lra. + field_simplify. + 2: lra. + apply Rdiv_lt_0_compat. + 2: lra. + rewrite <- mult_IZR. + rewrite <- minus_IZR. + apply IZR_lt. + lia. * symmetry. apply Znearest_imp with (n := zm / p2p). apply Rabs_lt. split. @@ -1103,7 +1132,7 @@ Proof. set (q := (zm / p2p)) in *. pose proof (Z_mod_lt zm p2p POS) as MOD. lia. -Admitted. +Qed. (** ** Algebraic identities *) -- cgit From 75ae1bbac93b1f3cdbdd16a266c5658a28bde90c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 19:53:15 +0100 Subject: ZofB_ne_correct --- lib/IEEE754_extra.v | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index d9329960..28da42ee 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1066,6 +1066,16 @@ Proof. rewrite <- minus_IZR. apply IZR_lt. lia. + ** apply Rcompare_Eq. + assert ((IZR q + IZR (q + 1))/2 - (IZR zm * / IZR p2p) = 0)%R; [idtac|lra]. + field_simplify; [idtac | lra]. + rewrite <- mult_IZR. + rewrite <- mult_IZR. + rewrite <- mult_IZR. + rewrite <- plus_IZR. + rewrite <- minus_IZR. + replace (q * p2p + (q + 1) * p2p - 2 * zm) with 0 by lia. + field. lra. * symmetry. apply Znearest_imp with (n := zm / p2p). apply Rabs_lt. split. -- cgit From 0be6b3d13d2e8920653dde8cf8f7e27af608d812 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 21:47:42 +0100 Subject: beautify proof --- lib/IEEE754_extra.v | 40 ++++++++++++++++------------------------ 1 file changed, 16 insertions(+), 24 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 28da42ee..2806e572 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1014,7 +1014,10 @@ Proof. rewrite plus_IZR. lra. Qed. - + +Ltac field_simplify_den := field_simplify ; [idtac | lra]. +Ltac Rdiv_lt_0_den := apply Rdiv_lt_0_compat ; [idtac | lra]. + Theorem ZofB_ne_correct: forall f, ZofB_ne f = if is_finite _ _ f then Some (ZnearestE (B2R _ _ f)) else None. @@ -1048,27 +1051,23 @@ Proof. split. ** assert (0 < IZR zm / IZR p2p - IZR q)%R. 2: lra. - field_simplify. - 2: lra. - apply Rdiv_lt_0_compat. - 2: lra. + field_simplify_den. + Rdiv_lt_0_den. rewrite <- mult_IZR. rewrite <- minus_IZR. apply IZR_lt. lia. ** assert (0 < IZR (q + 1) - (IZR zm * / IZR p2p))%R. 2: lra. - field_simplify. - 2: lra. - apply Rdiv_lt_0_compat. - 2: lra. + field_simplify_den. + Rdiv_lt_0_den. rewrite <- mult_IZR. rewrite <- minus_IZR. apply IZR_lt. lia. ** apply Rcompare_Eq. assert ((IZR q + IZR (q + 1))/2 - (IZR zm * / IZR p2p) = 0)%R; [idtac|lra]. - field_simplify; [idtac | lra]. + field_simplify_den. rewrite <- mult_IZR. rewrite <- mult_IZR. rewrite <- mult_IZR. @@ -1082,8 +1081,7 @@ Proof. ** pose proof (Z_mult_div_ge zm p2p POS). assert (0 <= IZR zm * / IZR p2p - IZR (zm / p2p))%R. 2: lra. - field_simplify. - 2: lra. + field_simplify_den. apply Rmult_le_pos. { rewrite <- mult_IZR. rewrite <- minus_IZR. @@ -1102,10 +1100,8 @@ Proof. lia. } assert (0 < -(IZR zm * / IZR p2p - IZR (zm / p2p) - / 2))%R as GT. 2: lra. - field_simplify. - 2: lra. - apply Rdiv_lt_0_compat. - 2: lra. + field_simplify_den. + Rdiv_lt_0_den. lra. * symmetry. apply Znearest_imp. @@ -1122,18 +1118,14 @@ Proof. } assert (0 < (/ 2) + IZR zm * / IZR p2p - IZR (zm / p2p + 1))%R. 2: lra. - field_simplify. - 2: lra. - apply Rdiv_lt_0_compat. - 2: lra. + field_simplify_den. + Rdiv_lt_0_den. rewrite plus_IZR. lra. ** assert (0 < IZR (zm / p2p + 1) - (IZR zm * / IZR p2p))%R. 2: lra. - field_simplify. - 2: lra. - apply Rdiv_lt_0_compat. - 2: lra. + field_simplify_den. + Rdiv_lt_0_den. rewrite <- mult_IZR. rewrite <- minus_IZR. apply IZR_lt. -- cgit From 9d4528e4edaa2b968d52628e6b5c10e6640f5e03 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 21:59:10 +0100 Subject: factorize proof --- lib/IEEE754_extra.v | 50 ++++++++++++++++++++++---------------------------- 1 file changed, 22 insertions(+), 28 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 2806e572..b8397210 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -899,19 +899,25 @@ Definition ZofB_ne (f: binary_float): option Z := | _ => None end. -Lemma ZnearestE_IZR: - forall n, (ZnearestE (IZR n)) = n. +Lemma Znearest_IZR : + forall choice n, (Znearest choice (IZR n)) = n. Proof. - intro n. - unfold ZnearestE. + intros. + unfold Znearest. case Rcompare_spec ; intro ORDER. - apply Zfloor_IZR. - - destruct negb. + - destruct choice. + apply Zceil_IZR. + apply Zfloor_IZR. - apply Zceil_IZR. Qed. +Lemma ZnearestE_IZR: + forall n, (ZnearestE (IZR n)) = n. +Proof. + apply Znearest_IZR. +Qed. + Lemma Zfloor_opp : forall x : R, (Zfloor (- x)) = - (Zceil x). Proof. @@ -1018,6 +1024,9 @@ Qed. Ltac field_simplify_den := field_simplify ; [idtac | lra]. Ltac Rdiv_lt_0_den := apply Rdiv_lt_0_compat ; [idtac | lra]. +Hint Rewrite <- plus_IZR minus_IZR opp_IZR mult_IZR : l_IZR. +Ltac l_IZR := autorewrite with l_IZR. + Theorem ZofB_ne_correct: forall f, ZofB_ne f = if is_finite _ _ f then Some (ZnearestE (B2R _ _ f)) else None. @@ -1053,28 +1062,22 @@ Proof. 2: lra. field_simplify_den. Rdiv_lt_0_den. - rewrite <- mult_IZR. - rewrite <- minus_IZR. + l_IZR. apply IZR_lt. lia. ** assert (0 < IZR (q + 1) - (IZR zm * / IZR p2p))%R. 2: lra. field_simplify_den. Rdiv_lt_0_den. - rewrite <- mult_IZR. - rewrite <- minus_IZR. + l_IZR. apply IZR_lt. lia. ** apply Rcompare_Eq. assert ((IZR q + IZR (q + 1))/2 - (IZR zm * / IZR p2p) = 0)%R; [idtac|lra]. field_simplify_den. - rewrite <- mult_IZR. - rewrite <- mult_IZR. - rewrite <- mult_IZR. - rewrite <- plus_IZR. - rewrite <- minus_IZR. + l_IZR. replace (q * p2p + (q + 1) * p2p - 2 * zm) with 0 by lia. - field. lra. + field. apply IZR_neq. lia. * symmetry. apply Znearest_imp with (n := zm / p2p). apply Rabs_lt. split. @@ -1083,8 +1086,7 @@ Proof. 2: lra. field_simplify_den. apply Rmult_le_pos. - { rewrite <- mult_IZR. - rewrite <- minus_IZR. + { l_IZR. apply IZR_le. lia. } @@ -1092,10 +1094,7 @@ Proof. 2: lra. apply Rinv_0_lt_compat. assumption. ** assert (0 < 2*(IZR p2p * IZR (zm / p2p) - IZR zm) + (IZR p2p))%R as LT. - { rewrite <- mult_IZR. - rewrite <- minus_IZR. - rewrite <- mult_IZR. - rewrite <- plus_IZR. + { l_IZR. apply IZR_lt. lia. } assert (0 < -(IZR zm * / IZR p2p - IZR (zm / p2p) - / 2))%R as GT. @@ -1108,11 +1107,7 @@ Proof. apply Rabs_lt. split. ** assert (0 < (IZR zm - IZR p2p * IZR (zm / p2p)) - (IZR p2p * (IZR (zm / p2p) + 1) - IZR zm))%R. { ring_simplify. - rewrite <- mult_IZR. - rewrite <- mult_IZR. - rewrite <- mult_IZR. - rewrite <- minus_IZR. - rewrite <- minus_IZR. + l_IZR. apply IZR_lt. lia. } @@ -1126,8 +1121,7 @@ Proof. 2: lra. field_simplify_den. Rdiv_lt_0_den. - rewrite <- mult_IZR. - rewrite <- minus_IZR. + l_IZR. apply IZR_lt. pose proof (Z_div_mod_eq_full zm p2p) as DECOMPOSE. ring_simplify. -- cgit From d2669ab4948104c85b795c15a47df98f0b81a40c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 22:37:29 +0100 Subject: factor proofs --- lib/IEEE754_extra.v | 21 ++++++++++++++++++++- 1 file changed, 20 insertions(+), 1 deletion(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index b8397210..761f218a 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -24,6 +24,8 @@ Require Import Psatz. Require Import Bool. Require Import Eqdep_dec. +Require Import Coq.Logic.FunctionalExtensionality. + Local Open Scope Z_scope. Section Extra_ops. @@ -936,6 +938,23 @@ Proof. reflexivity. Qed. +Lemma ZnearestE_opp + : forall x : R, ZnearestE (- x) = - ZnearestE x. +Proof. + intro. + rewrite Znearest_opp. + f_equal. + f_equal. + apply functional_extensionality. + intro. + rewrite Z.even_opp. + fold (Z.succ x0). + rewrite Z.even_succ. + f_equal. + apply Z.negb_odd. +Qed. + +(** more complicated way of proving Lemma Zceil_non_floor: forall x : R, (x > IZR(Zfloor x))%R -> Zceil x = Z.succ(Zfloor x). Proof. @@ -969,7 +988,6 @@ Proof. lra. Qed. -(* Can be also done with Znearest_opp and functional extensionality *) Lemma ZnearestE_opp : forall x : R, ZnearestE (- x) = - ZnearestE x. Proof. @@ -1020,6 +1038,7 @@ Proof. rewrite plus_IZR. lra. Qed. + *) Ltac field_simplify_den := field_simplify ; [idtac | lra]. Ltac Rdiv_lt_0_den := apply Rdiv_lt_0_compat ; [idtac | lra]. -- cgit From 4fb3896a10dac9137157262f532a6a5b6e390cf4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 23:24:28 +0100 Subject: more properties on Znearest --- lib/IEEE754_extra.v | 37 +++++++++++++++++++++++++++++++++++-- 1 file changed, 35 insertions(+), 2 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 761f218a..ed3d9f1f 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -954,7 +954,6 @@ Proof. apply Z.negb_odd. Qed. -(** more complicated way of proving Lemma Zceil_non_floor: forall x : R, (x > IZR(Zfloor x))%R -> Zceil x = Z.succ(Zfloor x). Proof. @@ -971,6 +970,7 @@ Proof. lra. Qed. +(** more complicated way of proving Lemma Zceil_non_ceil: forall x : R, (x < IZR(Zceil x))%R -> Zceil x = Z.succ(Zfloor x). Proof. @@ -1148,7 +1148,40 @@ Proof. pose proof (Z_mod_lt zm p2p POS) as MOD. lia. Qed. - + + +Lemma Znearest_imp2: + forall choice x, (Rabs (IZR (Znearest choice x) - x) <= /2)%R. +Proof. + intros. + unfold Znearest. + pose proof (Zfloor_lb x) as FL. + pose proof (Zceil_ub x) as CU. + pose proof (Zceil_non_floor x) as NF. + case Rcompare_spec; intro CMP; apply Rabs_le; split; try lra. + - destruct choice; lra. + - destruct choice. 2: lra. + rewrite NF. 2: lra. + unfold Z.succ. rewrite plus_IZR. lra. + - rewrite NF. 2: lra. + unfold Z.succ. rewrite plus_IZR. lra. +Qed. + +Lemma Rabs_le_rev : forall {a} {b}, (Rabs a <= b -> -b <= a <= b)%R. +Proof. + intros a b ABS. + unfold Rabs in ABS. + destruct Rcase_abs in ABS; lra. +Qed. + +Theorem ZofB_ne_range: + forall f n, ZofB_ne f = Some n -> (IZR n-1/2 <= B2R _ _ f <= IZR n+1/2)%R. +Proof. + intros. rewrite ZofB_ne_correct in H. destruct (is_finite prec emax f) eqn:FIN; inversion H. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) (B2R prec emax f)) as ABS. + pose proof (Rabs_le_rev ABS). + lra. +Qed. (** ** Algebraic identities *) -- cgit From 1c77deb420e626aa79d5b7db15231c887cd7d870 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Dec 2021 23:29:25 +0100 Subject: minor implicit issue --- lib/IEEE754_extra.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index ed3d9f1f..a76daab4 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1167,7 +1167,7 @@ Proof. unfold Z.succ. rewrite plus_IZR. lra. Qed. -Lemma Rabs_le_rev : forall {a} {b}, (Rabs a <= b -> -b <= a <= b)%R. +Lemma Rabs_le_rev : forall a b, (Rabs a <= b -> -b <= a <= b)%R. Proof. intros a b ABS. unfold Rabs in ABS. @@ -1179,7 +1179,7 @@ Theorem ZofB_ne_range: Proof. intros. rewrite ZofB_ne_correct in H. destruct (is_finite prec emax f) eqn:FIN; inversion H. pose proof (Znearest_imp2 (fun x => negb (Z.even x)) (B2R prec emax f)) as ABS. - pose proof (Rabs_le_rev ABS). + pose proof (Rabs_le_rev _ _ ABS). lra. Qed. -- cgit From 841fff988e26eb44f7aceeab1d77be5833873625 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 12 Dec 2021 13:11:59 +0100 Subject: progress on lemmas --- lib/IEEE754_extra.v | 82 ++++++++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 81 insertions(+), 1 deletion(-) diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index a76daab4..7e0e7260 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -1174,7 +1174,7 @@ Proof. destruct Rcase_abs in ABS; lra. Qed. -Theorem ZofB_ne_range: +Theorem ZofB_ne_ball: forall f n, ZofB_ne f = Some n -> (IZR n-1/2 <= B2R _ _ f <= IZR n+1/2)%R. Proof. intros. rewrite ZofB_ne_correct in H. destruct (is_finite prec emax f) eqn:FIN; inversion H. @@ -1183,6 +1183,86 @@ Proof. lra. Qed. +(* +Theorem ZofB_ne_minus: + forall minus_nan m f p q, + ZofB_ne f = Some p -> 0 <= p < 2*q -> q <= 2^prec -> (IZR q <= B2R _ _ f)%R -> + ZofB_ne (Bminus _ _ _ Hmax minus_nan m f (BofZ q)) = Some (p - q). +Proof. + intros. + assert (Q: -2^prec <= q <= 2^prec). + { split; auto. generalize (Zpower_ge_0 radix2 prec); simpl; lia. } + assert (RANGE: (IZR p -1/2 <= B2R _ _ f <= IZR p + 1/2)%R) by ( apply ZofB_ne_ball; auto ). + rewrite ZofB_ne_correct in H. destruct (is_finite prec emax f) eqn:FIN; try discriminate. + assert (PQ2: (IZR p + 1 <= IZR q * 2)%R). + { l_IZR. apply IZR_le. lia. } + assert (EXACT: round radix2 fexp (round_mode m) (B2R _ _ f - IZR q)%R = (B2R _ _ f - IZR q)%R). + { apply round_generic. apply valid_rnd_round_mode. + apply sterbenz_aux. now apply FLT_exp_valid. apply FLT_exp_monotone. apply generic_format_B2R. + apply integer_representable_n. auto. lra. } + 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_ne_correct. rewrite E. rewrite D. rewrite A. rewrite EXACT. + inversion H. f_equal. + rewrite ! Ztrunc_floor. apply Zfloor_minus. + lra. lra. +- rewrite A. fold emin; fold fexp. rewrite EXACT. + apply Rle_lt_trans with (bpow radix2 prec). + apply Rle_trans with (IZR q). apply Rabs_le. lra. + rewrite <- IZR_Zpower. apply IZR_le; auto. red in prec_gt_0_; lia. + apply bpow_lt. auto. +Qed. + *) + +Definition ZofB_ne_range (f: binary_float) (zmin zmax: Z): option Z := + match ZofB_ne f with + | None => None + | Some z => if Z.leb zmin z && Z.leb z zmax then Some z else None + end. + +Theorem ZofB_ne_range_correct: + forall f min max, + let n := ZnearestE (B2R _ _ f) in + ZofB_ne_range f min max = + if is_finite _ _ f && Z.leb min n && Z.leb n max then Some n else None. +Proof. + intros. unfold ZofB_ne_range. rewrite ZofB_ne_correct. fold n. + destruct (is_finite prec emax f); auto. +Qed. + +Lemma ZofB_ne_range_inversion: + forall f min max n, + ZofB_ne_range f min max = Some n -> + min <= n /\ n <= max /\ ZofB_ne f = Some n. +Proof. + intros. rewrite ZofB_ne_range_correct in H. rewrite ZofB_ne_correct. + destruct (is_finite prec emax f); try discriminate. + set (n1 := ZnearestE (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. + split. apply Zle_bool_imp_le; auto. + auto. +Qed. + + +(* +Theorem ZofB_ne_range_minus: + forall minus_nan m f p q, + ZofB_ne_range f 0 (2 * q - 1) = Some p -> q <= 2^prec -> (IZR q <= B2R _ _ f)%R -> + ZofB_ne_range (Bminus _ _ _ Hmax minus_nan m f (BofZ q)) (-q) (q - 1) = Some (p - q). +Proof. + intros. destruct (ZofB_ne_range_inversion _ _ _ _ H) as (A & B & C). + set (f' := Bminus prec emax prec_gt_0_ Hmax minus_nan m f (BofZ q)). + assert (D: ZofB_ne f' = Some (p - q)). + { apply ZofB_ne_minus. auto. lia. auto. auto. } + unfold ZofB_range. rewrite D. rewrite Zle_bool_true by lia. rewrite Zle_bool_true by lia. auto. +Qed. + *) + (** ** Algebraic identities *) (** Commutativity of addition and multiplication *) -- cgit From 5e83a412e68916f02d86f6e7b719cd93d50c152d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 12 Dec 2021 13:13:09 +0100 Subject: _ne conversions --- lib/Floats.v | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/lib/Floats.v b/lib/Floats.v index 9ee5302d..bc066b20 100644 --- a/lib/Floats.v +++ b/lib/Floats.v @@ -360,6 +360,15 @@ Definition to_long (f:float): option int64 := (**r conversion to signed 64-bit i Definition to_longu (f:float): option int64 := (**r conversion to unsigned 64-bit int *) option_map Int64.repr (ZofB_range _ _ f 0 Int64.max_unsigned). +Definition to_int_ne (f:float): option int := (**r conversion to signed 32-bit int *) + option_map Int.repr (ZofB_ne_range _ _ f Int.min_signed Int.max_signed). +Definition to_intu_ne (f:float): option int := (**r conversion to unsigned 32-bit int *) + option_map Int.repr (ZofB_ne_range _ _ f 0 Int.max_unsigned). +Definition to_long_ne (f:float): option int64 := (**r conversion to signed 64-bit int *) + option_map Int64.repr (ZofB_ne_range _ _ f Int64.min_signed Int64.max_signed). +Definition to_longu_ne (f:float): option int64 := (**r conversion to unsigned 64-bit int *) + option_map Int64.repr (ZofB_ne_range _ _ f 0 Int64.max_unsigned). + Definition of_int (n:int): float := (**r conversion from signed 32-bit int *) BofZ 53 1024 __ __ (Int.signed n). Definition of_intu (n:int): float:= (**r conversion from unsigned 32-bit int *) -- cgit From 1ecddb62bc5aa8e6a1f6c1a1a2da2e2a8e2b100f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 12 Dec 2021 15:18:18 +0100 Subject: Values: conversions to nearest int --- common/Values.v | 48 ++++++++++++++++++++++++++++++++++++++++++++++++ lib/Floats.v | 9 +++++++++ 2 files changed, 57 insertions(+) diff --git a/common/Values.v b/common/Values.v index 87ebea00..cf5a1255 100644 --- a/common/Values.v +++ b/common/Values.v @@ -244,6 +244,18 @@ Definition intuoffloat (v: val) : option val := | _ => None end. +Definition intoffloat_ne (v: val) : option val := + match v with + | Vfloat f => option_map Vint (Float.to_int_ne f) + | _ => None + end. + +Definition intuoffloat_ne (v: val) : option val := + match v with + | Vfloat f => option_map Vint (Float.to_intu_ne f) + | _ => None + end. + Definition floatofint (v: val) : option val := match v with | Vint n => Some (Vfloat (Float.of_int n)) @@ -268,6 +280,18 @@ Definition intuofsingle (v: val) : option val := | _ => None end. +Definition intofsingle_ne (v: val) : option val := + match v with + | Vsingle f => option_map Vint (Float32.to_int_ne f) + | _ => None + end. + +Definition intuofsingle_ne (v: val) : option val := + match v with + | Vsingle f => option_map Vint (Float32.to_intu_ne f) + | _ => None + end. + Definition singleofint (v: val) : option val := match v with | Vint n => Some (Vsingle (Float32.of_int n)) @@ -623,6 +647,30 @@ Definition longuofsingle (v: val) : option val := | _ => None end. +Definition longoffloat_ne (v: val) : option val := + match v with + | Vfloat f => option_map Vlong (Float.to_long_ne f) + | _ => None + end. + +Definition longuoffloat_ne (v: val) : option val := + match v with + | Vfloat f => option_map Vlong (Float.to_longu_ne f) + | _ => None + end. + +Definition longofsingle_ne (v: val) : option val := + match v with + | Vsingle f => option_map Vlong (Float32.to_long_ne f) + | _ => None + end. + +Definition longuofsingle_ne (v: val) : option val := + match v with + | Vsingle f => option_map Vlong (Float32.to_longu_ne f) + | _ => None + end. + Definition floatoflong (v: val) : option val := match v with | Vlong n => Some (Vfloat (Float.of_long n)) diff --git a/lib/Floats.v b/lib/Floats.v index bc066b20..b12c6585 100644 --- a/lib/Floats.v +++ b/lib/Floats.v @@ -1145,6 +1145,15 @@ Definition to_long (f:float32): option int64 := (**r conversion to signed 64-bit Definition to_longu (f:float32): option int64 := (**r conversion to unsigned 64-bit int *) option_map Int64.repr (ZofB_range _ _ f 0 Int64.max_unsigned). +Definition to_int_ne (f:float32): option int := (**r conversion to signed 32-bit int *) + option_map Int.repr (ZofB_ne_range _ _ f Int.min_signed Int.max_signed). +Definition to_intu_ne (f:float32): option int := (**r conversion to unsigned 32-bit int *) + option_map Int.repr (ZofB_ne_range _ _ f 0 Int.max_unsigned). +Definition to_long_ne (f:float32): option int64 := (**r conversion to signed 64-bit int *) + option_map Int64.repr (ZofB_ne_range _ _ f Int64.min_signed Int64.max_signed). +Definition to_longu_ne (f:float32): option int64 := (**r conversion to unsigned 64-bit int *) + option_map Int64.repr (ZofB_ne_range _ _ f 0 Int64.max_unsigned). + Definition of_int (n:int): float32 := (**r conversion from signed 32-bit int to single-precision float *) BofZ 24 128 __ __ (Int.signed n). Definition of_intu (n:int): float32 := (**r conversion from unsigned 32-bit int to single-precision float *) -- cgit From 4f7d6d6a081de52fe1151a29d44221f4fc35f7be Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 12 Dec 2021 17:31:41 +0100 Subject: Asm level --- kvx/Asm.v | 12 ++++++++++++ kvx/Asmvliw.v | 9 ++++++++- kvx/TargetPrinter.ml | 8 ++++++++ 3 files changed, 28 insertions(+), 1 deletion(-) diff --git a/kvx/Asm.v b/kvx/Asm.v index fd20316c..333854cf 100644 --- a/kvx/Asm.v +++ b/kvx/Asm.v @@ -163,6 +163,8 @@ Inductive instruction : Type := | Pfloatuwrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (u32 -> 32) *) | Pfloatudrnsz (rd rs: ireg) (**r Floating Point Conversion from unsigned integer (64 bits) *) | Pfloatdrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (64 bits) *) + + (* round to zero *) | Pfixedwrzz (rd rs: ireg) (**r Integer conversion from floating point *) | Pfixeduwrzz (rd rs: ireg) (**r Integer conversion from floating point (f32 -> 32 bits unsigned *) | Pfixeddrzz (rd rs: ireg) (**r Integer conversion from floating point (i64 -> 64 bits) *) @@ -170,6 +172,12 @@ Inductive instruction : Type := | Pfixedudrzz (rd rs: ireg) (**r unsigned Integer conversion from floating point (u64 -> 64 bits) *) | Pfixedudrzz_i32 (rd rs: ireg) (**r unsigned Integer conversion from floating point (u32 -> 64 bits) *) + (* round to nearest, prefer even numbers *) + | Pfixedwrne (rd rs: ireg) (**r Integer conversion from floating point *) + | Pfixeduwrne (rd rs: ireg) (**r Integer conversion from floating point (f32 -> 32 bits unsigned *) + | Pfixeddrne (rd rs: ireg) (**r Integer conversion from floating point (i64 -> 64 bits) *) + | Pfixedudrne (rd rs: ireg) (**r unsigned Integer conversion from floating point (u64 -> 64 bits) *) + (** Arith RI32 *) | Pmake (rd: ireg) (imm: int) (**r load immediate *) @@ -352,6 +360,10 @@ Definition basic_to_instruction (b: basic) := | PArithRR Asmvliw.Pfixedudrzz rd rs => Pfixedudrzz rd rs | PArithRR Asmvliw.Pfixeddrzz_i32 rd rs => Pfixeddrzz_i32 rd rs | PArithRR Asmvliw.Pfixedudrzz_i32 rd rs => Pfixedudrzz_i32 rd rs + | PArithRR Asmvliw.Pfixedwrne rd rs => Pfixedwrne rd rs + | PArithRR Asmvliw.Pfixeduwrne rd rs => Pfixeduwrne rd rs + | PArithRR Asmvliw.Pfixeddrne rd rs => Pfixeddrne rd rs + | PArithRR Asmvliw.Pfixedudrne rd rs => Pfixedudrne rd rs (* RI32 *) | PArithRI32 Asmvliw.Pmake rd imm => Pmake rd imm diff --git a/kvx/Asmvliw.v b/kvx/Asmvliw.v index 45b230e6..b3c0c8fa 100644 --- a/kvx/Asmvliw.v +++ b/kvx/Asmvliw.v @@ -402,7 +402,10 @@ Inductive arith_name_rr : Type := | Pfixedudrzz (**r Integer conversion from floating point (float -> unsigned long) *) | Pfixeddrzz_i32 (**r Integer conversion from floating point (float -> int) *) | Pfixedudrzz_i32 (**r Integer conversion from floating point (float -> unsigned int) *) -. + | Pfixedwrne (**r Integer conversion from floating point *) + | Pfixeduwrne (**r Integer conversion from floating point (f32 -> 32 bits unsigned *) + | Pfixeddrne (**r Integer conversion from floating point (i64 -> 64 bits) *) + | Pfixedudrne. (**r unsigned Integer conversion from floating point (u64 -> 64 bits) *) Inductive arith_name_ri32 : Type := | Pmake (**r load immediate *) @@ -955,6 +958,10 @@ Definition arith_eval_rr n v := | Pfixedudrzz => Val.maketotal (Val.longuoffloat v) | Pfixeddrzz_i32 => Val.maketotal (Val.intoffloat v) | Pfixedudrzz_i32 => Val.maketotal (Val.intuoffloat v) + | Pfixedudrne => Val.maketotal (Val.longuoffloat_ne v) + | Pfixeddrne => Val.maketotal (Val.longoffloat_ne v) + | Pfixeduwrne => Val.maketotal (Val.intuoffloat_ne v) + | Pfixedwrne => Val.maketotal (Val.intoffloat_ne v) end. Definition arith_eval_ri32 n i := diff --git a/kvx/TargetPrinter.ml b/kvx/TargetPrinter.ml index 9e2e3776..40e2be55 100644 --- a/kvx/TargetPrinter.ml +++ b/kvx/TargetPrinter.ml @@ -586,6 +586,14 @@ module Target (*: TARGET*) = fprintf oc " fixedd.rz %a = %a, 0\n" ireg rd ireg rs | Pfixedudrzz(rd, rs) | Pfixedudrzz_i32(rd, rs) -> fprintf oc " fixedud.rz %a = %a, 0\n" ireg rd ireg rs + | Pfixedudrne(rd, rs) -> + fprintf oc " fixedud.ne %a = %a, 0\n" ireg rd ireg rs + | Pfixeddrne(rd, rs) -> + fprintf oc " fixedd.ne %a = %a, 0\n" ireg rd ireg rs + | Pfixeduwrne(rd, rs) -> + fprintf oc " fixeduw.ne %a = %a, 0\n" ireg rd ireg rs + | Pfixedwrne(rd, rs) -> + fprintf oc " fixedw.ne %a = %a, 0\n" ireg rd ireg rs (* Arith RI32 instructions *) | Pmake (rd, imm) -> -- cgit From a3924f657b36abfad0448d0fe7d30fd40e0068de Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 12 Dec 2021 22:10:22 +0100 Subject: some more fixed etc. constructs --- backend/ValueDomain.v | 156 ++++++++++++++++++++++++++++++++++++++++ kvx/Asmblockdeps.v | 4 ++ kvx/NeedOp.v | 2 + kvx/Op.v | 30 ++++++++ kvx/PostpassSchedulingOracle.ml | 19 ++--- kvx/ValueAOp.v | 4 ++ 6 files changed, 206 insertions(+), 9 deletions(-) diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v index 5a7cfc12..8c58e32e 100644 --- a/backend/ValueDomain.v +++ b/backend/ValueDomain.v @@ -2269,6 +2269,24 @@ Proof. inv H; simpl; auto with va. rewrite E; constructor. Qed. +Definition intofsingle_ne (x: aval) := + match x with + | FS f => + match Float32.to_int_ne f with + | Some i => I i + | None => if va_strict tt then Vbot else ntop + end + | _ => ntop1 x + end. + +Lemma intofsingle_ne_sound: + forall v x w, vmatch v x -> Val.intofsingle_ne v = Some w -> vmatch w (intofsingle_ne x). +Proof. + unfold Val.intofsingle_ne; intros. destruct v; try discriminate. + destruct (Float32.to_int_ne f) as [i|] eqn:E; simpl in H0; inv H0. + inv H; simpl; auto with va. rewrite E; constructor. +Qed. + Definition intuofsingle (x: aval) := match x with | FS f => @@ -2287,6 +2305,24 @@ Proof. inv H; simpl; auto with va. rewrite E; constructor. Qed. +Definition intuofsingle_ne (x: aval) := + match x with + | FS f => + match Float32.to_intu_ne f with + | Some i => I i + | None => if va_strict tt then Vbot else ntop + end + | _ => ntop1 x + end. + +Lemma intuofsingle_ne_sound: + forall v x w, vmatch v x -> Val.intuofsingle_ne v = Some w -> vmatch w (intuofsingle_ne x). +Proof. + unfold Val.intuofsingle_ne; intros. destruct v; try discriminate. + destruct (Float32.to_intu_ne f) as [i|] eqn:E; simpl in H0; inv H0. + inv H; simpl; auto with va. rewrite E; constructor. +Qed. + Definition singleofint (x: aval) := match x with | I i => FS(Float32.of_int i) @@ -2349,6 +2385,42 @@ Proof. inv H; simpl; auto with va. rewrite E; constructor. Qed. +Definition longoffloat_ne (x: aval) := + match x with + | F f => + match Float.to_long_ne f with + | Some i => L i + | None => if va_strict tt then Vbot else ntop + end + | _ => ntop1 x + end. + +Lemma longoffloat_ne_sound: + forall v x w, vmatch v x -> Val.longoffloat_ne v = Some w -> vmatch w (longoffloat_ne x). +Proof. + unfold Val.longoffloat_ne; intros. destruct v; try discriminate. + destruct (Float.to_long_ne f) as [i|] eqn:E; simpl in H0; inv H0. + inv H; simpl; auto with va. rewrite E; constructor. +Qed. + +Definition longuoffloat_ne (x: aval) := + match x with + | F f => + match Float.to_longu_ne f with + | Some i => L i + | None => if va_strict tt then Vbot else ntop + end + | _ => ntop1 x + end. + +Lemma longuoffloat_ne_sound: + forall v x w, vmatch v x -> Val.longuoffloat_ne v = Some w -> vmatch w (longuoffloat_ne x). +Proof. + unfold Val.longuoffloat_ne; intros. destruct v; try discriminate. + destruct (Float.to_longu_ne f) as [i|] eqn:E; simpl in H0; inv H0. + inv H; simpl; auto with va. rewrite E; constructor. +Qed. + Definition floatoflong (x: aval) := match x with | L i => F(Float.of_long i) @@ -2566,6 +2638,46 @@ Definition longuofsingle_total (x: aval) := | _ => ntop1 x end. +Definition intofsingle_ne_total (x: aval) := + match x with + | FS f => + match Float32.to_int_ne f with + | Some i => I i + | None => ntop + end + | _ => ntop1 x + end. + +Definition intuofsingle_ne_total (x: aval) := + match x with + | FS f => + match Float32.to_intu_ne f with + | Some i => I i + | None => ntop + end + | _ => ntop1 x + end. + +Definition longoffloat_ne_total (x: aval) := + match x with + | F f => + match Float.to_long_ne f with + | Some i => L i + | None => ntop + end + | _ => ntop1 x + end. + +Definition longuoffloat_ne_total (x: aval) := + match x with + | F f => + match Float.to_longu_ne f with + | Some i => L i + | None => ntop + end + | _ => ntop1 x + end. + Lemma intoffloat_total_sound: forall v x (MATCH : vmatch v x), @@ -2606,6 +2718,26 @@ Proof. all: destruct (Float32.to_intu f) as [i|] eqn:E; simpl; [auto with va | constructor]. Qed. +Lemma intofsingle_ne_total_sound: + forall v x + (MATCH : vmatch v x), + vmatch (Val.maketotal (Val.intofsingle_ne v)) (intofsingle_ne_total x). +Proof. + unfold Val.intofsingle_ne, intofsingle_ne_total. intros. + inv MATCH; simpl in *; try constructor. + all: destruct (Float32.to_int_ne f) as [i|] eqn:E; simpl; [auto with va | constructor]. +Qed. + +Lemma intuofsingle_ne_total_sound: + forall v x + (MATCH : vmatch v x), + vmatch (Val.maketotal (Val.intuofsingle_ne v)) (intuofsingle_ne_total x). +Proof. + unfold Val.intofsingle, intofsingle_total. intros. + inv MATCH; simpl in *; try constructor. + all: destruct (Float32.to_intu_ne f) as [i|] eqn:E; simpl; [auto with va | constructor]. +Qed. + Lemma singleofint_total_sound: forall v x, vmatch v x -> vmatch (Val.maketotal (Val.singleofint v)) (singleofint x). @@ -2648,6 +2780,26 @@ Proof. all: destruct (Float.to_longu f) as [i|] eqn:E; simpl; [auto with va | constructor]. Qed. +Lemma longoffloat_ne_total_sound: + forall v x + (MATCH : vmatch v x), + vmatch (Val.maketotal (Val.longoffloat_ne v)) (longoffloat_ne_total x). +Proof. + unfold Val.longoffloat_ne, longoffloat_ne_total. intros. + inv MATCH; simpl in *; try constructor. + all: destruct (Float.to_long_ne f) as [i|] eqn:E; simpl; [auto with va | constructor]. +Qed. + +Lemma longuoffloat_ne_total_sound: + forall v x + (MATCH : vmatch v x), + vmatch (Val.maketotal (Val.longuoffloat_ne v)) (longuoffloat_ne_total x). +Proof. + unfold Val.longoffloat_ne, longoffloat_ne_total. intros. + inv MATCH; simpl in *; try constructor. + all: destruct (Float.to_longu_ne f) as [i|] eqn:E; simpl; [auto with va | constructor]. +Qed. + Lemma longofsingle_total_sound: forall v x (MATCH : vmatch v x), @@ -5200,6 +5352,10 @@ Global Hint Resolve cnot_sound symbol_address_sound intofsingle_sound intuofsingle_sound singleofint_sound singleofintu_sound longoffloat_sound longuoffloat_sound floatoflong_sound floatoflongu_sound longofsingle_sound longuofsingle_sound singleoflong_sound singleoflongu_sound + intofsingle_ne_sound intuofsingle_ne_sound + longoffloat_ne_sound longuoffloat_ne_sound + intofsingle_ne_total_sound intuofsingle_ne_total_sound + longoffloat_ne_total_sound longuoffloat_ne_total_sound longofwords_sound loword_sound hiword_sound intoffloat_total_sound intuoffloat_total_sound diff --git a/kvx/Asmblockdeps.v b/kvx/Asmblockdeps.v index a9786e0a..ac0b359e 100644 --- a/kvx/Asmblockdeps.v +++ b/kvx/Asmblockdeps.v @@ -1537,6 +1537,10 @@ Definition string_of_name_rr (n: arith_name_rr): pstring := | Pfixedudrzz => "Pfixedudrzz" | Pfixeddrzz_i32 => "Pfixeddrzz_i32" | Pfixedudrzz_i32 => "Pfixedudrzz_i32" + | Pfixedwrne => "Pfixedwrne" + | Pfixeduwrne => "Pfixeduwrne" + | Pfixeddrne => "Pfixeddrne" + | Pfixedudrne => "Pfixedudrne" end. Definition string_of_name_ri32 (n: arith_name_ri32): pstring := diff --git a/kvx/NeedOp.v b/kvx/NeedOp.v index 4578b4e8..455af70e 100644 --- a/kvx/NeedOp.v +++ b/kvx/NeedOp.v @@ -131,6 +131,8 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval := | Olongoffloat | Olonguoffloat | Ofloatoflong | Ofloatoflongu => op1 (default nv) | Ointofsingle | Ointuofsingle | Osingleofint | Osingleofintu => op1 (default nv) | Olongofsingle | Olonguofsingle | Osingleoflong | Osingleoflongu => op1 (default nv) + | Ointofsingle_ne | Ointuofsingle_ne => op1 (default nv) + | Olongoffloat_ne | Olonguoffloat_ne => op1 (default nv) | Ocmp c => needs_of_condition c | Oextfz _ _ | Oextfs _ _ | Oextfzl _ _ | Oextfsl _ _ => op1 (default nv) | Oinsf _ _ | Oinsfl _ _ => op2 (default nv) diff --git a/kvx/Op.v b/kvx/Op.v index 4458adb3..8549fc38 100644 --- a/kvx/Op.v +++ b/kvx/Op.v @@ -209,6 +209,11 @@ Inductive operation : Type := | Olonguofsingle (**r [rd = unsigned_long_of_float32(r1)] *) | Osingleoflong (**r [rd = float32_of_signed_long(r1)] *) | Osingleoflongu (**r [rd = float32_of_unsigned_int(r1)] *) + | Ointofsingle_ne (**r [rd = signed_int_of_float64(r1)] *) + | Ointuofsingle_ne (**r [rd = unsigned_int_of_float64(r1)] *) + | Olongoffloat_ne (**r [rd = signed_long_of_float64(r1)] *) + | Olonguoffloat_ne (**r [rd = unsigned_long_of_float64(r1)] *) + (*c Boolean tests: *) | Ocmp (cond: condition) (**r [rd = 1] if condition holds, [rd = 0] otherwise. *) | Oextfz (stop : Z) (start : Z) @@ -466,6 +471,11 @@ Definition eval_operation | Olonguofsingle, v1::nil => Some (Val.maketotal (Val.longuofsingle v1)) | Osingleoflong, v1::nil => Some (Val.maketotal (Val.singleoflong v1)) | Osingleoflongu, v1::nil => Some (Val.maketotal (Val.singleoflongu v1)) + | Ointofsingle_ne, v1::nil => Some (Val.maketotal (Val.intofsingle_ne v1)) + | Ointuofsingle_ne, v1::nil => Some (Val.maketotal (Val.intuofsingle_ne v1)) + | Olongoffloat_ne, v1::nil => Some (Val.maketotal (Val.longoffloat_ne v1)) + | Olonguoffloat_ne, v1::nil => Some (Val.maketotal (Val.longuoffloat_ne v1)) + | Ocmp c, _ => Some (Val.of_optbool (eval_condition c vl m)) | (Oextfz stop start), v0::nil => Some (extfz stop start v0) | (Oextfs stop start), v0::nil => Some (extfs stop start v0) @@ -683,6 +693,12 @@ Definition type_of_operation (op: operation) : list typ * typ := | Olonguofsingle => (Tsingle :: nil, Tlong) | Osingleoflong => (Tlong :: nil, Tsingle) | Osingleoflongu => (Tlong :: nil, Tsingle) + + | Ointofsingle_ne => (Tsingle :: nil, Tint) + | Ointuofsingle_ne => (Tsingle :: nil, Tint) + | Olongoffloat_ne => (Tfloat :: nil, Tlong) + | Olonguoffloat_ne => (Tfloat :: nil, Tlong) + | Ocmp c => (type_of_condition c, Tint) | Oextfz _ _ | Oextfs _ _ => (Tint :: nil, Tint) | Oextfzl _ _ | Oextfsl _ _ => (Tlong :: nil, Tlong) @@ -980,6 +996,12 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type). (* singleoflong, singleoflongu *) - destruct v0; cbn... - destruct v0; cbn... + (* intofsingle_ne, intuofsingle_ne *) + - destruct v0; cbn... destruct (Float32.to_int_ne f); cbn; trivial. + - destruct v0; cbn... destruct (Float32.to_intu_ne f); cbn; trivial. + (* longoffloat_ne, longuoffloat_ne *) + - destruct v0; cbn... destruct (Float.to_long_ne f); cbn; trivial. + - destruct v0; cbn... destruct (Float.to_longu_ne f); cbn; trivial. (* cmp *) - destruct (eval_condition cond vl m)... destruct b... (* extfz *) @@ -1669,6 +1691,14 @@ Proof. (* singleoflong, singleoflongu *) - inv H4; cbn; auto. - inv H4; cbn; auto. + + (* intofsingle_ne, intuofsingle_ne *) + - inv H4; cbn; auto. destruct (Float32.to_int_ne f0); cbn; auto. + - inv H4; cbn; auto. destruct (Float32.to_intu_ne f0); cbn; auto. + (* longoffloat_ne, longuoffloat_ne *) + - inv H4; cbn; auto. destruct (Float.to_long_ne f0); cbn; auto. + - inv H4; cbn; auto. destruct (Float.to_longu_ne f0); cbn; auto. + (* cmp *) - subst v1. destruct (eval_condition cond vl1 m1) eqn:?. exploit eval_condition_inj; eauto. intros EQ; rewrite EQ. diff --git a/kvx/PostpassSchedulingOracle.ml b/kvx/PostpassSchedulingOracle.ml index 3f4520a6..f53da14b 100644 --- a/kvx/PostpassSchedulingOracle.ml +++ b/kvx/PostpassSchedulingOracle.ml @@ -47,7 +47,8 @@ type real_instruction = | Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw | Fmind | Fminw | Fmaxd | Fmaxw | Finvw | Ffmaw | Ffmad | Ffmsw | Ffmsd - | Fnarrowdw | Fwidenlwd | Floatwz | Floatuwz | Floatdz | Floatudz | Fixedwz | Fixeduwz | Fixeddz | Fixedudz + | Fnarrowdw | Fwidenlwd | Floatwz | Floatuwz | Floatdz | Floatudz + | Fixedw | Fixeduw | Fixedd | Fixedud | Fcompw | Fcompd type ab_inst_rec = { @@ -86,12 +87,12 @@ let arith_rr_real = function | Pfloatuwrnsz -> Floatuwz | Pfloatudrnsz -> Floatudz | Pfloatdrnsz -> Floatdz - | Pfixedwrzz -> Fixedwz - | Pfixeduwrzz -> Fixeduwz - | Pfixeddrzz -> Fixeddz - | Pfixedudrzz -> Fixedudz - | Pfixeddrzz_i32 -> Fixeddz - | Pfixedudrzz_i32 -> Fixedudz + | Pfixedwrzz | Pfixedwrne -> Fixedw + | Pfixeduwrzz | Pfixeduwrne -> Fixeduw + | Pfixeddrzz | Pfixeddrne -> Fixedd + | Pfixedudrzz| Pfixedudrne -> Fixedud + | Pfixeddrzz_i32 -> Fixedd + | Pfixedudrzz_i32 -> Fixedud let arith_rrr_real = function | Pcompw it -> Compw @@ -643,7 +644,7 @@ let rec_to_usage r = (* TODO: check *) | Srsw | Srsd | Rorw -> (match encoding with None | Some U6 -> alu_lite | _ -> raise InvalidEncoding) | Extfz | Extfs | Insf -> (match encoding with None -> alu_lite | _ -> raise InvalidEncoding) - | Fixeduwz | Fixedwz | Floatwz | Floatuwz | Fixeddz | Fixedudz | Floatdz | Floatudz -> mau + | Fixeduw | Fixedw | Floatwz | Floatuwz | Fixedd | Fixedud | Floatdz | Floatudz -> mau | Lbs | Lbz | Lhs | Lhz | Lws | Ld | Lq | Lo -> (match encoding with None | Some U6 | Some S10 -> lsu_auxw | Some U27L5 | Some U27L10 -> lsu_auxw_x @@ -681,7 +682,7 @@ let real_inst_to_latency = function | Extfs | Extfz | Insf | Fcompw | Fcompd | Cmoved | Addxw | Addxd | Fmind | Fmaxd | Fminw | Fmaxw -> 1 - | Floatwz | Floatuwz | Fixeduwz | Fixedwz | Floatdz | Floatudz | Fixeddz | Fixedudz -> 4 + | Floatwz | Floatuwz | Fixeduw | Fixedw | Floatdz | Floatudz | Fixedd | Fixedud -> 4 | Mulw | Muld | Maddw | Maddd | Msbfw | Msbfd -> 2 (* FIXME - WORST CASE. If it's S10 then it's only 1 *) | Lbs | Lbz | Lhs | Lhz | Lws | Ld | Lq | Lo -> 3 | Sb | Sh | Sw | Sd | Sq | So -> 1 (* See kvx-Optimization.pdf page 19 *) diff --git a/kvx/ValueAOp.v b/kvx/ValueAOp.v index 87554258..4909a25b 100644 --- a/kvx/ValueAOp.v +++ b/kvx/ValueAOp.v @@ -298,6 +298,10 @@ Definition eval_static_operation (op: operation) (vl: list aval): aval := | Olonguofsingle, v1::nil => longuofsingle_total v1 | Osingleoflong, v1::nil => singleoflong v1 | Osingleoflongu, v1::nil => singleoflongu v1 + | Ointofsingle_ne, v1::nil => intofsingle_ne_total v1 + | Ointuofsingle_ne, v1::nil => intuofsingle_ne_total v1 + | Olongoffloat_ne, v1::nil => longoffloat_ne_total v1 + | Olonguoffloat_ne, v1::nil => longuoffloat_ne_total v1 | Ocmp c, _ => of_optbool (eval_static_condition c vl) | (Oextfz stop start), v0::nil => eval_static_extfz stop start v0 | (Oextfs stop start), v0::nil => eval_static_extfs stop start v0 -- cgit From 728c4dfe489b78b9816d2efb039618168eada263 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 12 Dec 2021 22:20:45 +0100 Subject: connect Asm to the rest for nearest rounding --- kvx/Asmblockgen.v | 12 ++++++++++++ kvx/Asmvliw.v | 4 ++-- 2 files changed, 14 insertions(+), 2 deletions(-) diff --git a/kvx/Asmblockgen.v b/kvx/Asmblockgen.v index ab827b1c..a8f1a045 100644 --- a/kvx/Asmblockgen.v +++ b/kvx/Asmblockgen.v @@ -799,6 +799,12 @@ Definition transl_op | Ointuofsingle, a1 :: nil => do rd <- ireg_of res; do rs <- freg_of a1; OK (Pfixeduwrzz rd rs ::i k) + | Ointofsingle_ne, a1 :: nil => + do rd <- ireg_of res; do rs <- freg_of a1; + OK (Pfixedwrne rd rs ::i k) + | Ointuofsingle_ne, a1 :: nil => + do rd <- ireg_of res; do rs <- freg_of a1; + OK (Pfixeduwrne rd rs ::i k) | Olongoffloat, a1 :: nil => do rd <- ireg_of res; do rs <- freg_of a1; OK (Pfixeddrzz rd rs ::i k) @@ -811,6 +817,12 @@ Definition transl_op | Olonguoffloat, a1 :: nil => do rd <- ireg_of res; do rs <- freg_of a1; OK (Pfixedudrzz rd rs ::i k) + | Olongoffloat_ne, a1 :: nil => + do rd <- ireg_of res; do rs <- freg_of a1; + OK (Pfixeddrne rd rs ::i k) + | Olonguoffloat_ne, a1 :: nil => + do rd <- ireg_of res; do rs <- freg_of a1; + OK (Pfixedudrne rd rs ::i k) | Ofloatofsingle, a1 :: nil => do rd <- freg_of res; do rs <- freg_of a1; diff --git a/kvx/Asmvliw.v b/kvx/Asmvliw.v index b3c0c8fa..6d60445a 100644 --- a/kvx/Asmvliw.v +++ b/kvx/Asmvliw.v @@ -960,8 +960,8 @@ Definition arith_eval_rr n v := | Pfixedudrzz_i32 => Val.maketotal (Val.intuoffloat v) | Pfixedudrne => Val.maketotal (Val.longuoffloat_ne v) | Pfixeddrne => Val.maketotal (Val.longoffloat_ne v) - | Pfixeduwrne => Val.maketotal (Val.intuoffloat_ne v) - | Pfixedwrne => Val.maketotal (Val.intoffloat_ne v) + | Pfixeduwrne => Val.maketotal (Val.intuofsingle_ne v) + | Pfixedwrne => Val.maketotal (Val.intofsingle_ne v) end. Definition arith_eval_ri32 n i := -- cgit From 761a1e046b4fb276ee0b63785ed20a35c26641c9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 12 Dec 2021 22:57:31 +0100 Subject: fix assembly syntax --- kvx/Builtins1.v | 10 +++++++++- kvx/CBuiltins.ml | 4 ++++ kvx/SelectOp.vp | 2 ++ kvx/SelectOpproof.v | 8 ++++++++ kvx/TargetPrinter.ml | 8 ++++---- 5 files changed, 27 insertions(+), 5 deletions(-) diff --git a/kvx/Builtins1.v b/kvx/Builtins1.v index 441345bf..de0d9885 100644 --- a/kvx/Builtins1.v +++ b/kvx/Builtins1.v @@ -25,7 +25,9 @@ Inductive platform_builtin : Type := | BI_fminf | BI_fmaxf | BI_fma -| BI_fmaf. +| BI_fmaf +| BI_lround_ne +| BI_luround_ne. Local Open Scope string_scope. @@ -36,6 +38,8 @@ Definition platform_builtin_table : list (string * platform_builtin) := :: ("__builtin_fmaxf", BI_fmaxf) :: ("__builtin_fma", BI_fma) :: ("__builtin_fmaf", BI_fmaf) + :: ("__builtin_lround_ne", BI_lround_ne) + :: ("__builtin_luround_ne", BI_luround_ne) :: nil. Definition platform_builtin_sig (b: platform_builtin) : signature := @@ -48,6 +52,8 @@ Definition platform_builtin_sig (b: platform_builtin) : signature := mksignature (Tfloat :: Tfloat :: Tfloat :: nil) Tfloat cc_default | BI_fmaf => mksignature (Tsingle :: Tsingle :: Tsingle :: nil) Tsingle cc_default + | BI_lround_ne | BI_luround_ne => + mksignature (Tfloat :: nil) Tlong cc_default end. Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (platform_builtin_sig b)) := @@ -58,4 +64,6 @@ Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (pl | BI_fmaxf => mkbuiltin_n2t Tsingle Tsingle Tsingle ExtFloat32.max | BI_fma => mkbuiltin_n3t Tfloat Tfloat Tfloat Tfloat Float.fma | BI_fmaf => mkbuiltin_n3t Tsingle Tsingle Tsingle Tsingle Float32.fma + | BI_lround_ne => mkbuiltin_n1p Tfloat Tlong Float.to_long_ne + | BI_luround_ne => mkbuiltin_n1p Tfloat Tlong Float.to_longu_ne end. diff --git a/kvx/CBuiltins.ml b/kvx/CBuiltins.ml index 7398e0f4..f2b7b09e 100644 --- a/kvx/CBuiltins.ml +++ b/kvx/CBuiltins.ml @@ -133,6 +133,10 @@ let builtins = { "__builtin_fmaf", (TFloat(FFloat, []), [TFloat(FFloat, []); TFloat(FFloat, []); TFloat(FFloat, [])], false); + "__builtin_lround_ne", + (TInt(ILong, []), [TFloat(FDouble, [])], false); + "__builtin_luround_ne", + (TInt(IULong, []), [TFloat(FDouble, [])], false); ] } diff --git a/kvx/SelectOp.vp b/kvx/SelectOp.vp index 4e1087f9..f529907d 100644 --- a/kvx/SelectOp.vp +++ b/kvx/SelectOp.vp @@ -750,6 +750,8 @@ Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr | BI_fmaxf => Some (Eop Omaxfs args) | BI_fma => gen_fma args | BI_fmaf => gen_fmaf args + | BI_lround_ne => Some (Eop Olongoffloat_ne args) + | BI_luround_ne => Some (Eop Olonguoffloat_ne args) end. End SELECT. diff --git a/kvx/SelectOpproof.v b/kvx/SelectOpproof.v index 0ede1e2d..a374ec54 100644 --- a/kvx/SelectOpproof.v +++ b/kvx/SelectOpproof.v @@ -1896,6 +1896,14 @@ Proof. repeat (try econstructor; try eassumption)). - apply eval_fma; assumption. - apply eval_fmaf; assumption. + - cbn in *; + destruct vl; trivial. destruct vl; trivial. + destruct v0; try discriminate; + cbn; rewrite H0; reflexivity. + - cbn in *; + destruct vl; trivial. destruct vl; trivial. + destruct v0; try discriminate; + cbn; rewrite H0; reflexivity. Qed. End CMCONSTR. diff --git a/kvx/TargetPrinter.ml b/kvx/TargetPrinter.ml index 40e2be55..01733858 100644 --- a/kvx/TargetPrinter.ml +++ b/kvx/TargetPrinter.ml @@ -587,13 +587,13 @@ module Target (*: TARGET*) = | Pfixedudrzz(rd, rs) | Pfixedudrzz_i32(rd, rs) -> fprintf oc " fixedud.rz %a = %a, 0\n" ireg rd ireg rs | Pfixedudrne(rd, rs) -> - fprintf oc " fixedud.ne %a = %a, 0\n" ireg rd ireg rs + fprintf oc " fixedud.rn %a = %a, 0\n" ireg rd ireg rs | Pfixeddrne(rd, rs) -> - fprintf oc " fixedd.ne %a = %a, 0\n" ireg rd ireg rs + fprintf oc " fixedd.rn %a = %a, 0\n" ireg rd ireg rs | Pfixeduwrne(rd, rs) -> - fprintf oc " fixeduw.ne %a = %a, 0\n" ireg rd ireg rs + fprintf oc " fixeduw.rn %a = %a, 0\n" ireg rd ireg rs | Pfixedwrne(rd, rs) -> - fprintf oc " fixedw.ne %a = %a, 0\n" ireg rd ireg rs + fprintf oc " fixedw.rn %a = %a, 0\n" ireg rd ireg rs (* Arith RI32 instructions *) | Pmake (rd, imm) -> -- cgit From d85683086bed8b76580616306166965079a02fb3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 13 Dec 2021 18:43:41 +0100 Subject: div_approx_reals_correct --- kvx/ExtFloats.v | 56 ++++++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 54 insertions(+), 2 deletions(-) diff --git a/kvx/ExtFloats.v b/kvx/ExtFloats.v index b08503a5..f6dae39e 100644 --- a/kvx/ExtFloats.v +++ b/kvx/ExtFloats.v @@ -13,7 +13,9 @@ (* *) (* *************************************************************) -Require Import Floats Integers ZArith. +From Flocq Require Import Core Digits Operations Round Bracket Sterbenz + Binary Round_odd. +Require Import Floats Integers ZArith IEEE754_extra Zdiv Psatz. Module ExtFloat. (** TODO check with the actual KVX; @@ -50,5 +52,55 @@ Definition max (x : float32) (y : float32) : float32 := Definition one := Float32.of_int (Int.repr (1%Z)). Definition inv (x : float32) : float32 := Float32.div one x. - End ExtFloat32. + + +Definition div_approx_reals (a b : Z) (x : R) := + let q:=ZnearestE x in + let r:=a-q*b in + if r Zfloor x = (y-1)%Z \/ Zfloor x = y. +Proof. + intros x y BALL. + apply Rabs_lt_inv in BALL. + case (Rcompare_spec x (IZR y)); intro CMP. + - left. apply Zfloor_imp. + replace (y-1+1) with y by ring. + rewrite minus_IZR. lra. + - subst. + rewrite Zfloor_IZR. right. reflexivity. + - right. apply Zfloor_imp. + rewrite plus_IZR. lra. +Qed. + +Theorem div_approx_reals_correct: + forall a b : Z, forall x : R, + b > 0 -> + (Rabs (x - IZR a/ IZR b) < 1/2)%R -> + div_approx_reals a b x = (a/b)%Z. +Proof. + intros a b x bPOS GAP. + assert (0 < IZR b)%R by (apply IZR_lt ; lia). + unfold div_approx_reals. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) x) as NEAR. + assert (Rabs (IZR (ZnearestE x) - IZR a/ IZR b) < 1)%R as BALL. + { pose proof (Rabs_triang (IZR (ZnearestE x) - x) + (x - IZR a/ IZR b)) as TRI. + ring_simplify (IZR (ZnearestE x) - x + (x - IZR a / IZR b))%R in TRI. + lra. + } + clear GAP NEAR. + rewrite Rabs_minus_sym in BALL. + pose proof (floor_ball1 _ _ BALL) as FLOOR. + clear BALL. + rewrite Zfloor_div in FLOOR by lia. + pose proof (Z_div_mod_eq_full a b) as DIV_MOD. + assert (0 < b) as bPOS' by lia. + pose proof (Z.mod_pos_bound a b bPOS') as MOD_BOUNDS. + case Z.ltb_spec; intro; destruct FLOOR; lia. +Qed. -- cgit From a90018e65444bcd7a4c85ddd815d315cd852e120 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 13 Dec 2021 18:56:49 +0100 Subject: div_approx_reals_correct simplified proof --- kvx/ExtFloats.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/ExtFloats.v b/kvx/ExtFloats.v index f6dae39e..ea663735 100644 --- a/kvx/ExtFloats.v +++ b/kvx/ExtFloats.v @@ -70,7 +70,7 @@ Proof. apply Rabs_lt_inv in BALL. case (Rcompare_spec x (IZR y)); intro CMP. - left. apply Zfloor_imp. - replace (y-1+1) with y by ring. + ring_simplify (y-1+1). rewrite minus_IZR. lra. - subst. rewrite Zfloor_IZR. right. reflexivity. -- cgit From 9007714f50a8ba49e2e6188cddada22a9fceed11 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 13 Dec 2021 20:40:25 +0100 Subject: begin div algo --- kvx/ExtFloats.v | 55 +++--------------------------------------------- kvx/ExtValues.v | 65 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 68 insertions(+), 52 deletions(-) diff --git a/kvx/ExtFloats.v b/kvx/ExtFloats.v index ea663735..332d3e3d 100644 --- a/kvx/ExtFloats.v +++ b/kvx/ExtFloats.v @@ -15,7 +15,7 @@ From Flocq Require Import Core Digits Operations Round Bracket Sterbenz Binary Round_odd. -Require Import Floats Integers ZArith IEEE754_extra Zdiv Psatz. +Require Import Floats Integers ZArith. Module ExtFloat. (** TODO check with the actual KVX; @@ -32,6 +32,8 @@ Definition max (x : float) (y : float) : float := | Some Eq | Some Gt => x | Some Lt | None => y end. + +Definition one := Float.of_int (Int.repr (1%Z)). End ExtFloat. Module ExtFloat32. @@ -53,54 +55,3 @@ Definition one := Float32.of_int (Int.repr (1%Z)). Definition inv (x : float32) : float32 := Float32.div one x. End ExtFloat32. - - -Definition div_approx_reals (a b : Z) (x : R) := - let q:=ZnearestE x in - let r:=a-q*b in - if r Zfloor x = (y-1)%Z \/ Zfloor x = y. -Proof. - intros x y BALL. - apply Rabs_lt_inv in BALL. - case (Rcompare_spec x (IZR y)); intro CMP. - - left. apply Zfloor_imp. - ring_simplify (y-1+1). - rewrite minus_IZR. lra. - - subst. - rewrite Zfloor_IZR. right. reflexivity. - - right. apply Zfloor_imp. - rewrite plus_IZR. lra. -Qed. - -Theorem div_approx_reals_correct: - forall a b : Z, forall x : R, - b > 0 -> - (Rabs (x - IZR a/ IZR b) < 1/2)%R -> - div_approx_reals a b x = (a/b)%Z. -Proof. - intros a b x bPOS GAP. - assert (0 < IZR b)%R by (apply IZR_lt ; lia). - unfold div_approx_reals. - pose proof (Znearest_imp2 (fun x => negb (Z.even x)) x) as NEAR. - assert (Rabs (IZR (ZnearestE x) - IZR a/ IZR b) < 1)%R as BALL. - { pose proof (Rabs_triang (IZR (ZnearestE x) - x) - (x - IZR a/ IZR b)) as TRI. - ring_simplify (IZR (ZnearestE x) - x + (x - IZR a / IZR b))%R in TRI. - lra. - } - clear GAP NEAR. - rewrite Rabs_minus_sym in BALL. - pose proof (floor_ball1 _ _ BALL) as FLOOR. - clear BALL. - rewrite Zfloor_div in FLOOR by lia. - pose proof (Z_div_mod_eq_full a b) as DIV_MOD. - assert (0 < b) as bPOS' by lia. - pose proof (Z.mod_pos_bound a b bPOS') as MOD_BOUNDS. - case Z.ltb_spec; intro; destruct FLOOR; lia. -Qed. diff --git a/kvx/ExtValues.v b/kvx/ExtValues.v index b4e14898..c478f70b 100644 --- a/kvx/ExtValues.v +++ b/kvx/ExtValues.v @@ -754,3 +754,68 @@ Definition fmaddfs := triple_op_single (fun f1 f2 f3 => Float32.fma f2 f3 f1). Definition fmsubf := triple_op_float (fun f1 f2 f3 => Float.fma (Float.neg f2) f3 f1). Definition fmsubfs := triple_op_single (fun f1 f2 f3 => Float32.fma (Float32.neg f2) f3 f1). + +From Flocq Require Import Core Digits Operations Round Bracket Sterbenz + Binary Round_odd. +Require Import IEEE754_extra Zdiv Psatz Floats ExtFloats. + +Definition div_approx_reals (a b : Z) (x : R) := + let q:=ZnearestE x in + let r:=a-q*b in + if r Zfloor x = (y-1)%Z \/ Zfloor x = y. +Proof. + intros x y BALL. + apply Rabs_lt_inv in BALL. + case (Rcompare_spec x (IZR y)); intro CMP. + - left. apply Zfloor_imp. + ring_simplify (y-1+1). + rewrite minus_IZR. lra. + - subst. + rewrite Zfloor_IZR. right. reflexivity. + - right. apply Zfloor_imp. + rewrite plus_IZR. lra. +Qed. + +Theorem div_approx_reals_correct: + forall a b : Z, forall x : R, + b > 0 -> + (Rabs (x - IZR a/ IZR b) < 1/2)%R -> + div_approx_reals a b x = (a/b)%Z. +Proof. + intros a b x bPOS GAP. + assert (0 < IZR b)%R by (apply IZR_lt ; lia). + unfold div_approx_reals. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) x) as NEAR. + assert (Rabs (IZR (ZnearestE x) - IZR a/ IZR b) < 1)%R as BALL. + { pose proof (Rabs_triang (IZR (ZnearestE x) - x) + (x - IZR a/ IZR b)) as TRI. + ring_simplify (IZR (ZnearestE x) - x + (x - IZR a / IZR b))%R in TRI. + lra. + } + clear GAP NEAR. + rewrite Rabs_minus_sym in BALL. + pose proof (floor_ball1 _ _ BALL) as FLOOR. + clear BALL. + rewrite Zfloor_div in FLOOR by lia. + pose proof (Z_div_mod_eq_full a b) as DIV_MOD. + assert (0 < b) as bPOS' by lia. + pose proof (Z.mod_pos_bound a b bPOS') as MOD_BOUNDS. + case Z.ltb_spec; intro; destruct FLOOR; lia. +Qed. + +Definition my_div (a b : val) := + let b_d := Val.maketotal (Val.floatofintu b) in + let invb_d := Val.floatofsingle (invfs (Val.maketotal (Val.singleofintu b))) in + let alpha := fmsubf (Vfloat ExtFloat.one) invb_d b_d in + let x := fmaddf invb_d alpha invb_d in + Val.mulf (Val.maketotal (Val.floatofintu a)) x. + +(* +Compute (my_div (Vint (Int.repr 3)) (Vint (Int.repr 5))). +*) -- cgit From 7f6cf940516fbdc769b12b61ede81ef710a21537 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 14 Dec 2021 14:26:14 +0100 Subject: begin work on fp division --- kvx/FPDivision.v | 43 +++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) create mode 100644 kvx/FPDivision.v diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v new file mode 100644 index 00000000..7fae685f --- /dev/null +++ b/kvx/FPDivision.v @@ -0,0 +1,43 @@ +Require Archi. +Require Import Coqlib. +Require Import Compopts. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Op. +Require Import CminorSel. +Require Import OpHelpers. +Require Import ExtValues ExtFloats. +Require Import DecBoolOps. +Require Import Chunks. +Require Import Builtins. +Require Import Values Globalenvs. +Require Compopts. + +Local Open Scope cminorsel_scope. + +Local Open Scope string_scope. +Local Open Scope error_monad_scope. + +Definition approx_inv b := + let invb_s := Eop Oinvfs ((Eop Osingleofintu ((Eletvar 0%nat):::Enil)):::Enil) in + let invb_d := Eop Ofloatofsingle (invb_s ::: Enil) in + let b_d := Eop Ofloatoflongu ((Eop Ocast32unsigned ((Eletvar 1%nat):::Enil)):::Enil) in + let invb_d_var := Eletvar (0%nat) in + let one := Eop (Ofloatconst ExtFloat.one) Enil in + let alpha := Eop Ofmsubf (one ::: invb_d_var ::: b_d ::: Enil) in + let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in + Elet b (Elet invb_d x). + +Theorem approx_inv_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_b : expr) b + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)), + eval_expr ge sp cmenv memenv le (approx_inv expr_b) + (Vfloat ExtFloat.one). +Proof. + intros. unfold approx_inv. + repeat econstructor. + { eassumption. } + cbn. +Abort. -- cgit From 478cabe3acf7ac68f4ab14d7211836ac23ba31ec Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 14 Dec 2021 14:57:49 +0100 Subject: approx_inv --- kvx/FPDivision.v | 14 ++++++++++---- 1 file changed, 10 insertions(+), 4 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 7fae685f..4e7fce31 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -1,3 +1,5 @@ +From Flocq Require Import Core Digits Operations Round Bracket Sterbenz + Binary Round_odd. Require Archi. Require Import Coqlib. Require Import Compopts. @@ -33,11 +35,15 @@ Theorem approx_inv_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_b : expr) b (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)), - eval_expr ge sp cmenv memenv le (approx_inv expr_b) - (Vfloat ExtFloat.one). + exists f : float, + eval_expr ge sp cmenv memenv le (approx_inv expr_b) (Vfloat f) /\ + is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int.unsigned b))) <= 0.1)%R. Proof. intros. unfold approx_inv. repeat econstructor. { eassumption. } - cbn. -Abort. + { reflexivity. } + all: set (invb_d := (Float.of_single (ExtFloat32.inv (Float32.of_intu b)))). + all: set (b_d := (Float.of_longu (Int64.repr (Int.unsigned b)))). + all: cbn. +Admitted. -- cgit From 4f4608206a72815f4c09952a16786435b0206ca0 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 14 Dec 2021 19:15:00 +0100 Subject: more on FPDivision --- kvx/FPDivision.v | 136 +++++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 133 insertions(+), 3 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 4e7fce31..ebf9136d 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -1,5 +1,5 @@ From Flocq Require Import Core Digits Operations Round Bracket Sterbenz - Binary Round_odd. + Binary Round_odd Bits. Require Archi. Require Import Coqlib. Require Import Compopts. @@ -9,12 +9,14 @@ Require Import Floats. Require Import Op. Require Import CminorSel. Require Import OpHelpers. -Require Import ExtValues ExtFloats. +Require Import ExtFloats. Require Import DecBoolOps. Require Import Chunks. Require Import Builtins. Require Import Values Globalenvs. Require Compopts. +Require Import Psatz. +Require Import IEEE754_extra. Local Open Scope cminorsel_scope. @@ -31,13 +33,15 @@ Definition approx_inv b := let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in Elet b (Elet invb_d x). +Definition approx_inv_thresh := (0.000000000000001)%R. + Theorem approx_inv_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_b : expr) b (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)), exists f : float, eval_expr ge sp cmenv memenv le (approx_inv expr_b) (Vfloat f) /\ - is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int.unsigned b))) <= 0.1)%R. + is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int.unsigned b))) <= approx_inv_thresh)%R. Proof. intros. unfold approx_inv. repeat econstructor. @@ -47,3 +51,129 @@ Proof. all: set (b_d := (Float.of_longu (Int64.repr (Int.unsigned b)))). all: cbn. Admitted. + +Definition approx_div a b := + let a_var := Eletvar (1%nat) in + let b_var := Eletvar (0%nat) in + let a_d := Eop Ofloatoflongu ((Eop Ocast32unsigned (a_var ::: Enil)) ::: Enil) in + let qr := Eop Olonguoffloat_ne ((Eop Omulf (a_d:::(approx_inv b_var):::Enil)):::Enil) in + let qr_var := Eletvar 0%nat in + let rem := Eop Omsubl ((Eop Ocast32unsigned ((Eletvar (2%nat)):::Enil))::: + qr_var ::: + (Eop Ocast32unsigned ((Eletvar (1%nat)):::Enil)):::Enil) in + let qr_m1 := Eop (Oaddlimm (Int64.repr (-1)%Z)) (qr_var:::Enil) in + let cases := Eop (Osel (Ccompl0 Clt) Tlong) + (qr_m1 ::: qr_var ::: rem ::: Enil) in (* (Elet qr cases) *) + Eop Olowlong ((Elet a (Elet (lift b) (Elet qr cases))) ::: Enil). + + +Definition div_approx_reals (a b : Z) (x : R) := + let q:=ZnearestE x in + let r:=a-q*b in + if r Zfloor x = (y-1)%Z \/ Zfloor x = y. +Proof. + intros x y BALL. + apply Rabs_lt_inv in BALL. + case (Rcompare_spec x (IZR y)); intro CMP. + - left. apply Zfloor_imp. + ring_simplify (y-1+1). + rewrite minus_IZR. lra. + - subst. + rewrite Zfloor_IZR. right. reflexivity. + - right. apply Zfloor_imp. + rewrite plus_IZR. lra. +Qed. + +Theorem div_approx_reals_correct: + forall a b : Z, forall x : R, + b > 0 -> + (Rabs (x - IZR a/ IZR b) < 1/2)%R -> + div_approx_reals a b x = (a/b)%Z. +Proof. + intros a b x bPOS GAP. + assert (0 < IZR b)%R by (apply IZR_lt ; lia). + unfold div_approx_reals. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) x) as NEAR. + assert (Rabs (IZR (ZnearestE x) - IZR a/ IZR b) < 1)%R as BALL. + { pose proof (Rabs_triang (IZR (ZnearestE x) - x) + (x - IZR a/ IZR b)) as TRI. + ring_simplify (IZR (ZnearestE x) - x + (x - IZR a / IZR b))%R in TRI. + lra. + } + clear GAP NEAR. + rewrite Rabs_minus_sym in BALL. + pose proof (floor_ball1 _ _ BALL) as FLOOR. + clear BALL. + rewrite Zfloor_div in FLOOR by lia. + pose proof (Z_div_mod_eq_full a b) as DIV_MOD. + assert (0 < b) as bPOS' by lia. + pose proof (Z.mod_pos_bound a b bPOS') as MOD_BOUNDS. + case Z.ltb_spec; intro; destruct FLOOR; lia. +Qed. + +Opaque approx_inv. + +Theorem approx_div_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : (Int.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (approx_div expr_a expr_b) + (Vint (Int.divu a b)). +Proof. + intros. + assert (eval_expr ge sp cmenv memenv (Vint b :: Vint a :: le) + (Eletvar 0) (Vint b)) as EVAL_b'. + { constructor. reflexivity. } + destruct (approx_inv_correct ge sp cmenv memenv (Vint b :: Vint a :: le) (Eletvar 0) b EVAL_b') as (inv_b & inv_b_eval & inv_b_finite & inv_b_correct). + unfold approx_div. + repeat econstructor. + { exact EVAL_a. } + { apply eval_lift. exact EVAL_b. } + exact inv_b_eval. + cbn. f_equal. + rewrite <- Float.of_intu_of_longu. + unfold Float.to_longu_ne. + rewrite ZofB_ne_range_correct by lia. + set (prod := (Float.mul (Float.of_intu a) inv_b)). + set (a' := Int.unsigned a) in *. + set (b' := Int.unsigned b) in *. + set (prod' := (B2R 53 1024 prod)). + set (prod_r := ZnearestE prod') in *. + replace (_ && _ && _) with true by admit. + cbn. + f_equal. + unfold Int.divu. + rewrite <- div_approx_reals_correct with (x := B2R _ _ prod). + 2: lia. + { + unfold div_approx_reals. + fold a' b' prod_r. + unfold Int64.mul. + rewrite Int64.unsigned_repr by admit. + rewrite Int64.unsigned_repr by admit. + unfold Int64.sub. + rewrite Int64.unsigned_repr by admit. + rewrite Int64.unsigned_repr by admit. + case Z.ltb_spec; intro CMP. + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true by admit. + cbn. + f_equal. + admit. + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false by admit. + cbn. + f_equal. + admit. + } + fold prod'. + admit. +Admitted. + + -- cgit From fb87921df88bf99385514fabb21458482e058260 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 14 Dec 2021 20:25:58 +0100 Subject: stuff --- kvx/FPDivision.v | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index ebf9136d..fad99f76 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -50,6 +50,8 @@ Proof. all: set (invb_d := (Float.of_single (ExtFloat32.inv (Float32.of_intu b)))). all: set (b_d := (Float.of_longu (Int64.repr (Int.unsigned b)))). all: cbn. + - admit. + - admit. Admitted. Definition approx_div a b := @@ -155,7 +157,7 @@ Proof. 2: lia. { unfold div_approx_reals. - fold a' b' prod_r. + fold a' b' prod' prod_r. unfold Int64.mul. rewrite Int64.unsigned_repr by admit. rewrite Int64.unsigned_repr by admit. -- cgit From e58c656f8223909cfcf531bbd19e32bf50bd162b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 14 Dec 2021 20:35:56 +0100 Subject: rm Flocq from RECDIRS --- Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Makefile b/Makefile index c9bf9ed4..8b50375f 100644 --- a/Makefile +++ b/Makefile @@ -33,7 +33,7 @@ BACKENDLIB?=Asmgenproof0.v Asmgenproof1.v DIRS := lib lib/Impure common $(ARCHDIRS) scheduling backend cfrontend driver \ export cparser -RECDIRS:=lib common $(ARCHDIRS) scheduling backend cfrontend driver flocq exportclight \ +RECDIRS:=lib common $(ARCHDIRS) scheduling backend cfrontend driver exportclight \ cparser COQINCLUDES := $(foreach d, $(DIRS), -R $(d) compcert.$(subst /,.,$d)) -- cgit From 823d9eb6a8a4f5c69def05c96d15658e93332e50 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 14 Dec 2021 21:38:24 +0100 Subject: correct pour Flocq externe --- kvx/Archi.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/Archi.v b/kvx/Archi.v index 6d59a3d1..b1739421 100644 --- a/kvx/Archi.v +++ b/kvx/Archi.v @@ -16,7 +16,7 @@ (** Architecture-dependent parameters for MPPA KVX. Mostly copied from the Risc-V backend *) Require Import ZArith List. -Require Import Binary Bits. +From Flocq Require Import Binary Bits. Definition ptr64 := true. -- cgit From 3e13e7c319815493a6925a1e5d5bd18e10e99f95 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 14 Dec 2021 21:41:41 +0100 Subject: command line for Flocq (temporarily) --- flocq.sh | 2 ++ 1 file changed, 2 insertions(+) create mode 100644 flocq.sh diff --git a/flocq.sh b/flocq.sh new file mode 100644 index 00000000..64b7534e --- /dev/null +++ b/flocq.sh @@ -0,0 +1,2 @@ +COQINCLUDES="-R /home/monniaux/.opam/4.12.0+flambda/lib/coq/user-contrib/Flocq flocq" +export COQINCLUDES -- cgit From 3043b099fc4837c1cb3727e5b317148154dcc364 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 11:18:18 +0100 Subject: progress --- kvx/FPDivision.v | 78 ++++++++++++++++++++++++++++++++++++++++++++++---------- 1 file changed, 64 insertions(+), 14 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index fad99f76..00e807d4 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -18,6 +18,8 @@ Require Compopts. Require Import Psatz. Require Import IEEE754_extra. +From Gappa Require Import Gappa_tactic. + Local Open Scope cminorsel_scope. Local Open Scope string_scope. @@ -35,23 +37,70 @@ Definition approx_inv b := Definition approx_inv_thresh := (0.000000000000001)%R. +Arguments is_finite {prec emax}. +Arguments B2R {prec emax}. +Arguments BofZ (prec emax) {prec_gt_0_ Hmax}. + +(* +Lemma BofZ_exact_zero: + forall (prec emax : Z) (prec_gt_0_ : Prec_gt_0 prec) + (Hmax : (prec < emax)%Z), + B2R (BofZ prec emax 0%Z (Hmax := Hmax)) = 0%R /\ + is_finite (BofZ prec emax 0%Z (Hmax := Hmax)) = true /\ + Bsign prec emax (BofZ prec emax 0%Z (Hmax := Hmax)) = false. +Proof. + intros. + apply BofZ_exact. + pose proof (Z.pow_nonneg 2 prec). + lia. +Qed. + *) + Theorem approx_inv_correct : forall (ge : genv) (sp: val) cmenv memenv - (le : letenv) (expr_b : expr) b - (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)), + (le : letenv) (expr_b : expr) (b : int) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : ((Int.unsigned b) > 0)%Z), exists f : float, eval_expr ge sp cmenv memenv le (approx_inv expr_b) (Vfloat f) /\ - is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int.unsigned b))) <= approx_inv_thresh)%R. + is_finite f = true /\ (Rabs((B2R f) - (1 / IZR (Int.unsigned b))) <= approx_inv_thresh)%R. Proof. intros. unfold approx_inv. - repeat econstructor. - { eassumption. } - { reflexivity. } - all: set (invb_d := (Float.of_single (ExtFloat32.inv (Float32.of_intu b)))). - all: set (b_d := (Float.of_longu (Int64.repr (Int.unsigned b)))). - all: cbn. - - admit. - - admit. + econstructor. constructor. + { repeat econstructor. + { eassumption. } + { reflexivity. } } + set (invb_d := (Float.of_single (ExtFloat32.inv (Float32.of_intu b)))). + set (b_d := (Float.of_longu (Int64.repr (Int.unsigned b)))). + cbn. + pose proof (Int.unsigned_range b) as RANGE. + change Int.modulus with 4294967296%Z in RANGE. + set (b' := Int.unsigned b) in *. + assert (0 <= IZR b' <= 4294967295)%R as RANGE'. + { split. + { apply IZR_le. lia. } + apply IZR_le. lia. + } + pose proof (BofZ_correct 24 128 eq_refl eq_refl b') as C1. + rewrite Rlt_bool_true in C1; cycle 1. + { clear C1. cbn. + admit. + } + rewrite Zlt_bool_false in C1 by lia. + destruct C1 as (C1E & C1F & C1S). + Local Transparent Float32.of_intu Float32.of_int Float32.div. + unfold ExtFloat32.inv, ExtFloat32.one, Float32.of_intu, Float32.of_int, Float32.div in invb_d. + fold b' in invb_d. + change (Int.signed (Int.repr 1%Z)) with 1%Z in invb_d. + pose proof (Bdiv_correct 24 128 eq_refl eq_refl Float32.binop_nan mode_NE + (BofZ 24 128 1 (Hmax := eq_refl)) + (BofZ 24 128 b' (Hmax := eq_refl))) as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. cbn. + admit. + } + + Search B2R (BofZ _). Admitted. Definition approx_div a b := @@ -68,6 +117,7 @@ Definition approx_div a b := (qr_m1 ::: qr_var ::: rem ::: Enil) in (* (Elet qr cases) *) Eop Olowlong ((Elet a (Elet (lift b) (Elet qr cases))) ::: Enil). +Open Scope Z. Definition div_approx_reals (a b : Z) (x : R) := let q:=ZnearestE x in @@ -134,7 +184,7 @@ Proof. assert (eval_expr ge sp cmenv memenv (Vint b :: Vint a :: le) (Eletvar 0) (Vint b)) as EVAL_b'. { constructor. reflexivity. } - destruct (approx_inv_correct ge sp cmenv memenv (Vint b :: Vint a :: le) (Eletvar 0) b EVAL_b') as (inv_b & inv_b_eval & inv_b_finite & inv_b_correct). + destruct (approx_inv_correct ge sp cmenv memenv (Vint b :: Vint a :: le) (Eletvar 0) b EVAL_b' b_nz) as (inv_b & inv_b_eval & inv_b_finite & inv_b_correct). unfold approx_div. repeat econstructor. { exact EVAL_a. } @@ -147,13 +197,13 @@ Proof. set (prod := (Float.mul (Float.of_intu a) inv_b)). set (a' := Int.unsigned a) in *. set (b' := Int.unsigned b) in *. - set (prod' := (B2R 53 1024 prod)). + set (prod' := (B2R prod)). set (prod_r := ZnearestE prod') in *. replace (_ && _ && _) with true by admit. cbn. f_equal. unfold Int.divu. - rewrite <- div_approx_reals_correct with (x := B2R _ _ prod). + rewrite <- div_approx_reals_correct with (x := B2R prod). 2: lia. { unfold div_approx_reals. -- cgit From 80c707afa9eb1042b7fca7367011e37343b89042 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 11:54:18 +0100 Subject: progress --- kvx/FPDivision.v | 18 +++++++++++++----- 1 file changed, 13 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 00e807d4..fdcd455e 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -55,7 +55,7 @@ Proof. lia. Qed. *) - + Theorem approx_inv_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_b : expr) (b : int) @@ -96,11 +96,19 @@ Proof. (BofZ 24 128 1 (Hmax := eq_refl)) (BofZ 24 128 b' (Hmax := eq_refl))) as C2. rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. cbn. - admit. - } + { clear C2. admit. } + assert (B2R (BofZ 24 128 b' (Hmax:=eq_refl)) <> 0%R) as NONZ. + { admit. } + destruct (C2 NONZ) as (C2E & C2F & C2S). + clear C2 NONZ. + Local Transparent Float.of_single. + unfold Float.of_single in invb_d. + pose proof (Bconv_correct 24 128 53 1024 eq_refl eq_refl Float.of_single_nan mode_NE (Bdiv 24 128 eq_refl eq_refl Float32.binop_nan mode_NE + (BofZ 24 128 1) (BofZ 24 128 b')) C2F) as C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. admit. } + destruct C3 as (C3E & C3F & C3S). - Search B2R (BofZ _). Admitted. Definition approx_div a b := -- cgit From 829e2de6eefc971051dc9a3315347348c8a93e00 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 12:48:31 +0100 Subject: progress --- kvx/FPDivision.v | 45 ++++++++++++++++++++++++++++++--------------- 1 file changed, 30 insertions(+), 15 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index fdcd455e..c6613fac 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -37,10 +37,6 @@ Definition approx_inv b := Definition approx_inv_thresh := (0.000000000000001)%R. -Arguments is_finite {prec emax}. -Arguments B2R {prec emax}. -Arguments BofZ (prec emax) {prec_gt_0_ Hmax}. - (* Lemma BofZ_exact_zero: forall (prec emax : Z) (prec_gt_0_ : Prec_gt_0 prec) @@ -55,7 +51,7 @@ Proof. lia. Qed. *) - + Theorem approx_inv_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_b : expr) (b : int) @@ -63,7 +59,7 @@ Theorem approx_inv_correct : (b_nz : ((Int.unsigned b) > 0)%Z), exists f : float, eval_expr ge sp cmenv memenv le (approx_inv expr_b) (Vfloat f) /\ - is_finite f = true /\ (Rabs((B2R f) - (1 / IZR (Int.unsigned b))) <= approx_inv_thresh)%R. + is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int.unsigned b))) <= approx_inv_thresh)%R. Proof. intros. unfold approx_inv. econstructor. constructor. @@ -91,24 +87,43 @@ Proof. Local Transparent Float32.of_intu Float32.of_int Float32.div. unfold ExtFloat32.inv, ExtFloat32.one, Float32.of_intu, Float32.of_int, Float32.div in invb_d. fold b' in invb_d. + Check BofZ. change (Int.signed (Int.repr 1%Z)) with 1%Z in invb_d. pose proof (Bdiv_correct 24 128 eq_refl eq_refl Float32.binop_nan mode_NE - (BofZ 24 128 1 (Hmax := eq_refl)) - (BofZ 24 128 b' (Hmax := eq_refl))) as C2. + (BofZ 24 128 eq_refl eq_refl 1) + (BofZ 24 128 eq_refl eq_refl b')) as C2. rewrite Rlt_bool_true in C2; cycle 1. { clear C2. admit. } - assert (B2R (BofZ 24 128 b' (Hmax:=eq_refl)) <> 0%R) as NONZ. + assert (B2R 24 128 (BofZ 24 128 eq_refl eq_refl b') <> 0%R) as NONZ. { admit. } destruct (C2 NONZ) as (C2E & C2F & C2S). clear C2 NONZ. Local Transparent Float.of_single. unfold Float.of_single in invb_d. - pose proof (Bconv_correct 24 128 53 1024 eq_refl eq_refl Float.of_single_nan mode_NE (Bdiv 24 128 eq_refl eq_refl Float32.binop_nan mode_NE - (BofZ 24 128 1) (BofZ 24 128 b')) C2F) as C3. + pose proof (Bconv_correct 24 128 53 1024 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) Float.of_single_nan mode_NE + (Bdiv 24 128 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) Float32.binop_nan mode_NE + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) 1) + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) b'))) as C3. + fold invb_d in C3. rewrite Rlt_bool_true in C3; cycle 1. { clear C3. admit. } - destruct C3 as (C3E & C3F & C3S). - + change (is_finite 24 128 (BofZ 24 128 eq_refl eq_refl 1)) with true in C2F. + destruct (C3 C2F) as (C3E & C3F & C3S). + unfold Float.fma. + assert (is_finite _ _ (Float.neg invb_d) = true) as invb_d_F. + { Local Transparent Float.neg. + unfold Float.neg. + rewrite is_finite_Bopp. + assumption. + } + Set Printing Implicit. + pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE + (Float.neg invb_d) b_d ExtFloat.one) as C4. Admitted. Definition approx_div a b := @@ -205,13 +220,13 @@ Proof. set (prod := (Float.mul (Float.of_intu a) inv_b)). set (a' := Int.unsigned a) in *. set (b' := Int.unsigned b) in *. - set (prod' := (B2R prod)). + set (prod' := (B2R _ _ prod)). set (prod_r := ZnearestE prod') in *. replace (_ && _ && _) with true by admit. cbn. f_equal. unfold Int.divu. - rewrite <- div_approx_reals_correct with (x := B2R prod). + rewrite <- div_approx_reals_correct with (x := B2R _ _ prod). 2: lia. { unfold div_approx_reals. -- cgit From 8c60bfedbad045a1d60994d46362d3c61e8b20a8 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 14:26:58 +0100 Subject: progress --- kvx/FPDivision.v | 45 ++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 40 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index c6613fac..be0fedfd 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -67,16 +67,25 @@ Proof. { eassumption. } { reflexivity. } } set (invb_d := (Float.of_single (ExtFloat32.inv (Float32.of_intu b)))). - set (b_d := (Float.of_longu (Int64.repr (Int.unsigned b)))). - cbn. + set (b' := Int.unsigned b) in *. pose proof (Int.unsigned_range b) as RANGE. + fold b' in RANGE. change Int.modulus with 4294967296%Z in RANGE. - set (b' := Int.unsigned b) in *. + assert (0 <= b' <= Int64.max_unsigned)%Z as b'RANGE. + { change Int64.max_unsigned with 18446744073709551615%Z. + lia. } assert (0 <= IZR b' <= 4294967295)%R as RANGE'. { split. { apply IZR_le. lia. } apply IZR_le. lia. } + cbn. + + set (b_d := (Float.of_longu (Int64.repr b'))) in *. + Local Transparent Float.of_longu. + unfold Float.of_longu in b_d. + + pose proof (BofZ_correct 24 128 eq_refl eq_refl b') as C1. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. cbn. @@ -84,6 +93,7 @@ Proof. } rewrite Zlt_bool_false in C1 by lia. destruct C1 as (C1E & C1F & C1S). + Local Transparent Float32.of_intu Float32.of_int Float32.div. unfold ExtFloat32.inv, ExtFloat32.one, Float32.of_intu, Float32.of_int, Float32.div in invb_d. fold b' in invb_d. @@ -120,10 +130,35 @@ Proof. rewrite is_finite_Bopp. assumption. } - Set Printing Implicit. + + pose proof (BofZ_correct 53 1024 eq_refl eq_refl b') as C4. + rewrite Rlt_bool_true in C4; cycle 1. + { clear C4. + admit. + } + rewrite Zlt_bool_false in C4 by lia. + destruct C4 as (C4E & C4F & C4S). + + assert (is_finite 53 1024 b_d = true) as b_d_F. + { unfold b_d. + rewrite Int64.unsigned_repr by lia. + assumption. + } + + assert (is_finite 53 1024 ExtFloat.one = true) as one_F by reflexivity. + pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE - (Float.neg invb_d) b_d ExtFloat.one) as C4. + (Float.neg invb_d) b_d ExtFloat.one invb_d_F b_d_F one_F) as C5. + + rewrite Rlt_bool_true in C5; cycle 1. + { admit. } + destruct C5 as (C5E & C5F & C5S). + + pose proof (Bfma_correct 53 1024 eq_refl eq_refl Float.fma_nan mode_NE + (Bfma 53 1024 eq_refl eq_refl Float.fma_nan mode_NE + (Float.neg invb_d) b_d ExtFloat.one) invb_d invb_d C5F C3F C3F) as C6. + Admitted. Definition approx_div a b := -- cgit From 3c48c9f0c53ca01cdee579cf1f2ca11879f8874f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 14:39:51 +0100 Subject: progress --- kvx/FPDivision.v | 16 ++++++++++------ 1 file changed, 10 insertions(+), 6 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index be0fedfd..28c7fe72 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -35,7 +35,7 @@ Definition approx_inv b := let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in Elet b (Elet invb_d x). -Definition approx_inv_thresh := (0.000000000000001)%R. +Definition approx_inv_thresh := (0.1)%R. (* Lemma BofZ_exact_zero: @@ -121,7 +121,7 @@ Proof. fold invb_d in C3. rewrite Rlt_bool_true in C3; cycle 1. { clear C3. admit. } - change (is_finite 24 128 (BofZ 24 128 eq_refl eq_refl 1)) with true in C2F. + change (is_finite 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1)) with true in C2F. destruct (C3 C2F) as (C3E & C3F & C3S). unfold Float.fma. assert (is_finite _ _ (Float.neg invb_d) = true) as invb_d_F. @@ -131,7 +131,7 @@ Proof. assumption. } - pose proof (BofZ_correct 53 1024 eq_refl eq_refl b') as C4. + pose proof (BofZ_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C4. rewrite Rlt_bool_true in C4; cycle 1. { clear C4. admit. @@ -155,10 +155,14 @@ Proof. { admit. } destruct C5 as (C5E & C5F & C5S). - pose proof (Bfma_correct 53 1024 eq_refl eq_refl Float.fma_nan mode_NE - (Bfma 53 1024 eq_refl eq_refl Float.fma_nan mode_NE + pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE + (Bfma 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE (Float.neg invb_d) b_d ExtFloat.one) invb_d invb_d C5F C3F C3F) as C6. - + rewrite Rlt_bool_true in C6; cycle 1. + { admit. } + destruct C6 as (C6E & C6F & C6S). + split. + { exact C6F. } Admitted. Definition approx_div a b := -- cgit From 10ac1b914a79d6d2f50f5432e5a86fb42d03c1d9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 14:50:10 +0100 Subject: progress --- kvx/FPDivision.v | 21 +++++++++++++++------ 1 file changed, 15 insertions(+), 6 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 28c7fe72..2c0739e1 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -84,9 +84,12 @@ Proof. set (b_d := (Float.of_longu (Int64.repr b'))) in *. Local Transparent Float.of_longu. unfold Float.of_longu in b_d. + + assert(SILLY : (- 2 ^ 24 <= 1 <= 2 ^ 24)%Z) by lia. + destruct (BofZ_exact 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C0E & C0F & C0S). + clear SILLY. - - pose proof (BofZ_correct 24 128 eq_refl eq_refl b') as C1. + pose proof (BofZ_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C1. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. cbn. admit. @@ -99,12 +102,12 @@ Proof. fold b' in invb_d. Check BofZ. change (Int.signed (Int.repr 1%Z)) with 1%Z in invb_d. - pose proof (Bdiv_correct 24 128 eq_refl eq_refl Float32.binop_nan mode_NE - (BofZ 24 128 eq_refl eq_refl 1) - (BofZ 24 128 eq_refl eq_refl b')) as C2. + pose proof (Bdiv_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float32.binop_nan mode_NE + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1) + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b')) as C2. rewrite Rlt_bool_true in C2; cycle 1. { clear C2. admit. } - assert (B2R 24 128 (BofZ 24 128 eq_refl eq_refl b') <> 0%R) as NONZ. + assert (B2R 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') <> 0%R) as NONZ. { admit. } destruct (C2 NONZ) as (C2E & C2F & C2S). clear C2 NONZ. @@ -163,6 +166,12 @@ Proof. destruct C6 as (C6E & C6F & C6S). split. { exact C6F. } + rewrite C6E. + rewrite C5E. + rewrite C3E. + rewrite C2E. + rewrite C1E. + rewrite C0E. Admitted. Definition approx_div a b := -- cgit From ec82303f25209670189572f07865430255da0fd5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 14:59:03 +0100 Subject: progress --- kvx/FPDivision.v | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 2c0739e1..86c21fe0 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -88,6 +88,10 @@ Proof. assert(SILLY : (- 2 ^ 24 <= 1 <= 2 ^ 24)%Z) by lia. destruct (BofZ_exact 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C0E & C0F & C0S). clear SILLY. + + assert(SILLY : (- 2 ^ 53 <= 1 <= 2 ^ 53)%Z) by lia. + destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C9E & C9F & C9S). + clear SILLY. pose proof (BofZ_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C1. rewrite Rlt_bool_true in C1; cycle 1. @@ -172,6 +176,13 @@ Proof. rewrite C2E. rewrite C1E. rewrite C0E. + unfold Float.neg. + rewrite B2R_Bopp. + unfold ExtFloat.one. + Local Transparent Float.of_int. + unfold Float.of_int. + rewrite (Int.signed_repr 1) by (cbn ; lia). + rewrite C9E. Admitted. Definition approx_div a b := -- cgit From c254d0b3935fa956ea220731b6ee209ab4cc641a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 15:57:15 +0100 Subject: progress --- kvx/FPDivision.v | 35 ++++++++++++++++++++++++++++------- 1 file changed, 28 insertions(+), 7 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 86c21fe0..46598501 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -51,7 +51,7 @@ Proof. lia. Qed. *) - + Theorem approx_inv_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_b : expr) (b : int) @@ -99,7 +99,7 @@ Proof. admit. } rewrite Zlt_bool_false in C1 by lia. - destruct C1 as (C1E & C1F & C1S). + destruct C1 as (C1E & C1F & _). Local Transparent Float32.of_intu Float32.of_int Float32.div. unfold ExtFloat32.inv, ExtFloat32.one, Float32.of_intu, Float32.of_int, Float32.div in invb_d. @@ -113,7 +113,7 @@ Proof. { clear C2. admit. } assert (B2R 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') <> 0%R) as NONZ. { admit. } - destruct (C2 NONZ) as (C2E & C2F & C2S). + destruct (C2 NONZ) as (C2E & C2F & _). clear C2 NONZ. Local Transparent Float.of_single. unfold Float.of_single in invb_d. @@ -129,7 +129,7 @@ Proof. rewrite Rlt_bool_true in C3; cycle 1. { clear C3. admit. } change (is_finite 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1)) with true in C2F. - destruct (C3 C2F) as (C3E & C3F & C3S). + destruct (C3 C2F) as (C3E & C3F & _). unfold Float.fma. assert (is_finite _ _ (Float.neg invb_d) = true) as invb_d_F. { Local Transparent Float.neg. @@ -144,7 +144,7 @@ Proof. admit. } rewrite Zlt_bool_false in C4 by lia. - destruct C4 as (C4E & C4F & C4S). + destruct C4 as (C4E & C4F & _). assert (is_finite 53 1024 b_d = true) as b_d_F. { unfold b_d. @@ -160,14 +160,14 @@ Proof. rewrite Rlt_bool_true in C5; cycle 1. { admit. } - destruct C5 as (C5E & C5F & C5S). + destruct C5 as (C5E & C5F & _). pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE (Bfma 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE (Float.neg invb_d) b_d ExtFloat.one) invb_d invb_d C5F C3F C3F) as C6. rewrite Rlt_bool_true in C6; cycle 1. { admit. } - destruct C6 as (C6E & C6F & C6S). + destruct C6 as (C6E & C6F & _). split. { exact C6F. } rewrite C6E. @@ -183,6 +183,27 @@ Proof. unfold Float.of_int. rewrite (Int.signed_repr 1) by (cbn ; lia). rewrite C9E. + rewrite C3E. + rewrite C2E. + rewrite C0E. + rewrite C1E. + unfold b_d. + rewrite Int64.unsigned_repr by lia. + rewrite C4E. + cbn. + set (rnd_d := round radix2 (FLT_exp (-1074) 53) ZnearestE). + set (rnd_s := round radix2 (FLT_exp (-149) 24) ZnearestE). + set (bi := IZR b'). + replace (- rnd_d (rnd_s (1 / rnd_s bi)) * rnd_d bi + 1)%R + with ((- rnd_d (rnd_s (1 / rnd_s bi)) + 1 / rnd_d bi) / (1 /rnd_d bi))%R; + cycle 1. + { field. + cut (1 <= bi <= 4294967295)%R. + unfold rnd_d. + gappa. + split; apply IZR_le; lia. + } + Admitted. Definition approx_div a b := -- cgit From e71543cef7a06dbe7091fac15aba7ca4d7ab32fd Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Dec 2021 18:19:02 +0100 Subject: patine --- kvx/FPDivision.v | 15 +++------------ 1 file changed, 3 insertions(+), 12 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 46598501..89b54346 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -191,19 +191,10 @@ Proof. rewrite Int64.unsigned_repr by lia. rewrite C4E. cbn. - set (rnd_d := round radix2 (FLT_exp (-1074) 53) ZnearestE). - set (rnd_s := round radix2 (FLT_exp (-149) 24) ZnearestE). + set (rd := round radix2 (FLT_exp (-1074) 53) ZnearestE). + set (rs := round radix2 (FLT_exp (-149) 24) ZnearestE). set (bi := IZR b'). - replace (- rnd_d (rnd_s (1 / rnd_s bi)) * rnd_d bi + 1)%R - with ((- rnd_d (rnd_s (1 / rnd_s bi)) + 1 / rnd_d bi) / (1 /rnd_d bi))%R; - cycle 1. - { field. - cut (1 <= bi <= 4294967295)%R. - unfold rnd_d. - gappa. - split; apply IZR_le; lia. - } - + set (invb := rd (rs (1/ rs bi))%R). Admitted. Definition approx_div a b := -- cgit From cb47d75e20231f06b3e5928b003dc3b5e265b6f8 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 19:36:11 +0100 Subject: test gappa --- kvx/FPDivision.v | 32 +++++++++++++++++--------------- 1 file changed, 17 insertions(+), 15 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 89b54346..afa31a72 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -35,7 +35,7 @@ Definition approx_inv b := let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in Elet b (Elet invb_d x). -Definition approx_inv_thresh := (0.1)%R. +Definition approx_inv_thresh := (1/10)%R. (* Lemma BofZ_exact_zero: @@ -74,7 +74,7 @@ Proof. assert (0 <= b' <= Int64.max_unsigned)%Z as b'RANGE. { change Int64.max_unsigned with 18446744073709551615%Z. lia. } - assert (0 <= IZR b' <= 4294967295)%R as RANGE'. + assert (1 <= IZR b' <= 4294967295)%R as RANGE'. { split. { apply IZR_le. lia. } apply IZR_le. lia. @@ -86,11 +86,11 @@ Proof. unfold Float.of_longu in b_d. assert(SILLY : (- 2 ^ 24 <= 1 <= 2 ^ 24)%Z) by lia. - destruct (BofZ_exact 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C0E & C0F & C0S). + destruct (BofZ_exact 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C0E & C0F & _). clear SILLY. assert(SILLY : (- 2 ^ 53 <= 1 <= 2 ^ 53)%Z) by lia. - destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C9E & C9F & C9S). + destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C9E & C9F & _). clear SILLY. pose proof (BofZ_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C1. @@ -138,13 +138,9 @@ Proof. assumption. } - pose proof (BofZ_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C4. - rewrite Rlt_bool_true in C4; cycle 1. - { clear C4. - admit. - } - rewrite Zlt_bool_false in C4 by lia. - destruct C4 as (C4E & C4F & _). + assert(SILLY : (- 2 ^ 53 <= b' <= 2 ^ 53)%Z) by lia. + destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b' SILLY) as (C4E & C4F & _). + clear SILLY. assert (is_finite 53 1024 b_d = true) as b_d_F. { unfold b_d. @@ -191,10 +187,16 @@ Proof. rewrite Int64.unsigned_repr by lia. rewrite C4E. cbn. - set (rd := round radix2 (FLT_exp (-1074) 53) ZnearestE). - set (rs := round radix2 (FLT_exp (-149) 24) ZnearestE). - set (bi := IZR b'). - set (invb := rd (rs (1/ rs bi))%R). + set (rd := round radix2 (FLT_exp (-1074) 53) ZnearestE) in *. + set (rs := round radix2 (FLT_exp (-149) 24) ZnearestE) in *. + set (bi := IZR b') in *. + set (invb0 := rd (rs (1/ rs bi))%R) in *. + set (alpha := (- invb0 * bi + 1)%R) in *. + assert (-1/IZR (2^21)%Z <= ((1 / bi) - rd (rs(1/ rs bi)))/(1/bi) <= 1/IZR(2^21)%Z)%R. + { unfold rd, rs. + cbn. + gappa. + } Admitted. Definition approx_div a b := -- cgit From 7b0c81a907ea5b772561b910acdf1efd615451b6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 21:14:01 +0100 Subject: finish proof but some admitted stuff --- kvx/FPDivision.v | 31 +++++++++++++++++++++++++++---- 1 file changed, 27 insertions(+), 4 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index afa31a72..b71c8990 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -35,7 +35,7 @@ Definition approx_inv b := let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in Elet b (Elet invb_d x). -Definition approx_inv_thresh := (1/10)%R. +Definition approx_inv_thresh := (1/4294967296)%R. (* Lemma BofZ_exact_zero: @@ -192,11 +192,34 @@ Proof. set (bi := IZR b') in *. set (invb0 := rd (rs (1/ rs bi))%R) in *. set (alpha := (- invb0 * bi + 1)%R) in *. - assert (-1/IZR (2^21)%Z <= ((1 / bi) - rd (rs(1/ rs bi)))/(1/bi) <= 1/IZR(2^21)%Z)%R. - { unfold rd, rs. - cbn. + set (alpha' := ((1/bi - rd (rs (1/ rs bi)))/(1/bi))%R) in *. + assert (alpha = alpha')%R as expand_alpha. + { + unfold alpha, alpha', invb0. + field. + lra. + } + assert(-1/2097152 <= alpha' <= 1/2097152)%R as alpha_BOUND. + { unfold alpha', rd, rs. gappa. } + set (delta := (rd (rd alpha * invb0 + invb0) - (alpha * invb0 + invb0))%R). + assert(-1/1125899906842624 <= delta <= 1/1125899906842624)%R as delta_BOUND. + { unfold delta, invb0. rewrite expand_alpha. unfold rd, rs. + gappa. + } + replace (rd (rd alpha * invb0 + invb0) - 1/bi)%R with + (delta + ((alpha * invb0 + invb0)-1/bi))%R by (unfold delta; ring). + replace (alpha * invb0 + invb0 - 1 / bi)%R with (alpha * (invb0 - 1/bi) + (alpha * (1/bi) + invb0 - 1 / bi))%R by ring. + replace (alpha * (1 / bi) + invb0 - 1 / bi)%R with 0%R; cycle 1. + { unfold alpha. + field. + lra. + } + rewrite expand_alpha. + unfold invb0, rd, rs, approx_inv_thresh. + apply Rabs_le. + gappa. Admitted. Definition approx_div a b := -- cgit From de209fbba248e47f703f90c90817041cfdbf3d06 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 21:20:44 +0100 Subject: one admit less --- kvx/FPDivision.v | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index b71c8990..37aa76ea 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -96,7 +96,10 @@ Proof. pose proof (BofZ_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C1. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. cbn. - admit. + assert (1 <= (round radix2 (FLT_exp (-149) 24) ZnearestE (IZR b')) <= 4294967296)%R. + gappa. + apply Rabs_lt. + lra. } rewrite Zlt_bool_false in C1 by lia. destruct C1 as (C1E & C1F & _). -- cgit From 3ae38acc45f9a9cddab0be68058f28743f8c7a47 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 21:50:49 +0100 Subject: some more gappa --- kvx/FPDivision.v | 27 +++++++++++++++++++++------ 1 file changed, 21 insertions(+), 6 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 37aa76ea..311632fd 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -51,7 +51,16 @@ Proof. lia. Qed. *) - + +Lemma Rabs_relax: + forall b b' (INEQ : (b < b')%R) x, + (-b <= x <= b)%R -> (Rabs x < b')%R. +Proof. + intros. + apply Rabs_lt. + lra. +Qed. + Theorem approx_inv_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_b : expr) (b : int) @@ -96,10 +105,9 @@ Proof. pose proof (BofZ_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C1. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. cbn. - assert (1 <= (round radix2 (FLT_exp (-149) 24) ZnearestE (IZR b')) <= 4294967296)%R. - gappa. - apply Rabs_lt. + eapply (Rabs_relax (IZR 4294967296)). lra. + gappa. } rewrite Zlt_bool_false in C1 by lia. destruct C1 as (C1E & C1F & _). @@ -113,7 +121,12 @@ Proof. (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1) (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b')) as C2. rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. admit. } + { clear C2. rewrite C1E. + apply (Rabs_relax (bpow radix2 24) (bpow radix2 128)). + { cbn; lra. } + unfold F2R. cbn. unfold F2R. cbn. + gappa. + } assert (B2R 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') <> 0%R) as NONZ. { admit. } destruct (C2 NONZ) as (C2E & C2F & _). @@ -130,7 +143,9 @@ Proof. (@eq_refl Datatypes.comparison Lt) b'))) as C3. fold invb_d in C3. rewrite Rlt_bool_true in C3; cycle 1. - { clear C3. admit. } + { clear C3. + admit. + } change (is_finite 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1)) with true in C2F. destruct (C3 C2F) as (C3E & C3F & _). unfold Float.fma. -- cgit From f3510994276cfc73fb2f8441f652c97259753d81 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 21:52:51 +0100 Subject: one admit less --- kvx/FPDivision.v | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 311632fd..2b99f392 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -128,7 +128,12 @@ Proof. gappa. } assert (B2R 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') <> 0%R) as NONZ. - { admit. } + { clear C2. + rewrite C1E. + cbn. + assert (1 <= round radix2 (FLT_exp (-149) 24) ZnearestE (IZR b'))%R by gappa. + lra. + } destruct (C2 NONZ) as (C2E & C2F & _). clear C2 NONZ. Local Transparent Float.of_single. -- cgit From 18ec7b4072aef0160e96915d13218381f4420c4c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 21:57:25 +0100 Subject: one admit less --- kvx/FPDivision.v | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 2b99f392..fed327b7 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -122,7 +122,7 @@ Proof. (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b')) as C2. rewrite Rlt_bool_true in C2; cycle 1. { clear C2. rewrite C1E. - apply (Rabs_relax (bpow radix2 24) (bpow radix2 128)). + apply (Rabs_relax (bpow radix2 10)). { cbn; lra. } unfold F2R. cbn. unfold F2R. cbn. gappa. @@ -149,7 +149,13 @@ Proof. fold invb_d in C3. rewrite Rlt_bool_true in C3; cycle 1. { clear C3. - admit. + rewrite C2E. + rewrite C1E. + rewrite C0E. + apply (Rabs_relax (bpow radix2 10)). + { cbn; lra. } + cbn. + gappa. } change (is_finite 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1)) with true in C2F. destruct (C3 C2F) as (C3E & C3F & _). -- cgit From b0f18b3de1c530a03a90c0618c39a3d26704f49e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 22:07:26 +0100 Subject: one admit less --- kvx/FPDivision.v | 22 ++++++++++++++++++++-- 1 file changed, 20 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index fed327b7..5d1ce1ac 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -153,12 +153,13 @@ Proof. rewrite C1E. rewrite C0E. apply (Rabs_relax (bpow radix2 10)). - { cbn; lra. } + { apply bpow_lt; lia. } cbn. gappa. } change (is_finite 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1)) with true in C2F. destruct (C3 C2F) as (C3E & C3F & _). + clear C3. unfold Float.fma. assert (is_finite _ _ (Float.neg invb_d) = true) as invb_d_F. { Local Transparent Float.neg. @@ -184,7 +185,24 @@ Proof. (Float.neg invb_d) b_d ExtFloat.one invb_d_F b_d_F one_F) as C5. rewrite Rlt_bool_true in C5; cycle 1. - { admit. } + { clear C5. + unfold Float.neg. + rewrite B2R_Bopp. + rewrite C3E. + rewrite C2E. + rewrite C0E. + rewrite C1E. + unfold ExtFloat.one. + change (Float.of_int (Int.repr 1)) with (BofZ 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1). + rewrite C9E. + unfold b_d. + rewrite Int64.unsigned_repr by lia. + rewrite C4E. + apply (Rabs_relax (bpow radix2 10)). + { apply bpow_lt; lia. } + cbn. + gappa. + } destruct C5 as (C5E & C5F & _). pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE -- cgit From 3fb606b155f0ed443bb99b0a94b2b32ca3ad4276 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 22:14:44 +0100 Subject: Qed --- kvx/FPDivision.v | 26 ++++++++++++++++++++++++-- 1 file changed, 24 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 5d1ce1ac..a5be3d0a 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -209,7 +209,29 @@ Proof. (Bfma 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE (Float.neg invb_d) b_d ExtFloat.one) invb_d invb_d C5F C3F C3F) as C6. rewrite Rlt_bool_true in C6; cycle 1. - { admit. } + { clear C6. + rewrite C3E. + rewrite C2E. + rewrite C1E. + rewrite C0E. + rewrite C5E. + unfold Float.neg. + rewrite B2R_Bopp. + rewrite C3E. + rewrite C2E. + rewrite C0E. + rewrite C1E. + unfold b_d. + rewrite Int64.unsigned_repr by lia. + rewrite C4E. + unfold ExtFloat.one. + change (Float.of_int (Int.repr 1)) with (BofZ 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1). + rewrite C9E. + apply (Rabs_relax (bpow radix2 10)). + { apply bpow_lt; lia. } + cbn. + gappa. + } destruct C6 as (C6E & C6F & _). split. { exact C6F. } @@ -267,7 +289,7 @@ Proof. unfold invb0, rd, rs, approx_inv_thresh. apply Rabs_le. gappa. -Admitted. +Qed. Definition approx_div a b := let a_var := Eletvar (1%nat) in -- cgit From e1eed52ee24859e2fbc6c0f88af5920889df204b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 22:37:50 +0100 Subject: progress --- kvx/FPDivision.v | 35 +++++++++++++++++++++++++++++++---- 1 file changed, 31 insertions(+), 4 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index a5be3d0a..a95072a5 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -35,7 +35,7 @@ Definition approx_inv b := let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in Elet b (Elet invb_d x). -Definition approx_inv_thresh := (1/4294967296)%R. +Definition approx_inv_thresh := (1/17179869184)%R. (* Lemma BofZ_exact_zero: @@ -383,11 +383,38 @@ Proof. unfold Float.to_longu_ne. rewrite ZofB_ne_range_correct by lia. set (prod := (Float.mul (Float.of_intu a) inv_b)). - set (a' := Int.unsigned a) in *. - set (b' := Int.unsigned b) in *. + pose proof (Int.unsigned_range a) as a_range. + change Int.modulus with 4294967296 in a_range. set (prod' := (B2R _ _ prod)). set (prod_r := ZnearestE prod') in *. - replace (_ && _ && _) with true by admit. + + Local Transparent Float.mul Float.of_intu. + unfold Float.mul, Float.of_intu in prod. + set (a' := Int.unsigned a) in *. + set (b' := Int.unsigned b) in *. + assert (SILLY : -2^53 <= a' <= 2^53). + { cbn; lia. } + destruct (BofZ_exact 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a' SILLY) as (C0E & C0F & _). + clear SILLY. + pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE + (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt)a') inv_b) as C1. + rewrite Rlt_bool_true in C1; cycle 1. + { admit. } + rewrite C0F in C1. + cbn in C1. + rewrite C0E in C1. + rewrite inv_b_finite in C1. + fold prod in C1. + fold prod' in C1. + destruct C1 as (C1E & C1F & _). + rewrite C1F. cbn. + + replace (_ && _ ) with true; cycle 1. + { + symmetry. + rewrite andb_true_iff. + admit. + } cbn. f_equal. unfold Int.divu. -- cgit From 18bbbf846dcfbc4454ccbbf0ff837024bcd383e3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 23:13:05 +0100 Subject: progress --- kvx/FPDivision.v | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index a95072a5..4ea3f5a3 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -384,7 +384,9 @@ Proof. rewrite ZofB_ne_range_correct by lia. set (prod := (Float.mul (Float.of_intu a) inv_b)). pose proof (Int.unsigned_range a) as a_range. + pose proof (Int.unsigned_range b) as b_range. change Int.modulus with 4294967296 in a_range. + change Int.modulus with 4294967296 in b_range. set (prod' := (B2R _ _ prod)). set (prod_r := ZnearestE prod') in *. @@ -409,6 +411,30 @@ Proof. destruct C1 as (C1E & C1F & _). rewrite C1F. cbn. + assert(prod'_range : (0 <= prod' <= 4294967296)%R). + { + rewrite C1E. + apply Rabs_def2b in inv_b_correct. + set (inv_b_r := B2R 53 1024 inv_b) in *. + assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R). + { + unfold Rdiv. split. + - apply Rmult_le_compat_l. lra. apply Rinv_le. + apply IZR_lt. lia. apply IZR_le. lia. + - replace 1%R with (1 * / 1)%R at 2 by field. + apply Rmult_le_compat_l. lra. + apply Rinv_le. lra. apply IZR_le. lia. + } + replace inv_b_r with (1/IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. + assert (a'_RANGE : (0 <= IZR a' <= 4294967295)%R). + { split; apply IZR_le; lia. + } + unfold approx_inv_thresh in inv_b_correct. + set (true_inv := (1 / IZR b')%R) in *. + set (delta := (inv_b_r - true_inv)%R) in *. + gappa. + } + replace (_ && _ ) with true; cycle 1. { symmetry. -- cgit From e88eb7bdf74f68c4577477650332240dabfee88e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 23:19:17 +0100 Subject: better --- kvx/FPDivision.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 4ea3f5a3..6fd7e14c 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -411,7 +411,7 @@ Proof. destruct C1 as (C1E & C1F & _). rewrite C1F. cbn. - assert(prod'_range : (0 <= prod' <= 4294967296)%R). + assert(prod'_range : (0 <= prod' <= 17179869181/4)%R). { rewrite C1E. apply Rabs_def2b in inv_b_correct. -- cgit From dab5ceabe0b11daa0199a6a48c73adbee9fafb4c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 23:27:40 +0100 Subject: progress --- kvx/FPDivision.v | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 6fd7e14c..1b81b632 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -435,11 +435,15 @@ Proof. gappa. } + assert(0 <= prod_r <= 4294967295) as prod_r_range. + { admit. + } + replace (_ && _ ) with true; cycle 1. { symmetry. rewrite andb_true_iff. - admit. + split; apply Zle_imp_le_bool; lia. } cbn. f_equal. -- cgit From ed7f246f082cc92283e773b6b6b4eb4828a3c2ca Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 23:38:15 +0100 Subject: glue in ranges --- kvx/FPDivision.v | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 1b81b632..7790d31e 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -359,6 +359,10 @@ Qed. Opaque approx_inv. +Theorem Znearest_le + : forall choice (x y : R), (x <= y)%R -> Znearest choice x <= Znearest choice y. +Admitted. + Theorem approx_div_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_a expr_b : expr) (a b : int) @@ -436,7 +440,10 @@ Proof. } assert(0 <= prod_r <= 4294967295) as prod_r_range. - { admit. + { unfold prod_r. + rewrite <- (Znearest_IZR (fun x => negb (Z.even x)) 0). + replace 4294967295 with (ZnearestE (17179869181 / 4)%R) by admit. + split; apply Znearest_le; lra. } replace (_ && _ ) with true; cycle 1. -- cgit From ddc5dbba8f6e6eaf0bb08a401a679e6c8fdfe07b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 23:41:37 +0100 Subject: removing admits --- kvx/FPDivision.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 7790d31e..b23e3e57 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -461,11 +461,11 @@ Proof. unfold div_approx_reals. fold a' b' prod' prod_r. unfold Int64.mul. - rewrite Int64.unsigned_repr by admit. - rewrite Int64.unsigned_repr by admit. + rewrite Int64.unsigned_repr by (cbn; lia). + rewrite Int64.unsigned_repr by (cbn; lia). unfold Int64.sub. - rewrite Int64.unsigned_repr by admit. - rewrite Int64.unsigned_repr by admit. + rewrite Int64.unsigned_repr by (cbn; lia). + rewrite Int64.unsigned_repr by (cbn; nia). case Z.ltb_spec; intro CMP. - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true by admit. cbn. -- cgit From 226b4b55b3518fc2e4e2627cd368725488e3f8c3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 23:50:57 +0100 Subject: less admit --- kvx/FPDivision.v | 23 +++++++++++++++++++---- 1 file changed, 19 insertions(+), 4 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index b23e3e57..41bb27d5 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -404,8 +404,26 @@ Proof. clear SILLY. pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt)a') inv_b) as C1. + set (inv_b_r := B2R 53 1024 inv_b) in *. + assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R). + apply Rabs_def2b in inv_b_correct. rewrite Rlt_bool_true in C1; cycle 1. - { admit. } + { clear C1. + rewrite C0E. + apply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. + set (delta := (inv_b_r - 1 / IZR b')%R) in *. + unfold approx_inv_thresh in inv_b_correct. + cbn. + assert (b'_RANGE : (1 <= IZR b' <= 4294967295)%R). + { split; apply IZR_le; lia. + } + assert (a'_RANGE : (0 <= IZR a' <= 4294967295)%R). + { split; apply IZR_le; lia. + } + gappa. + } rewrite C0F in C1. cbn in C1. rewrite C0E in C1. @@ -418,9 +436,6 @@ Proof. assert(prod'_range : (0 <= prod' <= 17179869181/4)%R). { rewrite C1E. - apply Rabs_def2b in inv_b_correct. - set (inv_b_r := B2R 53 1024 inv_b) in *. - assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R). { unfold Rdiv. split. - apply Rmult_le_compat_l. lra. apply Rinv_le. -- cgit From 03e9c06bc2f53212b596b4c61c047b1116f9408c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 16 Dec 2021 23:57:31 +0100 Subject: goes to end still some admit --- kvx/FPDivision.v | 10 +--------- 1 file changed, 1 insertion(+), 9 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 41bb27d5..257e22b9 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -405,7 +405,7 @@ Proof. pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt)a') inv_b) as C1. set (inv_b_r := B2R 53 1024 inv_b) in *. - assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R). + assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R) by admit. apply Rabs_def2b in inv_b_correct. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. @@ -436,14 +436,6 @@ Proof. assert(prod'_range : (0 <= prod' <= 17179869181/4)%R). { rewrite C1E. - { - unfold Rdiv. split. - - apply Rmult_le_compat_l. lra. apply Rinv_le. - apply IZR_lt. lia. apply IZR_le. lia. - - replace 1%R with (1 * / 1)%R at 2 by field. - apply Rmult_le_compat_l. lra. - apply Rinv_le. lra. apply IZR_le. lia. - } replace inv_b_r with (1/IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. assert (a'_RANGE : (0 <= IZR a' <= 4294967295)%R). { split; apply IZR_le; lia. -- cgit From e185e82170dad225355cc4472ab341938e718b79 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 07:42:03 +0100 Subject: one admit less --- kvx/FPDivision.v | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 257e22b9..06d833bd 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -449,7 +449,11 @@ Proof. assert(0 <= prod_r <= 4294967295) as prod_r_range. { unfold prod_r. rewrite <- (Znearest_IZR (fun x => negb (Z.even x)) 0). - replace 4294967295 with (ZnearestE (17179869181 / 4)%R) by admit. + replace 4294967295 with (ZnearestE (17179869181 / 4)%R); cycle 1. + { apply Znearest_imp. + apply Rabs_lt. + lra. + } split; apply Znearest_le; lra. } -- cgit From 449165cb8b108e15c60ccb76821f0208adac8056 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 07:50:46 +0100 Subject: Znearest_le --- kvx/FPDivision.v | 20 +++++++++++++++++++- 1 file changed, 19 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 06d833bd..72c2c8ca 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -359,9 +359,27 @@ Qed. Opaque approx_inv. +Search ({_ <= _} + {_ > _})%Z. Theorem Znearest_le : forall choice (x y : R), (x <= y)%R -> Znearest choice x <= Znearest choice y. -Admitted. +Proof. + intros. + destruct (Z_le_gt_dec (Znearest choice x) (Znearest choice y)) as [LE | GT]. + assumption. + exfalso. + assert (1 <= IZR (Znearest choice x) - IZR(Znearest choice y))%R as GAP. + { rewrite <- minus_IZR. + apply IZR_le. + lia. + } + pose proof (Znearest_imp2 choice x) as Rx. + pose proof (Znearest_imp2 choice y) as Ry. + apply Rabs_def2b in Rx. + apply Rabs_def2b in Ry. + assert (x = y) by lra. + subst y. + lia. +Qed. Theorem approx_div_correct : forall (ge : genv) (sp: val) cmenv memenv -- cgit From 94dff310377cf75e9457dbf67fbc39fc2ab82c7b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 09:24:14 +0100 Subject: one admit less --- kvx/FPDivision.v | 11 ++++++++++- 1 file changed, 10 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 72c2c8ca..573a8e18 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -423,7 +423,16 @@ Proof. pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt)a') inv_b) as C1. set (inv_b_r := B2R 53 1024 inv_b) in *. - assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R) by admit. + assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R). + { split; unfold Rdiv. + - apply Rmult_le_compat_l. lra. + apply Rinv_le. apply IZR_lt. lia. + apply IZR_le. lia. + - replace 1%R with (1 / 1)%R at 2 by field. + apply Rmult_le_compat_l. lra. + apply Rinv_le. apply IZR_lt. lia. + apply IZR_le. lia. + } apply Rabs_def2b in inv_b_correct. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. -- cgit From 6d75ed4e081bab3eb08b517a1a1a14131c22bcf1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 09:29:48 +0100 Subject: some admit less --- kvx/FPDivision.v | 18 ++++++++++++++++-- 1 file changed, 16 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 573a8e18..9866ff13 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -505,11 +505,25 @@ Proof. rewrite Int64.unsigned_repr by (cbn; lia). rewrite Int64.unsigned_repr by (cbn; nia). case Z.ltb_spec; intro CMP. - - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true by admit. + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. + { unfold Int64.lt. + change (Int64.signed Int64.zero) with 0. + rewrite Int64.signed_repr; cycle 1. + admit. + rewrite zlt_true by lia. + reflexivity. + } cbn. f_equal. admit. - - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false by admit. + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. + { unfold Int64.lt. + change (Int64.signed Int64.zero) with 0. + rewrite Int64.signed_repr; cycle 1. + admit. + rewrite zlt_false by lia. + reflexivity. + } cbn. f_equal. admit. -- cgit From 26a2b82b9d59fbf546e7b43bc5b2ee00ce5e4e98 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 09:34:06 +0100 Subject: one admit less --- kvx/FPDivision.v | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 9866ff13..e4e881d0 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -525,8 +525,9 @@ Proof. reflexivity. } cbn. - f_equal. - admit. + unfold Int64.loword. + rewrite Int64.unsigned_repr by (cbn; lia). + reflexivity. } fold prod'. admit. -- cgit From 9198ccc44fc942ca4e8428739d178489213f4e1b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 10:08:09 +0100 Subject: one less admit --- kvx/FPDivision.v | 29 +++++++++++++++++++++++++++-- 1 file changed, 27 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index e4e881d0..38f30696 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -529,8 +529,33 @@ Proof. rewrite Int64.unsigned_repr by (cbn; lia). reflexivity. } - fold prod'. - admit. + unfold prod. + pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a') inv_b) as C2. + rewrite C0E in C2. + rewrite Rlt_bool_true in C2 by admit. + destruct C2 as (C2E & C2F & _). + rewrite C2E. + fold inv_b_r a' b'. + assert (IZR_a' : (0 <= IZR a' <= 4294967295)%R). + { split; apply IZR_le; lia. } + assert (IZR_b' : (1 <= IZR b' <= 4294967295)%R). + { split; apply IZR_le; lia. } + replace ((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - + (IZR a' / IZR b'))%R with + (((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - + (IZR a' * inv_b_r)) + + (IZR a' * (inv_b_r - 1 / IZR b')))%R by (field ; lra). + set(delta := (inv_b_r - 1 / IZR b')%R) in *. + cbn. + assert(Rabs + (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. + unfold approx_inv_thresh in *. + assert (Rabs(IZR a' * delta) <= 3/8)%R as R2. + { apply Rabs_le. + split; nra. + } + pose proof (Rabs_triang (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - IZR a' * inv_b_r) (IZR a' * delta))%R. + lra. Admitted. -- cgit From 7df7358e5e5ea0fc19ef61f017c47c62af576f63 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 10:17:23 +0100 Subject: one less admit --- kvx/FPDivision.v | 18 ++++++++++++++---- 1 file changed, 14 insertions(+), 4 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 38f30696..6b0a0ce8 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -532,14 +532,24 @@ Proof. unfold prod. pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a') inv_b) as C2. rewrite C0E in C2. - rewrite Rlt_bool_true in C2 by admit. - destruct C2 as (C2E & C2F & _). - rewrite C2E. - fold inv_b_r a' b'. assert (IZR_a' : (0 <= IZR a' <= 4294967295)%R). { split; apply IZR_le; lia. } assert (IZR_b' : (1 <= IZR b' <= 4294967295)%R). { split; apply IZR_le; lia. } + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. reflexivity. } + cbn. + fold inv_b_r. + replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. + set (delta := (inv_b_r - 1 / IZR b')%R) in *. + unfold approx_inv_thresh in *. + gappa. + } + destruct C2 as (C2E & C2F & _). + rewrite C2E. + fold inv_b_r a' b'. replace ((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - (IZR a' / IZR b'))%R with (((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - -- cgit From 6af877e7e6a8c2b98850f80fad35a0d03126b5fa Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 11:15:45 +0100 Subject: one less admit --- kvx/FPDivision.v | 42 +++++++++++++++++++++++++++++++++++------- 1 file changed, 35 insertions(+), 7 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 6b0a0ce8..05d6cee7 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -416,6 +416,10 @@ Proof. unfold Float.mul, Float.of_intu in prod. set (a' := Int.unsigned a) in *. set (b' := Int.unsigned b) in *. + assert (IZR_a' : (0 <= IZR a' <= 4294967295)%R). + { split; apply IZR_le; lia. } + assert (IZR_b' : (1 <= IZR b' <= 4294967295)%R). + { split; apply IZR_le; lia. } assert (SILLY : -2^53 <= a' <= 2^53). { cbn; lia. } destruct (BofZ_exact 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a' SILLY) as (C0E & C0F & _). @@ -504,12 +508,40 @@ Proof. unfold Int64.sub. rewrite Int64.unsigned_repr by (cbn; lia). rewrite Int64.unsigned_repr by (cbn; nia). + assert(Rabs + (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. case Z.ltb_spec; intro CMP. - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. { unfold Int64.lt. change (Int64.signed Int64.zero) with 0. rewrite Int64.signed_repr; cycle 1. - admit. + { cbn. + unfold prod_r. + rewrite <- C1E in R1. + assert (IZR (-9223372036854775808) <= IZR (a' - ZnearestE prod' * b') <= IZR 9223372036854775807)%R. + 2: split; apply le_IZR; lra. + rewrite minus_IZR. + rewrite mult_IZR. + replace (IZR (ZnearestE prod')) with (prod' + (IZR (ZnearestE prod') - prod'))%R by ring. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) prod') as R2. + set (delta1 := (IZR (ZnearestE prod') - prod')%R) in *. + replace prod' with ((prod' - IZR a' * inv_b_r) + IZR a' * (inv_b_r - 1 / IZR b') + + IZR a' / IZR b')%R by (field; lra). + set (delta2 := (inv_b_r - 1 / IZR b')%R) in *. + set (delta3 := (prod' - IZR a' * inv_b_r)%R) in *. + replace (IZR a' - (delta3 + IZR a' * delta2 + IZR a' / IZR b' + delta1) * IZR b')%R with + (- (delta3 + IZR a' * delta2 + delta1) * IZR b')%R by (field; lra). + unfold approx_inv_thresh in *. + assert (Rabs(IZR a' * delta2) <= 1/4)%R as R4. + { apply Rabs_le. + split; + nra. } + set (delta4 := (IZR a' * delta2)%R) in *. + apply Rabs_def2b in R1. + apply Rabs_def2b in R2. + apply Rabs_def2b in R4. + split; nra. + } rewrite zlt_true by lia. reflexivity. } @@ -532,10 +564,6 @@ Proof. unfold prod. pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a') inv_b) as C2. rewrite C0E in C2. - assert (IZR_a' : (0 <= IZR a' <= 4294967295)%R). - { split; apply IZR_le; lia. } - assert (IZR_b' : (1 <= IZR b' <= 4294967295)%R). - { split; apply IZR_le; lia. } rewrite Rlt_bool_true in C2; cycle 1. { clear C2. apply (Rabs_relax (bpow radix2 64)). @@ -557,8 +585,8 @@ Proof. (IZR a' * (inv_b_r - 1 / IZR b')))%R by (field ; lra). set(delta := (inv_b_r - 1 / IZR b')%R) in *. cbn. - assert(Rabs - (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. + (* assert(Rabs + (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. *) unfold approx_inv_thresh in *. assert (Rabs(IZR a' * delta) <= 3/8)%R as R2. { apply Rabs_le. -- cgit From 73aa78c4e5c037b76a56cc76738f315fc2f5fa0b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 11:28:55 +0100 Subject: fixup --- kvx/FPDivision.v | 69 +++++++++++++++++++++++++++++--------------------------- 1 file changed, 36 insertions(+), 33 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 05d6cee7..9bfacb80 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -497,6 +497,9 @@ Proof. cbn. f_equal. unfold Int.divu. + assert(Rabs + (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. + rewrite <- div_approx_reals_correct with (x := B2R _ _ prod). 2: lia. { @@ -508,40 +511,39 @@ Proof. unfold Int64.sub. rewrite Int64.unsigned_repr by (cbn; lia). rewrite Int64.unsigned_repr by (cbn; nia). - assert(Rabs - (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. + assert (FMA_RANGE : Int64.min_signed <= a' - prod_r * b' <= Int64.max_signed). + { cbn. + unfold prod_r. + rewrite <- C1E in R1. + assert (IZR (-9223372036854775808) <= IZR (a' - ZnearestE prod' * b') <= IZR 9223372036854775807)%R. + 2: split; apply le_IZR; lra. + rewrite minus_IZR. + rewrite mult_IZR. + replace (IZR (ZnearestE prod')) with (prod' + (IZR (ZnearestE prod') - prod'))%R by ring. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) prod') as R2. + set (delta1 := (IZR (ZnearestE prod') - prod')%R) in *. + replace prod' with ((prod' - IZR a' * inv_b_r) + IZR a' * (inv_b_r - 1 / IZR b') + + IZR a' / IZR b')%R by (field; lra). + set (delta2 := (inv_b_r - 1 / IZR b')%R) in *. + set (delta3 := (prod' - IZR a' * inv_b_r)%R) in *. + replace (IZR a' - (delta3 + IZR a' * delta2 + IZR a' / IZR b' + delta1) * IZR b')%R with + (- (delta3 + IZR a' * delta2 + delta1) * IZR b')%R by (field; lra). + unfold approx_inv_thresh in *. + assert (Rabs(IZR a' * delta2) <= 1/4)%R as R4. + { apply Rabs_le. + split; + nra. } + set (delta4 := (IZR a' * delta2)%R) in *. + apply Rabs_def2b in R1. + apply Rabs_def2b in R2. + apply Rabs_def2b in R4. + split; nra. + } case Z.ltb_spec; intro CMP. - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. { unfold Int64.lt. change (Int64.signed Int64.zero) with 0. - rewrite Int64.signed_repr; cycle 1. - { cbn. - unfold prod_r. - rewrite <- C1E in R1. - assert (IZR (-9223372036854775808) <= IZR (a' - ZnearestE prod' * b') <= IZR 9223372036854775807)%R. - 2: split; apply le_IZR; lra. - rewrite minus_IZR. - rewrite mult_IZR. - replace (IZR (ZnearestE prod')) with (prod' + (IZR (ZnearestE prod') - prod'))%R by ring. - pose proof (Znearest_imp2 (fun x => negb (Z.even x)) prod') as R2. - set (delta1 := (IZR (ZnearestE prod') - prod')%R) in *. - replace prod' with ((prod' - IZR a' * inv_b_r) + IZR a' * (inv_b_r - 1 / IZR b') - + IZR a' / IZR b')%R by (field; lra). - set (delta2 := (inv_b_r - 1 / IZR b')%R) in *. - set (delta3 := (prod' - IZR a' * inv_b_r)%R) in *. - replace (IZR a' - (delta3 + IZR a' * delta2 + IZR a' / IZR b' + delta1) * IZR b')%R with - (- (delta3 + IZR a' * delta2 + delta1) * IZR b')%R by (field; lra). - unfold approx_inv_thresh in *. - assert (Rabs(IZR a' * delta2) <= 1/4)%R as R4. - { apply Rabs_le. - split; - nra. } - set (delta4 := (IZR a' * delta2)%R) in *. - apply Rabs_def2b in R1. - apply Rabs_def2b in R2. - apply Rabs_def2b in R4. - split; nra. - } + rewrite Int64.signed_repr by exact FMA_RANGE. rewrite zlt_true by lia. reflexivity. } @@ -551,8 +553,7 @@ Proof. - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. { unfold Int64.lt. change (Int64.signed Int64.zero) with 0. - rewrite Int64.signed_repr; cycle 1. - admit. + rewrite Int64.signed_repr by exact FMA_RANGE. rewrite zlt_false by lia. reflexivity. } @@ -592,7 +593,9 @@ Proof. { apply Rabs_le. split; nra. } - pose proof (Rabs_triang (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - IZR a' * inv_b_r) (IZR a' * delta))%R. + rewrite <- C1E. + rewrite <- C1E in R1. + pose proof (Rabs_triang (prod' - IZR a' * inv_b_r) (IZR a' * delta))%R. lra. Admitted. -- cgit From 54cf78c0d3579dced79d846ee111b8c5edfee615 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 11:52:13 +0100 Subject: reorganize proof --- kvx/FPDivision.v | 179 ++++++++++++++++++++++++++++--------------------------- 1 file changed, 90 insertions(+), 89 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index 9bfacb80..f22387d6 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -500,103 +500,104 @@ Proof. assert(Rabs (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. - rewrite <- div_approx_reals_correct with (x := B2R _ _ prod). - 2: lia. + assert ( (Rabs (B2R 53 1024 prod - IZR (Int.unsigned a) / IZR (Int.unsigned b)) < 1 / 2)%R) as NEAR. { - unfold div_approx_reals. - fold a' b' prod' prod_r. - unfold Int64.mul. - rewrite Int64.unsigned_repr by (cbn; lia). - rewrite Int64.unsigned_repr by (cbn; lia). - unfold Int64.sub. - rewrite Int64.unsigned_repr by (cbn; lia). - rewrite Int64.unsigned_repr by (cbn; nia). - assert (FMA_RANGE : Int64.min_signed <= a' - prod_r * b' <= Int64.max_signed). - { cbn. - unfold prod_r. - rewrite <- C1E in R1. - assert (IZR (-9223372036854775808) <= IZR (a' - ZnearestE prod' * b') <= IZR 9223372036854775807)%R. - 2: split; apply le_IZR; lra. - rewrite minus_IZR. - rewrite mult_IZR. - replace (IZR (ZnearestE prod')) with (prod' + (IZR (ZnearestE prod') - prod'))%R by ring. - pose proof (Znearest_imp2 (fun x => negb (Z.even x)) prod') as R2. - set (delta1 := (IZR (ZnearestE prod') - prod')%R) in *. - replace prod' with ((prod' - IZR a' * inv_b_r) + IZR a' * (inv_b_r - 1 / IZR b') - + IZR a' / IZR b')%R by (field; lra). - set (delta2 := (inv_b_r - 1 / IZR b')%R) in *. - set (delta3 := (prod' - IZR a' * inv_b_r)%R) in *. - replace (IZR a' - (delta3 + IZR a' * delta2 + IZR a' / IZR b' + delta1) * IZR b')%R with - (- (delta3 + IZR a' * delta2 + delta1) * IZR b')%R by (field; lra). + unfold prod. + pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a') inv_b) as C2. + rewrite C0E in C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. reflexivity. } + cbn. + fold inv_b_r. + replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. + set (delta := (inv_b_r - 1 / IZR b')%R) in *. unfold approx_inv_thresh in *. - assert (Rabs(IZR a' * delta2) <= 1/4)%R as R4. - { apply Rabs_le. - split; - nra. } - set (delta4 := (IZR a' * delta2)%R) in *. - apply Rabs_def2b in R1. - apply Rabs_def2b in R2. - apply Rabs_def2b in R4. + gappa. + } + destruct C2 as (C2E & C2F & _). + rewrite C2E. + fold inv_b_r a' b'. + replace ((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - + (IZR a' / IZR b'))%R with + (((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - + (IZR a' * inv_b_r)) + + (IZR a' * (inv_b_r - 1 / IZR b')))%R by (field ; lra). + set(delta := (inv_b_r - 1 / IZR b')%R) in *. + cbn. + (* assert(Rabs + (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. *) + unfold approx_inv_thresh in *. + assert (Rabs(IZR a' * delta) <= 3/8)%R as R2. + { apply Rabs_le. split; nra. } - case Z.ltb_spec; intro CMP. - - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. - { unfold Int64.lt. - change (Int64.signed Int64.zero) with 0. - rewrite Int64.signed_repr by exact FMA_RANGE. - rewrite zlt_true by lia. - reflexivity. - } - cbn. - f_equal. - admit. - - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. - { unfold Int64.lt. - change (Int64.signed Int64.zero) with 0. - rewrite Int64.signed_repr by exact FMA_RANGE. - rewrite zlt_false by lia. - reflexivity. - } - cbn. - unfold Int64.loword. - rewrite Int64.unsigned_repr by (cbn; lia). - reflexivity. + rewrite <- C1E. + rewrite <- C1E in R1. + pose proof (Rabs_triang (prod' - IZR a' * inv_b_r) (IZR a' * delta))%R. + lra. } - unfold prod. - pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a') inv_b) as C2. - rewrite C0E in C2. - rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. - apply (Rabs_relax (bpow radix2 64)). - { apply bpow_lt. reflexivity. } - cbn. - fold inv_b_r. - replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. - set (delta := (inv_b_r - 1 / IZR b')%R) in *. + rewrite <- div_approx_reals_correct with (x := B2R _ _ prod); cycle 1. lia. exact NEAR. + + unfold div_approx_reals. + fold a' b' prod' prod_r. + unfold Int64.mul. + rewrite Int64.unsigned_repr by (cbn; lia). + rewrite Int64.unsigned_repr by (cbn; lia). + unfold Int64.sub. + rewrite Int64.unsigned_repr by (cbn; lia). + rewrite Int64.unsigned_repr by (cbn; nia). + assert (FMA_RANGE : Int64.min_signed <= a' - prod_r * b' <= Int64.max_signed). + { cbn. + unfold prod_r. + rewrite <- C1E in R1. + assert (IZR (-9223372036854775808) <= IZR (a' - ZnearestE prod' * b') <= IZR 9223372036854775807)%R. + 2: split; apply le_IZR; lra. + rewrite minus_IZR. + rewrite mult_IZR. + replace (IZR (ZnearestE prod')) with (prod' + (IZR (ZnearestE prod') - prod'))%R by ring. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) prod') as R2. + set (delta1 := (IZR (ZnearestE prod') - prod')%R) in *. + replace prod' with ((prod' - IZR a' * inv_b_r) + IZR a' * (inv_b_r - 1 / IZR b') + + IZR a' / IZR b')%R by (field; lra). + set (delta2 := (inv_b_r - 1 / IZR b')%R) in *. + set (delta3 := (prod' - IZR a' * inv_b_r)%R) in *. + replace (IZR a' - (delta3 + IZR a' * delta2 + IZR a' / IZR b' + delta1) * IZR b')%R with + (- (delta3 + IZR a' * delta2 + delta1) * IZR b')%R by (field; lra). unfold approx_inv_thresh in *. - gappa. - } - destruct C2 as (C2E & C2F & _). - rewrite C2E. - fold inv_b_r a' b'. - replace ((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - - (IZR a' / IZR b'))%R with - (((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - - (IZR a' * inv_b_r)) + - (IZR a' * (inv_b_r - 1 / IZR b')))%R by (field ; lra). - set(delta := (inv_b_r - 1 / IZR b')%R) in *. - cbn. - (* assert(Rabs - (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. *) - unfold approx_inv_thresh in *. - assert (Rabs(IZR a' * delta) <= 3/8)%R as R2. - { apply Rabs_le. + assert (Rabs(IZR a' * delta2) <= 1/4)%R as R4. + { apply Rabs_le. + split; + nra. } + set (delta4 := (IZR a' * delta2)%R) in *. + apply Rabs_def2b in R1. + apply Rabs_def2b in R2. + apply Rabs_def2b in R4. split; nra. } - rewrite <- C1E. - rewrite <- C1E in R1. - pose proof (Rabs_triang (prod' - IZR a' * inv_b_r) (IZR a' * delta))%R. - lra. + case Z.ltb_spec; intro CMP. + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. + { unfold Int64.lt. + change (Int64.signed Int64.zero) with 0. + rewrite Int64.signed_repr by exact FMA_RANGE. + rewrite zlt_true by lia. + reflexivity. + } + cbn. + f_equal. + admit. + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. + { unfold Int64.lt. + change (Int64.signed Int64.zero) with 0. + rewrite Int64.signed_repr by exact FMA_RANGE. + rewrite zlt_false by lia. + reflexivity. + } + cbn. + unfold Int64.loword. + rewrite Int64.unsigned_repr by (cbn; lia). + reflexivity. Admitted. -- cgit From 1e3340be6913b8988cde998c165e77eea480783c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 12:10:38 +0100 Subject: reorganize proof --- kvx/FPDivision.v | 15 ++++++++++++++- 1 file changed, 14 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index f22387d6..db44c30a 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -538,7 +538,8 @@ Proof. pose proof (Rabs_triang (prod' - IZR a' * inv_b_r) (IZR a' * delta))%R. lra. } - rewrite <- div_approx_reals_correct with (x := B2R _ _ prod); cycle 1. lia. exact NEAR. + pose proof (div_approx_reals_correct (Int.unsigned a) (Int.unsigned b) prod') as DIV_CORRECT. + rewrite <- DIV_CORRECT; cycle 1. lia. exact NEAR. unfold div_approx_reals. fold a' b' prod' prod_r. @@ -586,6 +587,18 @@ Proof. } cbn. f_equal. + rewrite Int64.add_signed. + rewrite (Int64.signed_repr prod_r) by (cbn ; lia). + rewrite Int64.signed_repr by (cbn ; lia). + unfold Int64.loword. + rewrite Int64.unsigned_repr. reflexivity. + fold prod' in NEAR. + fold a' in NEAR. + fold b' in NEAR. + split. + 2: cbn; lia. + apply le_IZR. + rewrite plus_IZR. admit. - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. { unfold Int64.lt. -- cgit From cae3d77b314778bd6b0dcc679ac7833c38ce7aeb Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 12:27:15 +0100 Subject: no more admitted --- kvx/FPDivision.v | 16 ++++++++++------ 1 file changed, 10 insertions(+), 6 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index db44c30a..d9c108a7 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -541,7 +541,7 @@ Proof. pose proof (div_approx_reals_correct (Int.unsigned a) (Int.unsigned b) prod') as DIV_CORRECT. rewrite <- DIV_CORRECT; cycle 1. lia. exact NEAR. - unfold div_approx_reals. + unfold div_approx_reals in *. fold a' b' prod' prod_r. unfold Int64.mul. rewrite Int64.unsigned_repr by (cbn; lia). @@ -577,7 +577,10 @@ Proof. apply Rabs_def2b in R4. split; nra. } - case Z.ltb_spec; intro CMP. + fold a' b' prod_r in DIV_CORRECT. + + pose proof (Zlt_cases (a' - prod_r * b') 0) as CMP. + destruct (Z.ltb (a' - prod_r * b') 0). - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. { unfold Int64.lt. change (Int64.signed Int64.zero) with 0. @@ -597,9 +600,10 @@ Proof. fold b' in NEAR. split. 2: cbn; lia. - apply le_IZR. - rewrite plus_IZR. - admit. + replace (prod_r + (-1)) with (prod_r - 1) by lia. + rewrite (DIV_CORRECT b_nz NEAR). + apply Z.div_pos; lia. + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. { unfold Int64.lt. change (Int64.signed Int64.zero) with 0. @@ -611,6 +615,6 @@ Proof. unfold Int64.loword. rewrite Int64.unsigned_repr by (cbn; lia). reflexivity. -Admitted. +Qed. -- cgit From 16b875d816a0b62037f4ffd1a0233f516529e5ec Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 12:45:34 +0100 Subject: simplify proof --- kvx/FPDivision.v | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index d9c108a7..a7ed34fc 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -526,8 +526,6 @@ Proof. (IZR a' * (inv_b_r - 1 / IZR b')))%R by (field ; lra). set(delta := (inv_b_r - 1 / IZR b')%R) in *. cbn. - (* assert(Rabs - (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. *) unfold approx_inv_thresh in *. assert (Rabs(IZR a' * delta) <= 3/8)%R as R2. { apply Rabs_le. @@ -538,8 +536,8 @@ Proof. pose proof (Rabs_triang (prod' - IZR a' * inv_b_r) (IZR a' * delta))%R. lra. } - pose proof (div_approx_reals_correct (Int.unsigned a) (Int.unsigned b) prod') as DIV_CORRECT. - rewrite <- DIV_CORRECT; cycle 1. lia. exact NEAR. + pose proof (div_approx_reals_correct (Int.unsigned a) (Int.unsigned b) prod' b_nz NEAR) as DIV_CORRECT. + rewrite <- DIV_CORRECT. unfold div_approx_reals in *. fold a' b' prod' prod_r. @@ -601,7 +599,7 @@ Proof. split. 2: cbn; lia. replace (prod_r + (-1)) with (prod_r - 1) by lia. - rewrite (DIV_CORRECT b_nz NEAR). + rewrite DIV_CORRECT. apply Z.div_pos; lia. - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. -- cgit From 784976ab246e87657c483388edac46208c4c283b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 12:48:14 +0100 Subject: simplify proof --- kvx/FPDivision.v | 3 --- 1 file changed, 3 deletions(-) diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v index a7ed34fc..0ba9a39f 100644 --- a/kvx/FPDivision.v +++ b/kvx/FPDivision.v @@ -593,9 +593,6 @@ Proof. rewrite Int64.signed_repr by (cbn ; lia). unfold Int64.loword. rewrite Int64.unsigned_repr. reflexivity. - fold prod' in NEAR. - fold a' in NEAR. - fold b' in NEAR. split. 2: cbn; lia. replace (prod_r + (-1)) with (prod_r - 1) by lia. -- cgit From 0cb4f03fea5f28280d8ce066204f69146fab99b6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 13:34:26 +0100 Subject: cleanup --- kvx/FPDivision32.v | 613 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 613 insertions(+) create mode 100644 kvx/FPDivision32.v diff --git a/kvx/FPDivision32.v b/kvx/FPDivision32.v new file mode 100644 index 00000000..7c37c619 --- /dev/null +++ b/kvx/FPDivision32.v @@ -0,0 +1,613 @@ +From Flocq Require Import Core Digits Operations Round Bracket Sterbenz + Binary Round_odd Bits. +Require Archi. +Require Import Coqlib. +Require Import Compopts. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Op. +Require Import CminorSel. +Require Import OpHelpers. +Require Import ExtFloats. +Require Import DecBoolOps. +Require Import Chunks. +Require Import Builtins. +Require Import Values Globalenvs. +Require Compopts. +Require Import Psatz. +Require Import IEEE754_extra. + +From Gappa Require Import Gappa_tactic. + +Local Open Scope cminorsel_scope. + +Local Open Scope string_scope. +Local Open Scope error_monad_scope. + +Definition approx_inv b := + let invb_s := Eop Oinvfs ((Eop Osingleofintu ((Eletvar 0%nat):::Enil)):::Enil) in + let invb_d := Eop Ofloatofsingle (invb_s ::: Enil) in + let b_d := Eop Ofloatoflongu ((Eop Ocast32unsigned ((Eletvar 1%nat):::Enil)):::Enil) in + let invb_d_var := Eletvar (0%nat) in + let one := Eop (Ofloatconst ExtFloat.one) Enil in + let alpha := Eop Ofmsubf (one ::: invb_d_var ::: b_d ::: Enil) in + let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in + Elet b (Elet invb_d x). + +Definition approx_inv_thresh := (1/17179869184)%R. + +(* +Lemma BofZ_exact_zero: + forall (prec emax : Z) (prec_gt_0_ : Prec_gt_0 prec) + (Hmax : (prec < emax)%Z), + B2R (BofZ prec emax 0%Z (Hmax := Hmax)) = 0%R /\ + is_finite (BofZ prec emax 0%Z (Hmax := Hmax)) = true /\ + Bsign prec emax (BofZ prec emax 0%Z (Hmax := Hmax)) = false. +Proof. + intros. + apply BofZ_exact. + pose proof (Z.pow_nonneg 2 prec). + lia. +Qed. + *) + +Lemma Rabs_relax: + forall b b' (INEQ : (b < b')%R) x, + (-b <= x <= b)%R -> (Rabs x < b')%R. +Proof. + intros. + apply Rabs_lt. + lra. +Qed. + +Theorem approx_inv_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_b : expr) (b : int) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : ((Int.unsigned b) > 0)%Z), + exists f : float, + eval_expr ge sp cmenv memenv le (approx_inv expr_b) (Vfloat f) /\ + is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int.unsigned b))) <= approx_inv_thresh)%R. +Proof. + intros. unfold approx_inv. + econstructor. constructor. + { repeat econstructor. + { eassumption. } + { reflexivity. } } + set (invb_d := (Float.of_single (ExtFloat32.inv (Float32.of_intu b)))). + set (b' := Int.unsigned b) in *. + pose proof (Int.unsigned_range b) as RANGE. + fold b' in RANGE. + change Int.modulus with 4294967296%Z in RANGE. + assert (0 <= b' <= Int64.max_unsigned)%Z as b'RANGE. + { change Int64.max_unsigned with 18446744073709551615%Z. + lia. } + assert (1 <= IZR b' <= 4294967295)%R as RANGE'. + { split. + { apply IZR_le. lia. } + apply IZR_le. lia. + } + cbn. + + set (b_d := (Float.of_longu (Int64.repr b'))) in *. + Local Transparent Float.of_longu. + unfold Float.of_longu in b_d. + + assert(SILLY : (- 2 ^ 24 <= 1 <= 2 ^ 24)%Z) by lia. + destruct (BofZ_exact 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C0E & C0F & _). + clear SILLY. + + assert(SILLY : (- 2 ^ 53 <= 1 <= 2 ^ 53)%Z) by lia. + destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C9E & C9F & _). + clear SILLY. + + pose proof (BofZ_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C1. + rewrite Rlt_bool_true in C1; cycle 1. + { clear C1. cbn. + eapply (Rabs_relax (IZR 4294967296)). + lra. + gappa. + } + rewrite Zlt_bool_false in C1 by lia. + destruct C1 as (C1E & C1F & _). + + Local Transparent Float32.of_intu Float32.of_int Float32.div. + unfold ExtFloat32.inv, ExtFloat32.one, Float32.of_intu, Float32.of_int, Float32.div in invb_d. + fold b' in invb_d. + change (Int.signed (Int.repr 1%Z)) with 1%Z in invb_d. + pose proof (Bdiv_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float32.binop_nan mode_NE + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1) + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b')) as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. rewrite C1E. + apply (Rabs_relax (bpow radix2 10)). + { cbn; lra. } + unfold F2R. cbn. unfold F2R. cbn. + gappa. + } + assert (B2R 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') <> 0%R) as NONZ. + { clear C2. + rewrite C1E. + cbn. + assert (1 <= round radix2 (FLT_exp (-149) 24) ZnearestE (IZR b'))%R by gappa. + lra. + } + destruct (C2 NONZ) as (C2E & C2F & _). + clear C2 NONZ. + Local Transparent Float.of_single. + unfold Float.of_single in invb_d. + pose proof (Bconv_correct 24 128 53 1024 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) Float.of_single_nan mode_NE + (Bdiv 24 128 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) Float32.binop_nan mode_NE + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) 1) + (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) b'))) as C3. + fold invb_d in C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + rewrite C2E. + rewrite C1E. + rewrite C0E. + apply (Rabs_relax (bpow radix2 10)). + { apply bpow_lt; lia. } + cbn. + gappa. + } + change (is_finite 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1)) with true in C2F. + destruct (C3 C2F) as (C3E & C3F & _). + clear C3. + unfold Float.fma. + assert (is_finite _ _ (Float.neg invb_d) = true) as invb_d_F. + { Local Transparent Float.neg. + unfold Float.neg. + rewrite is_finite_Bopp. + assumption. + } + + assert(SILLY : (- 2 ^ 53 <= b' <= 2 ^ 53)%Z) by lia. + destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b' SILLY) as (C4E & C4F & _). + clear SILLY. + + assert (is_finite 53 1024 b_d = true) as b_d_F. + { unfold b_d. + rewrite Int64.unsigned_repr by lia. + assumption. + } + + assert (is_finite 53 1024 ExtFloat.one = true) as one_F by reflexivity. + + pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) + (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE + (Float.neg invb_d) b_d ExtFloat.one invb_d_F b_d_F one_F) as C5. + + rewrite Rlt_bool_true in C5; cycle 1. + { clear C5. + unfold Float.neg. + rewrite B2R_Bopp. + rewrite C3E. + rewrite C2E. + rewrite C0E. + rewrite C1E. + unfold ExtFloat.one. + change (Float.of_int (Int.repr 1)) with (BofZ 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1). + rewrite C9E. + unfold b_d. + rewrite Int64.unsigned_repr by lia. + rewrite C4E. + apply (Rabs_relax (bpow radix2 10)). + { apply bpow_lt; lia. } + cbn. + gappa. + } + destruct C5 as (C5E & C5F & _). + + pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE + (Bfma 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE + (Float.neg invb_d) b_d ExtFloat.one) invb_d invb_d C5F C3F C3F) as C6. + rewrite Rlt_bool_true in C6; cycle 1. + { clear C6. + rewrite C3E. + rewrite C2E. + rewrite C1E. + rewrite C0E. + rewrite C5E. + unfold Float.neg. + rewrite B2R_Bopp. + rewrite C3E. + rewrite C2E. + rewrite C0E. + rewrite C1E. + unfold b_d. + rewrite Int64.unsigned_repr by lia. + rewrite C4E. + unfold ExtFloat.one. + change (Float.of_int (Int.repr 1)) with (BofZ 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1). + rewrite C9E. + apply (Rabs_relax (bpow radix2 10)). + { apply bpow_lt; lia. } + cbn. + gappa. + } + destruct C6 as (C6E & C6F & _). + split. + { exact C6F. } + rewrite C6E. + rewrite C5E. + rewrite C3E. + rewrite C2E. + rewrite C1E. + rewrite C0E. + unfold Float.neg. + rewrite B2R_Bopp. + unfold ExtFloat.one. + Local Transparent Float.of_int. + unfold Float.of_int. + rewrite (Int.signed_repr 1) by (cbn ; lia). + rewrite C9E. + rewrite C3E. + rewrite C2E. + rewrite C0E. + rewrite C1E. + unfold b_d. + rewrite Int64.unsigned_repr by lia. + rewrite C4E. + cbn. + set (rd := round radix2 (FLT_exp (-1074) 53) ZnearestE) in *. + set (rs := round radix2 (FLT_exp (-149) 24) ZnearestE) in *. + set (bi := IZR b') in *. + set (invb0 := rd (rs (1/ rs bi))%R) in *. + set (alpha := (- invb0 * bi + 1)%R) in *. + set (alpha' := ((1/bi - rd (rs (1/ rs bi)))/(1/bi))%R) in *. + assert (alpha = alpha')%R as expand_alpha. + { + unfold alpha, alpha', invb0. + field. + lra. + } + assert(-1/2097152 <= alpha' <= 1/2097152)%R as alpha_BOUND. + { unfold alpha', rd, rs. + gappa. + } + set (delta := (rd (rd alpha * invb0 + invb0) - (alpha * invb0 + invb0))%R). + assert(-1/1125899906842624 <= delta <= 1/1125899906842624)%R as delta_BOUND. + { unfold delta, invb0. rewrite expand_alpha. unfold rd, rs. + gappa. + } + replace (rd (rd alpha * invb0 + invb0) - 1/bi)%R with + (delta + ((alpha * invb0 + invb0)-1/bi))%R by (unfold delta; ring). + replace (alpha * invb0 + invb0 - 1 / bi)%R with (alpha * (invb0 - 1/bi) + (alpha * (1/bi) + invb0 - 1 / bi))%R by ring. + replace (alpha * (1 / bi) + invb0 - 1 / bi)%R with 0%R; cycle 1. + { unfold alpha. + field. + lra. + } + rewrite expand_alpha. + unfold invb0, rd, rs, approx_inv_thresh. + apply Rabs_le. + gappa. +Qed. + +Definition fp_divu32 a b := + let a_var := Eletvar (1%nat) in + let b_var := Eletvar (0%nat) in + let a_d := Eop Ofloatoflongu ((Eop Ocast32unsigned (a_var ::: Enil)) ::: Enil) in + let qr := Eop Olonguoffloat_ne ((Eop Omulf (a_d:::(approx_inv b_var):::Enil)):::Enil) in + let qr_var := Eletvar 0%nat in + let rem := Eop Omsubl ((Eop Ocast32unsigned ((Eletvar (2%nat)):::Enil))::: + qr_var ::: + (Eop Ocast32unsigned ((Eletvar (1%nat)):::Enil)):::Enil) in + let qr_m1 := Eop (Oaddlimm (Int64.repr (-1)%Z)) (qr_var:::Enil) in + let cases := Eop (Osel (Ccompl0 Clt) Tlong) + (qr_m1 ::: qr_var ::: rem ::: Enil) in (* (Elet qr cases) *) + Eop Olowlong ((Elet a (Elet (lift b) (Elet qr cases))) ::: Enil). + +Open Scope Z. + +Definition div_approx_reals (a b : Z) (x : R) := + let q:=ZnearestE x in + let r:=a-q*b in + if r Zfloor x = (y-1)%Z \/ Zfloor x = y. +Proof. + intros x y BALL. + apply Rabs_lt_inv in BALL. + case (Rcompare_spec x (IZR y)); intro CMP. + - left. apply Zfloor_imp. + ring_simplify (y-1+1). + rewrite minus_IZR. lra. + - subst. + rewrite Zfloor_IZR. right. reflexivity. + - right. apply Zfloor_imp. + rewrite plus_IZR. lra. +Qed. + +Theorem div_approx_reals_correct: + forall a b : Z, forall x : R, + b > 0 -> + (Rabs (x - IZR a/ IZR b) < 1/2)%R -> + div_approx_reals a b x = (a/b)%Z. +Proof. + intros a b x bPOS GAP. + assert (0 < IZR b)%R by (apply IZR_lt ; lia). + unfold div_approx_reals. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) x) as NEAR. + assert (Rabs (IZR (ZnearestE x) - IZR a/ IZR b) < 1)%R as BALL. + { pose proof (Rabs_triang (IZR (ZnearestE x) - x) + (x - IZR a/ IZR b)) as TRI. + ring_simplify (IZR (ZnearestE x) - x + (x - IZR a / IZR b))%R in TRI. + lra. + } + clear GAP NEAR. + rewrite Rabs_minus_sym in BALL. + pose proof (floor_ball1 _ _ BALL) as FLOOR. + clear BALL. + rewrite Zfloor_div in FLOOR by lia. + pose proof (Z_div_mod_eq_full a b) as DIV_MOD. + assert (0 < b) as bPOS' by lia. + pose proof (Z.mod_pos_bound a b bPOS') as MOD_BOUNDS. + case Z.ltb_spec; intro; destruct FLOOR; lia. +Qed. + +Opaque approx_inv. + +Theorem Znearest_le + : forall choice (x y : R), (x <= y)%R -> Znearest choice x <= Znearest choice y. +Proof. + intros. + destruct (Z_le_gt_dec (Znearest choice x) (Znearest choice y)) as [LE | GT]. + assumption. + exfalso. + assert (1 <= IZR (Znearest choice x) - IZR(Znearest choice y))%R as GAP. + { rewrite <- minus_IZR. + apply IZR_le. + lia. + } + pose proof (Znearest_imp2 choice x) as Rx. + pose proof (Znearest_imp2 choice y) as Ry. + apply Rabs_def2b in Rx. + apply Rabs_def2b in Ry. + assert (x = y) by lra. + subst y. + lia. +Qed. + +Theorem fp_divu32_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : (Int.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_divu32 expr_a expr_b) + (Vint (Int.divu a b)). +Proof. + intros. + assert (eval_expr ge sp cmenv memenv (Vint b :: Vint a :: le) + (Eletvar 0) (Vint b)) as EVAL_b'. + { constructor. reflexivity. } + destruct (approx_inv_correct ge sp cmenv memenv (Vint b :: Vint a :: le) (Eletvar 0) b EVAL_b' b_nz) as (inv_b & inv_b_eval & inv_b_finite & inv_b_correct). + unfold fp_divu32. + repeat econstructor. + { exact EVAL_a. } + { apply eval_lift. exact EVAL_b. } + exact inv_b_eval. + cbn. f_equal. + rewrite <- Float.of_intu_of_longu. + unfold Float.to_longu_ne. + rewrite ZofB_ne_range_correct by lia. + set (prod := (Float.mul (Float.of_intu a) inv_b)). + pose proof (Int.unsigned_range a) as a_range. + pose proof (Int.unsigned_range b) as b_range. + change Int.modulus with 4294967296 in a_range. + change Int.modulus with 4294967296 in b_range. + set (prod' := (B2R _ _ prod)). + set (prod_r := ZnearestE prod') in *. + + Local Transparent Float.mul Float.of_intu. + unfold Float.mul, Float.of_intu in prod. + set (a' := Int.unsigned a) in *. + set (b' := Int.unsigned b) in *. + assert (IZR_a' : (0 <= IZR a' <= 4294967295)%R). + { split; apply IZR_le; lia. } + assert (IZR_b' : (1 <= IZR b' <= 4294967295)%R). + { split; apply IZR_le; lia. } + assert (SILLY : -2^53 <= a' <= 2^53). + { cbn; lia. } + destruct (BofZ_exact 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a' SILLY) as (C0E & C0F & _). + clear SILLY. + pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE + (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt)a') inv_b) as C1. + set (inv_b_r := B2R 53 1024 inv_b) in *. + assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R). + { split; unfold Rdiv. + - apply Rmult_le_compat_l. lra. + apply Rinv_le. apply IZR_lt. lia. + apply IZR_le. lia. + - replace 1%R with (1 / 1)%R at 2 by field. + apply Rmult_le_compat_l. lra. + apply Rinv_le. apply IZR_lt. lia. + apply IZR_le. lia. + } + apply Rabs_def2b in inv_b_correct. + rewrite Rlt_bool_true in C1; cycle 1. + { clear C1. + rewrite C0E. + apply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. + set (delta := (inv_b_r - 1 / IZR b')%R) in *. + unfold approx_inv_thresh in inv_b_correct. + cbn. + assert (b'_RANGE : (1 <= IZR b' <= 4294967295)%R). + { split; apply IZR_le; lia. + } + assert (a'_RANGE : (0 <= IZR a' <= 4294967295)%R). + { split; apply IZR_le; lia. + } + gappa. + } + rewrite C0F in C1. + cbn in C1. + rewrite C0E in C1. + rewrite inv_b_finite in C1. + fold prod in C1. + fold prod' in C1. + destruct C1 as (C1E & C1F & _). + rewrite C1F. cbn. + + assert(prod'_range : (0 <= prod' <= 17179869181/4)%R). + { + rewrite C1E. + replace inv_b_r with (1/IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. + assert (a'_RANGE : (0 <= IZR a' <= 4294967295)%R). + { split; apply IZR_le; lia. + } + unfold approx_inv_thresh in inv_b_correct. + set (true_inv := (1 / IZR b')%R) in *. + set (delta := (inv_b_r - true_inv)%R) in *. + gappa. + } + + assert(0 <= prod_r <= 4294967295) as prod_r_range. + { unfold prod_r. + rewrite <- (Znearest_IZR (fun x => negb (Z.even x)) 0). + replace 4294967295 with (ZnearestE (17179869181 / 4)%R); cycle 1. + { apply Znearest_imp. + apply Rabs_lt. + lra. + } + split; apply Znearest_le; lra. + } + + replace (_ && _ ) with true; cycle 1. + { + symmetry. + rewrite andb_true_iff. + split; apply Zle_imp_le_bool; lia. + } + cbn. + f_equal. + unfold Int.divu. + assert(Rabs + (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. + + assert ( (Rabs (B2R 53 1024 prod - IZR (Int.unsigned a) / IZR (Int.unsigned b)) < 1 / 2)%R) as NEAR. + { + unfold prod. + pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a') inv_b) as C2. + rewrite C0E in C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. reflexivity. } + cbn. + fold inv_b_r. + replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. + set (delta := (inv_b_r - 1 / IZR b')%R) in *. + unfold approx_inv_thresh in *. + gappa. + } + destruct C2 as (C2E & C2F & _). + rewrite C2E. + fold inv_b_r a' b'. + replace ((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - + (IZR a' / IZR b'))%R with + (((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - + (IZR a' * inv_b_r)) + + (IZR a' * (inv_b_r - 1 / IZR b')))%R by (field ; lra). + set(delta := (inv_b_r - 1 / IZR b')%R) in *. + cbn. + unfold approx_inv_thresh in *. + assert (Rabs(IZR a' * delta) <= 3/8)%R as R2. + { apply Rabs_le. + split; nra. + } + rewrite <- C1E. + rewrite <- C1E in R1. + pose proof (Rabs_triang (prod' - IZR a' * inv_b_r) (IZR a' * delta))%R. + lra. + } + pose proof (div_approx_reals_correct (Int.unsigned a) (Int.unsigned b) prod' b_nz NEAR) as DIV_CORRECT. + rewrite <- DIV_CORRECT. + + unfold div_approx_reals in *. + fold a' b' prod' prod_r. + unfold Int64.mul. + rewrite Int64.unsigned_repr by (cbn; lia). + rewrite Int64.unsigned_repr by (cbn; lia). + unfold Int64.sub. + rewrite Int64.unsigned_repr by (cbn; lia). + rewrite Int64.unsigned_repr by (cbn; nia). + assert (FMA_RANGE : Int64.min_signed <= a' - prod_r * b' <= Int64.max_signed). + { cbn. + unfold prod_r. + rewrite <- C1E in R1. + assert (IZR (-9223372036854775808) <= IZR (a' - ZnearestE prod' * b') <= IZR 9223372036854775807)%R. + 2: split; apply le_IZR; lra. + rewrite minus_IZR. + rewrite mult_IZR. + replace (IZR (ZnearestE prod')) with (prod' + (IZR (ZnearestE prod') - prod'))%R by ring. + pose proof (Znearest_imp2 (fun x => negb (Z.even x)) prod') as R2. + set (delta1 := (IZR (ZnearestE prod') - prod')%R) in *. + replace prod' with ((prod' - IZR a' * inv_b_r) + IZR a' * (inv_b_r - 1 / IZR b') + + IZR a' / IZR b')%R by (field; lra). + set (delta2 := (inv_b_r - 1 / IZR b')%R) in *. + set (delta3 := (prod' - IZR a' * inv_b_r)%R) in *. + replace (IZR a' - (delta3 + IZR a' * delta2 + IZR a' / IZR b' + delta1) * IZR b')%R with + (- (delta3 + IZR a' * delta2 + delta1) * IZR b')%R by (field; lra). + unfold approx_inv_thresh in *. + assert (Rabs(IZR a' * delta2) <= 1/4)%R as R4. + { apply Rabs_le. + split; + nra. } + set (delta4 := (IZR a' * delta2)%R) in *. + apply Rabs_def2b in R1. + apply Rabs_def2b in R2. + apply Rabs_def2b in R4. + split; nra. + } + fold a' b' prod_r in DIV_CORRECT. + + pose proof (Zlt_cases (a' - prod_r * b') 0) as CMP. + destruct (Z.ltb (a' - prod_r * b') 0). + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. + { unfold Int64.lt. + change (Int64.signed Int64.zero) with 0. + rewrite Int64.signed_repr by exact FMA_RANGE. + rewrite zlt_true by lia. + reflexivity. + } + cbn. + f_equal. + rewrite Int64.add_signed. + rewrite (Int64.signed_repr prod_r) by (cbn ; lia). + rewrite Int64.signed_repr by (cbn ; lia). + unfold Int64.loword. + rewrite Int64.unsigned_repr. reflexivity. + split. + 2: cbn; lia. + replace (prod_r + (-1)) with (prod_r - 1) by lia. + rewrite DIV_CORRECT. + apply Z.div_pos; lia. + + - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. + { unfold Int64.lt. + change (Int64.signed Int64.zero) with 0. + rewrite Int64.signed_repr by exact FMA_RANGE. + rewrite zlt_false by lia. + reflexivity. + } + cbn. + unfold Int64.loword. + rewrite Int64.unsigned_repr by (cbn; lia). + reflexivity. +Qed. + + -- cgit From 6ee0923290cca4782e6c863d752fb74c2c01df4f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 14:28:00 +0100 Subject: fpdivu --- kvx/Builtins1.v | 10 +- kvx/CBuiltins.ml | 4 +- kvx/FPDivision.v | 615 ---------------------------------------------------- kvx/SelectOp.vp | 6 + kvx/SelectOpproof.v | 27 ++- 5 files changed, 44 insertions(+), 618 deletions(-) delete mode 100644 kvx/FPDivision.v diff --git a/kvx/Builtins1.v b/kvx/Builtins1.v index de0d9885..163dcbd8 100644 --- a/kvx/Builtins1.v +++ b/kvx/Builtins1.v @@ -27,7 +27,8 @@ Inductive platform_builtin : Type := | BI_fma | BI_fmaf | BI_lround_ne -| BI_luround_ne. +| BI_luround_ne +| BI_fp_udiv32. Local Open Scope string_scope. @@ -40,6 +41,7 @@ Definition platform_builtin_table : list (string * platform_builtin) := :: ("__builtin_fmaf", BI_fmaf) :: ("__builtin_lround_ne", BI_lround_ne) :: ("__builtin_luround_ne", BI_luround_ne) + :: ("__builtin_fp_udiv32", BI_fp_udiv32) :: nil. Definition platform_builtin_sig (b: platform_builtin) : signature := @@ -54,6 +56,8 @@ Definition platform_builtin_sig (b: platform_builtin) : signature := mksignature (Tsingle :: Tsingle :: Tsingle :: nil) Tsingle cc_default | BI_lround_ne | BI_luround_ne => mksignature (Tfloat :: nil) Tlong cc_default + | BI_fp_udiv32 => + mksignature (Tint :: Tint :: nil) Tint cc_default end. Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (platform_builtin_sig b)) := @@ -66,4 +70,8 @@ Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (pl | BI_fmaf => mkbuiltin_n3t Tsingle Tsingle Tsingle Tsingle Float32.fma | BI_lround_ne => mkbuiltin_n1p Tfloat Tlong Float.to_long_ne | BI_luround_ne => mkbuiltin_n1p Tfloat Tlong Float.to_longu_ne + | BI_fp_udiv32 => mkbuiltin_n2p Tint Tint Tint + (fun n1 n2 => if Int.eq n2 Int.zero + then None + else Some (Int.divu n1 n2)) end. diff --git a/kvx/CBuiltins.ml b/kvx/CBuiltins.ml index f2b7b09e..fad86d3c 100644 --- a/kvx/CBuiltins.ml +++ b/kvx/CBuiltins.ml @@ -136,7 +136,9 @@ let builtins = { "__builtin_lround_ne", (TInt(ILong, []), [TFloat(FDouble, [])], false); "__builtin_luround_ne", - (TInt(IULong, []), [TFloat(FDouble, [])], false); + (TInt(IULong, []), [TFloat(FDouble, [])], false); + "__builtin_fp_udiv32", + (TInt(IUInt, []), [TInt(IUInt, []); TInt(IUInt, [])], false); ] } diff --git a/kvx/FPDivision.v b/kvx/FPDivision.v deleted file mode 100644 index 0ba9a39f..00000000 --- a/kvx/FPDivision.v +++ /dev/null @@ -1,615 +0,0 @@ -From Flocq Require Import Core Digits Operations Round Bracket Sterbenz - Binary Round_odd Bits. -Require Archi. -Require Import Coqlib. -Require Import Compopts. -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Op. -Require Import CminorSel. -Require Import OpHelpers. -Require Import ExtFloats. -Require Import DecBoolOps. -Require Import Chunks. -Require Import Builtins. -Require Import Values Globalenvs. -Require Compopts. -Require Import Psatz. -Require Import IEEE754_extra. - -From Gappa Require Import Gappa_tactic. - -Local Open Scope cminorsel_scope. - -Local Open Scope string_scope. -Local Open Scope error_monad_scope. - -Definition approx_inv b := - let invb_s := Eop Oinvfs ((Eop Osingleofintu ((Eletvar 0%nat):::Enil)):::Enil) in - let invb_d := Eop Ofloatofsingle (invb_s ::: Enil) in - let b_d := Eop Ofloatoflongu ((Eop Ocast32unsigned ((Eletvar 1%nat):::Enil)):::Enil) in - let invb_d_var := Eletvar (0%nat) in - let one := Eop (Ofloatconst ExtFloat.one) Enil in - let alpha := Eop Ofmsubf (one ::: invb_d_var ::: b_d ::: Enil) in - let x := Eop Ofmaddf (invb_d_var ::: alpha ::: invb_d_var ::: Enil) in - Elet b (Elet invb_d x). - -Definition approx_inv_thresh := (1/17179869184)%R. - -(* -Lemma BofZ_exact_zero: - forall (prec emax : Z) (prec_gt_0_ : Prec_gt_0 prec) - (Hmax : (prec < emax)%Z), - B2R (BofZ prec emax 0%Z (Hmax := Hmax)) = 0%R /\ - is_finite (BofZ prec emax 0%Z (Hmax := Hmax)) = true /\ - Bsign prec emax (BofZ prec emax 0%Z (Hmax := Hmax)) = false. -Proof. - intros. - apply BofZ_exact. - pose proof (Z.pow_nonneg 2 prec). - lia. -Qed. - *) - -Lemma Rabs_relax: - forall b b' (INEQ : (b < b')%R) x, - (-b <= x <= b)%R -> (Rabs x < b')%R. -Proof. - intros. - apply Rabs_lt. - lra. -Qed. - -Theorem approx_inv_correct : - forall (ge : genv) (sp: val) cmenv memenv - (le : letenv) (expr_b : expr) (b : int) - (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) - (b_nz : ((Int.unsigned b) > 0)%Z), - exists f : float, - eval_expr ge sp cmenv memenv le (approx_inv expr_b) (Vfloat f) /\ - is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int.unsigned b))) <= approx_inv_thresh)%R. -Proof. - intros. unfold approx_inv. - econstructor. constructor. - { repeat econstructor. - { eassumption. } - { reflexivity. } } - set (invb_d := (Float.of_single (ExtFloat32.inv (Float32.of_intu b)))). - set (b' := Int.unsigned b) in *. - pose proof (Int.unsigned_range b) as RANGE. - fold b' in RANGE. - change Int.modulus with 4294967296%Z in RANGE. - assert (0 <= b' <= Int64.max_unsigned)%Z as b'RANGE. - { change Int64.max_unsigned with 18446744073709551615%Z. - lia. } - assert (1 <= IZR b' <= 4294967295)%R as RANGE'. - { split. - { apply IZR_le. lia. } - apply IZR_le. lia. - } - cbn. - - set (b_d := (Float.of_longu (Int64.repr b'))) in *. - Local Transparent Float.of_longu. - unfold Float.of_longu in b_d. - - assert(SILLY : (- 2 ^ 24 <= 1 <= 2 ^ 24)%Z) by lia. - destruct (BofZ_exact 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C0E & C0F & _). - clear SILLY. - - assert(SILLY : (- 2 ^ 53 <= 1 <= 2 ^ 53)%Z) by lia. - destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1 SILLY) as (C9E & C9F & _). - clear SILLY. - - pose proof (BofZ_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') as C1. - rewrite Rlt_bool_true in C1; cycle 1. - { clear C1. cbn. - eapply (Rabs_relax (IZR 4294967296)). - lra. - gappa. - } - rewrite Zlt_bool_false in C1 by lia. - destruct C1 as (C1E & C1F & _). - - Local Transparent Float32.of_intu Float32.of_int Float32.div. - unfold ExtFloat32.inv, ExtFloat32.one, Float32.of_intu, Float32.of_int, Float32.div in invb_d. - fold b' in invb_d. - Check BofZ. - change (Int.signed (Int.repr 1%Z)) with 1%Z in invb_d. - pose proof (Bdiv_correct 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float32.binop_nan mode_NE - (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1) - (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b')) as C2. - rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. rewrite C1E. - apply (Rabs_relax (bpow radix2 10)). - { cbn; lra. } - unfold F2R. cbn. unfold F2R. cbn. - gappa. - } - assert (B2R 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b') <> 0%R) as NONZ. - { clear C2. - rewrite C1E. - cbn. - assert (1 <= round radix2 (FLT_exp (-149) 24) ZnearestE (IZR b'))%R by gappa. - lra. - } - destruct (C2 NONZ) as (C2E & C2F & _). - clear C2 NONZ. - Local Transparent Float.of_single. - unfold Float.of_single in invb_d. - pose proof (Bconv_correct 24 128 53 1024 (@eq_refl Datatypes.comparison Lt) - (@eq_refl Datatypes.comparison Lt) Float.of_single_nan mode_NE - (Bdiv 24 128 (@eq_refl Datatypes.comparison Lt) - (@eq_refl Datatypes.comparison Lt) Float32.binop_nan mode_NE - (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) - (@eq_refl Datatypes.comparison Lt) 1) - (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) - (@eq_refl Datatypes.comparison Lt) b'))) as C3. - fold invb_d in C3. - rewrite Rlt_bool_true in C3; cycle 1. - { clear C3. - rewrite C2E. - rewrite C1E. - rewrite C0E. - apply (Rabs_relax (bpow radix2 10)). - { apply bpow_lt; lia. } - cbn. - gappa. - } - change (is_finite 24 128 (BofZ 24 128 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1)) with true in C2F. - destruct (C3 C2F) as (C3E & C3F & _). - clear C3. - unfold Float.fma. - assert (is_finite _ _ (Float.neg invb_d) = true) as invb_d_F. - { Local Transparent Float.neg. - unfold Float.neg. - rewrite is_finite_Bopp. - assumption. - } - - assert(SILLY : (- 2 ^ 53 <= b' <= 2 ^ 53)%Z) by lia. - destruct (BofZ_exact 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) b' SILLY) as (C4E & C4F & _). - clear SILLY. - - assert (is_finite 53 1024 b_d = true) as b_d_F. - { unfold b_d. - rewrite Int64.unsigned_repr by lia. - assumption. - } - - assert (is_finite 53 1024 ExtFloat.one = true) as one_F by reflexivity. - - pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) - (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE - (Float.neg invb_d) b_d ExtFloat.one invb_d_F b_d_F one_F) as C5. - - rewrite Rlt_bool_true in C5; cycle 1. - { clear C5. - unfold Float.neg. - rewrite B2R_Bopp. - rewrite C3E. - rewrite C2E. - rewrite C0E. - rewrite C1E. - unfold ExtFloat.one. - change (Float.of_int (Int.repr 1)) with (BofZ 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1). - rewrite C9E. - unfold b_d. - rewrite Int64.unsigned_repr by lia. - rewrite C4E. - apply (Rabs_relax (bpow radix2 10)). - { apply bpow_lt; lia. } - cbn. - gappa. - } - destruct C5 as (C5E & C5F & _). - - pose proof (Bfma_correct 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE - (Bfma 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) Float.fma_nan mode_NE - (Float.neg invb_d) b_d ExtFloat.one) invb_d invb_d C5F C3F C3F) as C6. - rewrite Rlt_bool_true in C6; cycle 1. - { clear C6. - rewrite C3E. - rewrite C2E. - rewrite C1E. - rewrite C0E. - rewrite C5E. - unfold Float.neg. - rewrite B2R_Bopp. - rewrite C3E. - rewrite C2E. - rewrite C0E. - rewrite C1E. - unfold b_d. - rewrite Int64.unsigned_repr by lia. - rewrite C4E. - unfold ExtFloat.one. - change (Float.of_int (Int.repr 1)) with (BofZ 53 1024 (@eq_refl Datatypes.comparison Lt) (@eq_refl Datatypes.comparison Lt) 1). - rewrite C9E. - apply (Rabs_relax (bpow radix2 10)). - { apply bpow_lt; lia. } - cbn. - gappa. - } - destruct C6 as (C6E & C6F & _). - split. - { exact C6F. } - rewrite C6E. - rewrite C5E. - rewrite C3E. - rewrite C2E. - rewrite C1E. - rewrite C0E. - unfold Float.neg. - rewrite B2R_Bopp. - unfold ExtFloat.one. - Local Transparent Float.of_int. - unfold Float.of_int. - rewrite (Int.signed_repr 1) by (cbn ; lia). - rewrite C9E. - rewrite C3E. - rewrite C2E. - rewrite C0E. - rewrite C1E. - unfold b_d. - rewrite Int64.unsigned_repr by lia. - rewrite C4E. - cbn. - set (rd := round radix2 (FLT_exp (-1074) 53) ZnearestE) in *. - set (rs := round radix2 (FLT_exp (-149) 24) ZnearestE) in *. - set (bi := IZR b') in *. - set (invb0 := rd (rs (1/ rs bi))%R) in *. - set (alpha := (- invb0 * bi + 1)%R) in *. - set (alpha' := ((1/bi - rd (rs (1/ rs bi)))/(1/bi))%R) in *. - assert (alpha = alpha')%R as expand_alpha. - { - unfold alpha, alpha', invb0. - field. - lra. - } - assert(-1/2097152 <= alpha' <= 1/2097152)%R as alpha_BOUND. - { unfold alpha', rd, rs. - gappa. - } - set (delta := (rd (rd alpha * invb0 + invb0) - (alpha * invb0 + invb0))%R). - assert(-1/1125899906842624 <= delta <= 1/1125899906842624)%R as delta_BOUND. - { unfold delta, invb0. rewrite expand_alpha. unfold rd, rs. - gappa. - } - replace (rd (rd alpha * invb0 + invb0) - 1/bi)%R with - (delta + ((alpha * invb0 + invb0)-1/bi))%R by (unfold delta; ring). - replace (alpha * invb0 + invb0 - 1 / bi)%R with (alpha * (invb0 - 1/bi) + (alpha * (1/bi) + invb0 - 1 / bi))%R by ring. - replace (alpha * (1 / bi) + invb0 - 1 / bi)%R with 0%R; cycle 1. - { unfold alpha. - field. - lra. - } - rewrite expand_alpha. - unfold invb0, rd, rs, approx_inv_thresh. - apply Rabs_le. - gappa. -Qed. - -Definition approx_div a b := - let a_var := Eletvar (1%nat) in - let b_var := Eletvar (0%nat) in - let a_d := Eop Ofloatoflongu ((Eop Ocast32unsigned (a_var ::: Enil)) ::: Enil) in - let qr := Eop Olonguoffloat_ne ((Eop Omulf (a_d:::(approx_inv b_var):::Enil)):::Enil) in - let qr_var := Eletvar 0%nat in - let rem := Eop Omsubl ((Eop Ocast32unsigned ((Eletvar (2%nat)):::Enil))::: - qr_var ::: - (Eop Ocast32unsigned ((Eletvar (1%nat)):::Enil)):::Enil) in - let qr_m1 := Eop (Oaddlimm (Int64.repr (-1)%Z)) (qr_var:::Enil) in - let cases := Eop (Osel (Ccompl0 Clt) Tlong) - (qr_m1 ::: qr_var ::: rem ::: Enil) in (* (Elet qr cases) *) - Eop Olowlong ((Elet a (Elet (lift b) (Elet qr cases))) ::: Enil). - -Open Scope Z. - -Definition div_approx_reals (a b : Z) (x : R) := - let q:=ZnearestE x in - let r:=a-q*b in - if r Zfloor x = (y-1)%Z \/ Zfloor x = y. -Proof. - intros x y BALL. - apply Rabs_lt_inv in BALL. - case (Rcompare_spec x (IZR y)); intro CMP. - - left. apply Zfloor_imp. - ring_simplify (y-1+1). - rewrite minus_IZR. lra. - - subst. - rewrite Zfloor_IZR. right. reflexivity. - - right. apply Zfloor_imp. - rewrite plus_IZR. lra. -Qed. - -Theorem div_approx_reals_correct: - forall a b : Z, forall x : R, - b > 0 -> - (Rabs (x - IZR a/ IZR b) < 1/2)%R -> - div_approx_reals a b x = (a/b)%Z. -Proof. - intros a b x bPOS GAP. - assert (0 < IZR b)%R by (apply IZR_lt ; lia). - unfold div_approx_reals. - pose proof (Znearest_imp2 (fun x => negb (Z.even x)) x) as NEAR. - assert (Rabs (IZR (ZnearestE x) - IZR a/ IZR b) < 1)%R as BALL. - { pose proof (Rabs_triang (IZR (ZnearestE x) - x) - (x - IZR a/ IZR b)) as TRI. - ring_simplify (IZR (ZnearestE x) - x + (x - IZR a / IZR b))%R in TRI. - lra. - } - clear GAP NEAR. - rewrite Rabs_minus_sym in BALL. - pose proof (floor_ball1 _ _ BALL) as FLOOR. - clear BALL. - rewrite Zfloor_div in FLOOR by lia. - pose proof (Z_div_mod_eq_full a b) as DIV_MOD. - assert (0 < b) as bPOS' by lia. - pose proof (Z.mod_pos_bound a b bPOS') as MOD_BOUNDS. - case Z.ltb_spec; intro; destruct FLOOR; lia. -Qed. - -Opaque approx_inv. - -Search ({_ <= _} + {_ > _})%Z. -Theorem Znearest_le - : forall choice (x y : R), (x <= y)%R -> Znearest choice x <= Znearest choice y. -Proof. - intros. - destruct (Z_le_gt_dec (Znearest choice x) (Znearest choice y)) as [LE | GT]. - assumption. - exfalso. - assert (1 <= IZR (Znearest choice x) - IZR(Znearest choice y))%R as GAP. - { rewrite <- minus_IZR. - apply IZR_le. - lia. - } - pose proof (Znearest_imp2 choice x) as Rx. - pose proof (Znearest_imp2 choice y) as Ry. - apply Rabs_def2b in Rx. - apply Rabs_def2b in Ry. - assert (x = y) by lra. - subst y. - lia. -Qed. - -Theorem approx_div_correct : - forall (ge : genv) (sp: val) cmenv memenv - (le : letenv) (expr_a expr_b : expr) (a b : int) - (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) - (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) - (b_nz : (Int.unsigned b > 0)%Z), - eval_expr ge sp cmenv memenv le (approx_div expr_a expr_b) - (Vint (Int.divu a b)). -Proof. - intros. - assert (eval_expr ge sp cmenv memenv (Vint b :: Vint a :: le) - (Eletvar 0) (Vint b)) as EVAL_b'. - { constructor. reflexivity. } - destruct (approx_inv_correct ge sp cmenv memenv (Vint b :: Vint a :: le) (Eletvar 0) b EVAL_b' b_nz) as (inv_b & inv_b_eval & inv_b_finite & inv_b_correct). - unfold approx_div. - repeat econstructor. - { exact EVAL_a. } - { apply eval_lift. exact EVAL_b. } - exact inv_b_eval. - cbn. f_equal. - rewrite <- Float.of_intu_of_longu. - unfold Float.to_longu_ne. - rewrite ZofB_ne_range_correct by lia. - set (prod := (Float.mul (Float.of_intu a) inv_b)). - pose proof (Int.unsigned_range a) as a_range. - pose proof (Int.unsigned_range b) as b_range. - change Int.modulus with 4294967296 in a_range. - change Int.modulus with 4294967296 in b_range. - set (prod' := (B2R _ _ prod)). - set (prod_r := ZnearestE prod') in *. - - Local Transparent Float.mul Float.of_intu. - unfold Float.mul, Float.of_intu in prod. - set (a' := Int.unsigned a) in *. - set (b' := Int.unsigned b) in *. - assert (IZR_a' : (0 <= IZR a' <= 4294967295)%R). - { split; apply IZR_le; lia. } - assert (IZR_b' : (1 <= IZR b' <= 4294967295)%R). - { split; apply IZR_le; lia. } - assert (SILLY : -2^53 <= a' <= 2^53). - { cbn; lia. } - destruct (BofZ_exact 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a' SILLY) as (C0E & C0F & _). - clear SILLY. - pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE - (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt)a') inv_b) as C1. - set (inv_b_r := B2R 53 1024 inv_b) in *. - assert (INV_RANGE : (1/4294967296 <= 1/IZR b' <= 1)%R). - { split; unfold Rdiv. - - apply Rmult_le_compat_l. lra. - apply Rinv_le. apply IZR_lt. lia. - apply IZR_le. lia. - - replace 1%R with (1 / 1)%R at 2 by field. - apply Rmult_le_compat_l. lra. - apply Rinv_le. apply IZR_lt. lia. - apply IZR_le. lia. - } - apply Rabs_def2b in inv_b_correct. - rewrite Rlt_bool_true in C1; cycle 1. - { clear C1. - rewrite C0E. - apply (Rabs_relax (bpow radix2 64)). - { apply bpow_lt. lia. } - replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. - set (delta := (inv_b_r - 1 / IZR b')%R) in *. - unfold approx_inv_thresh in inv_b_correct. - cbn. - assert (b'_RANGE : (1 <= IZR b' <= 4294967295)%R). - { split; apply IZR_le; lia. - } - assert (a'_RANGE : (0 <= IZR a' <= 4294967295)%R). - { split; apply IZR_le; lia. - } - gappa. - } - rewrite C0F in C1. - cbn in C1. - rewrite C0E in C1. - rewrite inv_b_finite in C1. - fold prod in C1. - fold prod' in C1. - destruct C1 as (C1E & C1F & _). - rewrite C1F. cbn. - - assert(prod'_range : (0 <= prod' <= 17179869181/4)%R). - { - rewrite C1E. - replace inv_b_r with (1/IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. - assert (a'_RANGE : (0 <= IZR a' <= 4294967295)%R). - { split; apply IZR_le; lia. - } - unfold approx_inv_thresh in inv_b_correct. - set (true_inv := (1 / IZR b')%R) in *. - set (delta := (inv_b_r - true_inv)%R) in *. - gappa. - } - - assert(0 <= prod_r <= 4294967295) as prod_r_range. - { unfold prod_r. - rewrite <- (Znearest_IZR (fun x => negb (Z.even x)) 0). - replace 4294967295 with (ZnearestE (17179869181 / 4)%R); cycle 1. - { apply Znearest_imp. - apply Rabs_lt. - lra. - } - split; apply Znearest_le; lra. - } - - replace (_ && _ ) with true; cycle 1. - { - symmetry. - rewrite andb_true_iff. - split; apply Zle_imp_le_bool; lia. - } - cbn. - f_equal. - unfold Int.divu. - assert(Rabs - (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a' * inv_b_r) - (IZR a' * inv_b_r)) <= 1/512)%R as R1 by gappa. - - assert ( (Rabs (B2R 53 1024 prod - IZR (Int.unsigned a) / IZR (Int.unsigned b)) < 1 / 2)%R) as NEAR. - { - unfold prod. - pose proof (Bmult_correct 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) Float.binop_nan mode_NE (BofZ 53 1024 (@eq_refl _ Lt) (@eq_refl _ Lt) a') inv_b) as C2. - rewrite C0E in C2. - rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. - apply (Rabs_relax (bpow radix2 64)). - { apply bpow_lt. reflexivity. } - cbn. - fold inv_b_r. - replace inv_b_r with (1 / IZR b' + (inv_b_r - 1 / IZR b'))%R by ring. - set (delta := (inv_b_r - 1 / IZR b')%R) in *. - unfold approx_inv_thresh in *. - gappa. - } - destruct C2 as (C2E & C2F & _). - rewrite C2E. - fold inv_b_r a' b'. - replace ((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - - (IZR a' / IZR b'))%R with - (((round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) (IZR a' * inv_b_r)) - - (IZR a' * inv_b_r)) + - (IZR a' * (inv_b_r - 1 / IZR b')))%R by (field ; lra). - set(delta := (inv_b_r - 1 / IZR b')%R) in *. - cbn. - unfold approx_inv_thresh in *. - assert (Rabs(IZR a' * delta) <= 3/8)%R as R2. - { apply Rabs_le. - split; nra. - } - rewrite <- C1E. - rewrite <- C1E in R1. - pose proof (Rabs_triang (prod' - IZR a' * inv_b_r) (IZR a' * delta))%R. - lra. - } - pose proof (div_approx_reals_correct (Int.unsigned a) (Int.unsigned b) prod' b_nz NEAR) as DIV_CORRECT. - rewrite <- DIV_CORRECT. - - unfold div_approx_reals in *. - fold a' b' prod' prod_r. - unfold Int64.mul. - rewrite Int64.unsigned_repr by (cbn; lia). - rewrite Int64.unsigned_repr by (cbn; lia). - unfold Int64.sub. - rewrite Int64.unsigned_repr by (cbn; lia). - rewrite Int64.unsigned_repr by (cbn; nia). - assert (FMA_RANGE : Int64.min_signed <= a' - prod_r * b' <= Int64.max_signed). - { cbn. - unfold prod_r. - rewrite <- C1E in R1. - assert (IZR (-9223372036854775808) <= IZR (a' - ZnearestE prod' * b') <= IZR 9223372036854775807)%R. - 2: split; apply le_IZR; lra. - rewrite minus_IZR. - rewrite mult_IZR. - replace (IZR (ZnearestE prod')) with (prod' + (IZR (ZnearestE prod') - prod'))%R by ring. - pose proof (Znearest_imp2 (fun x => negb (Z.even x)) prod') as R2. - set (delta1 := (IZR (ZnearestE prod') - prod')%R) in *. - replace prod' with ((prod' - IZR a' * inv_b_r) + IZR a' * (inv_b_r - 1 / IZR b') - + IZR a' / IZR b')%R by (field; lra). - set (delta2 := (inv_b_r - 1 / IZR b')%R) in *. - set (delta3 := (prod' - IZR a' * inv_b_r)%R) in *. - replace (IZR a' - (delta3 + IZR a' * delta2 + IZR a' / IZR b' + delta1) * IZR b')%R with - (- (delta3 + IZR a' * delta2 + delta1) * IZR b')%R by (field; lra). - unfold approx_inv_thresh in *. - assert (Rabs(IZR a' * delta2) <= 1/4)%R as R4. - { apply Rabs_le. - split; - nra. } - set (delta4 := (IZR a' * delta2)%R) in *. - apply Rabs_def2b in R1. - apply Rabs_def2b in R2. - apply Rabs_def2b in R4. - split; nra. - } - fold a' b' prod_r in DIV_CORRECT. - - pose proof (Zlt_cases (a' - prod_r * b') 0) as CMP. - destruct (Z.ltb (a' - prod_r * b') 0). - - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with true; cycle 1. - { unfold Int64.lt. - change (Int64.signed Int64.zero) with 0. - rewrite Int64.signed_repr by exact FMA_RANGE. - rewrite zlt_true by lia. - reflexivity. - } - cbn. - f_equal. - rewrite Int64.add_signed. - rewrite (Int64.signed_repr prod_r) by (cbn ; lia). - rewrite Int64.signed_repr by (cbn ; lia). - unfold Int64.loword. - rewrite Int64.unsigned_repr. reflexivity. - split. - 2: cbn; lia. - replace (prod_r + (-1)) with (prod_r - 1) by lia. - rewrite DIV_CORRECT. - apply Z.div_pos; lia. - - - replace (Int64.lt (Int64.repr (a' - prod_r * b')) Int64.zero) with false; cycle 1. - { unfold Int64.lt. - change (Int64.signed Int64.zero) with 0. - rewrite Int64.signed_repr by exact FMA_RANGE. - rewrite zlt_false by lia. - reflexivity. - } - cbn. - unfold Int64.loword. - rewrite Int64.unsigned_repr by (cbn; lia). - reflexivity. -Qed. - - diff --git a/kvx/SelectOp.vp b/kvx/SelectOp.vp index f529907d..70941c48 100644 --- a/kvx/SelectOp.vp +++ b/kvx/SelectOp.vp @@ -742,6 +742,8 @@ Nondetfunction gen_fmaf args := | _ => None end. +Require FPDivision32. + Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr := match b with | BI_fmin => Some (Eop Ominf args) @@ -752,6 +754,10 @@ Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr | BI_fmaf => gen_fmaf args | BI_lround_ne => Some (Eop Olongoffloat_ne args) | BI_luround_ne => Some (Eop Olonguoffloat_ne args) + | BI_fp_udiv32 => (match args with + | a:::b:::Enil => Some (FPDivision32.fp_divu32 a b) + | _ => None + end) end. End SELECT. diff --git a/kvx/SelectOpproof.v b/kvx/SelectOpproof.v index a374ec54..658bf0b3 100644 --- a/kvx/SelectOpproof.v +++ b/kvx/SelectOpproof.v @@ -1883,6 +1883,9 @@ Proof. destruct v2; simpl; trivial. Qed. + +Require FPDivision32. + Theorem eval_platform_builtin: forall bf al a vl v le, platform_builtin bf al = Some a -> @@ -1903,7 +1906,29 @@ Proof. - cbn in *; destruct vl; trivial. destruct vl; trivial. destruct v0; try discriminate; - cbn; rewrite H0; reflexivity. + cbn; rewrite H0; reflexivity. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int.eq in EVAL. + change (Int.unsigned Int.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vint (Int.divu i i0)). split. + { + apply FPDivision32.fp_divu32_correct; auto. + pose proof (Int.unsigned_range i0). + lia. + } + inv EVAL. + constructor. Qed. End CMCONSTR. -- cgit From d308a98fcabbbd821c3b6f168a61b77497bdad54 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 17 Dec 2021 14:58:30 +0100 Subject: fix configure for extra files --- configure | 1 + 1 file changed, 1 insertion(+) diff --git a/configure b/configure index 3da00fb3..9cc6fbe0 100755 --- a/configure +++ b/configure @@ -822,6 +822,7 @@ BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v Asmblockprops.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ + FPDivision32.v \\ AbstractBasicBlocksDef.v SeqSimuTheory.v ImpSimuTest.v Parallelizability.v EOF fi -- cgit From ed48745ff6040f4e4ff97e04f0f6bdf214efe64f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 21 Dec 2021 23:13:52 +0100 Subject: fix configure for kvx --- configure | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/configure b/configure index 3da00fb3..e0b1ea7d 100755 --- a/configure +++ b/configure @@ -821,7 +821,8 @@ SIMU=kvx-cluster -- BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v Asmblockprops.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ - Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ + Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v\\ + FPDivision32.v ExtValues.v ExtFloats.v\\ AbstractBasicBlocksDef.v SeqSimuTheory.v ImpSimuTest.v Parallelizability.v EOF fi -- cgit From 2ad0d6e859cb90c0ae26652c71bf82e8f24ab49f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 21 Dec 2021 23:19:28 +0100 Subject: compile & install gappa for KVX --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index c8ccedb8..096d8959 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -276,7 +276,7 @@ build_kvx: - rm -rf download - eval `opam config env` - opam update - - opam install -y menhir + - opam pin add coq-flocq 3.4.0 && opam install -y menhir gappa coq-gappa -coq-flocq script: - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh - source /opt/kalray/accesscore/kalray.sh && make -j "$NJOBS" -- cgit From 9790ad0c32ff9ae1e08d974c0a410954bccf24b6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 21 Dec 2021 23:36:45 +0100 Subject: use option -n for pin add --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 096d8959..9c4379fd 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -276,7 +276,7 @@ build_kvx: - rm -rf download - eval `opam config env` - opam update - - opam pin add coq-flocq 3.4.0 && opam install -y menhir gappa coq-gappa -coq-flocq + - opam pin add -n coq-flocq 3.4.0 && opam install -y menhir gappa coq-gappa -coq-flocq script: - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh - source /opt/kalray/accesscore/kalray.sh && make -j "$NJOBS" -- cgit From 7063cdbba6a38b01429cb68389b22ea7b4b6aea2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 21 Dec 2021 23:40:38 +0100 Subject: fix opam pin --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 9c4379fd..79f755d7 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -276,7 +276,7 @@ build_kvx: - rm -rf download - eval `opam config env` - opam update - - opam pin add -n coq-flocq 3.4.0 && opam install -y menhir gappa coq-gappa -coq-flocq + - opam pin add coq-flocq 3.4.0 --no-action && opam install -y menhir gappa coq-gappa -coq-flocq script: - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh - source /opt/kalray/accesscore/kalray.sh && make -j "$NJOBS" -- cgit From 7b04543056d381c6bdb4ac256cc14ec9c31fcddd Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 21 Dec 2021 23:44:17 +0100 Subject: fix again opam --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 79f755d7..82690410 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -276,7 +276,7 @@ build_kvx: - rm -rf download - eval `opam config env` - opam update - - opam pin add coq-flocq 3.4.0 --no-action && opam install -y menhir gappa coq-gappa -coq-flocq + - opam pin add coq-flocq 3.4.0 --no-action && opam install -y menhir gappa coq-gappa coq-flocq script: - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh - source /opt/kalray/accesscore/kalray.sh && make -j "$NJOBS" -- cgit From 7875c804ceb0aaa054d9487866496a32a5787b13 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 21 Dec 2021 23:51:53 +0100 Subject: bison for gappa --- .gitlab-ci.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 82690410..e18f9e56 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -269,7 +269,8 @@ build_kvx: image: coqorg/coq:8.12.2-ocaml-4.11.2-flambda before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 + - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison + # bison for gappa - ./.download_from_Kalray.sh - (cd download ; rm -f *dkms*.deb *eclipse*.deb *llvm*.deb *board-mgmt* *oce-host* *pocl* *flash-util* *barebox* *-kann-* *-kaf-* *-stb-* *-opencv* *-eigen* *-task* *-blis* *-lz4*) - sudo dpkg -i download/*.deb -- cgit From ea0bb074c4edf530f8e7bf4065db4ee3cc15994a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 21 Dec 2021 23:57:51 +0100 Subject: add flex --- .gitlab-ci.yml | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index e18f9e56..b8b28eb6 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -269,15 +269,15 @@ build_kvx: image: coqorg/coq:8.12.2-ocaml-4.11.2-flambda before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison - # bison for gappa + - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison flex + # bison and flex for gappa - ./.download_from_Kalray.sh - (cd download ; rm -f *dkms*.deb *eclipse*.deb *llvm*.deb *board-mgmt* *oce-host* *pocl* *flash-util* *barebox* *-kann-* *-kaf-* *-stb-* *-opencv* *-eigen* *-task* *-blis* *-lz4*) - sudo dpkg -i download/*.deb - rm -rf download - eval `opam config env` - - opam update - - opam pin add coq-flocq 3.4.0 --no-action && opam install -y menhir gappa coq-gappa coq-flocq + - opam update && opam pin add coq-flocq 3.4.0 --no-action + - opam install -y menhir gappa coq-gappa coq-flocq script: - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh - source /opt/kalray/accesscore/kalray.sh && make -j "$NJOBS" -- cgit From e2d328f18030a00207f2610a159177060df857ec Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 22 Dec 2021 00:07:11 +0100 Subject: add libmpfr --- .gitlab-ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index b8b28eb6..b03119a7 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -269,8 +269,8 @@ build_kvx: image: coqorg/coq:8.12.2-ocaml-4.11.2-flambda before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison flex - # bison and flex for gappa + - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison flex libmpfr-dev + # bison, flex, libmpfr-dev for gappa - ./.download_from_Kalray.sh - (cd download ; rm -f *dkms*.deb *eclipse*.deb *llvm*.deb *board-mgmt* *oce-host* *pocl* *flash-util* *barebox* *-kann-* *-kaf-* *-stb-* *-opencv* *-eigen* *-task* *-blis* *-lz4*) - sudo dpkg -i download/*.deb -- cgit From dd1bb6c83309ef4c1a9a2f960c5dd935fc6882dd Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 22 Dec 2021 00:17:05 +0100 Subject: +Boost --- .gitlab-ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index b03119a7..d678aab7 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -269,8 +269,8 @@ build_kvx: image: coqorg/coq:8.12.2-ocaml-4.11.2-flambda before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison flex libmpfr-dev - # bison, flex, libmpfr-dev for gappa + - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison flex libmpfr-dev libboost-dev + # bison, flex, libmpfr-dev, libboost-dev for gappa - ./.download_from_Kalray.sh - (cd download ; rm -f *dkms*.deb *eclipse*.deb *llvm*.deb *board-mgmt* *oce-host* *pocl* *flash-util* *barebox* *-kann-* *-kaf-* *-stb-* *-opencv* *-eigen* *-task* *-blis* *-lz4*) - sudo dpkg -i download/*.deb -- cgit From bfe75ff0614ef6b93313c77fba44d12cb68bd94b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 22 Dec 2021 00:47:23 +0100 Subject: external Flocq --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index d678aab7..724d8543 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -279,7 +279,7 @@ build_kvx: - opam update && opam pin add coq-flocq 3.4.0 --no-action - opam install -y menhir gappa coq-gappa coq-flocq script: - - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh + - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh --use-external-Flocq - source /opt/kalray/accesscore/kalray.sh && make -j "$NJOBS" - make -j "$NJOBS" clightgen - source /opt/kalray/accesscore/kalray.sh && make -C test CCOMPOPTS=-static SIMU='kvx-cluster -- ' EXECUTE='kvx-cluster -- ' all test -- cgit From d51e5088717ef047902cca02f0ddd5e36fa50432 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 22 Dec 2021 09:22:26 +0100 Subject: rm flocq --- .gitlab-ci.yml | 1 + 1 file changed, 1 insertion(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 724d8543..41df7d8e 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -279,6 +279,7 @@ build_kvx: - opam update && opam pin add coq-flocq 3.4.0 --no-action - opam install -y menhir gappa coq-gappa coq-flocq script: + - rm -rf flocq - source /opt/kalray/accesscore/kalray.sh && ./config_kvx.sh --use-external-Flocq - source /opt/kalray/accesscore/kalray.sh && make -j "$NJOBS" - make -j "$NJOBS" clightgen -- cgit From 0123035a0a6f05a75f64b13ddd16c78842ccec74 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 09:06:58 +0100 Subject: we need to use external Flocq, the one that goes with gappa lib --- config_kvx.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/config_kvx.sh b/config_kvx.sh index 9040c23b..289f6f0a 100755 --- a/config_kvx.sh +++ b/config_kvx.sh @@ -1 +1 @@ -exec ./config_simple.sh kvx-cos "$@" +exec ./config_simple.sh kvx-cos -use-external-Flocq "$@" -- cgit From ba7a3de74288f8721699291d8fc2464c452498e1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 15:02:58 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 183 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 183 insertions(+) create mode 100644 kvx/FPDivision64.v diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v new file mode 100644 index 00000000..33a980c6 --- /dev/null +++ b/kvx/FPDivision64.v @@ -0,0 +1,183 @@ +From Flocq Require Import Core Digits Operations Round Bracket Sterbenz + Binary Round_odd Bits. +Require Archi. +Require Import Coqlib. +Require Import Compopts. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Op. +Require Import CminorSel. +Require Import OpHelpers. +Require Import ExtFloats. +Require Import DecBoolOps. +Require Import Chunks. +Require Import Builtins. +Require Import Values Globalenvs. +Require Compopts. +Require Import Psatz. +Require Import IEEE754_extra. + +From Gappa Require Import Gappa_tactic. + +Definition approx_inv_longu b := + let invb_s := ExtValues.invfs (Val.maketotal (Val.singleoflongu b)) in + let invb_d := Val.floatofsingle invb_s in + let b_d := Val.maketotal (Val.floatoflongu b) in + let one := Vfloat (ExtFloat.one) in + let alpha := ExtValues.fmsubf one invb_d b_d in + ExtValues.fmaddf invb_d alpha invb_d. + +Definition approx_inv_thresh := (1/70368744177664)%R. + +Lemma Rabs_relax: + forall b b' (INEQ : (b < b')%R) x, + (-b <= x <= b)%R -> (Rabs x < b')%R. +Proof. + intros. + apply Rabs_lt. + lra. +Qed. + +Theorem approx_inv_longu_correct : + forall b, + (1 <= Int64.unsigned b <= 9223372036854775808)%Z -> + exists (f : float), + (approx_inv_longu (Vlong b)) = Vfloat f /\ + is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= approx_inv_thresh)%R. +Proof. + intros b NONZ. + unfold approx_inv_longu. + cbn. + econstructor. + split. + reflexivity. + Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int. + unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int. + set (re := (@eq_refl Datatypes.comparison Lt)). + change (Int.signed (Int.repr 1)) with 1%Z. + set (b' := Int64.unsigned b) in *. + assert(1 <= IZR b' <= 9223372036854775808)%R as RANGE. + { split; apply IZR_le; lia. + } + + assert (-16777216 <= 1 <= 16777216)%Z as SILLY by lia. + destruct (BofZ_exact 24 128 re re 1 SILLY) as (C0R & C0F & _). + clear SILLY. + set (one_s := (BofZ 24 128 re re 1)) in *. + + pose proof (BofZ_correct 24 128 re re b') as C1. + cbn in C1. + rewrite Rlt_bool_true in C1; cycle 1. + { clear C1. + eapply (Rabs_relax (IZR 9223372036854775808)). + lra. + set (b'' := IZR b') in *. + gappa. + } + destruct C1 as (C1R & C1F & _). + set (b_s := (BofZ 24 128 re re b')) in *. + + assert(1 <= B2R 24 128 b_s <= 9223372036854775808)%R as b_s_RANGE. + { rewrite C1R. + gappa. + } + assert(B2R 24 128 b_s <> 0)%R as b_s_NONZ by lra. + + pose proof (Bdiv_correct 24 128 re re Float32.binop_nan mode_NE one_s b_s b_s_NONZ) as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply Rabs_relax with (b := 1%R). + { cbn; lra. } + rewrite C0R. + set (r_b_s := B2R 24 128 b_s) in *. + cbn. + gappa. + } + + destruct C2 as (C2R & C2F & _). + set (invb_s := (Bdiv 24 128 re re Float32.binop_nan mode_NE one_s b_s)) in *. + rewrite C0F in C2F. + + assert ((1/9223372036854775808 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. + { rewrite C2R. + set (r_b_s := B2R 24 128 b_s) in *. + rewrite C0R. + cbn. + gappa. + } + + pose proof (Bconv_correct 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s C2F) as C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + set (r_invb_s := (B2R 24 128 invb_s)) in *. + apply Rabs_relax with (b := 1%R). + { replace 1%R with (bpow radix2 0)%R by reflexivity. + apply bpow_lt. + lia. + } + cbn. + gappa. + } + + destruct C3 as (C3R & C3F & _). + set (invb_d := (Bconv 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s)) in *. + assert ((1/9223372036854775808 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. + { + rewrite C3R. + set (r_invb_s := B2R 24 128 invb_s) in *. + cbn. + gappa. + } + + pose proof (is_finite_Bopp 53 1024 Float.neg_nan invb_d) as opp_finite. + rewrite C3F in opp_finite. + + pose proof (BofZ_correct 53 1024 re re 1) as C4. + rewrite Rlt_bool_true in C4; cycle 1. + { clear C4. + cbn. + eapply (Rabs_relax (IZR 9223372036854775808)). + lra. + set (b'' := IZR b') in *. + gappa. + } + destruct C4 as (C4R & C4F & _). + + pose proof (BofZ_correct 53 1024 re re b') as C5. + cbn in C5. + rewrite Rlt_bool_true in C5; cycle 1. + { clear C5. + eapply (Rabs_relax (IZR 9223372036854775808)). + lra. + set (b'' := IZR b') in *. + gappa. + } + destruct C5 as (C5R & C5F & _). + set (b_d := (BofZ 53 1024 re re b')) in *. + + assert(1 <= B2R 53 1024 b_d <= 9223372036854775808)%R as b_d_RANGE. + { rewrite C5R. + gappa. + } + + pose proof (Bfma_correct 53 1024 re re Float.fma_nan mode_NE + (Bopp 53 1024 Float.neg_nan invb_d) (BofZ 53 1024 re re b') + (BofZ 53 1024 re re 1) opp_finite C5F C4F) as C6. + rewrite Rlt_bool_true in C6; cycle 1. + { clear C6. + rewrite C4R. + rewrite B2R_Bopp. + cbn. + eapply (Rabs_relax (IZR 9223372036854775808)). + { lra. } + fold invb_d. + fold b_d. + set (r_invb_d := B2R 53 1024 invb_d) in *. + set (r_b_d := B2R 53 1024 b_d) in *. + + unfold invb_d. + unfold invb_s. + gappa. + } +Admitted. -- cgit From 4f8bd24d05bcea0a8b9793ee61450b1e0c7234c1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 15:17:12 +0100 Subject: this works --- kvx/FPDivision64.v | 31 ++++++++++++++++++++----------- 1 file changed, 20 insertions(+), 11 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 33a980c6..424dfc77 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1,3 +1,12 @@ +(* +This needs a special gappa script + +#!/bin/sh +/home/monniaux/.opam/4.12.0+flambda/bin/gappa -Eprecision=100 "$@" + +in PATH before the normal gappa + *) + From Flocq Require Import Core Digits Operations Round Bracket Sterbenz Binary Round_odd Bits. Require Archi. @@ -38,10 +47,10 @@ Proof. apply Rabs_lt. lra. Qed. - + Theorem approx_inv_longu_correct : forall b, - (1 <= Int64.unsigned b <= 9223372036854775808)%Z -> + (1 <= Int64.unsigned b <= 18446744073709551616)%Z -> exists (f : float), (approx_inv_longu (Vlong b)) = Vfloat f /\ is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= approx_inv_thresh)%R. @@ -57,7 +66,7 @@ Proof. set (re := (@eq_refl Datatypes.comparison Lt)). change (Int.signed (Int.repr 1)) with 1%Z. set (b' := Int64.unsigned b) in *. - assert(1 <= IZR b' <= 9223372036854775808)%R as RANGE. + assert(1 <= IZR b' <= 18446744073709551616)%R as RANGE. { split; apply IZR_le; lia. } @@ -70,7 +79,7 @@ Proof. cbn in C1. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. - eapply (Rabs_relax (IZR 9223372036854775808)). + eapply (Rabs_relax (IZR 18446744073709551616)). lra. set (b'' := IZR b') in *. gappa. @@ -78,7 +87,7 @@ Proof. destruct C1 as (C1R & C1F & _). set (b_s := (BofZ 24 128 re re b')) in *. - assert(1 <= B2R 24 128 b_s <= 9223372036854775808)%R as b_s_RANGE. + assert(1 <= B2R 24 128 b_s <= 18446744073709551616)%R as b_s_RANGE. { rewrite C1R. gappa. } @@ -99,7 +108,7 @@ Proof. set (invb_s := (Bdiv 24 128 re re Float32.binop_nan mode_NE one_s b_s)) in *. rewrite C0F in C2F. - assert ((1/9223372036854775808 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. + assert ((1/18446744073709551616 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. { rewrite C2R. set (r_b_s := B2R 24 128 b_s) in *. rewrite C0R. @@ -122,7 +131,7 @@ Proof. destruct C3 as (C3R & C3F & _). set (invb_d := (Bconv 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s)) in *. - assert ((1/9223372036854775808 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. + assert ((1/18446744073709551616 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. { rewrite C3R. set (r_invb_s := B2R 24 128 invb_s) in *. @@ -137,7 +146,7 @@ Proof. rewrite Rlt_bool_true in C4; cycle 1. { clear C4. cbn. - eapply (Rabs_relax (IZR 9223372036854775808)). + eapply (Rabs_relax (IZR 18446744073709551616)). lra. set (b'' := IZR b') in *. gappa. @@ -148,7 +157,7 @@ Proof. cbn in C5. rewrite Rlt_bool_true in C5; cycle 1. { clear C5. - eapply (Rabs_relax (IZR 9223372036854775808)). + eapply (Rabs_relax (IZR 18446744073709551616)). lra. set (b'' := IZR b') in *. gappa. @@ -156,7 +165,7 @@ Proof. destruct C5 as (C5R & C5F & _). set (b_d := (BofZ 53 1024 re re b')) in *. - assert(1 <= B2R 53 1024 b_d <= 9223372036854775808)%R as b_d_RANGE. + assert(1 <= B2R 53 1024 b_d <= 18446744073709551616)%R as b_d_RANGE. { rewrite C5R. gappa. } @@ -169,7 +178,7 @@ Proof. rewrite C4R. rewrite B2R_Bopp. cbn. - eapply (Rabs_relax (IZR 9223372036854775808)). + eapply (Rabs_relax (IZR 18446744073709551616)). { lra. } fold invb_d. fold b_d. -- cgit From af9f4682a8e7f06fdf9eba707d008bc21873d1ce Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 15:27:56 +0100 Subject: progress in proof --- kvx/FPDivision64.v | 23 ++++++++++++++++++++--- 1 file changed, 20 insertions(+), 3 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 424dfc77..a44de06e 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -184,9 +184,26 @@ Proof. fold b_d. set (r_invb_d := B2R 53 1024 invb_d) in *. set (r_b_d := B2R 53 1024 b_d) in *. - - unfold invb_d. - unfold invb_s. gappa. } + fold b_d in C6. + destruct C6 as (C6R & C6F & _). + + pose proof (Bfma_correct 53 1024 re re Float.fma_nan mode_NE + (Bfma 53 1024 re re Float.fma_nan mode_NE + (Bopp 53 1024 Float.neg_nan invb_d) b_d (BofZ 53 1024 re re 1)) + invb_d invb_d C6F C3F C3F) as C7. + rewrite Rlt_bool_true in C7; cycle 1. + { clear C7. + rewrite C6R. + rewrite B2R_Bopp. + eapply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + rewrite C4R. + cbn. + set (r_invb_d := B2R 53 1024 invb_d) in *. + set (r_b_d := B2R 53 1024 b_d) in *. + gappa. + } + Admitted. -- cgit From b11fd5400cb042f4b0727e08d2fa53929092f74c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 15:32:51 +0100 Subject: progress --- kvx/FPDivision64.v | 3 +++ 1 file changed, 3 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index a44de06e..851a6cc4 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -205,5 +205,8 @@ Proof. set (r_b_d := B2R 53 1024 b_d) in *. gappa. } + destruct C7 as (C7R & C7F & _). + + split. assumption. Admitted. -- cgit From cd444fa2200427a1e792716ad5cbac9c1eac872e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 19:12:03 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 34 ++++++++++++++++++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 851a6cc4..f61cc48c 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -208,5 +208,39 @@ Proof. destruct C7 as (C7R & C7F & _). split. assumption. + rewrite C7R. + rewrite C6R. + rewrite C5R. + rewrite C4R. + rewrite B2R_Bopp. + rewrite C3R. + rewrite C2R. + rewrite C1R. + rewrite C0R. + cbn. + set(b1 := IZR b') in *. + replace (round radix2 (FLT_exp (-1074) 53) ZnearestE 1) with 1%R by gappa. + set (bd := round radix2 (FLT_exp (-1074) 53) ZnearestE b1). + set (x0 := round radix2 (FLT_exp (-1074) 53) ZnearestE + (round radix2 (FLT_exp (-149) 24) ZnearestE + (1 / round radix2 (FLT_exp (-149) 24) ZnearestE b1))). + set (alpha0 := (- x0 * bd + 1)%R). + set (y1 := (round radix2 (FLT_exp (-1074) 53) ZnearestE alpha0 * x0 + x0)%R). + set (x1 := round radix2 (FLT_exp (-1074) 53) ZnearestE y1). + replace (x1 - 1/b1)%R with ((y1-1/b1)+(x1-y1))%R by ring. + + assert(alpha0 = -((x0-1/bd)/(1/bd)))%R as alpha0_EQ. + { unfold alpha0. + field. + unfold bd. + gappa. + } + assert(y1-1/b1 = ((round radix2 (FLT_exp (-1074) 53) ZnearestE alpha0) + - alpha0) * x0 + + alpha0*(x0-1/b1) - ((bd-b1)/b1) * x0)%R as y1_EQ. + { unfold y1, alpha0. + field. + lra. + } Admitted. -- cgit From cf534a7a5144f4dadc0b98fc078cd2b11530650f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 23:02:33 +0100 Subject: qed for proof --- kvx/FPDivision64.v | 39 +++++++++++++++++++++++++++++++++++++-- 1 file changed, 37 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index f61cc48c..608b3710 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -242,5 +242,40 @@ Proof. field. lra. } - -Admitted. + assert(Rabs alpha0 <= 257/2147483648)%R as alpha0_ABS. + { rewrite alpha0_EQ. + unfold x0, bd. + gappa. + } + assert (Rabs (x0 - 1 / b1) <= 3/33554432)%R as x0_delta_ABS. + { unfold x0. + gappa. + } + set (x0_delta := (x0 - 1 / b1)%R) in *. + assert (Rabs ((bd - b1) / b1) <= 1/9007199254740992)%R as bd_delta_ABS. + { unfold bd. + gappa. + } + set (bd_delta := ((bd - b1) / b1)%R) in *. + set (rnd_alpha0_delta := (round radix2 (FLT_exp (-1074) 53) ZnearestE alpha0 - alpha0)%R) in *. + assert (Rabs rnd_alpha0_delta <= 1/75557863725914323419136)%R as rnd_alpha0_delta_ABS. + { unfold rnd_alpha0_delta. + gappa. + } + assert (1/18446744073709551616 <= x0 <= 1)%R as x0_RANGE. + { unfold x0. + gappa. + } + assert (Rabs (y1 - 1 / b1) <= 49/4503599627370496)%R as y1_delta_ABS. + { rewrite y1_EQ. + gappa. + } + assert (Rabs(x1 - y1) <= 1/9007199254740992)%R as x1_delta_ABS. + { unfold x1. + gappa. + } + set (y1_delta := (y1 - 1 / b1)%R) in *. + set (x1_delta := (x1-y1)%R) in *. + unfold approx_inv_thresh. + gappa. +Qed. -- cgit From bfbd4b7235fe67cc3fa023bcdec88d89ca6f3c99 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 23:05:19 +0100 Subject: nice intervals --- kvx/FPDivision64.v | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 608b3710..ffcc1ef9 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -50,7 +50,7 @@ Qed. Theorem approx_inv_longu_correct : forall b, - (1 <= Int64.unsigned b <= 18446744073709551616)%Z -> + (0 < Int64.unsigned b)%Z -> exists (f : float), (approx_inv_longu (Vlong b)) = Vfloat f /\ is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= approx_inv_thresh)%R. @@ -66,7 +66,9 @@ Proof. set (re := (@eq_refl Datatypes.comparison Lt)). change (Int.signed (Int.repr 1)) with 1%Z. set (b' := Int64.unsigned b) in *. - assert(1 <= IZR b' <= 18446744073709551616)%R as RANGE. + pose proof (Int64.unsigned_range b) as RANGE. + change Int64.modulus with 18446744073709551616%Z in RANGE. + assert(1 <= IZR b' <= 18446744073709551616)%R as RANGE'. { split; apply IZR_le; lia. } -- cgit From 41838c656dbbf817446af24b01eac8c071bafda7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Jan 2022 23:09:14 +0100 Subject: better bound --- kvx/FPDivision64.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index ffcc1ef9..ab593591 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -37,7 +37,8 @@ Definition approx_inv_longu b := let alpha := ExtValues.fmsubf one invb_d b_d in ExtValues.fmaddf invb_d alpha invb_d. -Definition approx_inv_thresh := (1/70368744177664)%R. +Definition approx_inv_thresh := (25/2251799813685248)%R. +(* 1.11022302462516e-14 *) Lemma Rabs_relax: forall b b' (INEQ : (b < b')%R) x, -- cgit From 071d3fdf1f864c0395e95a8134f319cc9ac0372e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 13:56:50 +0100 Subject: add FPDivision64 --- configure | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/configure b/configure index df031c1e..d71b4120 100755 --- a/configure +++ b/configure @@ -823,7 +823,7 @@ BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v Asmblockprops.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v\\ - FPDivision32.v ExtValues.v ExtFloats.v\\ + FPDivision32.v FPDivision64.v ExtValues.v ExtFloats.v\\ AbstractBasicBlocksDef.v SeqSimuTheory.v ImpSimuTest.v Parallelizability.v EOF fi -- cgit From 120446974d1a83e6f9ef99f049d597166be09271 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 14:18:07 +0100 Subject: rough approx --- kvx/FPDivision64.v | 116 +++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 113 insertions(+), 3 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index ab593591..6c053d7c 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -37,9 +37,6 @@ Definition approx_inv_longu b := let alpha := ExtValues.fmsubf one invb_d b_d in ExtValues.fmaddf invb_d alpha invb_d. -Definition approx_inv_thresh := (25/2251799813685248)%R. -(* 1.11022302462516e-14 *) - Lemma Rabs_relax: forall b b' (INEQ : (b < b')%R) x, (-b <= x <= b)%R -> (Rabs x < b')%R. @@ -48,6 +45,9 @@ Proof. apply Rabs_lt. lra. Qed. + +Definition approx_inv_thresh := (25/2251799813685248)%R. +(* 1.11022302462516e-14 *) Theorem approx_inv_longu_correct : forall b, @@ -282,3 +282,113 @@ Proof. unfold approx_inv_thresh. gappa. Qed. + +Definition rough_approx_inv_longu b := + let invb_s := ExtValues.invfs (Val.maketotal (Val.singleoflongu b)) in + Val.floatofsingle invb_s. + +Definition rough_approx_inv_thresh := (3/33554432)%R. +(* 8.94069671630859e-8 *) + +Theorem rough_approx_inv_longu_correct : + forall b, + (0 < Int64.unsigned b)%Z -> + exists (f : float), + (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ + is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= rough_approx_inv_thresh)%R. +Proof. + intros b NONZ. + unfold rough_approx_inv_longu. + cbn. + econstructor. + split. + reflexivity. + Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int. + unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int. + set (re := (@eq_refl Datatypes.comparison Lt)). + change (Int.signed (Int.repr 1)) with 1%Z. + set (b' := Int64.unsigned b) in *. + pose proof (Int64.unsigned_range b) as RANGE. + change Int64.modulus with 18446744073709551616%Z in RANGE. + assert(1 <= IZR b' <= 18446744073709551616)%R as RANGE'. + { split; apply IZR_le; lia. + } + + assert (-16777216 <= 1 <= 16777216)%Z as SILLY by lia. + destruct (BofZ_exact 24 128 re re 1 SILLY) as (C0R & C0F & _). + clear SILLY. + set (one_s := (BofZ 24 128 re re 1)) in *. + + pose proof (BofZ_correct 24 128 re re b') as C1. + cbn in C1. + rewrite Rlt_bool_true in C1; cycle 1. + { clear C1. + eapply (Rabs_relax (IZR 18446744073709551616)). + lra. + set (b'' := IZR b') in *. + gappa. + } + destruct C1 as (C1R & C1F & _). + set (b_s := (BofZ 24 128 re re b')) in *. + + assert(1 <= B2R 24 128 b_s <= 18446744073709551616)%R as b_s_RANGE. + { rewrite C1R. + gappa. + } + assert(B2R 24 128 b_s <> 0)%R as b_s_NONZ by lra. + + pose proof (Bdiv_correct 24 128 re re Float32.binop_nan mode_NE one_s b_s b_s_NONZ) as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply Rabs_relax with (b := 1%R). + { cbn; lra. } + rewrite C0R. + set (r_b_s := B2R 24 128 b_s) in *. + cbn. + gappa. + } + + destruct C2 as (C2R & C2F & _). + set (invb_s := (Bdiv 24 128 re re Float32.binop_nan mode_NE one_s b_s)) in *. + rewrite C0F in C2F. + + assert ((1/18446744073709551616 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. + { rewrite C2R. + set (r_b_s := B2R 24 128 b_s) in *. + rewrite C0R. + cbn. + gappa. + } + + pose proof (Bconv_correct 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s C2F) as C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + set (r_invb_s := (B2R 24 128 invb_s)) in *. + apply Rabs_relax with (b := 1%R). + { replace 1%R with (bpow radix2 0)%R by reflexivity. + apply bpow_lt. + lia. + } + cbn. + gappa. + } + + destruct C3 as (C3R & C3F & _). + set (invb_d := (Bconv 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s)) in *. + assert ((1/18446744073709551616 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. + { + rewrite C3R. + set (r_invb_s := B2R 24 128 invb_s) in *. + cbn. + gappa. + } + + split. assumption. + unfold rough_approx_inv_thresh. + rewrite C3R. + rewrite C2R. + rewrite C1R. + rewrite C0R. + cbn. + gappa. +Qed. -- cgit From f411b8eb07064070440169269d9f402b3cdbb009 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 14:55:20 +0100 Subject: state theorem --- kvx/FPDivision64.v | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 6c053d7c..9baa3f94 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -392,3 +392,18 @@ Proof. cbn. gappa. Qed. + +Definition rough_approx_div_longu a b := + Val.maketotal (Val.longuoffloat + (Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b))). + +Definition rough_approx_div_thresh := 1649267441663%Z. + +Theorem rough_approx_div_longu_correct : + forall a b, + (0 < Int64.unsigned b)%Z -> + exists (q : int64), + (rough_approx_div_longu (Vlong a) (Vlong b)) = Vlong q /\ + (Z.abs ((Int64.unsigned a) - (Int64.unsigned q) * (Int64.unsigned b)) <= rough_approx_div_thresh)%Z. +Proof. +Admitted. -- cgit From c31acf1e0448b9c6d0a6b175d0ac605a9b1b82ba Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 15:17:10 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 9baa3f94..e025fa94 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -406,4 +406,29 @@ Theorem rough_approx_div_longu_correct : (rough_approx_div_longu (Vlong a) (Vlong b)) = Vlong q /\ (Z.abs ((Int64.unsigned a) - (Int64.unsigned q) * (Int64.unsigned b)) <= rough_approx_div_thresh)%Z. Proof. + intros a b b_RANGE. + unfold rough_approx_div_longu. + destruct (rough_approx_inv_longu_correct b b_RANGE) as (f & C1R & C1F & C1D). + rewrite C1R. + cbn. + Local Transparent Float.to_longu Float.mul. + unfold Float.to_longu, Float.mul, Float.of_longu. + set (re := (@eq_refl Datatypes.comparison Lt)). + + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in a_RANGE. + set (a' := Int64.unsigned a) in *. + assert (0 <= IZR a' <= 18446744073709551615)%R as IZR_a_RANGE. + { split ; apply IZR_le; lia. + } + + pose proof (BofZ_correct 53 1024 re re a') as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt. lia. } + cbn. + gappa. + } + destruct C2 as (C2R & C2F & _). Admitted. -- cgit From b7875749ae445c43130fabd9eb9099eddea48318 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 15:31:39 +0100 Subject: proof progress --- kvx/FPDivision64.v | 26 ++++++++++++++++++++++++-- 1 file changed, 24 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index e025fa94..9463313a 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -406,11 +406,12 @@ Theorem rough_approx_div_longu_correct : (rough_approx_div_longu (Vlong a) (Vlong b)) = Vlong q /\ (Z.abs ((Int64.unsigned a) - (Int64.unsigned q) * (Int64.unsigned b)) <= rough_approx_div_thresh)%Z. Proof. - intros a b b_RANGE. + intros a b b_NONZ. unfold rough_approx_div_longu. - destruct (rough_approx_inv_longu_correct b b_RANGE) as (f & C1R & C1F & C1D). + destruct (rough_approx_inv_longu_correct b b_NONZ) as (f & C1R & C1F & C1D). rewrite C1R. cbn. + Local Transparent Float.to_longu Float.mul. unfold Float.to_longu, Float.mul, Float.of_longu. set (re := (@eq_refl Datatypes.comparison Lt)). @@ -422,6 +423,13 @@ Proof. { split ; apply IZR_le; lia. } + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in b_RANGE. + set (b' := Int64.unsigned b) in *. + assert (1 <= IZR b' <= 18446744073709551615)%R as IZR_b_RANGE. + { split ; apply IZR_le; lia. + } + pose proof (BofZ_correct 53 1024 re re a') as C2. rewrite Rlt_bool_true in C2; cycle 1. { clear C2. @@ -431,4 +439,18 @@ Proof. gappa. } destruct C2 as (C2R & C2F & _). + + pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re a') f) as C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + rewrite C2R. + replace (B2R 53 1024 f)%R with ((B2R 53 1024 f - 1 / IZR b') + 1 / IZR b')%R by ring. + unfold rough_approx_inv_thresh in C1D. + apply Rabs_relax with (b := bpow radix2 65). + { apply bpow_lt ; lia. } + cbn. + gappa. + } + Admitted. -- cgit From f42abd286e90cad92174cd77b2dd068ed8dd11ae Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 15:49:39 +0100 Subject: wrong operator --- kvx/FPDivision64.v | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 9463313a..9cde7d97 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -394,7 +394,7 @@ Proof. Qed. Definition rough_approx_div_longu a b := - Val.maketotal (Val.longuoffloat + Val.maketotal (Val.longuoffloat_ne (Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b))). Definition rough_approx_div_thresh := 1649267441663%Z. @@ -413,7 +413,7 @@ Proof. cbn. Local Transparent Float.to_longu Float.mul. - unfold Float.to_longu, Float.mul, Float.of_longu. + unfold Float.to_longu_ne, Float.mul, Float.of_longu. set (re := (@eq_refl Datatypes.comparison Lt)). pose proof (Int64.unsigned_range a) as a_RANGE. @@ -452,5 +452,15 @@ Proof. cbn. gappa. } + rewrite C1F in C3. + rewrite C2F in C3. + cbn in C3. + destruct C3 as (C3R & C3F & _). + + pose proof (ZofB_ne_range_correct 53 1024 + (Bmult 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re a') f) 0 Int64.max_unsigned) as C4. + rewrite C3F in C4. + rewrite C3R in C4. Admitted. -- cgit From a0303be2ae9b7b53f012a74c3e42020acb225fb9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 17:22:18 +0100 Subject: signs --- kvx/FPDivision64.v | 36 +++++++++++++++++++++++++++++++----- 1 file changed, 31 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 9cde7d97..1dbf2782 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -295,7 +295,9 @@ Theorem rough_approx_inv_longu_correct : (0 < Int64.unsigned b)%Z -> exists (f : float), (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ - is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= rough_approx_inv_thresh)%R. + is_finite _ _ f = true /\ + (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= rough_approx_inv_thresh)%R (* /\ + Bsign _ _ f = false. *) Proof. intros b NONZ. unfold rough_approx_inv_longu. @@ -328,7 +330,8 @@ Proof. set (b'' := IZR b') in *. gappa. } - destruct C1 as (C1R & C1F & _). + rewrite (Zlt_bool_false b' 0) in C1 by lia. + destruct C1 as (C1R & C1F & C1S). set (b_s := (BofZ 24 128 re re b')) in *. assert(1 <= B2R 24 128 b_s <= 18446744073709551616)%R as b_s_RANGE. @@ -347,10 +350,19 @@ Proof. cbn. gappa. } - - destruct C2 as (C2R & C2F & _). + destruct C2 as (C2R & C2F & C2Sz). + rewrite C1S in C2Sz. + change (xorb _ _) with false in C2Sz. set (invb_s := (Bdiv 24 128 re re Float32.binop_nan mode_NE one_s b_s)) in *. rewrite C0F in C2F. + assert (is_nan 24 128 invb_s = false) as NAN. + { apply is_finite_not_is_nan. + assumption. + } + pose proof (C2Sz NAN) as C2S. + clear C2Sz. + + apply is_finite_not_is_nan in C2S. assert ((1/18446744073709551616 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. { rewrite C2R. @@ -462,5 +474,19 @@ Proof. (BofZ 53 1024 re re a') f) 0 Int64.max_unsigned) as C4. rewrite C3F in C4. rewrite C3R in C4. - + rewrite C2R in C4. + + assert(0 <= + (round radix2 (FLT_exp (-1074) 53) ZnearestE + (round radix2 (FLT_exp (3 - 1024 - 53) 53) + (round_mode mode_NE) (IZR a') * B2R 53 1024 f)) + <= IZR Int64.max_unsigned)%R as q_RANGE. + { clear C4. + change (IZR Int64.max_unsigned) with 18446744073709551615%R. + cbn. + replace (B2R 53 1024 f)%R with ((B2R 53 1024 f - 1 / IZR b') + 1 / IZR b')%R by ring. + unfold rough_approx_inv_thresh in C1D. + set (delta := (B2R 53 1024 f - 1 / IZR b')%R) in *. + gappa. + } Admitted. -- cgit From e1e99fcf500c8c9822a017258acb9c63fa40229f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 18:23:42 +0100 Subject: some more proof --- kvx/FPDivision64.v | 54 +++++++++++++++++++++++++++++++++++++++++++++++------- 1 file changed, 47 insertions(+), 7 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 1dbf2782..4d9f0537 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -296,8 +296,8 @@ Theorem rough_approx_inv_longu_correct : exists (f : float), (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ is_finite _ _ f = true /\ - (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= rough_approx_inv_thresh)%R (* /\ - Bsign _ _ f = false. *) + (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= rough_approx_inv_thresh)%R /\ + Bsign _ _ f = false. Proof. intros b NONZ. unfold rough_approx_inv_longu. @@ -362,8 +362,6 @@ Proof. pose proof (C2Sz NAN) as C2S. clear C2Sz. - apply is_finite_not_is_nan in C2S. - assert ((1/18446744073709551616 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. { rewrite C2R. set (r_b_s := B2R 24 128 b_s) in *. @@ -385,7 +383,7 @@ Proof. gappa. } - destruct C3 as (C3R & C3F & _). + destruct C3 as (C3R & C3F & C3S). set (invb_d := (Bconv 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s)) in *. assert ((1/18446744073709551616 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. { @@ -394,8 +392,9 @@ Proof. cbn. gappa. } + rewrite C2S in C3S. - split. assumption. + split. assumption. split. 2: assumption. unfold rough_approx_inv_thresh. rewrite C3R. rewrite C2R. @@ -411,6 +410,22 @@ Definition rough_approx_div_longu a b := Definition rough_approx_div_thresh := 1649267441663%Z. +Lemma Bsign_false_nonneg: + forall prec emax f, + (Bsign prec emax f) = false -> (0 <= (B2R prec emax f))%R. +Proof. + intros until f. intro SIGN. + destruct f. + 1, 2, 3: cbn; lra. + cbn. + apply F2R_ge_0. + cbn. + cbn in SIGN. + rewrite SIGN. + cbn. + lia. +Qed. + Theorem rough_approx_div_longu_correct : forall a b, (0 < Int64.unsigned b)%Z -> @@ -420,7 +435,7 @@ Theorem rough_approx_div_longu_correct : Proof. intros a b b_NONZ. unfold rough_approx_div_longu. - destruct (rough_approx_inv_longu_correct b b_NONZ) as (f & C1R & C1F & C1D). + destruct (rough_approx_inv_longu_correct b b_NONZ) as (f & C1R & C1F & C1D & C1S). rewrite C1R. cbn. @@ -435,6 +450,12 @@ Proof. { split ; apply IZR_le; lia. } + set (a_d := (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a'))). + assert (0 <= a_d <= 18446744073709551616)%R as a_d_RANGE. + { unfold a_d. + gappa. + } + pose proof (Int64.unsigned_range b) as b_RANGE. change Int64.modulus with 18446744073709551616%Z in b_RANGE. set (b' := Int64.unsigned b) in *. @@ -442,6 +463,19 @@ Proof. { split ; apply IZR_le; lia. } + set (f_r := (B2R _ _ f)) in *. + assert (0 <= f_r <= 1)%R as f_RANGE. + { split. + { apply Bsign_false_nonneg. trivial .} + unfold rough_approx_inv_thresh in C1D. + replace f_r with ((f_r - 1 / IZR b') + 1/ IZR b')%R by (field; lra). + set (delta := (f_r - 1 / IZR b')%R) in *. + destruct (Z_le_gt_dec b' 1) as [LE | GT]. + { admit. } + assert (2 <= IZR b')%R as NOT1. + { apply IZR_le. lia. } + gappa. + } pose proof (BofZ_correct 53 1024 re re a') as C2. rewrite Rlt_bool_true in C2; cycle 1. { clear C2. @@ -482,7 +516,13 @@ Proof. (round_mode mode_NE) (IZR a') * B2R 53 1024 f)) <= IZR Int64.max_unsigned)%R as q_RANGE. { clear C4. + cbn. change (IZR Int64.max_unsigned) with 18446744073709551615%R. + fold a_d. + gappa. + set (y := ((IZR a') * B2R 53 1024 f)%R). + assert ((round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a') * B2R 53 1024 f) >= 0)%R. + cbn. replace (B2R 53 1024 f)%R with ((B2R 53 1024 f - 1 / IZR b') + 1 / IZR b')%R by ring. unfold rough_approx_inv_thresh in C1D. -- cgit From c6e3d6b38c7cc62877d85f561b69f8651a5f4a3f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 19:07:15 +0100 Subject: some more proofs? --- kvx/FPDivision64.v | 67 ++++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 65 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 4d9f0537..521cbe0f 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -404,6 +404,62 @@ Proof. gappa. Qed. +Theorem rough_approx_inv_longu_correct1 : + forall b, + (Int64.unsigned b = 1%Z) -> + exists f, + (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ + is_finite _ _ f = true /\ + (B2R _ _ f) = 1%R /\ + Bsign _ _ f = false. +Proof. + intros b ONE. + cbn. + unfold Float.of_single, ExtFloat32.inv, Float32.div, ExtFloat32.one, Float32.of_longu, Float32.of_int. + set (re := (@eq_refl Datatypes.comparison Lt)). + rewrite ONE. + change (Int.signed (Int.repr 1)) with 1%Z. + + econstructor. + split. reflexivity. + split. reflexivity. + split. 2: reflexivity. + destruct (BofZ_exact 24 128 re re 1%Z) as (C1R & C1F & C1S). lia. + + pose proof (Bdiv_correct 24 128 re re Float32.binop_nan mode_NE + (BofZ 24 128 re re 1) (BofZ 24 128 re re 1)) as C2. + rewrite C1R in C2. + rewrite C1F in C2. + rewrite C1S in C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply Rabs_relax with (b := bpow radix2 0%Z). + { apply bpow_lt. lia. } + cbn. + gappa. + } + assert (1 <> 0)%R as OBVIOUS by lra. + destruct (C2 OBVIOUS) as (C2R & C2F & C2S). + clear C2. + + pose proof (Bconv_correct 24 128 53 1024 re re Float.of_single_nan mode_NE + (Bdiv 24 128 re re Float32.binop_nan mode_NE + (BofZ 24 128 re re 1) (BofZ 24 128 re re 1)) C2F) as C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + rewrite C2R. + apply Rabs_relax with (b := bpow radix2 0%Z). + { apply bpow_lt. lia. } + cbn. + gappa. + } + destruct C3 as (C3R & C3F & C3S). + rewrite C3R. + rewrite C2R. + cbn. + gappa. +Qed. + Definition rough_approx_div_longu a b := Val.maketotal (Val.longuoffloat_ne (Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b))). @@ -468,14 +524,21 @@ Proof. { split. { apply Bsign_false_nonneg. trivial .} unfold rough_approx_inv_thresh in C1D. + destruct (Z_le_gt_dec b' 1) as [LE | GT]. + { unfold f_r. + assert (b' = 1%Z) as EQ by lia. + destruct (rough_approx_inv_longu_correct1 b EQ) as + (f2 & OE & OF & OR & OS). + replace f2 with f in * by congruence. + lra. + } replace f_r with ((f_r - 1 / IZR b') + 1/ IZR b')%R by (field; lra). set (delta := (f_r - 1 / IZR b')%R) in *. - destruct (Z_le_gt_dec b' 1) as [LE | GT]. - { admit. } assert (2 <= IZR b')%R as NOT1. { apply IZR_le. lia. } gappa. } + pose proof (BofZ_correct 53 1024 re re a') as C2. rewrite Rlt_bool_true in C2; cycle 1. { clear C2. -- cgit From d1754fa46e98ef709274c7528b0c6453a6050d9d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 20:30:35 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 71 ++++++++++++++++++++++++++++++++---------------------- 1 file changed, 42 insertions(+), 29 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 521cbe0f..a241e4f0 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -482,15 +482,37 @@ Proof. lia. Qed. +Lemma Znearest_IZR_le : + forall rnd n x, (IZR n <= x)%R -> (n <= Znearest rnd x)%Z. +Proof. + intros until x. intro ORDER. + pose proof (Znearest_ge_floor rnd x). + pose proof (Zfloor_le _ _ ORDER) as KK. + rewrite Zfloor_IZR in KK. + lia. +Qed. + +Lemma Znearest_le_IZR : + forall rnd n x, (x <= IZR n)%R -> (Znearest rnd x <= n)%Z. +Proof. + intros until x. intro ORDER. + pose proof (Znearest_le_ceil rnd x). + pose proof (Zceil_le _ _ ORDER) as KK. + rewrite Zceil_IZR in KK. + lia. +Qed. + + Theorem rough_approx_div_longu_correct : forall a b, - (0 < Int64.unsigned b)%Z -> + (2 <= Int64.unsigned b)%Z -> exists (q : int64), (rough_approx_div_longu (Vlong a) (Vlong b)) = Vlong q /\ (Z.abs ((Int64.unsigned a) - (Int64.unsigned q) * (Int64.unsigned b)) <= rough_approx_div_thresh)%Z. Proof. - intros a b b_NONZ. + intros a b b_NOT01. unfold rough_approx_div_longu. + assert (0 < Int64.unsigned b)%Z as b_NONZ by lia. destruct (rough_approx_inv_longu_correct b b_NONZ) as (f & C1R & C1F & C1D & C1S). rewrite C1R. cbn. @@ -515,27 +537,17 @@ Proof. pose proof (Int64.unsigned_range b) as b_RANGE. change Int64.modulus with 18446744073709551616%Z in b_RANGE. set (b' := Int64.unsigned b) in *. - assert (1 <= IZR b' <= 18446744073709551615)%R as IZR_b_RANGE. + assert (2 <= IZR b' <= 18446744073709551615)%R as IZR_b_RANGE. { split ; apply IZR_le; lia. } set (f_r := (B2R _ _ f)) in *. - assert (0 <= f_r <= 1)%R as f_RANGE. + assert (0 <= f_r <= 4194305/8388608)%R as f_RANGE. { split. { apply Bsign_false_nonneg. trivial .} unfold rough_approx_inv_thresh in C1D. - destruct (Z_le_gt_dec b' 1) as [LE | GT]. - { unfold f_r. - assert (b' = 1%Z) as EQ by lia. - destruct (rough_approx_inv_longu_correct1 b EQ) as - (f2 & OE & OF & OR & OS). - replace f2 with f in * by congruence. - lra. - } replace f_r with ((f_r - 1 / IZR b') + 1/ IZR b')%R by (field; lra). set (delta := (f_r - 1 / IZR b')%R) in *. - assert (2 <= IZR b')%R as NOT1. - { apply IZR_le. lia. } gappa. } @@ -554,8 +566,8 @@ Proof. rewrite Rlt_bool_true in C3; cycle 1. { clear C3. rewrite C2R. - replace (B2R 53 1024 f)%R with ((B2R 53 1024 f - 1 / IZR b') + 1 / IZR b')%R by ring. - unfold rough_approx_inv_thresh in C1D. + fold f_r. + apply Rabs_relax with (b := bpow radix2 65). { apply bpow_lt ; lia. } cbn. @@ -573,23 +585,24 @@ Proof. rewrite C3R in C4. rewrite C2R in C4. - assert(0 <= - (round radix2 (FLT_exp (-1074) 53) ZnearestE + set (y := (round radix2 (FLT_exp (-1074) 53) ZnearestE (round radix2 (FLT_exp (3 - 1024 - 53) 53) - (round_mode mode_NE) (IZR a') * B2R 53 1024 f)) - <= IZR Int64.max_unsigned)%R as q_RANGE. + (round_mode mode_NE) (IZR a') * B2R 53 1024 f))) in *. + assert(0 <= y <= IZR Int64.max_unsigned)%R as q_RANGE. { clear C4. + unfold y. cbn. change (IZR Int64.max_unsigned) with 18446744073709551615%R. - fold a_d. - gappa. - set (y := ((IZR a') * B2R 53 1024 f)%R). - assert ((round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a') * B2R 53 1024 f) >= 0)%R. - - cbn. - replace (B2R 53 1024 f)%R with ((B2R 53 1024 f - 1 / IZR b') + 1 / IZR b')%R by ring. - unfold rough_approx_inv_thresh in C1D. - set (delta := (B2R 53 1024 f - 1 / IZR b')%R) in *. + fold f_r. gappa. } + cbn in C4. + rewrite Zle_imp_le_bool in C4; cycle 1. + { apply Znearest_IZR_le. intuition. + } + rewrite Zle_imp_le_bool in C4; cycle 1. + { apply Znearest_le_IZR. intuition. + } + + Admitted. -- cgit From f09e226b326185e940517c6940e5f71abf7d4fac Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 10 Jan 2022 20:37:20 +0100 Subject: progress in proof --- kvx/FPDivision64.v | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index a241e4f0..906ad650 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -603,6 +603,16 @@ Proof. rewrite Zle_imp_le_bool in C4; cycle 1. { apply Znearest_le_IZR. intuition. } + cbn in C4. + change Int64.max_unsigned with 18446744073709551615%Z. + rewrite C4. + cbn. + exists (Int64.repr (ZnearestE y)). + split. reflexivity. + rewrite Int64.unsigned_repr; cycle 1. + { split. + - apply Znearest_IZR_le. intuition. + - apply Znearest_le_IZR. intuition. + } - Admitted. -- cgit From 76917c124e7c7de216acd99555209f18c1ecd16b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 11 Jan 2022 15:22:43 +0100 Subject: progress in proof --- kvx/FPDivision64.v | 86 +++++++++++++++++++++--------------------------------- 1 file changed, 34 insertions(+), 52 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 906ad650..76e21e54 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -461,10 +461,7 @@ Proof. Qed. Definition rough_approx_div_longu a b := - Val.maketotal (Val.longuoffloat_ne - (Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b))). - -Definition rough_approx_div_thresh := 1649267441663%Z. + Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b). Lemma Bsign_false_nonneg: forall prec emax f, @@ -502,15 +499,19 @@ Proof. lia. Qed. +Definition rough_approx_div_thresh := 1%R. Theorem rough_approx_div_longu_correct : forall a b, (2 <= Int64.unsigned b)%Z -> - exists (q : int64), - (rough_approx_div_longu (Vlong a) (Vlong b)) = Vlong q /\ - (Z.abs ((Int64.unsigned a) - (Int64.unsigned q) * (Int64.unsigned b)) <= rough_approx_div_thresh)%Z. + exists (q : float), + (rough_approx_div_longu (Vlong a) (Vlong b)) = Vfloat q /\ + (Rabs (B2R _ _ q - (IZR (Int64.unsigned a)) / (IZR (Int64.unsigned b))) + <= rough_approx_div_thresh)%R /\ + is_finite _ _ q = true /\ + Bsign _ _ q = false. Proof. - intros a b b_NOT01. + intros a b b_NON01. unfold rough_approx_div_longu. assert (0 < Int64.unsigned b)%Z as b_NONZ by lia. destruct (rough_approx_inv_longu_correct b b_NONZ) as (f & C1R & C1F & C1D & C1S). @@ -528,12 +529,6 @@ Proof. { split ; apply IZR_le; lia. } - set (a_d := (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a'))). - assert (0 <= a_d <= 18446744073709551616)%R as a_d_RANGE. - { unfold a_d. - gappa. - } - pose proof (Int64.unsigned_range b) as b_RANGE. change Int64.modulus with 18446744073709551616%Z in b_RANGE. set (b' := Int64.unsigned b) in *. @@ -541,6 +536,12 @@ Proof. { split ; apply IZR_le; lia. } + set (a_d := (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a'))). + assert (0 <= a_d <= 18446744073709551616)%R as a_d_RANGE. + { unfold a_d. + gappa. + } + set (f_r := (B2R _ _ f)) in *. assert (0 <= f_r <= 4194305/8388608)%R as f_RANGE. { split. @@ -559,7 +560,7 @@ Proof. cbn. gappa. } - destruct C2 as (C2R & C2F & _). + destruct C2 as (C2R & C2F & C2S). pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE (BofZ 53 1024 re re a') f) as C3. @@ -576,43 +577,24 @@ Proof. rewrite C1F in C3. rewrite C2F in C3. cbn in C3. - destruct C3 as (C3R & C3F & _). + destruct C3 as (C3R & C3F & C3S). + + econstructor. split. reflexivity. + rewrite C3R. + rewrite C2R. + rewrite C3S. + rewrite C2S. + rewrite C1S. + rewrite C3F. + rewrite Zlt_bool_false by lia. + split. 2: auto; fail. - pose proof (ZofB_ne_range_correct 53 1024 - (Bmult 53 1024 re re Float.binop_nan mode_NE - (BofZ 53 1024 re re a') f) 0 Int64.max_unsigned) as C4. - rewrite C3F in C4. - rewrite C3R in C4. - rewrite C2R in C4. - - set (y := (round radix2 (FLT_exp (-1074) 53) ZnearestE - (round radix2 (FLT_exp (3 - 1024 - 53) 53) - (round_mode mode_NE) (IZR a') * B2R 53 1024 f))) in *. - assert(0 <= y <= IZR Int64.max_unsigned)%R as q_RANGE. - { clear C4. - unfold y. - cbn. - change (IZR Int64.max_unsigned) with 18446744073709551615%R. - fold f_r. - gappa. - } - cbn in C4. - rewrite Zle_imp_le_bool in C4; cycle 1. - { apply Znearest_IZR_le. intuition. - } - rewrite Zle_imp_le_bool in C4; cycle 1. - { apply Znearest_le_IZR. intuition. - } - cbn in C4. - change Int64.max_unsigned with 18446744073709551615%Z. - rewrite C4. cbn. - exists (Int64.repr (ZnearestE y)). - split. reflexivity. - rewrite Int64.unsigned_repr; cycle 1. - { split. - - apply Znearest_IZR_le. intuition. - - apply Znearest_le_IZR. intuition. - } - + unfold rough_approx_div_thresh. + fold f_r. + set (rd := round radix2 (FLT_exp (-1074) 53) ZnearestE). + replace (rd (rd (IZR a') * f_r) - IZR a' / IZR b')%R with + ((rd (rd (IZR a') * f_r) - (rd (IZR a') * f_r)) + + ((rd (IZR a') - IZR a') * f_r) + + (IZR a') * (f_r - 1 / IZR b'))%R by (field ; lra). Admitted. -- cgit From aa8f655361670b593f805b1e57a1089b8f924938 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 11 Jan 2022 15:31:02 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 76e21e54..d16417fb 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -582,7 +582,9 @@ Proof. econstructor. split. reflexivity. rewrite C3R. rewrite C2R. - rewrite C3S. + rewrite C3S; cycle 1. + { apply is_finite_not_is_nan. + assumption. } rewrite C2S. rewrite C1S. rewrite C3F. @@ -597,4 +599,12 @@ Proof. ((rd (rd (IZR a') * f_r) - (rd (IZR a') * f_r)) + ((rd (IZR a') - IZR a') * f_r) + (IZR a') * (f_r - 1 / IZR b'))%R by (field ; lra). + + assert (Rabs (rd (IZR a') - IZR a') <= 1024)%R as DELTA1. + { unfold rd. gappa. } + set (delta1 := (rd (IZR a') - IZR a')%R) in *. + + assert (Rabs (rd (rd (IZR a') * f_r) - rd (IZR a') * f_r) <= 1024)%R as DELTA2. + { unfold rd. gappa. } + set (delta2 := (rd (rd (IZR a') * f_r) - rd (IZR a') * f_r)%R) in *. Admitted. -- cgit From 6e01e627018c745c6fd490dd6ced63203d482bbd Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 11 Jan 2022 15:37:52 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index d16417fb..f20efa4e 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -499,7 +499,7 @@ Proof. lia. Qed. -Definition rough_approx_div_thresh := 1%R. +Definition rough_approx_div_thresh := 1683627180032%R. Theorem rough_approx_div_longu_correct : forall a b, @@ -607,4 +607,7 @@ Proof. assert (Rabs (rd (rd (IZR a') * f_r) - rd (IZR a') * f_r) <= 1024)%R as DELTA2. { unfold rd. gappa. } set (delta2 := (rd (rd (IZR a') * f_r) - rd (IZR a') * f_r)%R) in *. -Admitted. + unfold rough_approx_inv_thresh in C1D. + set (delta3 := (f_r - 1 / IZR b')%R) in *. + gappa. +Qed. -- cgit From 185846ddf99f5c9f17b74e40fc9ad1deb91cb1ff Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 11:35:05 +0100 Subject: finer proof --- kvx/FPDivision64.v | 105 +++++++++++++++++++++++++++++++++++++++++++++++------ 1 file changed, 94 insertions(+), 11 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index f20efa4e..425e11da 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -289,14 +289,17 @@ Definition rough_approx_inv_longu b := Definition rough_approx_inv_thresh := (3/33554432)%R. (* 8.94069671630859e-8 *) - + +Local Notation "'rd'" := (round radix2 (FLT_exp (-1074) 53) ZnearestE). +Local Notation "'rs'" := (round radix2 (FLT_exp (-149) 24) ZnearestE). + Theorem rough_approx_inv_longu_correct : forall b, (0 < Int64.unsigned b)%Z -> exists (f : float), (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ + (B2R _ _ f) = (rd (rs (1 / rs (IZR (Int64.unsigned b))))) /\ is_finite _ _ f = true /\ - (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= rough_approx_inv_thresh)%R /\ Bsign _ _ f = false. Proof. intros b NONZ. @@ -350,6 +353,7 @@ Proof. cbn. gappa. } + rewrite C1R in C2. destruct C2 as (C2R & C2F & C2Sz). rewrite C1S in C2Sz. change (xorb _ _) with false in C2Sz. @@ -382,7 +386,6 @@ Proof. cbn. gappa. } - destruct C3 as (C3R & C3F & C3S). set (invb_d := (Bconv 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s)) in *. assert ((1/18446744073709551616 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. @@ -393,15 +396,10 @@ Proof. gappa. } rewrite C2S in C3S. + rewrite C2R in C3R. + rewrite C0R in C3R. - split. assumption. split. 2: assumption. - unfold rough_approx_inv_thresh. - rewrite C3R. - rewrite C2R. - rewrite C1R. - rewrite C0R. - cbn. - gappa. + auto. Qed. Theorem rough_approx_inv_longu_correct1 : @@ -611,3 +609,88 @@ Proof. set (delta3 := (f_r - 1 / IZR b')%R) in *. gappa. Qed. + +Definition rough_approx_euclid_div_longu a b := + Val.maketotal (Val.longuoffloat_ne (rough_approx_div_longu a b)). + +Definition smallb_threshold := 4398046511104%Z. +Definition rough_approx_div_longu_smallb_delta := 3917010173952%Z. + +Theorem rough_approx_div_longu_smallb : + forall a b, + (2 <= Int64.unsigned b <= smallb_threshold)%Z -> + exists (qr : int64), + (rough_approx_euclid_div_longu (Vlong a) (Vlong b)) = Vlong qr /\ + (Z.abs (Int64.unsigned a - Int64.unsigned b* Int64.unsigned qr) <= rough_approx_div_longu_smallb_delta)%Z /\ + (0 <= Int64.unsigned qr <= 11529215046068469760)%Z. +Proof. + intros a b b_RANGE. + unfold rough_approx_euclid_div_longu. + destruct (rough_approx_div_longu_correct a b) as (q & C1R & C1D & C1F & C1S). + lia. + rewrite C1R. + cbn. + set (qr := B2R 53 1024 q) in *. + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in a_RANGE. + set (a' := Int64.unsigned a) in *. + assert (0 <= IZR a' <= 18446744073709551615)%R as IZR_a_RANGE. + { split ; apply IZR_le; lia. + } + + set (b' := Int64.unsigned b) in *. + assert (2 <= IZR b' <= IZR smallb_threshold)%R as IZR_b_RANGE. + { split ; apply IZR_le; lia. + } + + assert (0 <= qr <= 11529215046068469760)%R as q_RANGE. + { split. + { unfold qr. + apply Bsign_false_nonneg. + assumption. + } + unfold rough_approx_div_thresh in C1D. + replace qr with ((qr - IZR a' / IZR b') + (IZR a' / IZR b'))%R by (field ; lra). + set (delta := (qr - IZR a' / IZR b')%R) in *. + unfold smallb_threshold in *. + gappa. + } + unfold Float.to_longu_ne. + pose proof (ZofB_ne_range_correct _ _ q 0 Int64.max_unsigned) as C2. + rewrite C1F in C2. + cbn in C2. + fold qr in C2. + assert (0 <= ZnearestE qr <= 11529215046068469760)%Z as round_RANGE. + { split. + { apply Znearest_IZR_le. + lra. + } + apply Znearest_le_IZR. + lra. + } + rewrite Zle_bool_true in C2 by lia. + rewrite Zle_bool_true in C2 by lia. + change Int64.max_unsigned with 18446744073709551615%Z. + cbn in C2. + rewrite C2. + cbn. + econstructor. split. reflexivity. + rewrite Int64.unsigned_repr; cycle 1. + { change Int64.max_unsigned with 18446744073709551615%Z. + lia. + } + split. + 2: assumption. + unfold rough_approx_div_thresh, rough_approx_div_longu_smallb_delta in *. + apply le_IZR. + rewrite <- Rabs_Zabs. + rewrite minus_IZR. + rewrite mult_IZR. + replace (IZR a' - IZR b' * IZR (ZnearestE qr))%R with + (- IZR b' * (IZR (ZnearestE qr) - qr) + - IZR b' * (qr - IZR a' / IZR b'))%R by (field ; lra). + pose proof (Znearest_imp2 (fun x : Z => negb (Z.even x)) qr) as ROUND. + set (delta1 := (IZR (ZnearestE qr) - qr)%R) in *. + set (delta2 := (qr - IZR a' / IZR b')%R) in *. + unfold smallb_threshold in *. + -- cgit From 7ff41e564bcc2e01d33b26419d3039d2f393440a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 11:41:27 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 53 ++++++++--------------------------------------------- 1 file changed, 8 insertions(+), 45 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 425e11da..b1ba6938 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -407,54 +407,17 @@ Theorem rough_approx_inv_longu_correct1 : (Int64.unsigned b = 1%Z) -> exists f, (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ - is_finite _ _ f = true /\ (B2R _ _ f) = 1%R /\ + is_finite _ _ f = true /\ Bsign _ _ f = false. Proof. - intros b ONE. - cbn. - unfold Float.of_single, ExtFloat32.inv, Float32.div, ExtFloat32.one, Float32.of_longu, Float32.of_int. - set (re := (@eq_refl Datatypes.comparison Lt)). - rewrite ONE. - change (Int.signed (Int.repr 1)) with 1%Z. - - econstructor. - split. reflexivity. - split. reflexivity. - split. 2: reflexivity. - destruct (BofZ_exact 24 128 re re 1%Z) as (C1R & C1F & C1S). lia. - - pose proof (Bdiv_correct 24 128 re re Float32.binop_nan mode_NE - (BofZ 24 128 re re 1) (BofZ 24 128 re re 1)) as C2. - rewrite C1R in C2. - rewrite C1F in C2. - rewrite C1S in C2. - rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. - apply Rabs_relax with (b := bpow radix2 0%Z). - { apply bpow_lt. lia. } - cbn. - gappa. - } - assert (1 <> 0)%R as OBVIOUS by lra. - destruct (C2 OBVIOUS) as (C2R & C2F & C2S). - clear C2. - - pose proof (Bconv_correct 24 128 53 1024 re re Float.of_single_nan mode_NE - (Bdiv 24 128 re re Float32.binop_nan mode_NE - (BofZ 24 128 re re 1) (BofZ 24 128 re re 1)) C2F) as C3. - rewrite Rlt_bool_true in C3; cycle 1. - { clear C3. - rewrite C2R. - apply Rabs_relax with (b := bpow radix2 0%Z). - { apply bpow_lt. lia. } - cbn. - gappa. - } - destruct C3 as (C3R & C3F & C3S). - rewrite C3R. - rewrite C2R. - cbn. + intros b EQ1. + assert (0 < Int64.unsigned b)%Z as b_RANGE by lia. + destruct (rough_approx_inv_longu_correct b b_RANGE) as (f & C1E & C1R & C1F & C1S). + rewrite EQ1 in C1R. + exists f. + repeat split; try assumption. + rewrite C1R. gappa. Qed. -- cgit From 41701fa253bf8adb790e0f27559bb4e6e6276ea8 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 13:10:49 +0100 Subject: more proofs --- kvx/FPDivision64.v | 82 ++++++++++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 77 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index b1ba6938..81d587f7 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -421,9 +421,6 @@ Proof. gappa. Qed. -Definition rough_approx_div_longu a b := - Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b). - Lemma Bsign_false_nonneg: forall prec emax f, (Bsign prec emax f) = false -> (0 <= (B2R prec emax f))%R. @@ -459,7 +456,83 @@ Proof. rewrite Zceil_IZR in KK. lia. Qed. + + +Definition rough_approx_div_longu a b := + Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b). +Definition smallb_threshold := 4398046511104%Z. +Definition rough_approx_div_longu_smallb_delta := 3917010173952%Z. + +Theorem rough_approx_div_longu_correct: + forall a b, + (1 < Int64.unsigned b)%Z -> + exists (q : float), + (rough_approx_div_longu (Vlong a) (Vlong b)) = Vfloat q /\ + (B2R _ _ q) = rd (rd (IZR (Int64.unsigned a)) * (rd (rs (1 / rs (IZR (Int64.unsigned b)))))) /\ + is_finite _ _ q = true /\ + Bsign _ _ q = false. +Proof. + intros a b b_NON01. + assert (0 < Int64.unsigned b)%Z as b_NON0 by lia. + destruct (rough_approx_inv_longu_correct b b_NON0) as (f & C1E & C1R & C1F & C1S). + unfold rough_approx_div_longu. + rewrite C1E. + cbn. + set (b' := Int64.unsigned b) in *. + Local Transparent Float.mul. + unfold Float.mul, Float.of_longu. + econstructor. + split. reflexivity. + set (a' := Int64.unsigned a) in *. + set (re := (@eq_refl Datatypes.comparison Lt)). + + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in a_RANGE. + assert (0 <= IZR a' <= 18446744073709551615)%R as IZR_a_RANGE. + { split; apply IZR_le; lia. } + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in b_RANGE. + assert (2 <= IZR b' <= 18446744073709551615)%R as IZR_b_RANGE. + { split; apply IZR_le; lia. } + pose proof (BofZ_correct 53 1024 re re a') as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt. lia. } + cbn. + gappa. + } + destruct C2 as (C2R & C2F & C2S). + rewrite Zlt_bool_false in C2S by lia. + + pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE (BofZ 53 1024 re re a') f) as C3. + rewrite C1S in C3. + rewrite C2S in C3. + rewrite C1F in C3. + rewrite C2F in C3. + rewrite C1R in C3. + rewrite C2R in C3. + rewrite Rlt_bool_true in C3; cycle 1. + { apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt ; lia. } + cbn. + gappa. + } + cbn in C3. + destruct C3 as (C3R & C3F & C3Sz). + assert (is_nan 53 1024 + (Bmult 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re a') f) = false) as NAN. + { apply is_finite_not_is_nan. + assumption. } + pose proof (C3Sz NAN) as C3S. + clear NAN C3Sz. + + auto. +Qed. + +(* Definition rough_approx_div_thresh := 1683627180032%R. Theorem rough_approx_div_longu_correct : @@ -572,12 +645,11 @@ Proof. set (delta3 := (f_r - 1 / IZR b')%R) in *. gappa. Qed. + *) Definition rough_approx_euclid_div_longu a b := Val.maketotal (Val.longuoffloat_ne (rough_approx_div_longu a b)). -Definition smallb_threshold := 4398046511104%Z. -Definition rough_approx_div_longu_smallb_delta := 3917010173952%Z. Theorem rough_approx_div_longu_smallb : forall a b, -- cgit From fc58331d1aa0c1379df8b17d52f71b07d80c3045 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 15:37:43 +0100 Subject: give a name --- kvx/FPDivision64.v | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 81d587f7..e8e7f8be 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -463,12 +463,16 @@ Definition rough_approx_div_longu a b := Definition smallb_threshold := 4398046511104%Z. Definition rough_approx_div_longu_smallb_delta := 3917010173952%Z. +Definition step1_real_quotient (a b : R) := + rd ((rd a) * (rd (rs (1 / rs b)))). + Theorem rough_approx_div_longu_correct: forall a b, (1 < Int64.unsigned b)%Z -> exists (q : float), (rough_approx_div_longu (Vlong a) (Vlong b)) = Vfloat q /\ - (B2R _ _ q) = rd (rd (IZR (Int64.unsigned a)) * (rd (rs (1 / rs (IZR (Int64.unsigned b)))))) /\ + (B2R _ _ q) = step1_real_quotient (IZR (Int64.unsigned a)) + (IZR (Int64.unsigned b)) /\ is_finite _ _ q = true /\ Bsign _ _ q = false. Proof. -- cgit From 9cacc4caa5710f35a7ab73de6aa0ddf849e60d73 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 16:06:35 +0100 Subject: proof of approx for a/b --- kvx/FPDivision64.v | 28 ++++++++++++++++++++++++++-- 1 file changed, 26 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index e8e7f8be..f4d727be 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -464,7 +464,7 @@ Definition smallb_threshold := 4398046511104%Z. Definition rough_approx_div_longu_smallb_delta := 3917010173952%Z. Definition step1_real_quotient (a b : R) := - rd ((rd a) * (rd (rs (1 / rs b)))). + rd ((rd (a)) * (rd (rs (1 / rs (b))))). Theorem rough_approx_div_longu_correct: forall a b, @@ -535,7 +535,31 @@ Proof. auto. Qed. - + +Definition smallb_thresh := 4398046511104%Z. + +Definition smallb_approx_range := 1649280000000%R. +Lemma step1_smallb : + forall a b + (a_RANGE : (0 <= a <= 18446744073709551615)%R) + (b_RANGE : (1 <= b <= IZR smallb_thresh)%R), + (Rabs((step1_real_quotient a b) - a/b) <= smallb_approx_range)%R. +Proof. + intros. + unfold smallb_thresh in b_RANGE. + unfold smallb_approx_range. + replace (step1_real_quotient a b - a/b)%R with + ((rd ((rd (a)) * (rd (rs (1 / rs (b))))) + - ((rd (a)) * (rd (rs (1 / rs (b)))))) + + (rd(a)) * (rd (rs (1 / rs (b))) - 1/b) + + (rd(a) - a)/b)%R; cycle 1. + { unfold step1_real_quotient. + field. + lra. + } + gappa. +Qed. + (* Definition rough_approx_div_thresh := 1683627180032%R. -- cgit From 82245a79c9bfd3400d309242377efba20882a8e3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 17:25:43 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 254 +++++++---------------------------------------------- 1 file changed, 30 insertions(+), 224 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index f4d727be..e7ba56e3 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -283,27 +283,27 @@ Proof. gappa. Qed. -Definition rough_approx_inv_longu b := +Definition step1_real_inv_longu b := let invb_s := ExtValues.invfs (Val.maketotal (Val.singleoflongu b)) in Val.floatofsingle invb_s. -Definition rough_approx_inv_thresh := (3/33554432)%R. +Definition step1_real_inv_thresh := (3/33554432)%R. (* 8.94069671630859e-8 *) Local Notation "'rd'" := (round radix2 (FLT_exp (-1074) 53) ZnearestE). Local Notation "'rs'" := (round radix2 (FLT_exp (-149) 24) ZnearestE). -Theorem rough_approx_inv_longu_correct : +Theorem step1_real_inv_longu_correct : forall b, (0 < Int64.unsigned b)%Z -> exists (f : float), - (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ + (step1_real_inv_longu (Vlong b)) = Vfloat f /\ (B2R _ _ f) = (rd (rs (1 / rs (IZR (Int64.unsigned b))))) /\ is_finite _ _ f = true /\ Bsign _ _ f = false. Proof. intros b NONZ. - unfold rough_approx_inv_longu. + unfold step1_real_inv_longu. cbn. econstructor. split. @@ -402,18 +402,18 @@ Proof. auto. Qed. -Theorem rough_approx_inv_longu_correct1 : +Theorem step1_real_inv_longu_correct1 : forall b, (Int64.unsigned b = 1%Z) -> exists f, - (rough_approx_inv_longu (Vlong b)) = Vfloat f /\ + (step1_real_inv_longu (Vlong b)) = Vfloat f /\ (B2R _ _ f) = 1%R /\ is_finite _ _ f = true /\ Bsign _ _ f = false. Proof. intros b EQ1. assert (0 < Int64.unsigned b)%Z as b_RANGE by lia. - destruct (rough_approx_inv_longu_correct b b_RANGE) as (f & C1E & C1R & C1F & C1S). + destruct (step1_real_inv_longu_correct b b_RANGE) as (f & C1E & C1R & C1F & C1S). rewrite EQ1 in C1R. exists f. repeat split; try assumption. @@ -457,20 +457,20 @@ Proof. lia. Qed. +Definition step1_real_div_longu a b := + Val.mulf (Val.maketotal (Val.floatoflongu a)) (step1_real_inv_longu b). -Definition rough_approx_div_longu a b := - Val.mulf (Val.maketotal (Val.floatoflongu a)) (rough_approx_inv_longu b). -Definition smallb_threshold := 4398046511104%Z. -Definition rough_approx_div_longu_smallb_delta := 3917010173952%Z. +Definition step1_div_longu a b := + Val.maketotal (Val.longuoffloat (Val.mulf (Val.maketotal (Val.floatoflongu a)) (step1_real_inv_longu b))). Definition step1_real_quotient (a b : R) := rd ((rd (a)) * (rd (rs (1 / rs (b))))). -Theorem rough_approx_div_longu_correct: +Theorem step1_real_div_longu_correct: forall a b, (1 < Int64.unsigned b)%Z -> exists (q : float), - (rough_approx_div_longu (Vlong a) (Vlong b)) = Vfloat q /\ + (step1_real_div_longu (Vlong a) (Vlong b)) = Vfloat q /\ (B2R _ _ q) = step1_real_quotient (IZR (Int64.unsigned a)) (IZR (Int64.unsigned b)) /\ is_finite _ _ q = true /\ @@ -478,8 +478,8 @@ Theorem rough_approx_div_longu_correct: Proof. intros a b b_NON01. assert (0 < Int64.unsigned b)%Z as b_NON0 by lia. - destruct (rough_approx_inv_longu_correct b b_NON0) as (f & C1E & C1R & C1F & C1S). - unfold rough_approx_div_longu. + destruct (step1_real_inv_longu_correct b b_NON0) as (f & C1E & C1R & C1F & C1S). + unfold step1_real_div_longu. rewrite C1E. cbn. set (b' := Int64.unsigned b) in *. @@ -536,224 +536,30 @@ Proof. auto. Qed. -Definition smallb_thresh := 4398046511104%Z. +Definition smallb_thresh := 4398046511104%Z. -Definition smallb_approx_range := 1649280000000%R. +Definition smallb_approx_range := 2200000000000%R. Lemma step1_smallb : forall a b - (a_RANGE : (0 <= a <= 18446744073709551615)%R) + (a_RANGE : (1 <= a <= 18446744073709551615)%R) (b_RANGE : (1 <= b <= IZR smallb_thresh)%R), - (Rabs((step1_real_quotient a b) - a/b) <= smallb_approx_range)%R. + (Rabs((step1_real_quotient a b) * b - a) <= smallb_approx_range)%R. Proof. intros. unfold smallb_thresh in b_RANGE. unfold smallb_approx_range. - replace (step1_real_quotient a b - a/b)%R with - ((rd ((rd (a)) * (rd (rs (1 / rs (b))))) - - ((rd (a)) * (rd (rs (1 / rs (b)))))) + - (rd(a)) * (rd (rs (1 / rs (b))) - 1/b) + - (rd(a) - a)/b)%R; cycle 1. - { unfold step1_real_quotient. + unfold step1_real_quotient. + set (q := ((rd (a)) * (rd (rs (1 / rs (b)))))%R) in *. + replace ((rd q) *b - a)%R with + ((rd(q)-q)/q * rd(a) * (1 + (rd (rs (1 / rs (b))) - 1/b)/(1/b)) + + (rd (a)) * ((rd (rs (1 / rs (b))) - 1 / b) / (1/b)) + + (rd(a) - a))%R; cycle 1. + { unfold q. field. - lra. - } - gappa. -Qed. - -(* -Definition rough_approx_div_thresh := 1683627180032%R. - -Theorem rough_approx_div_longu_correct : - forall a b, - (2 <= Int64.unsigned b)%Z -> - exists (q : float), - (rough_approx_div_longu (Vlong a) (Vlong b)) = Vfloat q /\ - (Rabs (B2R _ _ q - (IZR (Int64.unsigned a)) / (IZR (Int64.unsigned b))) - <= rough_approx_div_thresh)%R /\ - is_finite _ _ q = true /\ - Bsign _ _ q = false. -Proof. - intros a b b_NON01. - unfold rough_approx_div_longu. - assert (0 < Int64.unsigned b)%Z as b_NONZ by lia. - destruct (rough_approx_inv_longu_correct b b_NONZ) as (f & C1R & C1F & C1D & C1S). - rewrite C1R. - cbn. - - Local Transparent Float.to_longu Float.mul. - unfold Float.to_longu_ne, Float.mul, Float.of_longu. - set (re := (@eq_refl Datatypes.comparison Lt)). - - pose proof (Int64.unsigned_range a) as a_RANGE. - change Int64.modulus with 18446744073709551616%Z in a_RANGE. - set (a' := Int64.unsigned a) in *. - assert (0 <= IZR a' <= 18446744073709551615)%R as IZR_a_RANGE. - { split ; apply IZR_le; lia. - } - - pose proof (Int64.unsigned_range b) as b_RANGE. - change Int64.modulus with 18446744073709551616%Z in b_RANGE. - set (b' := Int64.unsigned b) in *. - assert (2 <= IZR b' <= 18446744073709551615)%R as IZR_b_RANGE. - { split ; apply IZR_le; lia. - } - - set (a_d := (round radix2 (FLT_exp (-1074) 53) ZnearestE (IZR a'))). - assert (0 <= a_d <= 18446744073709551616)%R as a_d_RANGE. - { unfold a_d. - gappa. - } - - set (f_r := (B2R _ _ f)) in *. - assert (0 <= f_r <= 4194305/8388608)%R as f_RANGE. - { split. - { apply Bsign_false_nonneg. trivial .} - unfold rough_approx_inv_thresh in C1D. - replace f_r with ((f_r - 1 / IZR b') + 1/ IZR b')%R by (field; lra). - set (delta := (f_r - 1 / IZR b')%R) in *. - gappa. - } - - pose proof (BofZ_correct 53 1024 re re a') as C2. - rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. - apply Rabs_relax with (b := bpow radix2 64). - { apply bpow_lt. lia. } - cbn. - gappa. - } - destruct C2 as (C2R & C2F & C2S). - - pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE - (BofZ 53 1024 re re a') f) as C3. - rewrite Rlt_bool_true in C3; cycle 1. - { clear C3. - rewrite C2R. - fold f_r. - - apply Rabs_relax with (b := bpow radix2 65). - { apply bpow_lt ; lia. } - cbn. + split. lra. + split. gappa. gappa. } - rewrite C1F in C3. - rewrite C2F in C3. - cbn in C3. - destruct C3 as (C3R & C3F & C3S). - - econstructor. split. reflexivity. - rewrite C3R. - rewrite C2R. - rewrite C3S; cycle 1. - { apply is_finite_not_is_nan. - assumption. } - rewrite C2S. - rewrite C1S. - rewrite C3F. - rewrite Zlt_bool_false by lia. - split. 2: auto; fail. - - cbn. - unfold rough_approx_div_thresh. - fold f_r. - set (rd := round radix2 (FLT_exp (-1074) 53) ZnearestE). - replace (rd (rd (IZR a') * f_r) - IZR a' / IZR b')%R with - ((rd (rd (IZR a') * f_r) - (rd (IZR a') * f_r)) - + ((rd (IZR a') - IZR a') * f_r) - + (IZR a') * (f_r - 1 / IZR b'))%R by (field ; lra). - - assert (Rabs (rd (IZR a') - IZR a') <= 1024)%R as DELTA1. - { unfold rd. gappa. } - set (delta1 := (rd (IZR a') - IZR a')%R) in *. - - assert (Rabs (rd (rd (IZR a') * f_r) - rd (IZR a') * f_r) <= 1024)%R as DELTA2. - { unfold rd. gappa. } - set (delta2 := (rd (rd (IZR a') * f_r) - rd (IZR a') * f_r)%R) in *. - unfold rough_approx_inv_thresh in C1D. - set (delta3 := (f_r - 1 / IZR b')%R) in *. + unfold q. gappa. Qed. - *) - -Definition rough_approx_euclid_div_longu a b := - Val.maketotal (Val.longuoffloat_ne (rough_approx_div_longu a b)). - - -Theorem rough_approx_div_longu_smallb : - forall a b, - (2 <= Int64.unsigned b <= smallb_threshold)%Z -> - exists (qr : int64), - (rough_approx_euclid_div_longu (Vlong a) (Vlong b)) = Vlong qr /\ - (Z.abs (Int64.unsigned a - Int64.unsigned b* Int64.unsigned qr) <= rough_approx_div_longu_smallb_delta)%Z /\ - (0 <= Int64.unsigned qr <= 11529215046068469760)%Z. -Proof. - intros a b b_RANGE. - unfold rough_approx_euclid_div_longu. - destruct (rough_approx_div_longu_correct a b) as (q & C1R & C1D & C1F & C1S). - lia. - rewrite C1R. - cbn. - set (qr := B2R 53 1024 q) in *. - pose proof (Int64.unsigned_range a) as a_RANGE. - change Int64.modulus with 18446744073709551616%Z in a_RANGE. - set (a' := Int64.unsigned a) in *. - assert (0 <= IZR a' <= 18446744073709551615)%R as IZR_a_RANGE. - { split ; apply IZR_le; lia. - } - - set (b' := Int64.unsigned b) in *. - assert (2 <= IZR b' <= IZR smallb_threshold)%R as IZR_b_RANGE. - { split ; apply IZR_le; lia. - } - - assert (0 <= qr <= 11529215046068469760)%R as q_RANGE. - { split. - { unfold qr. - apply Bsign_false_nonneg. - assumption. - } - unfold rough_approx_div_thresh in C1D. - replace qr with ((qr - IZR a' / IZR b') + (IZR a' / IZR b'))%R by (field ; lra). - set (delta := (qr - IZR a' / IZR b')%R) in *. - unfold smallb_threshold in *. - gappa. - } - unfold Float.to_longu_ne. - pose proof (ZofB_ne_range_correct _ _ q 0 Int64.max_unsigned) as C2. - rewrite C1F in C2. - cbn in C2. - fold qr in C2. - assert (0 <= ZnearestE qr <= 11529215046068469760)%Z as round_RANGE. - { split. - { apply Znearest_IZR_le. - lra. - } - apply Znearest_le_IZR. - lra. - } - rewrite Zle_bool_true in C2 by lia. - rewrite Zle_bool_true in C2 by lia. - change Int64.max_unsigned with 18446744073709551615%Z. - cbn in C2. - rewrite C2. - cbn. - econstructor. split. reflexivity. - rewrite Int64.unsigned_repr; cycle 1. - { change Int64.max_unsigned with 18446744073709551615%Z. - lia. - } - split. - 2: assumption. - unfold rough_approx_div_thresh, rough_approx_div_longu_smallb_delta in *. - apply le_IZR. - rewrite <- Rabs_Zabs. - rewrite minus_IZR. - rewrite mult_IZR. - replace (IZR a' - IZR b' * IZR (ZnearestE qr))%R with - (- IZR b' * (IZR (ZnearestE qr) - qr) - - IZR b' * (qr - IZR a' / IZR b'))%R by (field ; lra). - pose proof (Znearest_imp2 (fun x : Z => negb (Z.even x)) qr) as ROUND. - set (delta1 := (IZR (ZnearestE qr) - qr)%R) in *. - set (delta2 := (qr - IZR a' / IZR b')%R) in *. - unfold smallb_threshold in *. - -- cgit From 13876421ac06c904e9c8ea7636f5bef1fc4c8988 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 19:15:06 +0100 Subject: progress in proofs --- kvx/FPDivision64.v | 66 +++++++++++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 61 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index e7ba56e3..5da53535 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -461,7 +461,7 @@ Definition step1_real_div_longu a b := Val.mulf (Val.maketotal (Val.floatoflongu a)) (step1_real_inv_longu b). Definition step1_div_longu a b := - Val.maketotal (Val.longuoffloat (Val.mulf (Val.maketotal (Val.floatoflongu a)) (step1_real_inv_longu b))). + Val.maketotal (Val.longuoffloat_ne (step1_real_div_longu a b)). Definition step1_real_quotient (a b : R) := rd ((rd (a)) * (rd (rs (1 / rs (b))))). @@ -538,16 +538,16 @@ Qed. Definition smallb_thresh := 4398046511104%Z. -Definition smallb_approx_range := 2200000000000%R. -Lemma step1_smallb : +Definition smallb_approx_real_range := 2200000000000%R. +Lemma step1_smallb_real : forall a b (a_RANGE : (1 <= a <= 18446744073709551615)%R) (b_RANGE : (1 <= b <= IZR smallb_thresh)%R), - (Rabs((step1_real_quotient a b) * b - a) <= smallb_approx_range)%R. + (Rabs((step1_real_quotient a b) * b - a) <= smallb_approx_real_range)%R. Proof. intros. unfold smallb_thresh in b_RANGE. - unfold smallb_approx_range. + unfold smallb_approx_real_range. unfold step1_real_quotient. set (q := ((rd (a)) * (rd (rs (1 / rs (b)))))%R) in *. replace ((rd q) *b - a)%R with @@ -563,3 +563,59 @@ Proof. unfold q. gappa. Qed. + +Definition smallb_approx_range := 2200000000000%Z. +Lemma step1_div_longu_correct : + forall a b, + (1 < Int64.unsigned b <= smallb_thresh)%Z -> + exists (q : int64), + (step1_div_longu (Vlong a) (Vlong b)) = Vlong q /\ + (Z.abs (Int64.unsigned a - Int64.unsigned b*Int64.signed q) <= smallb_approx_range)%Z. +Proof. + intros a b b_RANGE. + unfold step1_div_longu. + assert (1 < Int64.unsigned b)%Z as b_NOT01 by lia. + destruct (step1_real_div_longu_correct a b b_NOT01) as (q & C1E & C1R & C1F & C1S). + rewrite C1E. cbn. + unfold Float.to_longu_ne. + pose proof (ZofB_ne_range_correct 53 1024 q 0 Int64.max_unsigned) as C2. + rewrite C1F in C2. + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in a_RANGE. + set (a' := Int64.unsigned a) in *. + assert (1 <= a')%Z by admit. + + set (b' := Int64.unsigned b) in *. + assert (1 <= IZR a' <= 18446744073709551615)%R as a_RANGE'. + { split; apply IZR_le; lia. } + assert (2 <= IZR b' <= IZR smallb_thresh)%R as b_RANGE'. + { split; apply IZR_le; lia. } + assert (1 <= IZR b' <= IZR smallb_thresh)%R as b_RANGE'' by lra. + pose proof (step1_smallb_real (IZR a') (IZR b') a_RANGE' b_RANGE'') as DELTA. + rewrite <- C1R in DELTA. + + assert (0 <= B2R _ _ q)%R as q_NONNEG. + { apply Bsign_false_nonneg. assumption. } + cbn in C2. + rewrite Zle_bool_true in C2; cycle 1. + { apply Znearest_IZR_le. assumption. } + assert (B2R _ _ q <= 18446744073709551615)%R as q_SMALL. + { replace (B2R _ _ q) with + ((IZR a' / IZR b') + (B2R _ _ q * IZR b' - IZR a') / IZR b')%R; cycle 1. + { field. lra. } + unfold smallb_approx_real_range in DELTA. + unfold smallb_thresh in b_RANGE'. + set (y := (B2R 53 1024 q * IZR b' - IZR a')%R) in *. + gappa. + } + rewrite Zle_bool_true in C2; cycle 1. + { apply Znearest_le_IZR. assumption. } + cbn in C2. + + change Int64.max_unsigned with 18446744073709551615%Z. + rewrite C2. + cbn. + + econstructor. split. reflexivity. + + -- cgit From 35b823b50952c5d46eaeb36f3c7781bbb2e91674 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 19:44:05 +0100 Subject: proof forward --- kvx/FPDivision64.v | 31 +++++++++++++++++++++++++------ 1 file changed, 25 insertions(+), 6 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 5da53535..76959118 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -564,13 +564,13 @@ Proof. gappa. Qed. -Definition smallb_approx_range := 2200000000000%Z. +Definition smallb_approx_range := 4400000000000%Z. Lemma step1_div_longu_correct : forall a b, (1 < Int64.unsigned b <= smallb_thresh)%Z -> exists (q : int64), (step1_div_longu (Vlong a) (Vlong b)) = Vlong q /\ - (Z.abs (Int64.unsigned a - Int64.unsigned b*Int64.signed q) <= smallb_approx_range)%Z. + (Z.abs (Int64.unsigned a - Int64.unsigned b*Int64.unsigned q) <= smallb_approx_range)%Z. Proof. intros a b b_RANGE. unfold step1_div_longu. @@ -599,7 +599,7 @@ Proof. cbn in C2. rewrite Zle_bool_true in C2; cycle 1. { apply Znearest_IZR_le. assumption. } - assert (B2R _ _ q <= 18446744073709551615)%R as q_SMALL. + assert (B2R _ _ q <= 9223376000000000000)%R as q_SMALL. { replace (B2R _ _ q) with ((IZR a' / IZR b') + (B2R _ _ q * IZR b' - IZR a') / IZR b')%R; cycle 1. { field. lra. } @@ -609,7 +609,7 @@ Proof. gappa. } rewrite Zle_bool_true in C2; cycle 1. - { apply Znearest_le_IZR. assumption. } + { apply Znearest_le_IZR. lra. } cbn in C2. change Int64.max_unsigned with 18446744073709551615%Z. @@ -617,5 +617,24 @@ Proof. cbn. econstructor. split. reflexivity. - - + rewrite Int64.unsigned_repr; cycle 1. + { split. + - apply Znearest_IZR_le. lra. + - apply Znearest_le_IZR. + change Int64.max_unsigned with 18446744073709551615%Z. + lra. + } + apply le_IZR. + rewrite abs_IZR. + unfold smallb_approx_real_range, smallb_approx_range, smallb_thresh in *. + rewrite minus_IZR. + rewrite mult_IZR. + set (q_r := B2R 53 1024 q) in *. + assert (Rabs (IZR (ZnearestE q_r) - q_r) <= / 2)%R as NEAR + by apply Znearest_imp2. + set (q_i := IZR (ZnearestE q_r)) in *. + replace (IZR a' - IZR b' * q_i)%R with + (-(IZR b' * (q_i - q_r)) - (q_r * IZR b' - IZR a'))%R by ring. + set (delta1 := (q_i - q_r)%R) in *. + set (delta2 := (q_r * IZR b' - IZR a')%R) in *. + gappa. -- cgit From c46cded7f85c692f12c1d5d912ba4a64150a0696 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 20:15:39 +0100 Subject: proof progresses --- kvx/FPDivision64.v | 44 ++++++++++++++++++++++++++++++++++++++------ 1 file changed, 38 insertions(+), 6 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 76959118..c17e4f1b 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -564,6 +564,12 @@ Proof. gappa. Qed. +Lemma step1_div_longu_a0 : + forall b, + (0 < Int64.unsigned b)%Z -> + (step1_div_longu (Vlong Int64.zero) (Vlong b)) = Vlong Int64.zero. +Admitted. + Definition smallb_approx_range := 4400000000000%Z. Lemma step1_div_longu_correct : forall a b, @@ -573,25 +579,50 @@ Lemma step1_div_longu_correct : (Z.abs (Int64.unsigned a - Int64.unsigned b*Int64.unsigned q) <= smallb_approx_range)%Z. Proof. intros a b b_RANGE. + + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in a_RANGE. + set (a' := Int64.unsigned a) in *. + set (b' := Int64.unsigned b) in *. + + destruct (Z_le_gt_dec a' 0). + { assert (a' = 0%Z) as ZERO by lia. + exists Int64.zero. + rewrite ZERO. + rewrite Int64.unsigned_zero. + replace (Z.abs (0 - b' * 0))%Z with 0%Z by lia. + replace a with Int64.zero; cycle 1. + { + unfold a' in ZERO. + unfold Int64.zero. + rewrite <- ZERO. + apply Int64.repr_unsigned. + } + split. + { apply step1_div_longu_a0. + lia. + } + unfold smallb_approx_range. + lia. + } + unfold step1_div_longu. - assert (1 < Int64.unsigned b)%Z as b_NOT01 by lia. + assert (1 < b')%Z as b_NOT01 by lia. destruct (step1_real_div_longu_correct a b b_NOT01) as (q & C1E & C1R & C1F & C1S). rewrite C1E. cbn. unfold Float.to_longu_ne. pose proof (ZofB_ne_range_correct 53 1024 q 0 Int64.max_unsigned) as C2. rewrite C1F in C2. - pose proof (Int64.unsigned_range a) as a_RANGE. - change Int64.modulus with 18446744073709551616%Z in a_RANGE. - set (a' := Int64.unsigned a) in *. - assert (1 <= a')%Z by admit. + - set (b' := Int64.unsigned b) in *. assert (1 <= IZR a' <= 18446744073709551615)%R as a_RANGE'. { split; apply IZR_le; lia. } assert (2 <= IZR b' <= IZR smallb_thresh)%R as b_RANGE'. { split; apply IZR_le; lia. } assert (1 <= IZR b' <= IZR smallb_thresh)%R as b_RANGE'' by lra. pose proof (step1_smallb_real (IZR a') (IZR b') a_RANGE' b_RANGE'') as DELTA. + fold a' in C1R. + fold b' in C1R. rewrite <- C1R in DELTA. assert (0 <= B2R _ _ q)%R as q_NONNEG. @@ -638,3 +669,4 @@ Proof. set (delta1 := (q_i - q_r)%R) in *. set (delta2 := (q_r * IZR b' - IZR a')%R) in *. gappa. +Qed. -- cgit From 5354fea25dd1537a9ce36fbebcc190f5dda241c1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 12 Jan 2022 20:51:23 +0100 Subject: more proofs on step1 small b --- kvx/FPDivision64.v | 59 +++++++++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 58 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index c17e4f1b..fc2050bb 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -568,7 +568,64 @@ Lemma step1_div_longu_a0 : forall b, (0 < Int64.unsigned b)%Z -> (step1_div_longu (Vlong Int64.zero) (Vlong b)) = Vlong Int64.zero. -Admitted. +Proof. + intros b b_NOT0. + unfold step1_div_longu. + unfold step1_real_div_longu. + destruct (step1_real_inv_longu_correct b b_NOT0) as + (f & C1E & C1R & C1F & C1S). + rewrite C1E. + cbn. + unfold Float.to_longu_ne, Float.of_longu, Float.mul. + rewrite Int64.unsigned_zero. + set (re := (@eq_refl Datatypes.comparison Lt)). + assert (- 2 ^ 53 <= 0 <= 2 ^ 53)%Z as SILLY by lia. + destruct (BofZ_exact 53 1024 re re 0 SILLY) as (C2R & C2F & C2S). + + pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re 0) f) as C3. + rewrite C1F in C3. + rewrite C2F in C3. + rewrite C1S in C3. + rewrite C2S in C3. + rewrite Z.ltb_irrefl in C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt. lia. } + cbn. + rewrite Rmult_0_l. + gappa. + } + rewrite C2R in C3. + rewrite Rmult_0_l in C3. + destruct C3 as (C3R & C3F & C3Sz). + change (true && true) with true in C3F. + change (xorb false false) with false in C3Sz. + assert (is_nan 53 1024 + (Bmult 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re 0) f) = false) as NAN. + { apply is_finite_not_is_nan. + assumption. + } + pose proof (C3Sz NAN) as C3S. + clear NAN C3Sz. + pose proof ((ZofB_ne_range_correct 53 1024 + (Bmult 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re 0) f) 0 Int64.max_unsigned)) as C4. + rewrite C3R in C4. + replace (round radix2 (FLT_exp (3 - 1024 - 53) 53) (round_mode mode_NE) 0) + with 0%R in C4 by (cbn ; gappa). + rewrite Znearest_IZR in C4. + cbn zeta in C4. + rewrite Z.leb_refl in C4. + change (0 <=? Int64.max_unsigned)%Z with true in C4. + rewrite andb_true_r in C4. + rewrite andb_true_r in C4. + rewrite C3F in C4. + rewrite C4. + reflexivity. +Qed. Definition smallb_approx_range := 4400000000000%Z. Lemma step1_div_longu_correct : -- cgit From ff28b8ca5249e8c4ff0ec519673f71a382e816ad Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 13 Jan 2022 00:07:39 +0100 Subject: moved functions to a more logical place --- kvx/FPDivision32.v | 27 +---- lib/IEEE754_extra.v | 344 +++++++++++++++++++++++++++------------------------- 2 files changed, 180 insertions(+), 191 deletions(-) diff --git a/kvx/FPDivision32.v b/kvx/FPDivision32.v index 7c37c619..5cda1077 100644 --- a/kvx/FPDivision32.v +++ b/kvx/FPDivision32.v @@ -21,10 +21,7 @@ Require Import IEEE754_extra. From Gappa Require Import Gappa_tactic. Local Open Scope cminorsel_scope. - -Local Open Scope string_scope. -Local Open Scope error_monad_scope. - + Definition approx_inv b := let invb_s := Eop Oinvfs ((Eop Osingleofintu ((Eletvar 0%nat):::Enil)):::Enil) in let invb_d := Eop Ofloatofsingle (invb_s ::: Enil) in @@ -358,27 +355,6 @@ Qed. Opaque approx_inv. -Theorem Znearest_le - : forall choice (x y : R), (x <= y)%R -> Znearest choice x <= Znearest choice y. -Proof. - intros. - destruct (Z_le_gt_dec (Znearest choice x) (Znearest choice y)) as [LE | GT]. - assumption. - exfalso. - assert (1 <= IZR (Znearest choice x) - IZR(Znearest choice y))%R as GAP. - { rewrite <- minus_IZR. - apply IZR_le. - lia. - } - pose proof (Znearest_imp2 choice x) as Rx. - pose proof (Znearest_imp2 choice y) as Ry. - apply Rabs_def2b in Rx. - apply Rabs_def2b in Ry. - assert (x = y) by lra. - subst y. - lia. -Qed. - Theorem fp_divu32_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_a expr_b : expr) (a b : int) @@ -610,4 +586,3 @@ Proof. reflexivity. Qed. - diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 7e0e7260..636ea1ff 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -28,6 +28,184 @@ Require Import Coq.Logic.FunctionalExtensionality. Local Open Scope Z_scope. + +Lemma Znearest_IZR : + forall choice n, (Znearest choice (IZR n)) = n. +Proof. + intros. + unfold Znearest. + case Rcompare_spec ; intro ORDER. + - apply Zfloor_IZR. + - destruct choice. + + apply Zceil_IZR. + + apply Zfloor_IZR. + - apply Zceil_IZR. +Qed. + +Lemma ZnearestE_IZR: + forall n, (ZnearestE (IZR n)) = n. +Proof. + apply Znearest_IZR. +Qed. + +Lemma Zfloor_opp : + forall x : R, (Zfloor (- x)) = - (Zceil x). +Proof. + unfold Zceil, Zfloor. + intro x. + rewrite Z.opp_involutive. + reflexivity. +Qed. + +Lemma Zceil_opp : + forall x : R, (Zceil (- x)) = - (Zfloor x). +Proof. + unfold Zceil, Zfloor. + intro x. + rewrite Ropp_involutive. + reflexivity. +Qed. + +Lemma ZnearestE_opp + : forall x : R, ZnearestE (- x) = - ZnearestE x. +Proof. + intro. + rewrite Znearest_opp. + f_equal. + f_equal. + apply functional_extensionality. + intro. + rewrite Z.even_opp. + fold (Z.succ x0). + rewrite Z.even_succ. + f_equal. + apply Z.negb_odd. +Qed. + +Lemma Zceil_non_floor: + forall x : R, (x > IZR(Zfloor x))%R -> Zceil x = Z.succ(Zfloor x). +Proof. + intros x BETWEEN. + unfold Z.succ. + apply Zceil_imp. + split. + { rewrite minus_IZR. + rewrite plus_IZR. + lra. + } + rewrite plus_IZR. + pose proof (Zfloor_ub x). + lra. +Qed. + +(** more complicated way of proving +Lemma Zceil_non_ceil: + forall x : R, (x < IZR(Zceil x))%R -> Zceil x = Z.succ(Zfloor x). +Proof. + intros x BETWEEN. + unfold Z.succ. + cut (Zfloor x = (Zceil x) - 1). { intros; lia. } + apply Zfloor_imp. + split. + { rewrite minus_IZR. + pose proof (Zceil_lb x). + lra. + } + rewrite plus_IZR. + rewrite minus_IZR. + lra. +Qed. + +Lemma ZnearestE_opp + : forall x : R, ZnearestE (- x) = - ZnearestE x. +Proof. + intro x. + unfold ZnearestE. + case (Rcompare_spec (x - IZR (Zfloor x)) (/ 2)); intro CMP. + - pose proof (Zfloor_lb x) as LB. + destruct (Rcompare_spec x (IZR (Zfloor x))) as [ ABSURD | EXACT | INEXACT]. + lra. + { set (n := Zfloor x) in *. + rewrite EXACT. + rewrite <- opp_IZR. + rewrite Zfloor_IZR. + rewrite opp_IZR. + rewrite Rcompare_Lt by lra. + reflexivity. + } + rewrite Rcompare_Gt. + { apply Zceil_opp. } + rewrite Zfloor_opp. + rewrite opp_IZR. + rewrite Zceil_non_floor by assumption. + unfold Z.succ. + rewrite plus_IZR. + lra. + - rewrite Rcompare_Eq. + { rewrite Zceil_opp. + rewrite Zfloor_opp. + rewrite Z.even_opp. + rewrite Zceil_non_floor by lra. + rewrite Z.even_succ. + rewrite Z.negb_odd. + destruct (Z.even (Zfloor x)); reflexivity. + } + rewrite Zfloor_opp. + rewrite opp_IZR. + ring_simplify. + rewrite Zceil_non_floor by lra. + unfold Z.succ. + rewrite plus_IZR. + lra. + - rewrite Rcompare_Lt. + { apply Zfloor_opp. } + rewrite Zfloor_opp. + rewrite opp_IZR. + rewrite Zceil_non_floor by lra. + unfold Z.succ. + rewrite plus_IZR. + lra. +Qed. + *) + +Lemma Znearest_imp2: + forall choice x, (Rabs (IZR (Znearest choice x) - x) <= /2)%R. +Proof. + intros. + unfold Znearest. + pose proof (Zfloor_lb x) as FL. + pose proof (Zceil_ub x) as CU. + pose proof (Zceil_non_floor x) as NF. + case Rcompare_spec; intro CMP; apply Rabs_le; split; try lra. + - destruct choice; lra. + - destruct choice. 2: lra. + rewrite NF. 2: lra. + unfold Z.succ. rewrite plus_IZR. lra. + - rewrite NF. 2: lra. + unfold Z.succ. rewrite plus_IZR. lra. +Qed. + +Theorem Znearest_le + : forall choice (x y : R), (x <= y)%R -> Znearest choice x <= Znearest choice y. +Proof. + intros. + destruct (Z_le_gt_dec (Znearest choice x) (Znearest choice y)) as [LE | GT]. + assumption. + exfalso. + assert (1 <= IZR (Znearest choice x) - IZR(Znearest choice y))%R as GAP. + { rewrite <- minus_IZR. + apply IZR_le. + lia. + } + pose proof (Znearest_imp2 choice x) as Rx. + pose proof (Znearest_imp2 choice y) as Ry. + apply Rabs_le_inv in Rx. + apply Rabs_le_inv in Ry. + assert (x = y) by lra. + subst y. + lia. +Qed. + Section Extra_ops. (** [prec] is the number of bits of the mantissa including the implicit one. @@ -901,145 +1079,6 @@ Definition ZofB_ne (f: binary_float): option Z := | _ => None end. -Lemma Znearest_IZR : - forall choice n, (Znearest choice (IZR n)) = n. -Proof. - intros. - unfold Znearest. - case Rcompare_spec ; intro ORDER. - - apply Zfloor_IZR. - - destruct choice. - + apply Zceil_IZR. - + apply Zfloor_IZR. - - apply Zceil_IZR. -Qed. - -Lemma ZnearestE_IZR: - forall n, (ZnearestE (IZR n)) = n. -Proof. - apply Znearest_IZR. -Qed. - -Lemma Zfloor_opp : - forall x : R, (Zfloor (- x)) = - (Zceil x). -Proof. - unfold Zceil, Zfloor. - intro x. - rewrite Z.opp_involutive. - reflexivity. -Qed. - -Lemma Zceil_opp : - forall x : R, (Zceil (- x)) = - (Zfloor x). -Proof. - unfold Zceil, Zfloor. - intro x. - rewrite Ropp_involutive. - reflexivity. -Qed. - -Lemma ZnearestE_opp - : forall x : R, ZnearestE (- x) = - ZnearestE x. -Proof. - intro. - rewrite Znearest_opp. - f_equal. - f_equal. - apply functional_extensionality. - intro. - rewrite Z.even_opp. - fold (Z.succ x0). - rewrite Z.even_succ. - f_equal. - apply Z.negb_odd. -Qed. - -Lemma Zceil_non_floor: - forall x : R, (x > IZR(Zfloor x))%R -> Zceil x = Z.succ(Zfloor x). -Proof. - intros x BETWEEN. - unfold Z.succ. - apply Zceil_imp. - split. - { rewrite minus_IZR. - rewrite plus_IZR. - lra. - } - rewrite plus_IZR. - pose proof (Zfloor_ub x). - lra. -Qed. - -(** more complicated way of proving -Lemma Zceil_non_ceil: - forall x : R, (x < IZR(Zceil x))%R -> Zceil x = Z.succ(Zfloor x). -Proof. - intros x BETWEEN. - unfold Z.succ. - cut (Zfloor x = (Zceil x) - 1). { intros; lia. } - apply Zfloor_imp. - split. - { rewrite minus_IZR. - pose proof (Zceil_lb x). - lra. - } - rewrite plus_IZR. - rewrite minus_IZR. - lra. -Qed. - -Lemma ZnearestE_opp - : forall x : R, ZnearestE (- x) = - ZnearestE x. -Proof. - intro x. - unfold ZnearestE. - case (Rcompare_spec (x - IZR (Zfloor x)) (/ 2)); intro CMP. - - pose proof (Zfloor_lb x) as LB. - destruct (Rcompare_spec x (IZR (Zfloor x))) as [ ABSURD | EXACT | INEXACT]. - lra. - { set (n := Zfloor x) in *. - rewrite EXACT. - rewrite <- opp_IZR. - rewrite Zfloor_IZR. - rewrite opp_IZR. - rewrite Rcompare_Lt by lra. - reflexivity. - } - rewrite Rcompare_Gt. - { apply Zceil_opp. } - rewrite Zfloor_opp. - rewrite opp_IZR. - rewrite Zceil_non_floor by assumption. - unfold Z.succ. - rewrite plus_IZR. - lra. - - rewrite Rcompare_Eq. - { rewrite Zceil_opp. - rewrite Zfloor_opp. - rewrite Z.even_opp. - rewrite Zceil_non_floor by lra. - rewrite Z.even_succ. - rewrite Z.negb_odd. - destruct (Z.even (Zfloor x)); reflexivity. - } - rewrite Zfloor_opp. - rewrite opp_IZR. - ring_simplify. - rewrite Zceil_non_floor by lra. - unfold Z.succ. - rewrite plus_IZR. - lra. - - rewrite Rcompare_Lt. - { apply Zfloor_opp. } - rewrite Zfloor_opp. - rewrite opp_IZR. - rewrite Zceil_non_floor by lra. - unfold Z.succ. - rewrite plus_IZR. - lra. -Qed. - *) - Ltac field_simplify_den := field_simplify ; [idtac | lra]. Ltac Rdiv_lt_0_den := apply Rdiv_lt_0_compat ; [idtac | lra]. @@ -1149,37 +1188,12 @@ Proof. lia. Qed. - -Lemma Znearest_imp2: - forall choice x, (Rabs (IZR (Znearest choice x) - x) <= /2)%R. -Proof. - intros. - unfold Znearest. - pose proof (Zfloor_lb x) as FL. - pose proof (Zceil_ub x) as CU. - pose proof (Zceil_non_floor x) as NF. - case Rcompare_spec; intro CMP; apply Rabs_le; split; try lra. - - destruct choice; lra. - - destruct choice. 2: lra. - rewrite NF. 2: lra. - unfold Z.succ. rewrite plus_IZR. lra. - - rewrite NF. 2: lra. - unfold Z.succ. rewrite plus_IZR. lra. -Qed. - -Lemma Rabs_le_rev : forall a b, (Rabs a <= b -> -b <= a <= b)%R. -Proof. - intros a b ABS. - unfold Rabs in ABS. - destruct Rcase_abs in ABS; lra. -Qed. - Theorem ZofB_ne_ball: forall f n, ZofB_ne f = Some n -> (IZR n-1/2 <= B2R _ _ f <= IZR n+1/2)%R. Proof. intros. rewrite ZofB_ne_correct in H. destruct (is_finite prec emax f) eqn:FIN; inversion H. pose proof (Znearest_imp2 (fun x => negb (Z.even x)) (B2R prec emax f)) as ABS. - pose proof (Rabs_le_rev _ _ ABS). + pose proof (Rabs_le_inv _ _ ABS). lra. Qed. -- cgit From 57780f0ec50350d7158c242ed77047ec6cbf5791 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 13 Jan 2022 20:52:22 +0100 Subject: relative bound of error --- kvx/FPDivision64.v | 233 ++++++++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 229 insertions(+), 4 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index fc2050bb..f9557335 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -49,7 +49,7 @@ Qed. Definition approx_inv_thresh := (25/2251799813685248)%R. (* 1.11022302462516e-14 *) -Theorem approx_inv_longu_correct : +Theorem approx_inv_longu_correct_abs : forall b, (0 < Int64.unsigned b)%Z -> exists (f : float), @@ -283,6 +283,234 @@ Proof. gappa. Qed. +Local Notation "'rd'" := (round radix2 (FLT_exp (-1074) 53) ZnearestE). +Local Notation "'rs'" := (round radix2 (FLT_exp (-149) 24) ZnearestE). + +Definition approx_inv_rel_thresh := (1049/72057594037927936)%R. +Theorem approx_inv_longu_correct_rel : + forall b, + (0 < Int64.unsigned b)%Z -> + exists (f : float), + (approx_inv_longu (Vlong b)) = Vfloat f /\ + is_finite _ _ f = true /\ (Rabs(IZR (Int64.unsigned b) * (B2R _ _ f) - 1) <= approx_inv_rel_thresh)%R. +Proof. + intros b NONZ. + unfold approx_inv_longu. + cbn. + econstructor. + split. + reflexivity. + Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int. + unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int. + set (re := (@eq_refl Datatypes.comparison Lt)). + change (Int.signed (Int.repr 1)) with 1%Z. + set (b' := Int64.unsigned b) in *. + pose proof (Int64.unsigned_range b) as RANGE. + change Int64.modulus with 18446744073709551616%Z in RANGE. + assert(1 <= IZR b' <= 18446744073709551616)%R as RANGE'. + { split; apply IZR_le; lia. + } + + assert (-16777216 <= 1 <= 16777216)%Z as SILLY by lia. + destruct (BofZ_exact 24 128 re re 1 SILLY) as (C0R & C0F & _). + clear SILLY. + set (one_s := (BofZ 24 128 re re 1)) in *. + + pose proof (BofZ_correct 24 128 re re b') as C1. + cbn in C1. + rewrite Rlt_bool_true in C1; cycle 1. + { clear C1. + eapply (Rabs_relax (IZR 18446744073709551616)). + lra. + set (b'' := IZR b') in *. + gappa. + } + destruct C1 as (C1R & C1F & _). + set (b_s := (BofZ 24 128 re re b')) in *. + + assert(1 <= B2R 24 128 b_s <= 18446744073709551616)%R as b_s_RANGE. + { rewrite C1R. + gappa. + } + assert(B2R 24 128 b_s <> 0)%R as b_s_NONZ by lra. + + pose proof (Bdiv_correct 24 128 re re Float32.binop_nan mode_NE one_s b_s b_s_NONZ) as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply Rabs_relax with (b := 1%R). + { cbn; lra. } + rewrite C0R. + set (r_b_s := B2R 24 128 b_s) in *. + cbn. + gappa. + } + + destruct C2 as (C2R & C2F & _). + set (invb_s := (Bdiv 24 128 re re Float32.binop_nan mode_NE one_s b_s)) in *. + rewrite C0F in C2F. + + assert ((1/18446744073709551616 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. + { rewrite C2R. + set (r_b_s := B2R 24 128 b_s) in *. + rewrite C0R. + cbn. + gappa. + } + + pose proof (Bconv_correct 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s C2F) as C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + set (r_invb_s := (B2R 24 128 invb_s)) in *. + apply Rabs_relax with (b := 1%R). + { replace 1%R with (bpow radix2 0)%R by reflexivity. + apply bpow_lt. + lia. + } + cbn. + gappa. + } + + destruct C3 as (C3R & C3F & _). + set (invb_d := (Bconv 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s)) in *. + assert ((1/18446744073709551616 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. + { + rewrite C3R. + set (r_invb_s := B2R 24 128 invb_s) in *. + cbn. + gappa. + } + + pose proof (is_finite_Bopp 53 1024 Float.neg_nan invb_d) as opp_finite. + rewrite C3F in opp_finite. + + pose proof (BofZ_correct 53 1024 re re 1) as C4. + rewrite Rlt_bool_true in C4; cycle 1. + { clear C4. + cbn. + eapply (Rabs_relax (IZR 18446744073709551616)). + lra. + set (b'' := IZR b') in *. + gappa. + } + destruct C4 as (C4R & C4F & _). + + pose proof (BofZ_correct 53 1024 re re b') as C5. + cbn in C5. + rewrite Rlt_bool_true in C5; cycle 1. + { clear C5. + eapply (Rabs_relax (IZR 18446744073709551616)). + lra. + set (b'' := IZR b') in *. + gappa. + } + destruct C5 as (C5R & C5F & _). + set (b_d := (BofZ 53 1024 re re b')) in *. + + assert(1 <= B2R 53 1024 b_d <= 18446744073709551616)%R as b_d_RANGE. + { rewrite C5R. + gappa. + } + + pose proof (Bfma_correct 53 1024 re re Float.fma_nan mode_NE + (Bopp 53 1024 Float.neg_nan invb_d) (BofZ 53 1024 re re b') + (BofZ 53 1024 re re 1) opp_finite C5F C4F) as C6. + rewrite Rlt_bool_true in C6; cycle 1. + { clear C6. + rewrite C4R. + rewrite B2R_Bopp. + cbn. + eapply (Rabs_relax (IZR 18446744073709551616)). + { lra. } + fold invb_d. + fold b_d. + set (r_invb_d := B2R 53 1024 invb_d) in *. + set (r_b_d := B2R 53 1024 b_d) in *. + gappa. + } + fold b_d in C6. + destruct C6 as (C6R & C6F & _). + + pose proof (Bfma_correct 53 1024 re re Float.fma_nan mode_NE + (Bfma 53 1024 re re Float.fma_nan mode_NE + (Bopp 53 1024 Float.neg_nan invb_d) b_d (BofZ 53 1024 re re 1)) + invb_d invb_d C6F C3F C3F) as C7. + rewrite Rlt_bool_true in C7; cycle 1. + { clear C7. + rewrite C6R. + rewrite B2R_Bopp. + eapply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + rewrite C4R. + cbn. + set (r_invb_d := B2R 53 1024 invb_d) in *. + set (r_b_d := B2R 53 1024 b_d) in *. + gappa. + } + destruct C7 as (C7R & C7F & _). + + split. assumption. + rewrite C7R. + rewrite C6R. + rewrite C5R. + rewrite C4R. + rewrite B2R_Bopp. + rewrite C3R. + rewrite C2R. + rewrite C1R. + rewrite C0R. + cbn. + set(b1 := IZR b') in *. + + replace (rd 1) with 1%R by gappa. + replace (rd (rs (1 / rs b1))) with + ((((rd (rs (1 / rs b1)) - (/b1))/(/b1))+1)*(/ b1))%R ; cycle 1. + { field. lra. } + set (er0 := ((rd (rs (1 / rs b1)) - (/b1))/(/b1))%R). + replace (rd b1) with ((((rd b1) - b1)/b1 + 1) * b1)%R; cycle 1. + { field. lra. } + set (er1 := (((rd b1) - b1)/b1)%R). + replace (- ((er0 + 1) * / b1) * ((er1 + 1) * b1) + 1)%R + with (1 - (er0 + 1)*(er1 + 1))%R ; cycle 1. + { field. lra. } + set (z0 := (1 - (er0 + 1) * (er1 + 1))%R). + assert (Rabs er0 <= 257/2147483648)%R as er0_ABS. + { unfold er0. + gappa. + } + assert (Rabs er1 <= 1/9007199254740992)%R as er1_ABS. + { unfold er1. + gappa. + } + replace (rd z0) with ((rd(z0)-z0)+z0)%R by ring. + set (ea0 := (rd(z0)-z0)%R). + assert (Rabs ea0 <= 1/75557863725914323419136)%R as ea0_ABS. + { unfold ea0. unfold z0. + gappa. + } + set (z1 := ((ea0 + z0) * ((er0 + 1) * / b1) + (er0 + 1) * / b1)%R). + replace (rd z1) with ((((rd z1)-z1)/z1+1)*z1)%R; cycle 1. + { field. + unfold z1. + unfold z0. + gappa. + } + set (er2 := ((rd z1 - z1) / z1)%R). + assert (Rabs er2 <= 1/9007199254740992)%R as er2_ABS. + { unfold er2. + unfold z1, z0. + gappa. + } + unfold z1, z0. + replace (b1 * + ((er2 + 1) * + ((ea0 + (1 - (er0 + 1) * (er1 + 1))) * ((er0 + 1) * / b1) + + (er0 + 1) * / b1)) - 1)%R + with (-er0*er0*er1*er2 - er0*er0*er1 + ea0*er0*er2 - er0*er0*er2 - 2*er0*er1*er2 + ea0*er0 - er0*er0 - 2*er0*er1 + ea0*er2 - er1*er2 + ea0 - er1 + er2)%R; cycle 1. + { field. lra. } + unfold approx_inv_rel_thresh. + gappa. +Qed. + Definition step1_real_inv_longu b := let invb_s := ExtValues.invfs (Val.maketotal (Val.singleoflongu b)) in Val.floatofsingle invb_s. @@ -290,9 +518,6 @@ Definition step1_real_inv_longu b := Definition step1_real_inv_thresh := (3/33554432)%R. (* 8.94069671630859e-8 *) -Local Notation "'rd'" := (round radix2 (FLT_exp (-1074) 53) ZnearestE). -Local Notation "'rs'" := (round radix2 (FLT_exp (-149) 24) ZnearestE). - Theorem step1_real_inv_longu_correct : forall b, (0 < Int64.unsigned b)%Z -> -- cgit From 0742df3a567f98445e95669f92a291fda733d286 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 14 Jan 2022 11:14:36 +0100 Subject: some progress --- kvx/FPDivision64.v | 75 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 75 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index f9557335..767e7ae0 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -952,3 +952,78 @@ Proof. set (delta2 := (q_r * IZR b' - IZR a')%R) in *. gappa. Qed. + +Lemma range_up_le : + forall a x b, + (IZR a <= IZR x <= (IZR b) - 1)%R -> + (a <= x < b)%Z. +Proof. + intros until b. intro RANGE. + split. + { apply le_IZR. lra. } + assert (x <= b-1)%Z. + { apply le_IZR. rewrite minus_IZR. lra. } + lia. +Qed. + +Lemma find_quotient: + forall (a b : Z) + (b_POS : (b > 0)%Z) + (invb : R) + (eps : R) + (invb_EPS : (Rabs (invb * IZR b - 1) <= eps)%R) + (SMALL : (IZR (Z.abs a)*eps < IZR b / 2)%R), + (a / b)%Z = + let q := ZnearestE ((IZR a) * invb) in + if (b*q >? a)%Z + then (q-1)%Z + else q. +Proof. + intros. + rewrite <- Rabs_Zabs in SMALL. + set (q := ZnearestE (IZR a * invb)%R). + cbn zeta. + set (b' := IZR b) in *. + set (a' := IZR a) in *. + assert (1 <= b')%R as b_POS'. + { apply IZR_le. + lia. + } + + pose proof (Znearest_imp2 (fun x : Z => negb (Z.even x)) (a' * invb)) as NEAR. + fold q in NEAR. + set (q' := IZR q) in *. + assert ((Rabs a') * Rabs (invb * b' - 1) <= (Rabs a') * eps)%R as S1. + { apply Rmult_le_compat_l. + apply Rabs_pos. + assumption. } + rewrite <- Rabs_mult in S1. + assert ((Rabs b') * Rabs (q' - a' * invb) <= (Rabs b') / 2)%R as S2. + { apply Rmult_le_compat_l. + apply Rabs_pos. + assumption. } + rewrite <- Rabs_mult in S2. + rewrite (Rabs_right b') in S2 by lra. + pose proof (Rabs_triang (a' * (invb * b' - 1)) + (b' * (q' - a' * invb))) as TRIANGLE. + replace (a' * (invb * b' - 1) + b' * (q' - a' * invb))%R with + (b' * q' - a')%R in TRIANGLE by ring. + assert (Rabs (b' * q' - a') < b')%R as DELTA by lra. + + pose proof (Zgt_cases (b * q) a)%Z as CASE. + destruct (_ >? _)%Z. + { admit. } + apply Zdiv_unique with (b := (a - q*b)%Z). + ring. + apply range_up_le. + rewrite minus_IZR. + rewrite mult_IZR. + fold a' b' q'. + + + Search (Rabs _ * Rabs _)%R. +Definition step2_real_div_long a b := + Val.mulf (Val.maketotal (Val.floatoflong a)) (step1_real_inv_longu b). + +Definition step2_div_long a b := + Val.maketotal (Val.longuoffloat_ne (step2_real_div_long a)). -- cgit From a5df0e05f97f5ce4ee35723089cbd4a83435d37f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 14 Jan 2022 11:39:22 +0100 Subject: some more proof --- kvx/FPDivision64.v | 21 +++++++++++++++------ 1 file changed, 15 insertions(+), 6 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 767e7ae0..6d82ae67 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -953,6 +953,7 @@ Proof. gappa. Qed. +(* Lemma range_up_le : forall a x b, (IZR a <= IZR x <= (IZR b) - 1)%R -> @@ -965,6 +966,7 @@ Proof. { apply le_IZR. rewrite minus_IZR. lra. } lia. Qed. + *) Lemma find_quotient: forall (a b : Z) @@ -1015,13 +1017,20 @@ Proof. { admit. } apply Zdiv_unique with (b := (a - q*b)%Z). ring. - apply range_up_le. - rewrite minus_IZR. - rewrite mult_IZR. - fold a' b' q'. - - Search (Rabs _ * Rabs _)%R. + split. + { rewrite Z.mul_comm. lia. } + rewrite <- Rabs_Ropp in DELTA. + unfold b', q', a' in DELTA. + rewrite <- mult_IZR in DELTA. + rewrite <- minus_IZR in DELTA. + rewrite <- opp_IZR in DELTA. + rewrite Rabs_Zabs in DELTA. + apply lt_IZR in DELTA. + rewrite Z.abs_eq in DELTA by lia. + lia. +Admitted. + Definition step2_real_div_long a b := Val.mulf (Val.maketotal (Val.floatoflong a)) (step1_real_inv_longu b). -- cgit From 0ca71551f9d47dcc04e7a97bb294a0379732bc04 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 14 Jan 2022 11:47:58 +0100 Subject: find_quotient --- kvx/FPDivision64.v | 24 +++++++++++++++++------- 1 file changed, 17 insertions(+), 7 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 6d82ae67..8f0d2fa1 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1014,12 +1014,18 @@ Proof. pose proof (Zgt_cases (b * q) a)%Z as CASE. destruct (_ >? _)%Z. - { admit. } - apply Zdiv_unique with (b := (a - q*b)%Z). - ring. + { unfold b', q', a' in DELTA. + rewrite <- mult_IZR in DELTA. + rewrite <- minus_IZR in DELTA. + rewrite Rabs_Zabs in DELTA. + apply lt_IZR in DELTA. + rewrite Z.abs_eq in DELTA by lia. + + apply Zdiv_unique with (b := (a - (q-1)*b)%Z). + ring. + split; lia. + } - split. - { rewrite Z.mul_comm. lia. } rewrite <- Rabs_Ropp in DELTA. unfold b', q', a' in DELTA. rewrite <- mult_IZR in DELTA. @@ -1028,8 +1034,12 @@ Proof. rewrite Rabs_Zabs in DELTA. apply lt_IZR in DELTA. rewrite Z.abs_eq in DELTA by lia. - lia. -Admitted. + + apply Zdiv_unique with (b := (a - q*b)%Z). + ring. + + split; lia. +Qed. Definition step2_real_div_long a b := Val.mulf (Val.maketotal (Val.floatoflong a)) (step1_real_inv_longu b). -- cgit From d403269fa27446c1f7281e35a51ea0eb6483c987 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 14 Jan 2022 21:10:30 +0100 Subject: find_quotient --- kvx/FPDivision64.v | 64 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 64 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 8f0d2fa1..b5052917 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -968,6 +968,7 @@ Proof. Qed. *) +(* Lemma find_quotient: forall (a b : Z) (b_POS : (b > 0)%Z) @@ -1040,6 +1041,69 @@ Proof. split; lia. Qed. + *) + +Lemma find_quotient: + forall (a b : Z) + (b_POS : (b > 0)%Z) + (invb : R) + (qr : R) + (GAP : (Rabs (IZR a / IZR b - qr) < /2)%R), + (a / b)%Z = + let q := ZnearestE qr in + if (b*q >? a)%Z + then (q-1)%Z + else q. +Proof. + intros. + set (q := ZnearestE qr). + cbn zeta. + set (b' := IZR b) in *. + set (a' := IZR a) in *. + assert (1 <= b')%R as b_POS'. + { apply IZR_le. + lia. + } + + pose proof (Znearest_imp2 (fun x : Z => negb (Z.even x)) qr) as ROUND. + fold q in ROUND. + set (q' := IZR q) in *. + + pose proof (Rabs_triang (a' / b' - qr) + (qr - q'))%R as TRIANGLE. + replace ((a' / b' - qr) + (qr - q'))%R with + (a' / b' - q')%R in TRIANGLE by ring. + rewrite <- Rabs_Ropp in ROUND. + replace (- (q' - qr))%R with (qr - q')%R in ROUND by ring. + assert (Z.abs (a - b*q) < b)%Z as DELTA. + { apply lt_IZR. + rewrite <- Rabs_Zabs. + rewrite minus_IZR. + rewrite mult_IZR. + fold a' q' b'. + apply Rmult_lt_reg_r with (r := (/b')%R). + { apply Rinv_0_lt_compat. lra. } + rewrite Rinv_r by lra. + replace (/ b')%R with (/ Rabs(b'))%R ; cycle 1. + { f_equal. + apply Rabs_pos_eq. lra. } + rewrite <- Rabs_Rinv by lra. + rewrite <- Rabs_mult. + replace ((a' - b' * q') * / b')%R with (a'/b' - q')%R by (field ; lra). + lra. + } + + pose proof (Zgt_cases (b * q) a)%Z as CASE. + destruct (_ >? _)%Z. + { apply Zdiv_unique with (b := (a - (q-1)*b)%Z). + ring. + split; lia. + } + + apply Zdiv_unique with (b := (a - q*b)%Z). + ring. + split; lia. +Qed. Definition step2_real_div_long a b := Val.mulf (Val.maketotal (Val.floatoflong a)) (step1_real_inv_longu b). -- cgit From 0284db0e9b9f73fa98f2c92defdf61ea90f91b1a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 19 Jan 2022 15:51:12 +0100 Subject: so that it compiles --- kvx/FPDivision64.v | 60 ++++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 58 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index b5052917..64ac77a1 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1106,7 +1106,63 @@ Proof. Qed. Definition step2_real_div_long a b := - Val.mulf (Val.maketotal (Val.floatoflong a)) (step1_real_inv_longu b). + Val.mulf (Val.maketotal (Val.floatoflong a)) (approx_inv_longu b). + +Definition smalla_thresh := 35184372088832%Z. + +Lemma step2_real_div_long_smalla_correct : + forall (a b : int64) + (a_SMALL : (Z.abs (Int64.signed a) <= smalla_thresh)%Z) + (b_NOT0 : (0 < Int64.unsigned b)%Z), + exists (q : float), + (step2_real_div_long (Vlong a) (Vlong b)) = Vfloat q /\ + (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) < (1/2))%R. +Proof. + intros. + unfold step2_real_div_long. + destruct (approx_inv_longu_correct_rel b b_NOT0) as (f & C0E & C0F & C0R). + rewrite C0E. + econstructor. + split. reflexivity. + Local Transparent Float.of_long. + unfold Float.mul, Float.of_long. + set (re := (@eq_refl Datatypes.comparison Lt)) in *. + set (a' := Int64.signed a) in *. + set (b' := Int64.unsigned b) in *. + assert(Rabs (IZR a') <= IZR smalla_thresh)%R as a_RANGE'. + { rewrite Rabs_Zabs. + apply IZR_le. + assumption. + } + assert (- 2 ^ 53 <= a' <= 2 ^ 53)%Z as SILLY. + { unfold smalla_thresh in a_SMALL. + apply Z.abs_le. + lia. + } + destruct (BofZ_exact 53 1024 re re (Int64.signed a) SILLY) as (C1R & C1F & C1S). + fold a' in C1R. + pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE (BofZ 53 1024 re re a') f) as C2. + rewrite Rlt_bool_true in C2 ; cycle 1. + { apply Rabs_relax with (b := bpow radix2 53). + { apply bpow_lt. lia. } + cbn. + rewrite C1R. + unfold approx_inv_rel_thresh in C0R. + admit. + } + cbn. +Admitted. + +Definition step2_div_long' a b := + Val.maketotal (Val.longuoffloat_ne (step2_real_div_long a b)). Definition step2_div_long a b := - Val.maketotal (Val.longuoffloat_ne (step2_real_div_long a)). + let q := Val.maketotal (Val.longuoffloat_ne (step2_div_long' a b)) in + Val.select (Val.cmpl_bool Clt (Val.subl a (Val.mull q b)) (Vlong Int64.zero)) + (Val.addl q (Vlong Int64.mone)) q Tlong. + +Lemma step2_div_long_smalla_correct : + forall a b, + (Z.abs (Int64.signed a) <= smalla_thresh)%Z -> + step2_div_long (Vlong a) (Vlong b) = Vlong (Int64.repr (Int64.signed a / Int64.unsigned b))%Z. +Admitted. -- cgit From a7db868181d74403449ce436028f0658f284088e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 09:23:43 +0100 Subject: progress --- kvx/FPDivision64.v | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 64ac77a1..0dfa9882 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1127,8 +1127,13 @@ Proof. Local Transparent Float.of_long. unfold Float.mul, Float.of_long. set (re := (@eq_refl Datatypes.comparison Lt)) in *. + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in b_RANGE. set (a' := Int64.signed a) in *. set (b' := Int64.unsigned b) in *. + assert (1 <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. + } assert(Rabs (IZR a') <= IZR smalla_thresh)%R as a_RANGE'. { rewrite Rabs_Zabs. apply IZR_le. @@ -1143,12 +1148,17 @@ Proof. fold a' in C1R. pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE (BofZ 53 1024 re re a') f) as C2. rewrite Rlt_bool_true in C2 ; cycle 1. - { apply Rabs_relax with (b := bpow radix2 53). + { clear C2. + apply Rabs_relax with (b := bpow radix2 53). { apply bpow_lt. lia. } cbn. rewrite C1R. unfold approx_inv_rel_thresh in C0R. - admit. + replace (B2R 53 1024 f) with + ((1/IZR b') * ((IZR b' * B2R 53 1024 f - 1) + 1))%R ; cycle 1. + { field. lra. } + unfold smalla_thresh in *. + gappa. } cbn. Admitted. -- cgit From fa79debdca4121f2435750b8090e83b541887979 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 09:35:23 +0100 Subject: progress --- kvx/FPDivision64.v | 11 +++++++++-- 1 file changed, 9 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 0dfa9882..8d40e88e 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1145,7 +1145,7 @@ Proof. lia. } destruct (BofZ_exact 53 1024 re re (Int64.signed a) SILLY) as (C1R & C1F & C1S). - fold a' in C1R. + fold a' in C1R, C1F, C1S. pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE (BofZ 53 1024 re re a') f) as C2. rewrite Rlt_bool_true in C2 ; cycle 1. { clear C2. @@ -1160,7 +1160,14 @@ Proof. unfold smalla_thresh in *. gappa. } - cbn. + rewrite C0F in C2. + rewrite C1R in C2. + rewrite C1F in C2. + rewrite C1S in C2. + cbn in C2. + destruct C2 as (C2R & C2F & _). + rewrite C2R. + admit. Admitted. Definition step2_div_long' a b := -- cgit From 3b36bd818ffc20f7a455da46a89641c02491a4e1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 10:02:43 +0100 Subject: progress --- kvx/FPDivision64.v | 20 +++++++++++++++++--- 1 file changed, 17 insertions(+), 3 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 8d40e88e..6894460f 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1116,7 +1116,7 @@ Lemma step2_real_div_long_smalla_correct : (b_NOT0 : (0 < Int64.unsigned b)%Z), exists (q : float), (step2_real_div_long (Vlong a) (Vlong b)) = Vfloat q /\ - (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) < (1/2))%R. + (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (255/256))%R. Proof. intros. unfold step2_real_div_long. @@ -1167,8 +1167,22 @@ Proof. cbn in C2. destruct C2 as (C2R & C2F & _). rewrite C2R. - admit. -Admitted. + replace (IZR a' * (B2R 53 1024 f))%R with + ((IZR a'/IZR b') * ((IZR b' * B2R 53 1024 f - 1) + 1))%R ; cycle 1. + { field. lra. } + set (delta1 := (IZR b' * B2R 53 1024 f - 1)%R) in *. + set (q1 := (IZR a' / IZR b' * (delta1 + 1))%R). + replace (rd q1) with (((rd q1) - q1) + q1)%R by ring. + set (delta2 := ((rd q1) - q1)%R). + unfold q1. + replace (delta2 + IZR a' / IZR b' * (delta1 + 1) - IZR a' / IZR b')%R with + (delta2 + (IZR a' / IZR b') * delta1)%R by ring. + unfold delta2. + unfold q1. + unfold approx_inv_rel_thresh in *. + unfold smalla_thresh in *. + gappa. +Qed. Definition step2_div_long' a b := Val.maketotal (Val.longuoffloat_ne (step2_real_div_long a b)). -- cgit From db85b47c5487ddfa051b01e444d2be555013e7f7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 10:04:21 +0100 Subject: progress --- kvx/FPDivision64.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 6894460f..9ed12064 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1108,7 +1108,7 @@ Qed. Definition step2_real_div_long a b := Val.mulf (Val.maketotal (Val.floatoflong a)) (approx_inv_longu b). -Definition smalla_thresh := 35184372088832%Z. +Definition smalla_thresh := 34184372088832%Z. Lemma step2_real_div_long_smalla_correct : forall (a b : int64) @@ -1116,7 +1116,7 @@ Lemma step2_real_div_long_smalla_correct : (b_NOT0 : (0 < Int64.unsigned b)%Z), exists (q : float), (step2_real_div_long (Vlong a) (Vlong b)) = Vfloat q /\ - (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (255/256))%R. + (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (32767/65536))%R. Proof. intros. unfold step2_real_div_long. -- cgit From 6e45966f1cd5ea6feadb044f2fda2e1b1c81e57c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 10:07:56 +0100 Subject: progress --- kvx/FPDivision64.v | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 9ed12064..72ad43d1 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1116,7 +1116,8 @@ Lemma step2_real_div_long_smalla_correct : (b_NOT0 : (0 < Int64.unsigned b)%Z), exists (q : float), (step2_real_div_long (Vlong a) (Vlong b)) = Vfloat q /\ - (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (32767/65536))%R. + (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (32767/65536))%R /\ + is_finite _ _ q = true. Proof. intros. unfold step2_real_div_long. @@ -1166,6 +1167,8 @@ Proof. rewrite C1S in C2. cbn in C2. destruct C2 as (C2R & C2F & _). + split. + 2: exact C2F. rewrite C2R. replace (IZR a' * (B2R 53 1024 f))%R with ((IZR a'/IZR b') * ((IZR b' * B2R 53 1024 f - 1) + 1))%R ; cycle 1. -- cgit From 64552329d9e6c1ae4071f2f144f7c767287f75c6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 11:31:04 +0100 Subject: fix issue --- kvx/FPDivision64.v | 71 ++++++++++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 66 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 72ad43d1..7f7aa71a 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1188,15 +1188,76 @@ Proof. Qed. Definition step2_div_long' a b := - Val.maketotal (Val.longuoffloat_ne (step2_real_div_long a b)). + Val.maketotal (Val.longoffloat_ne (step2_real_div_long a b)). Definition step2_div_long a b := - let q := Val.maketotal (Val.longuoffloat_ne (step2_div_long' a b)) in + let q := step2_div_long' a b in Val.select (Val.cmpl_bool Clt (Val.subl a (Val.mull q b)) (Vlong Int64.zero)) (Val.addl q (Vlong Int64.mone)) q Tlong. +Lemma Znearest_lub : + forall choice (n : Z) (x : R), (IZR n <= x)%R -> (n <= Znearest choice x)%Z. +Proof. + intros until x. intro BND. + pose proof (Zfloor_lub n x BND). + pose proof (Znearest_ge_floor choice x). + lia. +Qed. + +Lemma Znearest_glb : + forall choice (n : Z) (x : R), (x <= IZR n)%R -> (Znearest choice x <= n)%Z. +Proof. + intros until x. intro BND. + pose proof (Zceil_glb n x BND). + pose proof (Znearest_le_ceil choice x). + lia. +Qed. + Lemma step2_div_long_smalla_correct : - forall a b, - (Z.abs (Int64.signed a) <= smalla_thresh)%Z -> + forall a b + (a_SMALL : (Z.abs (Int64.signed a) <= smalla_thresh)%Z) + (b_NOT0 : (0 < Int64.unsigned b)%Z), step2_div_long (Vlong a) (Vlong b) = Vlong (Int64.repr (Int64.signed a / Int64.unsigned b))%Z. -Admitted. +Proof. + intros. + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in b_RANGE. + set (a' := (Int64.signed a)) in *. + set (b' := (Int64.unsigned b)) in *. + assert (Rabs (IZR a') <= IZR smalla_thresh)%R as a_RANGE'. + { rewrite Rabs_Zabs. + apply IZR_le. + assumption. + } + assert (1 <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. + } + destruct (step2_real_div_long_smalla_correct a b a_SMALL b_NOT0) as (q & C1R & C1E & C1F). + fold a' b' in C1E. + assert ((Int64.min_signed <=? ZnearestE (B2R 53 1024 q))=true)%Z as q_LOW. + { apply Zle_imp_le_bool. + change Int64.min_signed with (-9223372036854775808)%Z. + set (q' := B2R 53 1024 q) in *. + apply Znearest_lub. + replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + unfold smalla_thresh in a_RANGE'. + gappa. + } + assert ((ZnearestE (B2R 53 1024 q) <=? Int64.max_signed)=true)%Z as q_HIGH. + { apply Zle_imp_le_bool. + change Int64.max_signed with (9223372036854775807)%Z. + apply Znearest_glb. + set (q' := B2R 53 1024 q) in *. + replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + unfold smalla_thresh in a_RANGE'. + gappa. + } + unfold step2_div_long, step2_div_long'. + rewrite C1R. + cbn. + unfold Float.to_long_ne. + rewrite (ZofB_ne_range_correct _ _ q Int64.min_signed Int64.max_signed). + rewrite C1F. + rewrite q_LOW. + rewrite q_HIGH. + cbn. -- cgit From f2ea4988ec81d931ad035595855088c4cb667477 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 11:40:16 +0100 Subject: some progress --- kvx/FPDivision64.v | 26 +++++++++++++++++++++++++- 1 file changed, 25 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 7f7aa71a..63b48140 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1213,6 +1213,23 @@ Proof. lia. Qed. +Lemma function_ite : + forall {A B : Type} (fn : A->B) (b : bool) (x y: A), + fn (if b then x else y) = (if b then fn x else fn y). +Proof. + intros. + destruct b; reflexivity. +Qed. + +Lemma normalize_ite : + forall ty (b : bool) x y, + Val.normalize (if b then x else y) ty = + (if b then Val.normalize x ty else Val.normalize y ty). +Proof. + intros. + destruct b; reflexivity. +Qed. + Lemma step2_div_long_smalla_correct : forall a b (a_SMALL : (Z.abs (Int64.signed a) <= smalla_thresh)%Z) @@ -1237,8 +1254,8 @@ Proof. assert ((Int64.min_signed <=? ZnearestE (B2R 53 1024 q))=true)%Z as q_LOW. { apply Zle_imp_le_bool. change Int64.min_signed with (-9223372036854775808)%Z. - set (q' := B2R 53 1024 q) in *. apply Znearest_lub. + set (q' := B2R 53 1024 q) in *. replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. unfold smalla_thresh in a_RANGE'. gappa. @@ -1261,3 +1278,10 @@ Proof. rewrite q_LOW. rewrite q_HIGH. cbn. + rewrite normalize_ite. + cbn. + rewrite <- (function_ite Vlong). + f_equal. + set (q' := B2R 53 1024 q) in *. + + -- cgit From c61b77be54b4696e5c5d5ed2a60f00fbc3dedb56 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 14:35:52 +0100 Subject: progress --- kvx/FPDivision64.v | 70 ++++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 68 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 63b48140..766ff80f 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1192,7 +1192,7 @@ Definition step2_div_long' a b := Definition step2_div_long a b := let q := step2_div_long' a b in - Val.select (Val.cmpl_bool Clt (Val.subl a (Val.mull q b)) (Vlong Int64.zero)) + Val.select (Val.cmpl_bool Cgt (Val.subl (Val.mull q b) a) (Vlong Int64.zero)) (Val.addl q (Vlong Int64.mone)) q Tlong. Lemma Znearest_lub : @@ -1229,7 +1229,44 @@ Proof. intros. destruct b; reflexivity. Qed. - + + +Lemma int64_mul_signed_unsigned: + forall x y : int64, + Int64.mul x y = Int64.repr (Int64.signed x * Int64.unsigned y). +Proof. + intros. + unfold Int64.mul. + apply Int64.eqm_samerepr. + apply Int64.eqm_mult. + - apply Int64.eqm_sym. + apply Int64.eqm_signed_unsigned. + - apply Int64.eqm_refl. +Qed. + +Lemma int64_eqm_signed_repr: + forall z : Z, Int64.eqm z (Int64.signed (Int64.repr z)). +Proof. + intros. + apply Int64.eqm_trans with (y := Int64.unsigned (Int64.repr z)). + - apply Int64.eqm_unsigned_repr. + - apply Int64.eqm_sym. + apply Int64.eqm_signed_unsigned. +Qed. + +Lemma signed_repr_sub: + forall x y, + Int64.repr (Int64.signed (Int64.repr x) - y) = + Int64.repr (x - y). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_sub. + - apply Int64.eqm_sym. + apply int64_eqm_signed_repr. + - apply Int64.eqm_refl. +Qed. + Lemma step2_div_long_smalla_correct : forall a b (a_SMALL : (Z.abs (Int64.signed a) <= smalla_thresh)%Z) @@ -1282,6 +1319,35 @@ Proof. cbn. rewrite <- (function_ite Vlong). f_equal. + unfold Int64.lt. set (q' := B2R 53 1024 q) in *. + fold a'. + assert (Int64.signed (Int64.repr (ZnearestE q')) = ZnearestE q') as q_SIGNED. + { apply Int64.signed_repr. + split; lia. + } + rewrite Int64.add_signed. + rewrite q_SIGNED. + rewrite Int64.signed_mone. + rewrite Int64.signed_zero. + rewrite <- (function_ite Int64.repr). + f_equal. + replace (ZnearestE q' + -1)%Z with (ZnearestE q' - 1)%Z by ring. + set (q'' := (ZnearestE q')) in *. + Check Int64.sub_signed. + fold a'. + rewrite int64_mul_signed_unsigned. + rewrite q_SIGNED. + fold b'. + + rewrite Int64.sub_signed. + fold a'. + rewrite signed_repr_sub. + + + Check Int64.signed_repr. + + rewrite find_quotient with (qr := q'). + cbn. -- cgit From 74dec8e174e1a49facaf68d45b3c920c5547123d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 14:48:59 +0100 Subject: progress --- kvx/FPDivision64.v | 24 +++++++++++++++++------- 1 file changed, 17 insertions(+), 7 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 766ff80f..23cf73e3 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1045,8 +1045,7 @@ Qed. Lemma find_quotient: forall (a b : Z) - (b_POS : (b > 0)%Z) - (invb : R) + (b_POS : (0 < b)%Z) (qr : R) (GAP : (Rabs (IZR a / IZR b - qr) < /2)%R), (a / b)%Z = @@ -1345,9 +1344,20 @@ Proof. fold a'. rewrite signed_repr_sub. - - - Check Int64.signed_repr. + rewrite Int64.signed_repr; cycle 1. + { admit. } + + assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. + { replace (IZR a' / IZR b' - q')%R with + (-(q' - IZR a' / IZR b'))%R by ring. + rewrite Rabs_Ropp. + lra. + } + rewrite (find_quotient a' b' b_NOT0 q' HALF). + fold q''. + unfold zlt. - rewrite find_quotient with (qr := q'). - cbn. + destruct (Z_lt_dec 0). + - replace ( _ >? _ )%Z with true by lia. reflexivity. + - replace ( _ >? _ )%Z with false by lia. reflexivity. +Admitted. -- cgit From 9bb316c0f32cd88cd95895111401ff7995938fec Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 15:47:56 +0100 Subject: progress --- kvx/FPDivision64.v | 42 +++++++++++++++++++++++++++++++----------- 1 file changed, 31 insertions(+), 11 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 23cf73e3..0d11c3cb 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1269,7 +1269,8 @@ Qed. Lemma step2_div_long_smalla_correct : forall a b (a_SMALL : (Z.abs (Int64.signed a) <= smalla_thresh)%Z) - (b_NOT0 : (0 < Int64.unsigned b)%Z), + (b_NOT0 : (0 < Int64.unsigned b)%Z) + (b_NOT_VERY_BIG : (Int64.unsigned b <= Int64.max_signed)%Z), step2_div_long (Vlong a) (Vlong b) = Vlong (Int64.repr (Int64.signed a / Int64.unsigned b))%Z. Proof. intros. @@ -1344,20 +1345,39 @@ Proof. fold a'. rewrite signed_repr_sub. - rewrite Int64.signed_repr; cycle 1. - { admit. } - assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. { replace (IZR a' / IZR b' - q')%R with (-(q' - IZR a' / IZR b'))%R by ring. rewrite Rabs_Ropp. lra. } - rewrite (find_quotient a' b' b_NOT0 q' HALF). - fold q''. - unfold zlt. - - destruct (Z_lt_dec 0). - - replace ( _ >? _ )%Z with true by lia. reflexivity. - - replace ( _ >? _ )%Z with false by lia. reflexivity. + pose proof (find_quotient a' b' b_NOT0 q' HALF) as QUOTIENT. + fold q'' in QUOTIENT. + cbn zeta in QUOTIENT. + + assert (b' <> 0)%Z as NONZ by lia. + pose proof (Zmod_eq_full a' b' NONZ) as MOD. + assert (b' > 0)%Z as b_GT0 by lia. + pose proof (Z_mod_lt a' b' b_GT0) as MOD_LT. + destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. + { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. + replace q'' with (1 + (a' / b'))%Z by lia. + replace ((1 + a' / b') * b' - a')%Z + with (-(a' - a' / b' * b')+b')%Z by ring. + rewrite <- MOD. + rewrite Int64.signed_repr; cycle 1. + { admit. + } + rewrite zlt_true by lia. + ring. + } + replace (b' * q'' >? a')%Z with false in QUOTIENT by lia. + rewrite <- QUOTIENT. + replace (a' / b' * b' - a')%Z with (-(a' - a' / b' * b'))%Z by ring. + rewrite <- MOD. + rewrite Int64.signed_repr ; cycle 1. + { admit. + } + rewrite zlt_false by lia. + reflexivity. Admitted. -- cgit From d8f1751b88d0013166412677c965079700e36557 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 15:51:41 +0100 Subject: done for small a --- kvx/FPDivision64.v | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 0d11c3cb..a53169b8 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1366,7 +1366,9 @@ Proof. with (-(a' - a' / b' * b')+b')%Z by ring. rewrite <- MOD. rewrite Int64.signed_repr; cycle 1. - { admit. + { change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + lia. } rewrite zlt_true by lia. ring. @@ -1376,8 +1378,10 @@ Proof. replace (a' / b' * b' - a')%Z with (-(a' - a' / b' * b'))%Z by ring. rewrite <- MOD. rewrite Int64.signed_repr ; cycle 1. - { admit. + { change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + lia. } rewrite zlt_false by lia. reflexivity. -Admitted. +Qed. -- cgit From c5a426f410f95d6e410da5b2afd5a9225e902c3e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 19:21:56 +0100 Subject: twostep_div_longu_smallb_correct --- kvx/FPDivision64.v | 82 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 82 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index a53169b8..758061b5 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1385,3 +1385,85 @@ Proof. rewrite zlt_false by lia. reflexivity. Qed. + +Definition twostep_div_longu a b := + let q1 := step1_div_longu a b in + let q2 := step2_div_long (Val.subl a (Val.mull b q1)) b in + Val.addl q1 q2. + +Lemma unsigned_repr_sub : + forall a b, + Int64.repr (a - b) = Int64.repr (a - Int64.unsigned (Int64.repr b)). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_sub. + - apply Int64.eqm_refl. + - apply Int64.eqm_unsigned_repr. +Qed. + +Lemma unsigned_repr_add : + forall a b, + Int64.repr (a + b) = Int64.repr (a + Int64.unsigned (Int64.repr b)). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_add. + - apply Int64.eqm_refl. + - apply Int64.eqm_unsigned_repr. +Qed. + +Lemma twostep_div_longu_smallb_correct : + forall a b + (b_RANGE : (1 < Int64.unsigned b <= smallb_thresh)%Z), + (twostep_div_longu (Vlong a) (Vlong b)) = + (Val.maketotal (Val.divlu (Vlong a) (Vlong b))). +Proof. + intros. + unfold twostep_div_longu. + destruct (step1_div_longu_correct a b b_RANGE) as (q1 & C1R & C1E). + rewrite C1R. + set (q1' := Int64.unsigned q1) in *. + set (b' := Int64.unsigned b) in *. + set (a' := Int64.unsigned a) in *. + assert ( Z.abs (Int64.signed (Int64.sub a (Int64.mul b q1))) <= smalla_thresh)%Z as r1_SMALL. + { unfold smalla_thresh, smallb_approx_range in *. + unfold Int64.sub, Int64.mul. + fold q1' b' a'. + rewrite <- unsigned_repr_sub. + rewrite Int64.signed_repr ; cycle 1. + { change Int64.min_signed with (-9223372036854775808)%Z. + change Int64.max_signed with (9223372036854775807)%Z. + lia. + } + lia. + } + assert (0 < b')%Z as b_NOT0 by lia. + assert (b' <= Int64.max_signed)%Z as b_NOTBIG. + { change Int64.max_signed with (9223372036854775807)%Z. + unfold smallb_thresh in b_RANGE. + lia. + } + cbn. + rewrite (step2_div_long_smalla_correct (Int64.sub a (Int64.mul b q1)) b r1_SMALL b_NOT0 b_NOTBIG). + unfold Int64.add, Int64.sub, Int64.mul, Int64.divu. + fold q1' b' a'. + rewrite <- unsigned_repr_sub. + rewrite <- unsigned_repr_add. + rewrite Int64.signed_repr ; cycle 1. + { + change Int64.min_signed with (-9223372036854775808)%Z. + change Int64.max_signed with (9223372036854775807)%Z. + unfold smallb_approx_range in *. + lia. + } + rewrite Z.add_comm. + rewrite <- Z.div_add by lia. + replace (a' - b' * q1' + q1' * b')%Z with a' by ring. + rewrite Int64.eq_false ; cycle 1. + { intro Z. unfold b' in b_NOT0. rewrite Z in b_NOT0. + rewrite Int64.unsigned_zero in b_NOT0. + lia. + } + reflexivity. +Qed. -- cgit From 30c2e58daf62d53530505bea2401e646df79bd31 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 21:06:06 +0100 Subject: progress --- kvx/FPDivision64.v | 190 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 190 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 758061b5..580f4dd9 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1467,3 +1467,193 @@ Proof. } reflexivity. Qed. + + +Lemma step2_real_div_long_bigb_correct : + forall (a b : int64) + (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z), + exists (q : float), + (step2_real_div_long (Vlong a) (Vlong b)) = Vfloat q /\ + (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (32767/65536))%R /\ + is_finite _ _ q = true. +Proof. + intros. + unfold step2_real_div_long. + assert (0 < Int64.unsigned b)%Z as b_NOT0 by (unfold smallb_thresh in *; lia). + destruct (approx_inv_longu_correct_rel b b_NOT0) as (f & C0E & C0F & C0R). + rewrite C0E. + econstructor. + split. reflexivity. + Local Transparent Float.of_long. + unfold Float.mul, Float.of_long. + set (re := (@eq_refl Datatypes.comparison Lt)) in *. + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in b_RANGE. + pose proof (Int64.signed_range a) as a_RANGE. + set (a' := Int64.signed a) in *. + set (b' := Int64.unsigned b) in *. + assert (IZR (1 + smallb_thresh) <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. + } + assert(IZR Int64.min_signed <= IZR a' <= IZR Int64.max_signed)%R as a_RANGE'. + { split; apply IZR_le; lia. + } + change Int64.min_signed with (-9223372036854775808)%Z in a_RANGE'. + change Int64.max_signed with (9223372036854775807)%Z in a_RANGE'. + pose proof (BofZ_correct 53 1024 re re a') as C1. + rewrite Rlt_bool_true in C1 ; cycle 1. + { clear C1. + apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt; lia. } + cbn. + gappa. + } + cbn in C1. + destruct C1 as (C1R & C1F & C1S). + + unfold smallb_thresh in b_RANGE'; cbn in b_RANGE'. + + pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE (BofZ 53 1024 re re a') f) as C2. + rewrite Rlt_bool_true in C2 ; cycle 1. + { clear C2. + apply Rabs_relax with (b := bpow radix2 53). + { apply bpow_lt. lia. } + cbn. + rewrite C1R. + unfold approx_inv_rel_thresh in C0R. + replace (B2R 53 1024 f) with + ((1/IZR b') * ((IZR b' * B2R 53 1024 f - 1) + 1))%R ; cycle 1. + { field. lra. } + gappa. + } + rewrite C0F in C2. + rewrite C1R in C2. + rewrite C1F in C2. + rewrite C1S in C2. + cbn in C2. + destruct C2 as (C2R & C2F & _). + split. + 2: exact C2F. + rewrite C2R. + set (f' := (B2R 53 1024 f)) in *. + replace (rd(rd (IZR a') * f') - IZR a' / IZR b')%R with + ((rd(rd (IZR a') * f') - IZR a' * f') + IZR a' / IZR b' * (IZR b' * f' - 1))%R ; cycle 1. + { field. lra. } + unfold approx_inv_rel_thresh in *. + gappa. +Qed. + +Lemma step2_real_div_long_bigb_correct : + forall (a b : int64) + (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z), + exists (q : float), + (step2_real_div_long (Vlong a) (Vlong b)) = Vfloat q /\ + (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (32767/65536))%R /\ + is_finite _ _ q = true. +Proof. + intros. + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in b_RANGE. + set (a' := (Int64.signed a)) in *. + set (b' := (Int64.unsigned b)) in *. + assert (IZR (1 + smallb_thresh) <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. + } + destruct (step2_real_div_long_smalla_correct a b a_SMALL b_NOT0) as (q & C1R & C1E & C1F). + fold a' b' in C1E. + assert ((Int64.min_signed <=? ZnearestE (B2R 53 1024 q))=true)%Z as q_LOW. + { apply Zle_imp_le_bool. + change Int64.min_signed with (-9223372036854775808)%Z. + apply Znearest_lub. + set (q' := B2R 53 1024 q) in *. + replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + unfold smalla_thresh in a_RANGE'. + gappa. + } + assert ((ZnearestE (B2R 53 1024 q) <=? Int64.max_signed)=true)%Z as q_HIGH. + { apply Zle_imp_le_bool. + change Int64.max_signed with (9223372036854775807)%Z. + apply Znearest_glb. + set (q' := B2R 53 1024 q) in *. + replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + unfold smalla_thresh in a_RANGE'. + gappa. + } + unfold step2_div_long, step2_div_long'. + rewrite C1R. + cbn. + unfold Float.to_long_ne. + rewrite (ZofB_ne_range_correct _ _ q Int64.min_signed Int64.max_signed). + rewrite C1F. + rewrite q_LOW. + rewrite q_HIGH. + cbn. + rewrite normalize_ite. + cbn. + rewrite <- (function_ite Vlong). + f_equal. + unfold Int64.lt. + set (q' := B2R 53 1024 q) in *. + fold a'. + assert (Int64.signed (Int64.repr (ZnearestE q')) = ZnearestE q') as q_SIGNED. + { apply Int64.signed_repr. + split; lia. + } + rewrite Int64.add_signed. + rewrite q_SIGNED. + rewrite Int64.signed_mone. + rewrite Int64.signed_zero. + rewrite <- (function_ite Int64.repr). + f_equal. + replace (ZnearestE q' + -1)%Z with (ZnearestE q' - 1)%Z by ring. + + set (q'' := (ZnearestE q')) in *. + Check Int64.sub_signed. + fold a'. + rewrite int64_mul_signed_unsigned. + rewrite q_SIGNED. + fold b'. + + rewrite Int64.sub_signed. + fold a'. + rewrite signed_repr_sub. + + assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. + { replace (IZR a' / IZR b' - q')%R with + (-(q' - IZR a' / IZR b'))%R by ring. + rewrite Rabs_Ropp. + lra. + } + pose proof (find_quotient a' b' b_NOT0 q' HALF) as QUOTIENT. + fold q'' in QUOTIENT. + cbn zeta in QUOTIENT. + + assert (b' <> 0)%Z as NONZ by lia. + pose proof (Zmod_eq_full a' b' NONZ) as MOD. + assert (b' > 0)%Z as b_GT0 by lia. + pose proof (Z_mod_lt a' b' b_GT0) as MOD_LT. + destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. + { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. + replace q'' with (1 + (a' / b'))%Z by lia. + replace ((1 + a' / b') * b' - a')%Z + with (-(a' - a' / b' * b')+b')%Z by ring. + rewrite <- MOD. + rewrite Int64.signed_repr; cycle 1. + { change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + lia. + } + rewrite zlt_true by lia. + ring. + } + replace (b' * q'' >? a')%Z with false in QUOTIENT by lia. + rewrite <- QUOTIENT. + replace (a' / b' * b' - a')%Z with (-(a' - a' / b' * b'))%Z by ring. + rewrite <- MOD. + rewrite Int64.signed_repr ; cycle 1. + { change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + lia. + } + rewrite zlt_false by lia. + reflexivity. -- cgit From 6290759de7aec8ec914c16eac6882b353ad980f1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 21:07:29 +0100 Subject: rm spurious function --- kvx/FPDivision64.v | 115 ----------------------------------------------------- 1 file changed, 115 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 580f4dd9..1e5d8478 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1542,118 +1542,3 @@ Proof. unfold approx_inv_rel_thresh in *. gappa. Qed. - -Lemma step2_real_div_long_bigb_correct : - forall (a b : int64) - (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z), - exists (q : float), - (step2_real_div_long (Vlong a) (Vlong b)) = Vfloat q /\ - (Rabs ((B2R _ _ q) - (IZR (Int64.signed a)) / (IZR (Int64.unsigned b))) <= (32767/65536))%R /\ - is_finite _ _ q = true. -Proof. - intros. - pose proof (Int64.unsigned_range b) as b_RANGE. - change Int64.modulus with 18446744073709551616%Z in b_RANGE. - set (a' := (Int64.signed a)) in *. - set (b' := (Int64.unsigned b)) in *. - assert (IZR (1 + smallb_thresh) <= IZR b' <= 18446744073709551615)%R as b_RANGE'. - { split; apply IZR_le; lia. - } - destruct (step2_real_div_long_smalla_correct a b a_SMALL b_NOT0) as (q & C1R & C1E & C1F). - fold a' b' in C1E. - assert ((Int64.min_signed <=? ZnearestE (B2R 53 1024 q))=true)%Z as q_LOW. - { apply Zle_imp_le_bool. - change Int64.min_signed with (-9223372036854775808)%Z. - apply Znearest_lub. - set (q' := B2R 53 1024 q) in *. - replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. - unfold smalla_thresh in a_RANGE'. - gappa. - } - assert ((ZnearestE (B2R 53 1024 q) <=? Int64.max_signed)=true)%Z as q_HIGH. - { apply Zle_imp_le_bool. - change Int64.max_signed with (9223372036854775807)%Z. - apply Znearest_glb. - set (q' := B2R 53 1024 q) in *. - replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. - unfold smalla_thresh in a_RANGE'. - gappa. - } - unfold step2_div_long, step2_div_long'. - rewrite C1R. - cbn. - unfold Float.to_long_ne. - rewrite (ZofB_ne_range_correct _ _ q Int64.min_signed Int64.max_signed). - rewrite C1F. - rewrite q_LOW. - rewrite q_HIGH. - cbn. - rewrite normalize_ite. - cbn. - rewrite <- (function_ite Vlong). - f_equal. - unfold Int64.lt. - set (q' := B2R 53 1024 q) in *. - fold a'. - assert (Int64.signed (Int64.repr (ZnearestE q')) = ZnearestE q') as q_SIGNED. - { apply Int64.signed_repr. - split; lia. - } - rewrite Int64.add_signed. - rewrite q_SIGNED. - rewrite Int64.signed_mone. - rewrite Int64.signed_zero. - rewrite <- (function_ite Int64.repr). - f_equal. - replace (ZnearestE q' + -1)%Z with (ZnearestE q' - 1)%Z by ring. - - set (q'' := (ZnearestE q')) in *. - Check Int64.sub_signed. - fold a'. - rewrite int64_mul_signed_unsigned. - rewrite q_SIGNED. - fold b'. - - rewrite Int64.sub_signed. - fold a'. - rewrite signed_repr_sub. - - assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. - { replace (IZR a' / IZR b' - q')%R with - (-(q' - IZR a' / IZR b'))%R by ring. - rewrite Rabs_Ropp. - lra. - } - pose proof (find_quotient a' b' b_NOT0 q' HALF) as QUOTIENT. - fold q'' in QUOTIENT. - cbn zeta in QUOTIENT. - - assert (b' <> 0)%Z as NONZ by lia. - pose proof (Zmod_eq_full a' b' NONZ) as MOD. - assert (b' > 0)%Z as b_GT0 by lia. - pose proof (Z_mod_lt a' b' b_GT0) as MOD_LT. - destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. - { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. - replace q'' with (1 + (a' / b'))%Z by lia. - replace ((1 + a' / b') * b' - a')%Z - with (-(a' - a' / b' * b')+b')%Z by ring. - rewrite <- MOD. - rewrite Int64.signed_repr; cycle 1. - { change Int64.min_signed with (-9223372036854775808)%Z in *. - change Int64.max_signed with (9223372036854775807)%Z in *. - lia. - } - rewrite zlt_true by lia. - ring. - } - replace (b' * q'' >? a')%Z with false in QUOTIENT by lia. - rewrite <- QUOTIENT. - replace (a' / b' * b' - a')%Z with (-(a' - a' / b' * b'))%Z by ring. - rewrite <- MOD. - rewrite Int64.signed_repr ; cycle 1. - { change Int64.min_signed with (-9223372036854775808)%Z in *. - change Int64.max_signed with (9223372036854775807)%Z in *. - lia. - } - rewrite zlt_false by lia. - reflexivity. -- cgit From 56c5f324c1792df0a1c660f37bb9d5028c496a3f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 21:42:52 +0100 Subject: progress --- kvx/FPDivision64.v | 121 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 121 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 1e5d8478..74303a7a 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1542,3 +1542,124 @@ Proof. unfold approx_inv_rel_thresh in *. gappa. Qed. + +Lemma step2_div_long_bigb_correct : + forall a b + (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z) + (b_NOT_TOO_BIG : ((Int64.unsigned b) <= Int64.max_signed)%Z), + step2_div_long (Vlong a) (Vlong b) = Vlong (Int64.repr (Int64.signed a / Int64.unsigned b))%Z. +Proof. + intros. + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in b_RANGE. + pose proof (Int64.signed_range a) as a_RANGE. + set (a' := (Int64.signed a)) in *. + set (b' := (Int64.unsigned b)) in *. + assert (IZR (1 + smallb_thresh) <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. + } + assert(IZR Int64.min_signed <= IZR a' <= IZR Int64.max_signed)%R as a_RANGE'. + { split; apply IZR_le; lia. + } + unfold smallb_thresh in *; cbn in b_RANGE'. + change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + assert (0 < b')%Z as b_NOT0 by lia. + + destruct (step2_real_div_long_bigb_correct a b b_BIG) as (q & C1R & C1E & C1F). + fold a' b' in C1E. + assert ((Int64.min_signed <=? ZnearestE (B2R 53 1024 q))=true)%Z as q_LOW. + { apply Zle_imp_le_bool. + change Int64.min_signed with (-9223372036854775808)%Z. + apply Znearest_lub. + set (q' := B2R 53 1024 q) in *. + replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + gappa. + } + assert ((ZnearestE (B2R 53 1024 q) <=? Int64.max_signed)=true)%Z as q_HIGH. + { apply Zle_imp_le_bool. + change Int64.max_signed with (9223372036854775807)%Z. + apply Znearest_glb. + set (q' := B2R 53 1024 q) in *. + replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + gappa. + } + unfold step2_div_long, step2_div_long'. + rewrite C1R. + cbn. + unfold Float.to_long_ne. + rewrite (ZofB_ne_range_correct _ _ q Int64.min_signed Int64.max_signed). + rewrite C1F. + rewrite q_LOW. + rewrite q_HIGH. + cbn. + rewrite normalize_ite. + cbn. + rewrite <- (function_ite Vlong). + f_equal. + unfold Int64.lt. + set (q' := B2R 53 1024 q) in *. + fold a'. + assert (Int64.signed (Int64.repr (ZnearestE q')) = ZnearestE q') as q_SIGNED. + { apply Int64.signed_repr. + split; lia. + } + rewrite Int64.add_signed. + rewrite q_SIGNED. + rewrite Int64.signed_mone. + rewrite Int64.signed_zero. + rewrite <- (function_ite Int64.repr). + f_equal. + replace (ZnearestE q' + -1)%Z with (ZnearestE q' - 1)%Z by ring. + + set (q'' := (ZnearestE q')) in *. + Check Int64.sub_signed. + fold a'. + rewrite int64_mul_signed_unsigned. + rewrite q_SIGNED. + fold b'. + + rewrite Int64.sub_signed. + fold a'. + rewrite signed_repr_sub. + + assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. + { replace (IZR a' / IZR b' - q')%R with + (-(q' - IZR a' / IZR b'))%R by ring. + rewrite Rabs_Ropp. + lra. + } + pose proof (find_quotient a' b' b_NOT0 q' HALF) as QUOTIENT. + fold q'' in QUOTIENT. + cbn zeta in QUOTIENT. + + assert (b' <> 0)%Z as NONZ by lia. + pose proof (Zmod_eq_full a' b' NONZ) as MOD. + assert (b' > 0)%Z as b_GT0 by lia. + pose proof (Z_mod_lt a' b' b_GT0) as MOD_LT. + destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. + { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. + replace q'' with (1 + (a' / b'))%Z by lia. + replace ((1 + a' / b') * b' - a')%Z + with (-(a' - a' / b' * b')+b')%Z by ring. + rewrite <- MOD. + rewrite Int64.signed_repr; cycle 1. + { change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + lia. + } + rewrite zlt_true by lia. + ring. + } + replace (b' * q'' >? a')%Z with false in QUOTIENT by lia. + rewrite <- QUOTIENT. + replace (a' / b' * b' - a')%Z with (-(a' - a' / b' * b'))%Z by ring. + rewrite <- MOD. + rewrite Int64.signed_repr ; cycle 1. + { change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + lia. + } + rewrite zlt_false by lia. + reflexivity. +Qed. -- cgit From 0ec09c99af2071f52271dd5e14846f5da3bcb718 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 22:37:25 +0100 Subject: some progress --- kvx/FPDivision64.v | 176 +++++++++++++++++++++++++++++++++++++++++++++++------ 1 file changed, 157 insertions(+), 19 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 74303a7a..e4f883e7 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -29,6 +29,25 @@ Require Import IEEE754_extra. From Gappa Require Import Gappa_tactic. + +Lemma Znearest_lub : + forall choice (n : Z) (x : R), (IZR n <= x)%R -> (n <= Znearest choice x)%Z. +Proof. + intros until x. intro BND. + pose proof (Zfloor_lub n x BND). + pose proof (Znearest_ge_floor choice x). + lia. +Qed. + +Lemma Znearest_glb : + forall choice (n : Z) (x : R), (x <= IZR n)%R -> (Znearest choice x <= n)%Z. +Proof. + intros until x. intro BND. + pose proof (Zceil_glb n x BND). + pose proof (Znearest_le_ceil choice x). + lia. +Qed. + Definition approx_inv_longu b := let invb_s := ExtValues.invfs (Val.maketotal (Val.singleoflongu b)) in let invb_d := Val.floatofsingle invb_s in @@ -851,7 +870,99 @@ Proof. rewrite C4. reflexivity. Qed. - + +Lemma step1_div_longu_correct_anyb : + forall a b + (b_NOT01 : (1 < Int64.unsigned b)%Z), + exists (q : int64), + (step1_div_longu (Vlong a) (Vlong b)) = Vlong q. +Proof. + intros. + + pose proof (Int64.unsigned_range a) as a_RANGE. + pose proof (Int64.unsigned_range b) as b_RANGE. + change Int64.modulus with 18446744073709551616%Z in *. + set (a' := Int64.unsigned a) in *. + set (b' := Int64.unsigned b) in *. + assert (0 <= IZR a' <= 18446744073709551615)%R as a_RANGE'. + { split; apply IZR_le; lia. } + assert (2 <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. } + + assert (0 < b')%Z as b_NOT0 by lia. + + destruct (Z_le_gt_dec a' 0). + { assert (a' = 0%Z) as ZERO by lia. + replace a with Int64.zero; cycle 1. + { + unfold a' in ZERO. + unfold Int64.zero. + rewrite <- ZERO. + apply Int64.repr_unsigned. + } + exists Int64.zero. + apply step1_div_longu_a0. + exact b_NOT0. + } + + unfold step1_div_longu. + unfold step1_real_div_longu. + destruct (step1_real_inv_longu_correct b b_NOT0) as (f & C1E & C1R & C1F & C1S). + rewrite C1E. + cbn. + unfold Float.of_longu, Float.mul, Float.to_longu_ne. + set (re := (@eq_refl Datatypes.comparison Lt)). + fold a'. + fold b' in C1R. + pose proof (BofZ_correct 53 1024 re re a') as C2. + rewrite Rlt_bool_true in C2; cycle 1. + { clear C2. + apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt. lia. } + cbn. + gappa. + } + cbn in C2. + destruct C2 as (C2R & C2F & C2S). + pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re a') f) as C3. + rewrite C2R in C3. + rewrite C2F in C3. + rewrite C2S in C3. + rewrite C1R in C3. + rewrite C1F in C3. + rewrite C1S in C3. + rewrite Rlt_bool_true in C3; cycle 1. + { clear C3. + apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt. lia. } + cbn. + gappa. + } + cbn in C3. + destruct C3 as (C3R & C3F & _). + pose proof (ZofB_ne_range_correct 53 1024 + (Bmult 53 1024 re re Float.binop_nan mode_NE + (BofZ 53 1024 re re a') f) 0 Int64.max_unsigned) as C4. + rewrite C3R in C4. + rewrite C3F in C4. + cbn zeta in C4. + rewrite Zle_bool_true in C4 ; cycle 1. + { clear C4. + apply Znearest_lub. + gappa. + } + rewrite Zle_bool_true in C4 ; cycle 1. + { clear C4. + apply Znearest_glb. + cbn. + gappa. + } + rewrite C4. + cbn. + eauto. +Qed. + Definition smallb_approx_range := 4400000000000%Z. Lemma step1_div_longu_correct : forall a b, @@ -1194,24 +1305,6 @@ Definition step2_div_long a b := Val.select (Val.cmpl_bool Cgt (Val.subl (Val.mull q b) a) (Vlong Int64.zero)) (Val.addl q (Vlong Int64.mone)) q Tlong. -Lemma Znearest_lub : - forall choice (n : Z) (x : R), (IZR n <= x)%R -> (n <= Znearest choice x)%Z. -Proof. - intros until x. intro BND. - pose proof (Zfloor_lub n x BND). - pose proof (Znearest_ge_floor choice x). - lia. -Qed. - -Lemma Znearest_glb : - forall choice (n : Z) (x : R), (x <= IZR n)%R -> (Znearest choice x <= n)%Z. -Proof. - intros until x. intro BND. - pose proof (Zceil_glb n x BND). - pose proof (Znearest_le_ceil choice x). - lia. -Qed. - Lemma function_ite : forall {A B : Type} (fn : A->B) (b : bool) (x y: A), fn (if b then x else y) = (if b then fn x else fn y). @@ -1663,3 +1756,48 @@ Proof. rewrite zlt_false by lia. reflexivity. Qed. + +Lemma twostep_div_longu_bigb_correct : + forall a b + (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z) + (b_NOT_TOO_BIG : ((Int64.unsigned b) <= Int64.max_signed)%Z), + (twostep_div_longu (Vlong a) (Vlong b)) = + (Val.maketotal (Val.divlu (Vlong a) (Vlong b))). +Proof. + intros. + unfold twostep_div_longu. + set (b' := Int64.unsigned b) in *. + set (a' := Int64.unsigned a) in *. +Admitted. + +(* + assert (0 < b')%Z as b_NOT0 by lia. + assert (b' <= Int64.max_signed)%Z as b_NOTBIG. + { change Int64.max_signed with (9223372036854775807)%Z. + unfold smallb_thresh in b_RANGE. + lia. + } + cbn. + rewrite (step2_div_long_bigb_correct (Int64.sub a (Int64.mul b q1)) b r1_SMALL b_NOT0 b_NOTBIG). + unfold Int64.add, Int64.sub, Int64.mul, Int64.divu. + fold q1' b' a'. + rewrite <- unsigned_repr_sub. + rewrite <- unsigned_repr_add. + rewrite Int64.signed_repr ; cycle 1. + { + change Int64.min_signed with (-9223372036854775808)%Z. + change Int64.max_signed with (9223372036854775807)%Z. + unfold smallb_approx_range in *. + lia. + } + rewrite Z.add_comm. + rewrite <- Z.div_add by lia. + replace (a' - b' * q1' + q1' * b')%Z with a' by ring. + rewrite Int64.eq_false ; cycle 1. + { intro Z. unfold b' in b_NOT0. rewrite Z in b_NOT0. + rewrite Int64.unsigned_zero in b_NOT0. + lia. + } + reflexivity. +Qed. +*) -- cgit From 2bbe86d4efe2ae08444359008798c12adf0c9b60 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 24 Jan 2022 22:39:55 +0100 Subject: some progress --- kvx/FPDivision64.v | 2 -- 1 file changed, 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index e4f883e7..3685f601 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1428,7 +1428,6 @@ Proof. replace (ZnearestE q' + -1)%Z with (ZnearestE q' - 1)%Z by ring. set (q'' := (ZnearestE q')) in *. - Check Int64.sub_signed. fold a'. rewrite int64_mul_signed_unsigned. rewrite q_SIGNED. @@ -1706,7 +1705,6 @@ Proof. replace (ZnearestE q' + -1)%Z with (ZnearestE q' - 1)%Z by ring. set (q'' := (ZnearestE q')) in *. - Check Int64.sub_signed. fold a'. rewrite int64_mul_signed_unsigned. rewrite q_SIGNED. -- cgit From 0fb548cbf78c19430e60827f7c253fc42aa25e05 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 25 Jan 2022 10:54:19 +0100 Subject: some progress? --- kvx/FPDivision64.v | 311 +++++++++++++++++++++++++++++++---------------------- 1 file changed, 185 insertions(+), 126 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 3685f601..69e49f29 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1064,96 +1064,6 @@ Proof. gappa. Qed. -(* -Lemma range_up_le : - forall a x b, - (IZR a <= IZR x <= (IZR b) - 1)%R -> - (a <= x < b)%Z. -Proof. - intros until b. intro RANGE. - split. - { apply le_IZR. lra. } - assert (x <= b-1)%Z. - { apply le_IZR. rewrite minus_IZR. lra. } - lia. -Qed. - *) - -(* -Lemma find_quotient: - forall (a b : Z) - (b_POS : (b > 0)%Z) - (invb : R) - (eps : R) - (invb_EPS : (Rabs (invb * IZR b - 1) <= eps)%R) - (SMALL : (IZR (Z.abs a)*eps < IZR b / 2)%R), - (a / b)%Z = - let q := ZnearestE ((IZR a) * invb) in - if (b*q >? a)%Z - then (q-1)%Z - else q. -Proof. - intros. - rewrite <- Rabs_Zabs in SMALL. - set (q := ZnearestE (IZR a * invb)%R). - cbn zeta. - set (b' := IZR b) in *. - set (a' := IZR a) in *. - assert (1 <= b')%R as b_POS'. - { apply IZR_le. - lia. - } - - pose proof (Znearest_imp2 (fun x : Z => negb (Z.even x)) (a' * invb)) as NEAR. - fold q in NEAR. - set (q' := IZR q) in *. - assert ((Rabs a') * Rabs (invb * b' - 1) <= (Rabs a') * eps)%R as S1. - { apply Rmult_le_compat_l. - apply Rabs_pos. - assumption. } - rewrite <- Rabs_mult in S1. - assert ((Rabs b') * Rabs (q' - a' * invb) <= (Rabs b') / 2)%R as S2. - { apply Rmult_le_compat_l. - apply Rabs_pos. - assumption. } - rewrite <- Rabs_mult in S2. - rewrite (Rabs_right b') in S2 by lra. - pose proof (Rabs_triang (a' * (invb * b' - 1)) - (b' * (q' - a' * invb))) as TRIANGLE. - replace (a' * (invb * b' - 1) + b' * (q' - a' * invb))%R with - (b' * q' - a')%R in TRIANGLE by ring. - assert (Rabs (b' * q' - a') < b')%R as DELTA by lra. - - pose proof (Zgt_cases (b * q) a)%Z as CASE. - destruct (_ >? _)%Z. - { unfold b', q', a' in DELTA. - rewrite <- mult_IZR in DELTA. - rewrite <- minus_IZR in DELTA. - rewrite Rabs_Zabs in DELTA. - apply lt_IZR in DELTA. - rewrite Z.abs_eq in DELTA by lia. - - apply Zdiv_unique with (b := (a - (q-1)*b)%Z). - ring. - split; lia. - } - - rewrite <- Rabs_Ropp in DELTA. - unfold b', q', a' in DELTA. - rewrite <- mult_IZR in DELTA. - rewrite <- minus_IZR in DELTA. - rewrite <- opp_IZR in DELTA. - rewrite Rabs_Zabs in DELTA. - apply lt_IZR in DELTA. - rewrite Z.abs_eq in DELTA by lia. - - apply Zdiv_unique with (b := (a - q*b)%Z). - ring. - - split; lia. -Qed. - *) - Lemma find_quotient: forall (a b : Z) (b_POS : (0 < b)%Z) @@ -1754,48 +1664,197 @@ Proof. rewrite zlt_false by lia. reflexivity. Qed. - -Lemma twostep_div_longu_bigb_correct : - forall a b - (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z) - (b_NOT_TOO_BIG : ((Int64.unsigned b) <= Int64.max_signed)%Z), - (twostep_div_longu (Vlong a) (Vlong b)) = - (Val.maketotal (Val.divlu (Vlong a) (Vlong b))). + +Definition step2_real_div_longu a b := + Val.mulf (Val.maketotal (Val.floatoflongu a)) (approx_inv_longu b). + +Definition step2_div_longu' a b := + Val.maketotal (Val.longuoffloat_ne (step2_real_div_longu a b)). + +Definition step2_div_longu a b := + let q := step2_div_longu' a b in + Val.select (Val.cmpl_bool Cgt (Val.subl (Val.mull q b) a) (Vlong Int64.zero)) + (Val.addl q (Vlong Int64.mone)) q Tlong. + +Lemma step2_real_div_longu_bigb_correct : + forall (a b : int64) + (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z), + exists (q : float), + (step2_real_div_longu (Vlong a) (Vlong b)) = Vfloat q /\ + (Rabs ((B2R _ _ q) - (IZR (Int64.unsigned a)) / (IZR (Int64.unsigned b))) <= (32767/65536))%R /\ + is_finite _ _ q = true. Proof. intros. - unfold twostep_div_longu. - set (b' := Int64.unsigned b) in *. + unfold step2_real_div_longu. + assert (0 < Int64.unsigned b)%Z as b_NOT0 by (unfold smallb_thresh in *; lia). + destruct (approx_inv_longu_correct_rel b b_NOT0) as (f & C0E & C0F & C0R). + rewrite C0E. + econstructor. + split. reflexivity. + Local Transparent Float.of_longu. + unfold Float.mul, Float.of_longu. + set (re := (@eq_refl Datatypes.comparison Lt)) in *. + pose proof (Int64.unsigned_range b) as b_RANGE. + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in *. set (a' := Int64.unsigned a) in *. -Admitted. + set (b' := Int64.unsigned b) in *. + assert (IZR (1 + smallb_thresh) <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. + } + assert(0 <= IZR a' <= 18446744073709551615)%R as a_RANGE'. + { split; apply IZR_le; lia. + } + pose proof (BofZ_correct 53 1024 re re a') as C1. + rewrite Rlt_bool_true in C1 ; cycle 1. + { clear C1. + apply Rabs_relax with (b := bpow radix2 64). + { apply bpow_lt; lia. } + cbn. + gappa. + } + cbn in C1. + destruct C1 as (C1R & C1F & C1S). -(* + unfold smallb_thresh in b_RANGE'; cbn in b_RANGE'. + + pose proof (Bmult_correct 53 1024 re re Float.binop_nan mode_NE (BofZ 53 1024 re re a') f) as C2. + rewrite Rlt_bool_true in C2 ; cycle 1. + { clear C2. + apply Rabs_relax with (b := bpow radix2 53). + { apply bpow_lt. lia. } + cbn. + rewrite C1R. + unfold approx_inv_rel_thresh in C0R. + replace (B2R 53 1024 f) with + ((1/IZR b') * ((IZR b' * B2R 53 1024 f - 1) + 1))%R ; cycle 1. + { field. lra. } + gappa. + } + rewrite C0F in C2. + rewrite C1R in C2. + rewrite C1F in C2. + rewrite C1S in C2. + cbn in C2. + destruct C2 as (C2R & C2F & _). + split. + 2: exact C2F. + rewrite C2R. + set (f' := (B2R 53 1024 f)) in *. + replace (rd(rd (IZR a') * f') - IZR a' / IZR b')%R with + ((rd(rd (IZR a') * f') - IZR a' * f') + IZR a' / IZR b' * (IZR b' * f' - 1))%R ; cycle 1. + { field. lra. } + unfold approx_inv_rel_thresh in *. + gappa. +Qed. + +Lemma step2_div_longu_bigb_correct : + forall a b + (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z) + (b_NOT_TOO_BIG : ((Int64.unsigned b) <= Int64.max_signed)%Z), + step2_div_longu (Vlong a) (Vlong b) = Vlong (Int64.repr (Int64.unsigned a / Int64.unsigned b))%Z. +Proof. + intros. + pose proof (Int64.unsigned_range b) as b_RANGE. + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in *. + set (a' := (Int64.unsigned a)) in *. + set (b' := (Int64.unsigned b)) in *. + assert (IZR (1 + smallb_thresh) <= IZR b' <= 18446744073709551615)%R as b_RANGE'. + { split; apply IZR_le; lia. + } + assert(0 <= IZR a' <= 18446744073709551615)%R as a_RANGE'. + { split; apply IZR_le; lia. + } + unfold smallb_thresh in *; cbn in b_RANGE'. assert (0 < b')%Z as b_NOT0 by lia. - assert (b' <= Int64.max_signed)%Z as b_NOTBIG. - { change Int64.max_signed with (9223372036854775807)%Z. - unfold smallb_thresh in b_RANGE. - lia. + + destruct (step2_real_div_longu_bigb_correct a b b_BIG) as (q & C1R & C1E & C1F). + fold a' b' in C1E. + + assert ((0 <=? ZnearestE (B2R 53 1024 q))=true)%Z as q_LOW. + { apply Zle_imp_le_bool. + set (q' := B2R 53 1024 q) in *. + assert (-32767 / 65536 <= q')%R as LOWROUND. + { replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + gappa. + } + destruct (Rcase_abs q'). + { replace (ZnearestE q') with 0%Z. lia. + symmetry. + apply Znearest_imp. + apply Rabs_lt. + split; lra. + } + apply Znearest_lub. + lra. } + assert ((ZnearestE (B2R 53 1024 q) <=? Int64.max_unsigned)=true)%Z as q_HIGH. + { apply Zle_imp_le_bool. + change Int64.max_unsigned with (18446744073709551615)%Z. + apply Znearest_glb. + set (q' := B2R 53 1024 q) in *. + replace q' with (IZR a' / IZR b' + (q' - IZR a' / IZR b'))%R by ring. + gappa. + } + + unfold step2_div_longu, step2_div_longu'. + rewrite C1R. cbn. - rewrite (step2_div_long_bigb_correct (Int64.sub a (Int64.mul b q1)) b r1_SMALL b_NOT0 b_NOTBIG). - unfold Int64.add, Int64.sub, Int64.mul, Int64.divu. - fold q1' b' a'. - rewrite <- unsigned_repr_sub. - rewrite <- unsigned_repr_add. - rewrite Int64.signed_repr ; cycle 1. - { - change Int64.min_signed with (-9223372036854775808)%Z. - change Int64.max_signed with (9223372036854775807)%Z. - unfold smallb_approx_range in *. - lia. + unfold Float.to_longu_ne. + rewrite (ZofB_ne_range_correct _ _ q _ _). + rewrite C1F. + rewrite q_LOW. + rewrite q_HIGH. + cbn. + rewrite normalize_ite. + cbn. + rewrite <- (function_ite Vlong). + f_equal. + unfold Int64.lt. + set (q' := B2R 53 1024 q) in *. + fold a'. + rewrite Int64.signed_zero. + set (q'' := (ZnearestE q')) in *. + assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. + { replace (IZR a' / IZR b' - q')%R with + (-(q' - IZR a' / IZR b'))%R by ring. + rewrite Rabs_Ropp. + lra. } - rewrite Z.add_comm. - rewrite <- Z.div_add by lia. - replace (a' - b' * q1' + q1' * b')%Z with a' by ring. - rewrite Int64.eq_false ; cycle 1. - { intro Z. unfold b' in b_NOT0. rewrite Z in b_NOT0. - rewrite Int64.unsigned_zero in b_NOT0. - lia. + pose proof (find_quotient a' b' b_NOT0 q' HALF) as QUOTIENT. + fold q'' in QUOTIENT. + cbn zeta in QUOTIENT. + + assert (b' <> 0)%Z as NONZ by lia. + pose proof (Zmod_eq_full a' b' NONZ) as MOD. + assert (b' > 0)%Z as b_GT0 by lia. + pose proof (Z_mod_lt a' b' b_GT0) as MOD_LT. + assert ((Int64.signed (Int64.sub (Int64.mul (Int64.repr q'') b) a)) + = q''*b'-a')%Z as SIMPL. + { admit. } + rewrite SIMPL. + destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. + { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. + rewrite zlt_true by lia. + replace q'' with (1 + (a' / b'))%Z by lia. + unfold Int64.add. + apply Int64.eqm_samerepr. + apply Int64.eqm_trans with (y := ((1 + a' / b') + (-1))%Z). + { apply Int64.eqm_add. + apply Int64.eqm_sym. + apply Int64.eqm_unsigned_repr. + rewrite Int64.unsigned_mone. + replace (-1)%Z with (0 - 1)%Z by ring. + apply Int64.eqm_add. + exists 1%Z. + lia. + apply Int64.eqm_refl. + } + replace (1 + a' / b' + -1)%Z with (a'/b')%Z by ring. + apply Int64.eqm_refl. } - reflexivity. -Qed. -*) + rewrite zlt_false by lia. + replace (b' * q'' >? a')%Z with false in QUOTIENT by lia. + congruence. +Admitted. -- cgit From 32a69b2acc078ffc5078ca5f12747170b16d9d78 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 25 Jan 2022 11:28:55 +0100 Subject: big progress on longu --- kvx/FPDivision64.v | 70 +++++++++++++++++++++++++++++++++++++++++++++++------- 1 file changed, 61 insertions(+), 9 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 69e49f29..d96ced31 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1747,7 +1747,40 @@ Proof. unfold approx_inv_rel_thresh in *. gappa. Qed. - + +Lemma repr_unsigned_mul: + forall a b, + (Int64.repr (Int64.unsigned (Int64.repr a) * b)) = Int64.repr (a * b). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_mult. + - apply Int64.eqm_sym. apply Int64.eqm_unsigned_repr. + - apply Int64.eqm_refl. +Qed. + +Lemma repr_unsigned_sub: + forall a b, + (Int64.repr (Int64.unsigned (Int64.repr a) - b)) = Int64.repr (a - b). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_sub. + - apply Int64.eqm_sym. apply Int64.eqm_unsigned_repr. + - apply Int64.eqm_refl. +Qed. + +Lemma repr_unsigned_add: + forall a b, + (Int64.repr (Int64.unsigned (Int64.repr a) + b)) = Int64.repr (a + b). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_add. + - apply Int64.eqm_sym. apply Int64.eqm_unsigned_repr. + - apply Int64.eqm_refl. +Qed. + Lemma step2_div_longu_bigb_correct : forall a b (b_BIG : ((Int64.unsigned b) > smallb_thresh)%Z) @@ -1830,15 +1863,23 @@ Proof. pose proof (Zmod_eq_full a' b' NONZ) as MOD. assert (b' > 0)%Z as b_GT0 by lia. pose proof (Z_mod_lt a' b' b_GT0) as MOD_LT. - assert ((Int64.signed (Int64.sub (Int64.mul (Int64.repr q'') b) a)) - = q''*b'-a')%Z as SIMPL. - { admit. } - rewrite SIMPL. destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. + unfold Int64.sub, Int64.mul. + fold a' b'. + replace q'' with (1 + a'/b')%Z by lia. + rewrite repr_unsigned_mul. + rewrite repr_unsigned_sub. + + replace ((1 + a' / b') * b' - a')%Z with (b' - (a' - a' / b' * b'))%Z by ring. + rewrite <- MOD. + rewrite Int64.signed_repr ; cycle 1. + { change Int64.max_signed with 9223372036854775807%Z in *. + change Int64.min_signed with (-9223372036854775808)%Z in *. + lia. + } rewrite zlt_true by lia. replace q'' with (1 + (a' / b'))%Z by lia. - unfold Int64.add. apply Int64.eqm_samerepr. apply Int64.eqm_trans with (y := ((1 + a' / b') + (-1))%Z). { apply Int64.eqm_add. @@ -1854,7 +1895,18 @@ Proof. replace (1 + a' / b' + -1)%Z with (a'/b')%Z by ring. apply Int64.eqm_refl. } - rewrite zlt_false by lia. replace (b' * q'' >? a')%Z with false in QUOTIENT by lia. - congruence. -Admitted. + rewrite <- QUOTIENT. + unfold Int64.sub, Int64.mul. + fold a' b'. + rewrite repr_unsigned_mul. + rewrite repr_unsigned_sub. + replace (a' / b' * b' - a')%Z with (- (a' mod b'))%Z by lia. + rewrite Int64.signed_repr ; cycle 1. + { change Int64.max_signed with 9223372036854775807%Z in *. + change Int64.min_signed with (-9223372036854775808)%Z in *. + lia. + } + rewrite zlt_false by lia. + reflexivity. +Qed. -- cgit From a32b5bf92865e354f81330353857ae2d9ec80d57 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 25 Jan 2022 11:49:19 +0100 Subject: anystep_div_longu_correct --- kvx/FPDivision64.v | 38 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index d96ced31..e57f0d66 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -26,6 +26,7 @@ Require Import Values Globalenvs. Require Compopts. Require Import Psatz. Require Import IEEE754_extra. +Require Import Memory. From Gappa Require Import Gappa_tactic. @@ -1910,3 +1911,40 @@ Proof. rewrite zlt_false by lia. reflexivity. Qed. + +Definition anystep_div_longu a b m := + Val.select (Val.cmplu_bool (Mem.valid_pointer m) Cgt b (Vlong (Int64.repr smallb_thresh))) + (step2_div_longu a b) + (twostep_div_longu a b) Tlong. + +Lemma anystep_div_longu_correct : + forall a b m + (b_RANGE : (1 < Int64.unsigned b <= Int64.max_signed)%Z), + anystep_div_longu (Vlong a) (Vlong b) m = Vlong (Int64.repr (Int64.unsigned a / Int64.unsigned b))%Z. +Proof. + intros. + unfold anystep_div_longu. + set (a' := Int64.unsigned a). + set (b' := Int64.unsigned b). + assert (0 < b')%Z as b_NOT0 by lia. + cbn. + unfold Int64.ltu. + fold b'. + rewrite Int64.unsigned_repr; cycle 1. + { unfold smallb_thresh. + change Int64.max_unsigned with 18446744073709551615%Z. + lia. + } + destruct zlt. + { rewrite step2_div_longu_bigb_correct by lia. + reflexivity. + } + rewrite twostep_div_longu_smallb_correct by lia. + cbn. + rewrite Int64.eq_false ; cycle 1. + { intro Z. unfold b' in b_NOT0. rewrite Z in b_NOT0. + rewrite Int64.unsigned_zero in b_NOT0. + lia. + } + reflexivity. +Qed. -- cgit From 82d72ac1253b2c45dbc8f305349611ea2eae0675 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 25 Jan 2022 12:19:05 +0100 Subject: attempt at definition --- kvx/FPDivision64.v | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index e57f0d66..1739e5cf 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1948,3 +1948,16 @@ Proof. } reflexivity. Qed. + +Definition full_div_longu a b m := + let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in + let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in + let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in + let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Clt a b) + (Vlong Int64.zero) (Vlong Int64.one) Tlong in + let if_special := Val.select is_big if_big a Tlong in + Val.select (Val.cmp_bool Cne is_special (Vlong Int64.zero)) + if_special + (anystep_div_longu a b m) Tlong. + + -- cgit From 2d8a28a1db545c2fc6f25fc55057ea230decb5c9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 25 Jan 2022 13:06:18 +0100 Subject: some progress --- kvx/FPDivision64.v | 68 ++++++++++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 66 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 1739e5cf..bf9be23d 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1949,6 +1949,7 @@ Proof. reflexivity. Qed. + Definition full_div_longu a b m := let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in @@ -1956,8 +1957,71 @@ Definition full_div_longu a b m := let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Clt a b) (Vlong Int64.zero) (Vlong Int64.one) Tlong in let if_special := Val.select is_big if_big a Tlong in - Val.select (Val.cmp_bool Cne is_special (Vlong Int64.zero)) + Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) if_special (anystep_div_longu a b m) Tlong. - +Lemma one_bigb_div : + forall a b + (b_RANGE : (9223372036854775808 <= b < 18446744073709551616)%Z) + (ORDER : (b <= a < 18446744073709551616)%Z), + (a / b = 1)%Z. +Proof. + intros. + assert (((a - b) / b) = 0)%Z as ZERO. + { apply Zdiv_small. lia. + } + replace a with (1 * b + (a - b))%Z by ring. + rewrite Z.div_add_l by lia. + rewrite ZERO. + ring. +Qed. + +Lemma full_div_longu_correct : + forall a b m + (b_NOT0 : (0 < Int64.unsigned b)%Z), + full_div_longu (Vlong a) (Vlong b) m = Vlong (Int64.repr (Int64.unsigned a / Int64.unsigned b))%Z. +Proof. + intros. + unfold full_div_longu. + cbn. + unfold Int64.lt, Int64.ltu. + pose proof (Int64.unsigned_range a). + pose proof (Int64.unsigned_range b). + set (a' := Int64.unsigned a) in *. + set (b' := Int64.unsigned b) in *. + rewrite Int64.signed_zero. + rewrite Int64.unsigned_one. + destruct zlt. + { replace (Val.cmp_bool Cne + (Val.or Vtrue + (if negb (if zlt 1 b' then true else false) then Vtrue else Vfalse)) + (Vint Int.zero)) with (Some true) ; cycle 1. + { destruct zlt; reflexivity. + } + cbn. + destruct zlt; cbn. + { rewrite Zdiv_small by lia. + reflexivity. + } + rewrite one_bigb_div. + reflexivity. + { + change Int64.modulus with 18446744073709551616%Z in *. + split. 2: lia. + unfold b'. + rewrite Int64.unsigned_signed. + unfold Int64.lt. + rewrite Int64.signed_zero. + rewrite zlt_true by lia. + pose proof (Int64.signed_range b). + change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + +! -- cgit From 657a8d362bfdfdcd362749f68ec2b661d166df0b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 25 Jan 2022 13:14:51 +0100 Subject: full_div_longu_correct --- kvx/FPDivision64.v | 30 ++++++++++++++++++++++++++++-- 1 file changed, 28 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index bf9be23d..4e91b853 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1982,6 +1982,8 @@ Lemma full_div_longu_correct : (b_NOT0 : (0 < Int64.unsigned b)%Z), full_div_longu (Vlong a) (Vlong b) m = Vlong (Int64.repr (Int64.unsigned a / Int64.unsigned b))%Z. Proof. + + Local Opaque anystep_div_longu. intros. unfold full_div_longu. cbn. @@ -2023,5 +2025,29 @@ Proof. change Int64.modulus with 18446744073709551616%Z in *. lia. } - -! + destruct zlt; cbn. + { change (negb (Int.eq (Int.or Int.zero Int.zero) Int.zero)) with false. + cbn. + rewrite anystep_div_longu_correct. + reflexivity. + + change Int64.modulus with 18446744073709551616%Z in *. + split. lia. + rewrite Int64.unsigned_signed. + unfold Int64.lt. + rewrite Int64.signed_zero. + rewrite zlt_false by lia. + pose proof (Int64.signed_range b). + change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + change (negb (Int.eq (Int.or Int.zero Int.one) Int.zero)) with true. + cbn. + replace b' with 1%Z by lia. + rewrite Z.div_1_r. + unfold a'. + rewrite Int64.repr_unsigned. + reflexivity. +Qed. -- cgit From 2b05d4e7657a0d838e1003ababf19dca43029b64 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 2 Feb 2022 15:53:13 +0100 Subject: attempts at dealing with gappa --- Makefile | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/Makefile b/Makefile index 8b50375f..5ed2c580 100644 --- a/Makefile +++ b/Makefile @@ -77,7 +77,7 @@ COQCOPTS ?= \ cparser/Parser.vo: COQCOPTS += -w -deprecated-instance-without-locality -COQC="$(COQBIN)coqc" -q $(COQINCLUDES) $(COQCOPTS) +COQC=PATH=tools:$$PATH "$(COQBIN)coqc" -q $(COQINCLUDES) $(COQCOPTS) COQDEP="$(COQBIN)coqdep" $(COQINCLUDES) COQDOC="$(COQBIN)coqdoc" COQEXEC="$(COQBIN)coqtop" $(COQINCLUDES) -batch -load-vernac-source @@ -297,6 +297,13 @@ else ocamlc -o tools/modorder str.cma tools/modorder.ml endif +tools/gappa: + echo "#!/bin/sh" > $@ + echo -n "exec " >> $@ + which gappa | tr -d '\n' >> $@ + echo ' -Eprecision=100 "$$@"' >> $@ + chmod a+rx $@ + latexdoc: cd doc; $(COQDOC) --latex -o doc/doc.tex -g $(FILES) -- cgit From d8c61b62f1673ac36aca5584da5909c2a8994a9a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 2 Feb 2022 17:09:21 +0100 Subject: dependency to generate gappa script --- Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Makefile b/Makefile index 5ed2c580..a841569b 100644 --- a/Makefile +++ b/Makefile @@ -307,7 +307,7 @@ tools/gappa: latexdoc: cd doc; $(COQDOC) --latex -o doc/doc.tex -g $(FILES) -%.vo: %.v +%.vo: %.v tools/gappa @rm -f doc/$(*F).glob @echo "COQC $*.v" @$(COQC) -dump-glob doc/$(*F).glob $*.v -- cgit From e5a0de154dae052e0cddb82b4c116a1b14501a1e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 2 Feb 2022 18:48:28 +0100 Subject: update Coq --- .gitlab-ci.yml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 41df7d8e..378eb1fe 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -266,7 +266,7 @@ build_rv32: build_kvx: stage: build - image: coqorg/coq:8.12.2-ocaml-4.11.2-flambda + image: coq:8.13.2-ocaml-4.11.2-flambda before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison flex libmpfr-dev libboost-dev @@ -327,9 +327,9 @@ pages: # TODO: change to "deploy" when "build" succeeds (or integrate with "buil - if: '$CI_COMMIT_BRANCH == "master"' when: always -build_aarch64_coq13: +build_aarch64_coq14: stage: build - image: coqorg/coq:8.13.2-ocaml-4.11.2-flambda + image: coqorg/coq:8.14.1-ocaml-4.12.0-flambda before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install gcc-aarch64-linux-gnu qemu-user -- cgit From 760300e972f4c4b3eb14d90e79580bb1f2632197 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 2 Feb 2022 18:50:43 +0100 Subject: fix path to coq image --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 378eb1fe..c730a00e 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -266,7 +266,7 @@ build_rv32: build_kvx: stage: build - image: coq:8.13.2-ocaml-4.11.2-flambda + image: coqorg/coq:8.13.2-ocaml-4.11.2-flambda before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install sshpass openssh-client libzip4 lttng-tools liblttng-ctl-dev liblttng-ust-dev babeltrace libzip4 bison flex libmpfr-dev libboost-dev -- cgit From 607186f87c74817af529cd39d40390f0b86e59d4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 4 Feb 2022 20:17:14 +0100 Subject: some progress --- kvx/FPDivision64.v | 108 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 108 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 4e91b853..ddd96f8e 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1065,6 +1065,114 @@ Proof. gappa. Qed. +Lemma le_IZR3 : + forall n m p : Z, (IZR n <= IZR m <= IZR p)%R -> (n <= m <= p)%Z. +Proof. + intros ; split ; apply le_IZR ; lra. +Qed. + +(* 9223372036854775808%Z. *) +Definition bigb_thresh := 1000000000000000000%Z. +Lemma step1_div_longu_correct_anyb2 : + forall a b, + (1 < Int64.unsigned b <= bigb_thresh)%Z -> + exists (q : int64), + (step1_div_longu (Vlong a) (Vlong b)) = Vlong q /\ + (Int64.min_signed <= (Int64.unsigned a - Int64.unsigned b*Int64.unsigned q) <= Int64.max_signed)%Z. +Proof. + intros a b b_RANGE. + + pose proof (Int64.unsigned_range a) as a_RANGE. + change Int64.modulus with 18446744073709551616%Z in a_RANGE. + set (a' := Int64.unsigned a) in *. + set (b' := Int64.unsigned b) in *. + + destruct (Z_le_gt_dec a' 0). + { assert (a' = 0%Z) as ZERO by lia. + exists Int64.zero. + rewrite ZERO. + rewrite Int64.unsigned_zero. + replace (Z.abs (0 - b' * 0))%Z with 0%Z by lia. + replace a with Int64.zero; cycle 1. + { + unfold a' in ZERO. + unfold Int64.zero. + rewrite <- ZERO. + apply Int64.repr_unsigned. + } + split. + { apply step1_div_longu_a0. + lia. + } + change Int64.min_signed with (-9223372036854775808)%Z. + change Int64.max_signed with ( 9223372036854775807)%Z. + lia. + } + + unfold step1_div_longu. + assert (1 < b')%Z as b_NOT01 by lia. + destruct (step1_real_div_longu_correct a b b_NOT01) as (q & C1E & C1R & C1F & C1S). + rewrite C1E. cbn. + unfold Float.to_longu_ne. + pose proof (ZofB_ne_range_correct 53 1024 q 0 Int64.max_unsigned) as C2. + rewrite C1F in C2. + + + assert (1 <= IZR a' <= 18446744073709551615)%R as a_RANGE'. + { split; apply IZR_le; lia. } + assert (2 <= IZR b' <= IZR bigb_thresh)%R as b_RANGE'. + { split; apply IZR_le; lia. } + assert (1 <= IZR b' <= IZR bigb_thresh)%R as b_RANGE'' by lra. + cbn zeta in C2. + rewrite C2. + cbn. + rewrite C1R. + unfold step1_real_quotient. + fold a' b'. + unfold bigb_thresh in *. + + rewrite Zle_bool_true ; cycle 1. + { apply Znearest_IZR_le. + gappa. + } + rewrite Zle_bool_true ; cycle 1. + { apply Znearest_le_IZR. + gappa. + } + cbn. + econstructor; split. reflexivity. + set (q_r := (rd (rd (IZR a') * rd (rs (1 / rs (IZR b')))))%R). + assert (Rabs (IZR (ZnearestE q_r) - q_r) <= /2)%R as NEAR by apply Znearest_imp2. + set (delta1 := (IZR (ZnearestE q_r) - q_r)%R) in NEAR. + apply le_IZR3. + rewrite minus_IZR. + rewrite mult_IZR. + rewrite Int64.unsigned_repr ; cycle 1. + { split. + - apply Znearest_IZR_le. unfold q_r. + gappa. + - apply Znearest_le_IZR. unfold q_r. + change Int64.max_unsigned with 18446744073709551615%Z. + gappa. + } + replace (IZR (ZnearestE q_r)) with ((IZR (ZnearestE q_r) - q_r) + q_r)%R by ring. + fold delta1. + unfold q_r. + set (a1 := IZR a') in *. + set (b1 := IZR b') in *. + replace (rd (rd a1 * rd (rs (1 / rs b1))))%R with + ((((rd (rd a1 * rd (rs (1 / rs b1)))-(a1 * (1 / b1))) / (a1 * (1 / b1)))+1) * (a1 / b1))%R; + cycle 1. + { field. lra. } + set (delta2 := ((rd (rd a1 * rd (rs (1 / rs b1)))-(a1 * (1 / b1))) / (a1 * (1 / b1)))%R) in *. + assert (Rabs (delta2) <= 1/4194304)%R. + { unfold delta2. gappa. } + replace (a1 - b1 * (delta1 + (delta2 + 1) * (a1 / b1)))%R with + (-b1*delta1 - a1*delta2)%R; cycle 1. + { field. lra. } + gappa. +Qed. + Lemma find_quotient: forall (a b : Z) (b_POS : (0 < b)%Z) -- cgit From 9e37305cc4b817869de98b5cce49d3599d2c0aba Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 4 Feb 2022 20:19:17 +0100 Subject: progress --- kvx/FPDivision64.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index ddd96f8e..fc8e6173 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1071,8 +1071,7 @@ Proof. intros ; split ; apply le_IZR ; lra. Qed. -(* 9223372036854775808%Z. *) -Definition bigb_thresh := 1000000000000000000%Z. +Definition bigb_thresh := 9223372036854775808%Z. Lemma step1_div_longu_correct_anyb2 : forall a b, (1 < Int64.unsigned b <= bigb_thresh)%Z -> -- cgit From 6eac4cb3302f57d43e6b8d4c17688388db9619e5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 4 Feb 2022 20:35:11 +0100 Subject: progress --- kvx/FPDivision64.v | 15 ++++++++++----- 1 file changed, 10 insertions(+), 5 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index fc8e6173..79bfdc3b 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1071,15 +1071,16 @@ Proof. intros ; split ; apply le_IZR ; lra. Qed. -Definition bigb_thresh := 9223372036854775808%Z. +Definition bigb_thresh := 18446740000000000000%Z. Lemma step1_div_longu_correct_anyb2 : forall a b, - (1 < Int64.unsigned b <= bigb_thresh)%Z -> + (1 < Int64.unsigned b <= bigb_thresh)%Z -> + (Int64.unsigned b <= Int64.unsigned a)%Z -> exists (q : int64), (step1_div_longu (Vlong a) (Vlong b)) = Vlong q /\ (Int64.min_signed <= (Int64.unsigned a - Int64.unsigned b*Int64.unsigned q) <= Int64.max_signed)%Z. Proof. - intros a b b_RANGE. + intros a b b_RANGE LESS. pose proof (Int64.unsigned_range a) as a_RANGE. change Int64.modulus with 18446744073709551616%Z in a_RANGE. @@ -1164,11 +1165,15 @@ Proof. cycle 1. { field. lra. } set (delta2 := ((rd (rd a1 * rd (rs (1 / rs b1)))-(a1 * (1 / b1))) / (a1 * (1 / b1)))%R) in *. - assert (Rabs (delta2) <= 1/4194304)%R. - { unfold delta2. gappa. } + (* assert (Rabs (delta2) <= 1/4194304)%R. + { unfold delta2. gappa. } *) replace (a1 - b1 * (delta1 + (delta2 + 1) * (a1 / b1)))%R with (-b1*delta1 - a1*delta2)%R; cycle 1. { field. lra. } + assert (b1 <= a1)%R as LESS'. + { unfold a1, b1. + apply IZR_le. lia. } + unfold delta2. gappa. Qed. -- cgit From 7b4583e6c008469a4300b2fe2d405c8de3753238 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 9 Feb 2022 15:58:23 +0100 Subject: twostep_div_longu_mostb_correct --- kvx/FPDivision64.v | 86 +++++++++++++++++++++++++++++++++++++++++++++++------- 1 file changed, 75 insertions(+), 11 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 79bfdc3b..53961294 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1071,16 +1071,15 @@ Proof. intros ; split ; apply le_IZR ; lra. Qed. -Definition bigb_thresh := 18446740000000000000%Z. -Lemma step1_div_longu_correct_anyb2 : +Definition mostb_thresh := 18446740000000000000%Z. +Lemma step1_div_longu_correct_mostb : forall a b, - (1 < Int64.unsigned b <= bigb_thresh)%Z -> - (Int64.unsigned b <= Int64.unsigned a)%Z -> + (1 < Int64.unsigned b <= mostb_thresh)%Z -> exists (q : int64), (step1_div_longu (Vlong a) (Vlong b)) = Vlong q /\ (Int64.min_signed <= (Int64.unsigned a - Int64.unsigned b*Int64.unsigned q) <= Int64.max_signed)%Z. Proof. - intros a b b_RANGE LESS. + intros a b b_RANGE. pose proof (Int64.unsigned_range a) as a_RANGE. change Int64.modulus with 18446744073709551616%Z in a_RANGE. @@ -1120,16 +1119,16 @@ Proof. assert (1 <= IZR a' <= 18446744073709551615)%R as a_RANGE'. { split; apply IZR_le; lia. } - assert (2 <= IZR b' <= IZR bigb_thresh)%R as b_RANGE'. + assert (2 <= IZR b' <= IZR mostb_thresh)%R as b_RANGE'. { split; apply IZR_le; lia. } - assert (1 <= IZR b' <= IZR bigb_thresh)%R as b_RANGE'' by lra. + assert (1 <= IZR b' <= IZR mostb_thresh)%R as b_RANGE'' by lra. cbn zeta in C2. rewrite C2. cbn. rewrite C1R. unfold step1_real_quotient. fold a' b'. - unfold bigb_thresh in *. + unfold mostb_thresh in *. rewrite Zle_bool_true ; cycle 1. { apply Znearest_IZR_le. @@ -1170,9 +1169,6 @@ Proof. replace (a1 - b1 * (delta1 + (delta2 + 1) * (a1 / b1)))%R with (-b1*delta1 - a1*delta2)%R; cycle 1. { field. lra. } - assert (b1 <= a1)%R as LESS'. - { unfold a1, b1. - apply IZR_le. lia. } unfold delta2. gappa. Qed. @@ -2163,3 +2159,71 @@ Proof. rewrite Int64.repr_unsigned. reflexivity. Qed. + +Lemma repr_unsigned_sub2: + forall a b, + (Int64.repr (a - Int64.unsigned (Int64.repr b))) = Int64.repr (a - b). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_sub. + - apply Int64.eqm_refl. + - apply Int64.eqm_sym. apply Int64.eqm_unsigned_repr. +Qed. + +Lemma repr_unsigned_add2: + forall a b, + (Int64.repr (a + Int64.unsigned (Int64.repr b))) = Int64.repr (a + b). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_add. + - apply Int64.eqm_refl. + - apply Int64.eqm_sym. apply Int64.eqm_unsigned_repr. +Qed. + +Lemma twostep_div_longu_mostb_correct : + forall a b + (b_RANGE : (1 < Int64.unsigned b <= Int64.max_signed)%Z), + (twostep_div_longu (Vlong a) (Vlong b)) = + (Val.maketotal (Val.divlu (Vlong a) (Vlong b))). +Proof. + intros. + destruct (Z_le_gt_dec (Int64.unsigned b) smallb_thresh). + { apply twostep_div_longu_smallb_correct. lia. } + set (a' := Int64.unsigned a). + set (b' := Int64.unsigned b). + assert (0 < b')%Z as b_NOT0 by lia. + cbn. + rewrite Int64.eq_false ; cycle 1. + { intro Z. unfold b' in b_NOT0. rewrite Z in b_NOT0. + rewrite Int64.unsigned_zero in b_NOT0. + lia. + } + cbn. + + unfold twostep_div_longu. + assert (1 < Int64.unsigned b <= mostb_thresh)%Z as MOST_B. + { unfold mostb_thresh. + change Int64.max_signed with 9223372036854775807%Z in b_RANGE. + lia. + } + destruct (step1_div_longu_correct_mostb a b MOST_B) as + (q & step1_REW & step1_RANGE). + rewrite step1_REW. + cbn. + rewrite step2_div_long_bigb_correct; cycle 1. + 1, 2: lia. + f_equal. + + unfold Int64.sub, Int64.mul. + rewrite repr_unsigned_sub2. + rewrite Int64.signed_repr by lia. + unfold Int64.add, Int64.divu. + fold a' b'. + set (q' := Int64.unsigned q) in *. + rewrite repr_unsigned_add2. + rewrite <- Z.div_add_l by lia. + f_equal. f_equal. + ring. +Qed. -- cgit From 57c868e304df143918fa3a99d6272437a14299cc Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 9 Feb 2022 16:10:59 +0100 Subject: full2_div_longu_correct --- kvx/FPDivision64.v | 93 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 93 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 53961294..38eb6c1b 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2227,3 +2227,96 @@ Proof. f_equal. f_equal. ring. Qed. + +Definition full2_div_longu a b m := + let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in + let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in + let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in + let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Clt a b) + (Vlong Int64.zero) (Vlong Int64.one) Tlong in + let if_special := Val.select is_big if_big a Tlong in + Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) + if_special + (twostep_div_longu a b) Tlong. + +Lemma full2_div_longu_correct : + forall a b m + (b_NOT0 : (0 < Int64.unsigned b)%Z), + full2_div_longu (Vlong a) (Vlong b) m = Vlong (Int64.repr (Int64.unsigned a / Int64.unsigned b))%Z. +Proof. + + Local Opaque twostep_div_longu. + intros. + unfold full2_div_longu. + cbn. + unfold Int64.lt, Int64.ltu. + pose proof (Int64.unsigned_range a). + pose proof (Int64.unsigned_range b). + set (a' := Int64.unsigned a) in *. + set (b' := Int64.unsigned b) in *. + rewrite Int64.signed_zero. + rewrite Int64.unsigned_one. + destruct zlt. + { replace (Val.cmp_bool Cne + (Val.or Vtrue + (if negb (if zlt 1 b' then true else false) then Vtrue else Vfalse)) + (Vint Int.zero)) with (Some true) ; cycle 1. + { destruct zlt; reflexivity. + } + cbn. + destruct zlt; cbn. + { rewrite Zdiv_small by lia. + reflexivity. + } + rewrite one_bigb_div. + reflexivity. + { + change Int64.modulus with 18446744073709551616%Z in *. + split. 2: lia. + unfold b'. + rewrite Int64.unsigned_signed. + unfold Int64.lt. + rewrite Int64.signed_zero. + rewrite zlt_true by lia. + pose proof (Int64.signed_range b). + change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + destruct zlt; cbn. + { change (negb (Int.eq (Int.or Int.zero Int.zero) Int.zero)) with false. + cbn. + rewrite twostep_div_longu_mostb_correct. + { + cbn. + unfold Int64.eq. + fold b'. + rewrite Int64.unsigned_zero. + rewrite zeq_false by lia. + reflexivity. + } + + change Int64.modulus with 18446744073709551616%Z in *. + split. lia. + rewrite Int64.unsigned_signed. + unfold Int64.lt. + rewrite Int64.signed_zero. + rewrite zlt_false by lia. + pose proof (Int64.signed_range b). + change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + change (negb (Int.eq (Int.or Int.zero Int.one) Int.zero)) with true. + cbn. + replace b' with 1%Z by lia. + rewrite Z.div_1_r. + unfold a'. + rewrite Int64.repr_unsigned. + reflexivity. +Qed. -- cgit From dc50627d829a4600b3248c4b79ce20f87f3e6d21 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 9 Feb 2022 17:23:30 +0100 Subject: step1 expression is ok --- kvx/FPDivision64.v | 36 ++++++++++++++++++++++++++++++++++++ 1 file changed, 36 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 38eb6c1b..2c6a49c5 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2320,3 +2320,39 @@ Proof. rewrite Int64.repr_unsigned. reflexivity. Qed. + +Open Scope cminorsel_scope. +Definition e_invfs a := Eop Oinvfs (a ::: Enil). +Definition e_float_of_longu a := Eop Ofloatoflongu (a ::: Enil). +Definition e_single_of_longu a := Eop Osingleoflongu (a ::: Enil). +Definition e_float_of_single a := Eop Ofloatofsingle (a ::: Enil). +Definition e_longu_of_float_ne a := Eop Olonguoffloat_ne (a ::: Enil). +Definition e_mulf a b := Eop Omulf (a ::: b ::: Enil). + +Definition a_var1 := Eletvar (4%nat). +Definition a_d_var1 := Eletvar (3%nat). +Definition b_var1 := Eletvar (2%nat). +Definition b_d_var1 := Eletvar (1%nat). +Definition binv_d_var1 := Eletvar (0%nat). + +Definition e_setup1 a b rest := + Elet a (Elet (e_float_of_longu (Eletvar 0%nat)) + (Elet (lift (lift b)) (Elet (e_float_of_longu (Eletvar 0%nat)) + (Elet (e_float_of_single (e_invfs (e_single_of_longu (Eletvar 1%nat)))) + rest)))). +Definition e_step1 := e_longu_of_float_ne (e_mulf a_d_var1 binv_d_var1). + +Lemma e_step1_correct : + forall (ge : genv) (sp: val) cmenv memenv (le : letenv) + (expr_a : expr) (a : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (expr_b : expr) (b : int64) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)), + (eval_expr ge sp cmenv memenv le (e_setup1 expr_a expr_b (e_step1)) + (step1_div_longu (Vlong a) (Vlong b))). +Proof. + intros. + unfold e_setup1, step1_div_longu. + repeat econstructor. + { eassumption. } + { cbn. apply eval_lift. apply eval_lift. eassumption. } +Qed. + -- cgit From 68149e3a492be432561e61356f570ba394ca1e3a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 10:18:46 +0100 Subject: progress --- kvx/FPDivision64.v | 38 ++++++++++++++++++++++++++++++++++++-- 1 file changed, 36 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 2c6a49c5..619fcfa6 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2324,11 +2324,15 @@ Qed. Open Scope cminorsel_scope. Definition e_invfs a := Eop Oinvfs (a ::: Enil). Definition e_float_of_longu a := Eop Ofloatoflongu (a ::: Enil). +Definition e_float_of_long a := Eop Ofloatoflong (a ::: Enil). Definition e_single_of_longu a := Eop Osingleoflongu (a ::: Enil). Definition e_float_of_single a := Eop Ofloatofsingle (a ::: Enil). +Definition e_long_of_float_ne a := Eop Olongoffloat_ne (a ::: Enil). Definition e_longu_of_float_ne a := Eop Olonguoffloat_ne (a ::: Enil). Definition e_mulf a b := Eop Omulf (a ::: b ::: Enil). - +Definition e_float_const c := Eop (Ofloatconst c) Enil. +Definition e_fmaddf a b c := Eop Ofmaddf (a ::: b ::: c ::: Enil). +Definition e_fmsubf a b c := Eop Ofmsubf (a ::: b ::: c ::: Enil). Definition a_var1 := Eletvar (4%nat). Definition a_d_var1 := Eletvar (3%nat). Definition b_var1 := Eletvar (2%nat). @@ -2355,4 +2359,34 @@ Proof. { eassumption. } { cbn. apply eval_lift. apply eval_lift. eassumption. } Qed. - + +Definition e_setup2 a b rest := (e_setup1 a b (Elet e_step1 rest)). + +Definition a_var2 := Eletvar (5%nat). +Definition a_d_var2 := Eletvar (4%nat). +Definition b_var2 := Eletvar (3%nat). +Definition b_d_var2 := Eletvar (2%nat). +Definition binv_d_var2 := Eletvar (1%nat). +Definition step1_var2 := Eletvar (0%nat). + +Definition e_step2 := + e_long_of_float_ne + (e_mulf (e_float_of_long step1_var2) + (e_fmaddf + binv_d_var2 + (e_fmsubf (e_float_const ExtFloat.one) + binv_d_var2 + b_d_var2 ) + binv_d_var2)). + +Lemma e_step2_correct : + forall (ge : genv) (sp: val) cmenv memenv (le : letenv) + (expr_a : expr) (a : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (expr_b : expr) (b : int64) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)), + (eval_expr ge sp cmenv memenv le (e_setup2 expr_a expr_b (e_step2)) + (step2_div_long' (step1_div_longu (Vlong a) (Vlong b)) (Vlong b))). +Proof. +intros. +unfold e_setup2, e_setup1, e_step2, step2_div_long', step2_real_div_long, approx_inv_longu. +repeat (econstructor + apply eval_lift + eassumption). +Qed. -- cgit From ea53967476937f15ec4ceea1b31ece4d75b85f98 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 10:51:46 +0100 Subject: progress --- kvx/FPDivision64.v | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 619fcfa6..572df195 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2333,6 +2333,10 @@ Definition e_mulf a b := Eop Omulf (a ::: b ::: Enil). Definition e_float_const c := Eop (Ofloatconst c) Enil. Definition e_fmaddf a b c := Eop Ofmaddf (a ::: b ::: c ::: Enil). Definition e_fmsubf a b c := Eop Ofmsubf (a ::: b ::: c ::: Enil). +Definition e_addlimm a b := Eop (Oaddlimm b) (a ::: Enil). +Definition e_msub a b c := Eop Omsub (a ::: b ::: c ::: Enil). +Definition e_ite ty c vc v1 v2 := Eop (Osel c ty) (vc ::: v1 ::: v2 ::: Enil). + Definition a_var1 := Eletvar (4%nat). Definition a_d_var1 := Eletvar (3%nat). Definition b_var1 := Eletvar (2%nat). @@ -2390,3 +2394,17 @@ intros. unfold e_setup2, e_setup1, e_step2, step2_div_long', step2_real_div_long, approx_inv_longu. repeat (econstructor + apply eval_lift + eassumption). Qed. + +Definition e_setup3 a b rest := (e_setup2 a b (Elet e_step2 rest)). + +Definition a_var3 := Eletvar (6%nat). +Definition a_d_var3 := Eletvar (5%nat). +Definition b_var3 := Eletvar (4%nat). +Definition b_d_var3 := Eletvar (3%nat). +Definition binv_d_var3 := Eletvar (2%nat). +Definition step1_var3 := Eletvar (1%nat). +Definition step2_var3 := Eletvar (0%nat). + +Definition e_step3 := + e_ite Tlong (Ccompl0 Clt) (e_msub a_var3 step2_var3 b_var3) + (e_addlimm step2_var3 Int64.mone) step2_var3. -- cgit From 3cb1d2813f8ff6ac9cbad44974fee00b8a0236c9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 11:01:38 +0100 Subject: Cgt / Clt --- kvx/FPDivision64.v | 27 ++++++++++++++++++++------- 1 file changed, 20 insertions(+), 7 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 572df195..34dc108b 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1321,7 +1321,7 @@ Definition step2_div_long' a b := Definition step2_div_long a b := let q := step2_div_long' a b in - Val.select (Val.cmpl_bool Cgt (Val.subl (Val.mull q b) a) (Vlong Int64.zero)) + Val.select (Val.cmpl_bool Clt (Val.subl a (Val.mull q b)) (Vlong Int64.zero)) (Val.addl q (Vlong Int64.mone)) q Tlong. Lemma function_ite : @@ -1377,6 +1377,19 @@ Proof. apply int64_eqm_signed_repr. - apply Int64.eqm_refl. Qed. + +Lemma signed_repr_sub2: + forall x y, + Int64.repr (x - Int64.signed (Int64.repr y)) = + Int64.repr (x - y). +Proof. + intros. + apply Int64.eqm_samerepr. + apply Int64.eqm_sub. + - apply Int64.eqm_refl. + - apply Int64.eqm_sym. + apply int64_eqm_signed_repr. +Qed. Lemma step2_div_long_smalla_correct : forall a b @@ -1454,7 +1467,7 @@ Proof. rewrite Int64.sub_signed. fold a'. - rewrite signed_repr_sub. + rewrite signed_repr_sub2. assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. { replace (IZR a' / IZR b' - q')%R with @@ -1473,8 +1486,8 @@ Proof. destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. replace q'' with (1 + (a' / b'))%Z by lia. - replace ((1 + a' / b') * b' - a')%Z - with (-(a' - a' / b' * b')+b')%Z by ring. + replace (a' - (1 + a' / b') * b')%Z + with ((a' - a' / b' * b')-b')%Z by ring. rewrite <- MOD. rewrite Int64.signed_repr; cycle 1. { change Int64.min_signed with (-9223372036854775808)%Z in *. @@ -1731,7 +1744,7 @@ Proof. rewrite Int64.sub_signed. fold a'. - rewrite signed_repr_sub. + rewrite signed_repr_sub2. assert ((Rabs (IZR a' / IZR b' - q') < / 2)%R) as HALF. { replace (IZR a' / IZR b' - q')%R with @@ -1750,8 +1763,8 @@ Proof. destruct (Z_lt_dec a' (b' * q'')) as [LT | GE]. { replace (b' * q'' >? a')%Z with true in QUOTIENT by lia. replace q'' with (1 + (a' / b'))%Z by lia. - replace ((1 + a' / b') * b' - a')%Z - with (-(a' - a' / b' * b')+b')%Z by ring. + replace (a' - (1 + a' / b') * b')%Z + with ((a' - a' / b' * b')-b')%Z by ring. rewrite <- MOD. rewrite Int64.signed_repr; cycle 1. { change Int64.min_signed with (-9223372036854775808)%Z in *. -- cgit From aa7c85e476d1f63d834cd73e3a2da81b73afd32b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 11:49:59 +0100 Subject: progress --- kvx/FPDivision64.v | 18 +++++++++++++++--- 1 file changed, 15 insertions(+), 3 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 34dc108b..f991ff93 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2347,8 +2347,8 @@ Definition e_float_const c := Eop (Ofloatconst c) Enil. Definition e_fmaddf a b c := Eop Ofmaddf (a ::: b ::: c ::: Enil). Definition e_fmsubf a b c := Eop Ofmsubf (a ::: b ::: c ::: Enil). Definition e_addlimm a b := Eop (Oaddlimm b) (a ::: Enil). -Definition e_msub a b c := Eop Omsub (a ::: b ::: c ::: Enil). -Definition e_ite ty c vc v1 v2 := Eop (Osel c ty) (vc ::: v1 ::: v2 ::: Enil). +Definition e_msubl a b c := Eop Omsubl (a ::: b ::: c ::: Enil). +Definition e_ite ty c vc v1 v2 := Eop (Osel c ty) (v1 ::: v2 ::: vc ::: Enil). Definition a_var1 := Eletvar (4%nat). Definition a_d_var1 := Eletvar (3%nat). @@ -2419,5 +2419,17 @@ Definition step1_var3 := Eletvar (1%nat). Definition step2_var3 := Eletvar (0%nat). Definition e_step3 := - e_ite Tlong (Ccompl0 Clt) (e_msub a_var3 step2_var3 b_var3) + e_ite Tlong (Ccompl0 Clt) (e_msubl step1_var3 step2_var3 b_var3) (e_addlimm step2_var3 Int64.mone) step2_var3. + +Lemma e_step3_correct : + forall (ge : genv) (sp: val) cmenv memenv (le : letenv) + (expr_a : expr) (a : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (expr_b : expr) (b : int64) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)), + (eval_expr ge sp cmenv memenv le (e_setup3 expr_a expr_b (e_step3)) + (step2_div_long (step1_div_longu (Vlong a) (Vlong b)) (Vlong b))). +Proof. +intros. +unfold e_setup2, e_setup1, e_step2, step2_div_long, step2_div_long', step2_real_div_long, approx_inv_longu, step1_div_longu. +repeat (econstructor + apply eval_lift + eassumption). +Qed. -- cgit From 3e445a4f31b5b309b64cdf307830716f0001003f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 17:16:26 +0100 Subject: progress --- kvx/FPDivision64.v | 59 +++++++++++++++++++++++++++++++++--------------------- 1 file changed, 36 insertions(+), 23 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index f991ff93..457b280b 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2386,48 +2386,61 @@ Definition b_d_var2 := Eletvar (2%nat). Definition binv_d_var2 := Eletvar (1%nat). Definition step1_var2 := Eletvar (0%nat). -Definition e_step2 := +Definition e_step2 := e_msubl a_var2 b_var2 step1_var2. + +Definition e_setup3 a b rest := (e_setup2 a b (Elet e_step2 rest)). + +Definition a_var3 := Eletvar (6%nat). +Definition a_d_var3 := Eletvar (5%nat). +Definition b_var3 := Eletvar (4%nat). +Definition b_d_var3 := Eletvar (3%nat). +Definition binv_d_var3 := Eletvar (2%nat). +Definition step1_var3 := Eletvar (1%nat). +Definition step2_var3 := Eletvar (0%nat). + +Definition e_step3 := e_long_of_float_ne - (e_mulf (e_float_of_long step1_var2) + (e_mulf (e_float_of_long step2_var3) (e_fmaddf - binv_d_var2 + binv_d_var3 (e_fmsubf (e_float_const ExtFloat.one) - binv_d_var2 - b_d_var2 ) - binv_d_var2)). + binv_d_var3 + b_d_var3 ) + binv_d_var3)). -Lemma e_step2_correct : +Lemma e_step3_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_a : expr) (a : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) (expr_b : expr) (b : int64) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)), - (eval_expr ge sp cmenv memenv le (e_setup2 expr_a expr_b (e_step2)) - (step2_div_long' (step1_div_longu (Vlong a) (Vlong b)) (Vlong b))). + (eval_expr ge sp cmenv memenv le (e_setup3 expr_a expr_b (e_step3)) + (step2_div_long' (Val.subl (Vlong a) (Val.mull (Vlong b) (step1_div_longu (Vlong a) (Vlong b)))) (Vlong b))). Proof. intros. unfold e_setup2, e_setup1, e_step2, step2_div_long', step2_real_div_long, approx_inv_longu. repeat (econstructor + apply eval_lift + eassumption). Qed. -Definition e_setup3 a b rest := (e_setup2 a b (Elet e_step2 rest)). +Definition e_setup4 a b rest := (e_setup3 a b (Elet e_step3 rest)). -Definition a_var3 := Eletvar (6%nat). -Definition a_d_var3 := Eletvar (5%nat). -Definition b_var3 := Eletvar (4%nat). -Definition b_d_var3 := Eletvar (3%nat). -Definition binv_d_var3 := Eletvar (2%nat). -Definition step1_var3 := Eletvar (1%nat). -Definition step2_var3 := Eletvar (0%nat). +Definition a_var4 := Eletvar (7%nat). +Definition a_d_var4 := Eletvar (6%nat). +Definition b_var4 := Eletvar (5%nat). +Definition b_d_var4 := Eletvar (4%nat). +Definition binv_d_var4 := Eletvar (3%nat). +Definition step1_var4 := Eletvar (2%nat). +Definition step2_var4 := Eletvar (1%nat). +Definition step3_var4 := Eletvar (0%nat). -Definition e_step3 := - e_ite Tlong (Ccompl0 Clt) (e_msubl step1_var3 step2_var3 b_var3) - (e_addlimm step2_var3 Int64.mone) step2_var3. +Definition e_step4 := + e_ite Tlong (Ccompl0 Clt) (e_msubl step2_var4 step3_var4 b_var4) + (e_addlimm step3_var4 Int64.mone) step3_var4. -Lemma e_step3_correct : +Lemma e_step4_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_a : expr) (a : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) (expr_b : expr) (b : int64) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)), - (eval_expr ge sp cmenv memenv le (e_setup3 expr_a expr_b (e_step3)) - (step2_div_long (step1_div_longu (Vlong a) (Vlong b)) (Vlong b))). + (eval_expr ge sp cmenv memenv le (e_setup4 expr_a expr_b (e_step4)) + (step2_div_long (Val.subl (Vlong a) (Val.mull (Vlong b) (step1_div_longu (Vlong a) (Vlong b)))) (Vlong b))). Proof. intros. unfold e_setup2, e_setup1, e_step2, step2_div_long, step2_div_long', step2_real_div_long, approx_inv_longu, step1_div_longu. -- cgit From d8a3b4a525bcd3ea17d1585ec2e93f3d2dcacba6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 19:27:56 +0100 Subject: progress --- kvx/FPDivision64.v | 46 ++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 44 insertions(+), 2 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 457b280b..a1ab09b8 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2348,8 +2348,10 @@ Definition e_fmaddf a b c := Eop Ofmaddf (a ::: b ::: c ::: Enil). Definition e_fmsubf a b c := Eop Ofmsubf (a ::: b ::: c ::: Enil). Definition e_addlimm a b := Eop (Oaddlimm b) (a ::: Enil). Definition e_msubl a b c := Eop Omsubl (a ::: b ::: c ::: Enil). -Definition e_ite ty c vc v1 v2 := Eop (Osel c ty) (v1 ::: v2 ::: vc ::: Enil). - +Definition e_ite ty c vc v1 v2 := Eop (Osel c ty) (v1 ::: v2 ::: vc ::: Enil). +Definition e_cmplimm c v n := Eop (Ocmp (Ccomplimm c n)) (v ::: Enil). +Definition e_addl a b := Eop Oaddl (a ::: b ::: Enil). + Definition a_var1 := Eletvar (4%nat). Definition a_d_var1 := Eletvar (3%nat). Definition b_var1 := Eletvar (2%nat). @@ -2446,3 +2448,43 @@ intros. unfold e_setup2, e_setup1, e_step2, step2_div_long, step2_div_long', step2_real_div_long, approx_inv_longu, step1_div_longu. repeat (econstructor + apply eval_lift + eassumption). Qed. + +Definition e_setup5 a b rest := (e_setup4 a b (Elet e_step4 rest)). + + +Definition a_var5 := Eletvar (8%nat). +Definition a_d_var5 := Eletvar (7%nat). +Definition b_var5 := Eletvar (6%nat). +Definition b_d_var5 := Eletvar (5%nat). +Definition binv_d_var5 := Eletvar (4%nat). +Definition step1_var5 := Eletvar (3%nat). +Definition step2_var5 := Eletvar (2%nat). +Definition step3_var5 := Eletvar (1%nat). +Definition step4_var5 := Eletvar (0%nat). + +Definition e_step5 := e_addl step1_var5 step4_var5. + +Lemma e_step5_correct : + forall (ge : genv) (sp: val) cmenv memenv (le : letenv) + (expr_a : expr) (a : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (expr_b : expr) (b : int64) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)), + (eval_expr ge sp cmenv memenv le (e_setup5 expr_a expr_b (e_step5)) + (twostep_div_longu (Vlong a) (Vlong b))). +Proof. + intros. + Local Transparent twostep_div_longu. + repeat unfold e_setup2, e_setup1, e_step2, step2_div_long, step2_div_long', step2_real_div_long, approx_inv_longu, step1_div_longu, twostep_div_longu. +repeat (econstructor + apply eval_lift + eassumption). +Qed. + +(* + let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in + let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in + let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in + let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Clt a b) + (Vlong Int64.zero) (Vlong Int64.one) Tlong in + let if_special := Val.select is_big if_big a Tlong in + Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) + if_special + (twostep_div_longu a b) Tlong. + *) -- cgit From 6bd3c44a582d3d1f23fe4ed99dec8b59893e3a85 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 22:35:22 +0100 Subject: finished proof but still why need for decomposition --- kvx/FPDivision64.v | 72 ++++++++++++++++++++++++++++++++++++++++++------------ 1 file changed, 56 insertions(+), 16 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index a1ab09b8..c08103b1 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2075,8 +2075,7 @@ Definition full_div_longu a b m := let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in - let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Clt a b) - (Vlong Int64.zero) (Vlong Int64.one) Tlong in + let if_big := Val.longofintu (Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer m) Cge a b)) in let if_special := Val.select is_big if_big a Tlong in Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) if_special @@ -2245,8 +2244,7 @@ Definition full2_div_longu a b m := let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in - let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Clt a b) - (Vlong Int64.zero) (Vlong Int64.one) Tlong in + let if_big := Val.longofintu (Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer m) Cge a b)) in let if_special := Val.select is_big if_big a Tlong in Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) if_special @@ -2350,7 +2348,11 @@ Definition e_addlimm a b := Eop (Oaddlimm b) (a ::: Enil). Definition e_msubl a b c := Eop Omsubl (a ::: b ::: c ::: Enil). Definition e_ite ty c vc v1 v2 := Eop (Osel c ty) (v1 ::: v2 ::: vc ::: Enil). Definition e_cmplimm c v n := Eop (Ocmp (Ccomplimm c n)) (v ::: Enil). +Definition e_cmpluimm c v n := Eop (Ocmp (Ccompluimm c n)) (v ::: Enil). Definition e_addl a b := Eop Oaddl (a ::: b ::: Enil). +Definition e_or a b := Eop Oor (a ::: b ::: Enil). +Definition e_cast32unsigned a := Eop Ocast32unsigned (a ::: Enil). +Definition e_cmplu c a b := Eop (Ocmp (Ccomplu c)) (a ::: b ::: Enil). Definition a_var1 := Eletvar (4%nat). Definition a_d_var1 := Eletvar (3%nat). @@ -2451,7 +2453,6 @@ Qed. Definition e_setup5 a b rest := (e_setup4 a b (Elet e_step4 rest)). - Definition a_var5 := Eletvar (8%nat). Definition a_d_var5 := Eletvar (7%nat). Definition b_var5 := Eletvar (6%nat). @@ -2477,14 +2478,53 @@ Proof. repeat (econstructor + apply eval_lift + eassumption). Qed. -(* - let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in - let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in - let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in - let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Clt a b) - (Vlong Int64.zero) (Vlong Int64.one) Tlong in - let if_special := Val.select is_big if_big a Tlong in - Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) - if_special - (twostep_div_longu a b) Tlong. - *) +Definition e_setup6 a b rest := (e_setup5 a b (Elet e_step5 rest)). + +Definition a_var6 := Eletvar (9%nat). +Definition a_d_var6 := Eletvar (8%nat). +Definition b_var6 := Eletvar (7%nat). +Definition b_d_var6 := Eletvar (6%nat). +Definition binv_d_var6 := Eletvar (5%nat). +Definition step1_var6 := Eletvar (4%nat). +Definition step2_var6 := Eletvar (3%nat). +Definition step3_var6 := Eletvar (2%nat). +Definition step4_var6 := Eletvar (1%nat). +Definition twostep_var6 := Eletvar (0%nat). + +Definition e_step6 := e_cmplimm Clt b_var6 Int64.zero. + +Definition e_setup7 a b rest := e_setup6 a b (Elet e_step6 rest). + +Definition a_var7 := Eletvar (10%nat). +Definition a_d_var7 := Eletvar (9%nat). +Definition b_var7 := Eletvar (8%nat). +Definition b_d_var7 := Eletvar (7%nat). +Definition binv_d_var7 := Eletvar (6%nat). +Definition step1_var7 := Eletvar (5%nat). +Definition step2_var7 := Eletvar (5%nat). +Definition step3_var7 := Eletvar (3%nat). +Definition step4_var7 := Eletvar (2%nat). +Definition twostep_var7 := Eletvar (1%nat). +Definition is_big_var7 := Eletvar (0%nat). + +Definition e_is_one := e_cmpluimm Cle b_var7 Int64.one. +Definition e_is_special := e_or is_big_var7 e_is_one. +Definition e_if_big := e_cast32unsigned (e_cmplu Cge a_var7 b_var7). +Definition e_if_special := e_ite Tlong (Ccompu0 Cne) is_big_var7 e_if_big a_var7. +Definition e_step7 := e_ite Tlong (Ccompu0 Cne) e_is_special e_if_special twostep_var7. + +Lemma e_step7_correct : + forall (ge : genv) (sp: val) cmenv memenv (le : letenv) + (expr_a : expr) (a : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (expr_b : expr) (b : int64) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)), + (eval_expr ge sp cmenv memenv le (e_setup7 expr_a expr_b (e_step7)) + (full2_div_longu (Vlong a) (Vlong b) memenv)). +Proof. + intros. + Local Transparent full2_div_longu. + repeat unfold e_setup2, e_setup1, e_step2, step2_div_long, step2_div_long', step2_real_div_long, approx_inv_longu, step1_div_longu, twostep_div_longu, full2_div_longu. + repeat (econstructor + apply eval_lift + eassumption). + cbn. + repeat f_equal. + destruct (Int64.lt b Int64.zero); cbn; change (Int.eq Int.one Int.zero) with false; change (Int.eq Int.zero Int.zero) with true; cbn; reflexivity. +Qed. -- cgit From b45d9ac4b98ed9ce095a1ad64a4192fea89f7b5f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 10 Feb 2022 22:48:26 +0100 Subject: fp_divu64_correct --- kvx/FPDivision64.v | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index c08103b1..411865ba 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2528,3 +2528,20 @@ Proof. repeat f_equal. destruct (Int64.lt b Int64.zero); cbn; change (Int.eq Int.one Int.zero) with false; change (Int.eq Int.zero Int.zero) with true; cbn; reflexivity. Qed. + +Definition fp_divu64 a b := e_setup7 a b e_step7. + +Theorem fp_divu64_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int64) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)) + (b_nz : (Int64.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_divu64 expr_a expr_b) + (Vlong (Int64.divu a b)). +Proof. + intros. + unfold Int64.divu. + rewrite <- full2_div_longu_correct with (m := memenv) by lia. + apply e_step7_correct; assumption. +Qed. -- cgit From f8d32a19caf88733a9bbeee976f5c2fc549d4f92 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 10:32:15 +0100 Subject: rewrite with longu -> double -> single conversions --- kvx/FPDivision64.v | 61 ++++++++++++++++++++++++++++++++++-------------------- 1 file changed, 38 insertions(+), 23 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 411865ba..a4d609f0 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -532,7 +532,7 @@ Proof. Qed. Definition step1_real_inv_longu b := - let invb_s := ExtValues.invfs (Val.maketotal (Val.singleoflongu b)) in + let invb_s := ExtValues.invfs (Val.singleoffloat (Val.maketotal (Val.floatoflongu b))) in Val.floatofsingle invb_s. Definition step1_real_inv_thresh := (3/33554432)%R. @@ -543,7 +543,7 @@ Theorem step1_real_inv_longu_correct : (0 < Int64.unsigned b)%Z -> exists (f : float), (step1_real_inv_longu (Vlong b)) = Vfloat f /\ - (B2R _ _ f) = (rd (rs (1 / rs (IZR (Int64.unsigned b))))) /\ + (B2R _ _ f) = (rd (rs (1 / rs (rd (IZR (Int64.unsigned b)))))) /\ is_finite _ _ f = true /\ Bsign _ _ f = false. Proof. @@ -553,8 +553,8 @@ Proof. econstructor. split. reflexivity. - Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int. - unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int. + Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int Float.to_single. + unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int, Float.to_single. set (re := (@eq_refl Datatypes.comparison Lt)). change (Int.signed (Int.repr 1)) with 1%Z. set (b' := Int64.unsigned b) in *. @@ -565,25 +565,37 @@ Proof. } assert (-16777216 <= 1 <= 16777216)%Z as SILLY by lia. - destruct (BofZ_exact 24 128 re re 1 SILLY) as (C0R & C0F & _). + destruct (BofZ_exact 24 128 re re 1 SILLY) as (C0R & C0F & C0S). clear SILLY. set (one_s := (BofZ 24 128 re re 1)) in *. - - pose proof (BofZ_correct 24 128 re re b') as C1. - cbn in C1. + + pose proof (BofZ_correct 53 1024 re re b') as C0'. + rewrite Rlt_bool_true in C0'; cycle 1. + { apply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + cbn. + gappa. + } + cbn in C0'. + destruct C0' as (C0'R & C0'F & C0'S). + set (b_d := (BofZ 53 1024 re re b')) in *. + + pose proof (Bconv_correct 53 1024 24 128 re re Float.to_single_nan mode_NE b_d C0'F) as C1. + rewrite C0'R in C1. + rewrite C0'S in C1. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. - eapply (Rabs_relax (IZR 18446744073709551616)). - lra. - set (b'' := IZR b') in *. + eapply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + cbn. gappa. } - rewrite (Zlt_bool_false b' 0) in C1 by lia. destruct C1 as (C1R & C1F & C1S). - set (b_s := (BofZ 24 128 re re b')) in *. + set (b_s := (Bconv 53 1024 24 128 re re Float.to_single_nan mode_NE b_d)) in *. assert(1 <= B2R 24 128 b_s <= 18446744073709551616)%R as b_s_RANGE. { rewrite C1R. + cbn. gappa. } assert(B2R 24 128 b_s <> 0)%R as b_s_NONZ by lra. @@ -599,9 +611,12 @@ Proof. gappa. } rewrite C1R in C2. + rewrite C1S in C2. + rewrite C0S in C2. destruct C2 as (C2R & C2F & C2Sz). - rewrite C1S in C2Sz. - change (xorb _ _) with false in C2Sz. + change (1 Date: Fri, 11 Feb 2022 11:12:45 +0100 Subject: remove singleoflongu (does not exist) --- kvx/Builtins1.v | 10 ++++++++- kvx/CBuiltins.ml | 2 ++ kvx/FPDivision64.v | 65 ++++++++++++++++++++++++++++------------------------- kvx/SelectOp.vp | 6 ++++- kvx/SelectOpproof.v | 22 ++++++++++++++++++ 5 files changed, 73 insertions(+), 32 deletions(-) diff --git a/kvx/Builtins1.v b/kvx/Builtins1.v index 163dcbd8..9a04815b 100644 --- a/kvx/Builtins1.v +++ b/kvx/Builtins1.v @@ -28,7 +28,8 @@ Inductive platform_builtin : Type := | BI_fmaf | BI_lround_ne | BI_luround_ne -| BI_fp_udiv32. +| BI_fp_udiv32 +| BI_fp_udiv64. Local Open Scope string_scope. @@ -42,6 +43,7 @@ Definition platform_builtin_table : list (string * platform_builtin) := :: ("__builtin_lround_ne", BI_lround_ne) :: ("__builtin_luround_ne", BI_luround_ne) :: ("__builtin_fp_udiv32", BI_fp_udiv32) + :: ("__builtin_fp_udiv64", BI_fp_udiv64) :: nil. Definition platform_builtin_sig (b: platform_builtin) : signature := @@ -58,6 +60,8 @@ Definition platform_builtin_sig (b: platform_builtin) : signature := mksignature (Tfloat :: nil) Tlong cc_default | BI_fp_udiv32 => mksignature (Tint :: Tint :: nil) Tint cc_default + | BI_fp_udiv64 => + mksignature (Tlong :: Tlong :: nil) Tlong cc_default end. Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (platform_builtin_sig b)) := @@ -74,4 +78,8 @@ Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (pl (fun n1 n2 => if Int.eq n2 Int.zero then None else Some (Int.divu n1 n2)) + | BI_fp_udiv64 => mkbuiltin_n2p Tlong Tlong Tlong + (fun n1 n2 => if Int64.eq n2 Int64.zero + then None + else Some (Int64.divu n1 n2)) end. diff --git a/kvx/CBuiltins.ml b/kvx/CBuiltins.ml index fad86d3c..49534867 100644 --- a/kvx/CBuiltins.ml +++ b/kvx/CBuiltins.ml @@ -139,6 +139,8 @@ let builtins = { (TInt(IULong, []), [TFloat(FDouble, [])], false); "__builtin_fp_udiv32", (TInt(IUInt, []), [TInt(IUInt, []); TInt(IUInt, [])], false); + "__builtin_fp_udiv64", + (TInt(IULong, []), [TInt(IULong, []); TInt(IULong, [])], false); ] } diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index a4d609f0..bbd94321 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -50,7 +50,7 @@ Proof. Qed. Definition approx_inv_longu b := - let invb_s := ExtValues.invfs (Val.maketotal (Val.singleoflongu b)) in + let invb_s := ExtValues.invfs (Val.singleoffloat (Val.maketotal (Val.floatoflongu b))) in let invb_d := Val.floatofsingle invb_s in let b_d := Val.maketotal (Val.floatoflongu b) in let one := Vfloat (ExtFloat.one) in @@ -68,7 +68,8 @@ Qed. Definition approx_inv_thresh := (25/2251799813685248)%R. (* 1.11022302462516e-14 *) - + +(* Theorem approx_inv_longu_correct_abs : forall b, (0 < Int64.unsigned b)%Z -> @@ -302,6 +303,7 @@ Proof. unfold approx_inv_thresh. gappa. Qed. + *) Local Notation "'rd'" := (round radix2 (FLT_exp (-1074) 53) ZnearestE). Local Notation "'rs'" := (round radix2 (FLT_exp (-149) 24) ZnearestE). @@ -320,8 +322,8 @@ Proof. econstructor. split. reflexivity. - Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int. - unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int. + Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int Float.to_single. + unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int, Float.to_single. set (re := (@eq_refl Datatypes.comparison Lt)). change (Int.signed (Int.repr 1)) with 1%Z. set (b' := Int64.unsigned b) in *. @@ -336,20 +338,34 @@ Proof. clear SILLY. set (one_s := (BofZ 24 128 re re 1)) in *. - pose proof (BofZ_correct 24 128 re re b') as C1. - cbn in C1. + pose proof (BofZ_correct 53 1024 re re b') as C5. + rewrite Rlt_bool_true in C5; cycle 1. + { clear C5. + eapply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + cbn. + gappa. + } + cbn in C5. + destruct C5 as (C5R & C5F & C5S). + set (b_d := (BofZ 53 1024 re re b')) in *. + + pose proof (Bconv_correct 53 1024 24 128 re re Float.to_single_nan mode_NE b_d C5F) as C1. rewrite Rlt_bool_true in C1; cycle 1. { clear C1. - eapply (Rabs_relax (IZR 18446744073709551616)). - lra. - set (b'' := IZR b') in *. + apply (Rabs_relax (bpow radix2 64)). + { apply bpow_lt. lia. } + rewrite C5R. + cbn. gappa. } - destruct C1 as (C1R & C1F & _). - set (b_s := (BofZ 24 128 re re b')) in *. - + cbn in C1. + destruct C1 as (C1R & C1F & C1S). + set (b_s := (Bconv 53 1024 24 128 re re Float.to_single_nan mode_NE b_d)) in *. assert(1 <= B2R 24 128 b_s <= 18446744073709551616)%R as b_s_RANGE. { rewrite C1R. + rewrite C5R. + cbn. gappa. } assert(B2R 24 128 b_s <> 0)%R as b_s_NONZ by lra. @@ -413,19 +429,7 @@ Proof. gappa. } destruct C4 as (C4R & C4F & _). - - pose proof (BofZ_correct 53 1024 re re b') as C5. - cbn in C5. - rewrite Rlt_bool_true in C5; cycle 1. - { clear C5. - eapply (Rabs_relax (IZR 18446744073709551616)). - lra. - set (b'' := IZR b') in *. - gappa. - } - destruct C5 as (C5R & C5F & _). - set (b_d := (BofZ 53 1024 re re b')) in *. - + assert(1 <= B2R 53 1024 b_d <= 18446744073709551616)%R as b_d_RANGE. { rewrite C5R. gappa. @@ -477,15 +481,16 @@ Proof. rewrite C3R. rewrite C2R. rewrite C1R. + rewrite C5R. rewrite C0R. cbn. set(b1 := IZR b') in *. replace (rd 1) with 1%R by gappa. - replace (rd (rs (1 / rs b1))) with - ((((rd (rs (1 / rs b1)) - (/b1))/(/b1))+1)*(/ b1))%R ; cycle 1. + replace (rd (rs (1 / rs (rd b1)))) with + ((((rd (rs (1 / rs (rd b1))) - (/b1))/(/b1))+1)*(/ b1))%R ; cycle 1. { field. lra. } - set (er0 := ((rd (rs (1 / rs b1)) - (/b1))/(/b1))%R). + set (er0 := ((rd (rs (1 / rs (rd b1))) - (/b1))/(/b1))%R). replace (rd b1) with ((((rd b1) - b1)/b1 + 1) * b1)%R; cycle 1. { field. lra. } set (er1 := (((rd b1) - b1)/b1)%R). @@ -2351,8 +2356,8 @@ Open Scope cminorsel_scope. Definition e_invfs a := Eop Oinvfs (a ::: Enil). Definition e_float_of_longu a := Eop Ofloatoflongu (a ::: Enil). Definition e_float_of_long a := Eop Ofloatoflong (a ::: Enil). -Definition e_single_of_longu a := Eop Osingleoflongu (a ::: Enil). Definition e_float_of_single a := Eop Ofloatofsingle (a ::: Enil). +Definition e_single_of_float a := Eop Osingleoffloat (a ::: Enil). Definition e_long_of_float_ne a := Eop Olongoffloat_ne (a ::: Enil). Definition e_longu_of_float_ne a := Eop Olonguoffloat_ne (a ::: Enil). Definition e_mulf a b := Eop Omulf (a ::: b ::: Enil). @@ -2378,7 +2383,7 @@ Definition binv_d_var1 := Eletvar (0%nat). Definition e_setup1 a b rest := Elet a (Elet (e_float_of_longu (Eletvar 0%nat)) (Elet (lift (lift b)) (Elet (e_float_of_longu (Eletvar 0%nat)) - (Elet (e_float_of_single (e_invfs (e_single_of_longu (Eletvar 1%nat)))) + (Elet (e_float_of_single (e_invfs (e_single_of_float (Eletvar 0%nat)))) rest)))). Definition e_step1 := e_longu_of_float_ne (e_mulf a_d_var1 binv_d_var1). diff --git a/kvx/SelectOp.vp b/kvx/SelectOp.vp index 70941c48..72a6e4b3 100644 --- a/kvx/SelectOp.vp +++ b/kvx/SelectOp.vp @@ -742,7 +742,7 @@ Nondetfunction gen_fmaf args := | _ => None end. -Require FPDivision32. +Require FPDivision32 FPDivision64. Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr := match b with @@ -758,6 +758,10 @@ Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr | a:::b:::Enil => Some (FPDivision32.fp_divu32 a b) | _ => None end) + | BI_fp_udiv64 => (match args with + | a:::b:::Enil => Some (FPDivision64.fp_divu64 a b) + | _ => None + end) end. End SELECT. diff --git a/kvx/SelectOpproof.v b/kvx/SelectOpproof.v index 658bf0b3..e6cce0fa 100644 --- a/kvx/SelectOpproof.v +++ b/kvx/SelectOpproof.v @@ -1929,6 +1929,28 @@ Proof. } inv EVAL. constructor. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int64.eq in EVAL. + change (Int64.unsigned Int64.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int64.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vlong (Int64.divu i i0)). split. + { + apply FPDivision64.fp_divu64_correct; auto. + pose proof (Int64.unsigned_range i0). + lia. + } + inv EVAL. + constructor. Qed. End CMCONSTR. -- cgit From d9f17c66b52dc49ced37b0a792eb638d7124ffcd Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 14:26:28 +0100 Subject: experiments in division --- test/monniaux/division/compare_timings.c | 67 ++++++++++++++++++++++++++++++++ test/monniaux/division/cycles.h | 1 + 2 files changed, 68 insertions(+) create mode 100644 test/monniaux/division/compare_timings.c create mode 120000 test/monniaux/division/cycles.h diff --git a/test/monniaux/division/compare_timings.c b/test/monniaux/division/compare_timings.c new file mode 100644 index 00000000..15195d3e --- /dev/null +++ b/test/monniaux/division/compare_timings.c @@ -0,0 +1,67 @@ +#include +#include +#include +#include "cycles.h" + +#define CHECKS(mode, quotient) \ +void checks_##mode() { \ + uint64_t checksum=UINT64_C(0), \ + a=UINT64_C(0x10000000000), \ + b=UINT64_C(0x1000); \ + for(int i=0; i<10000; i++) { \ + uint64_t q = (quotient); \ + a += UINT64_C(0x36667); \ + b += UINT64_C(0x13); \ + checksum += q; \ + } \ + printf("checksum = %" PRIx64 "\n", checksum); \ +} + +#define CHECKS2(mode, quotient) \ +void checks2_##mode() { \ + uint64_t checksum=UINT64_C(0), \ + a=UINT64_C(0x10000000000), \ + b=UINT64_C(0x1000); \ + for(int i=0; i<5000; i++) { \ + uint64_t q = (quotient); \ + a += UINT64_C(0x36667); \ + b += UINT64_C(0x13); \ + checksum += q; \ + q = (quotient); \ + a += UINT64_C(0x36667); \ + b += UINT64_C(0x13); \ + checksum += q; \ + } \ + printf("checksum = %" PRIx64 "\n", checksum); \ +} + +CHECKS(normal, a/b) +CHECKS(fp, __builtin_fp_udiv64(a, b)) + +CHECKS2(normal, a/b) +CHECKS2(fp, __builtin_fp_udiv64(a, b)) + +int main() { + cycle_t start, stop; + cycle_count_config(); + + start = get_cycle(); + checks_normal(); + stop = get_cycle(); + printf("normal division: %" PRcycle " cycles\n", stop-start); + + start = get_cycle(); + checks_fp(); + stop = get_cycle(); + printf("fp division: %" PRcycle " cycles\n", stop-start); + + start = get_cycle(); + checks2_normal(); + stop = get_cycle(); + printf("normal division x2: %" PRcycle " cycles\n", stop-start); + + start = get_cycle(); + checks2_fp(); + stop = get_cycle(); + printf("fp division x2: %" PRcycle " cycles\n", stop-start); +} diff --git a/test/monniaux/division/cycles.h b/test/monniaux/division/cycles.h new file mode 120000 index 00000000..84e54d21 --- /dev/null +++ b/test/monniaux/division/cycles.h @@ -0,0 +1 @@ +../cycles.h \ No newline at end of file -- cgit From f6e1cd79c17461d50f5119818e788d1e07451ef6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 17:50:40 +0100 Subject: rm useless stuff --- kvx/FPDivision64.v | 123 ----------------------------------------------------- 1 file changed, 123 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index bbd94321..8cab70b8 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2053,54 +2053,6 @@ Proof. reflexivity. Qed. -Definition anystep_div_longu a b m := - Val.select (Val.cmplu_bool (Mem.valid_pointer m) Cgt b (Vlong (Int64.repr smallb_thresh))) - (step2_div_longu a b) - (twostep_div_longu a b) Tlong. - -Lemma anystep_div_longu_correct : - forall a b m - (b_RANGE : (1 < Int64.unsigned b <= Int64.max_signed)%Z), - anystep_div_longu (Vlong a) (Vlong b) m = Vlong (Int64.repr (Int64.unsigned a / Int64.unsigned b))%Z. -Proof. - intros. - unfold anystep_div_longu. - set (a' := Int64.unsigned a). - set (b' := Int64.unsigned b). - assert (0 < b')%Z as b_NOT0 by lia. - cbn. - unfold Int64.ltu. - fold b'. - rewrite Int64.unsigned_repr; cycle 1. - { unfold smallb_thresh. - change Int64.max_unsigned with 18446744073709551615%Z. - lia. - } - destruct zlt. - { rewrite step2_div_longu_bigb_correct by lia. - reflexivity. - } - rewrite twostep_div_longu_smallb_correct by lia. - cbn. - rewrite Int64.eq_false ; cycle 1. - { intro Z. unfold b' in b_NOT0. rewrite Z in b_NOT0. - rewrite Int64.unsigned_zero in b_NOT0. - lia. - } - reflexivity. -Qed. - - -Definition full_div_longu a b m := - let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in - let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in - let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in - let if_big := Val.longofintu (Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer m) Cge a b)) in - let if_special := Val.select is_big if_big a Tlong in - Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) - if_special - (anystep_div_longu a b m) Tlong. - Lemma one_bigb_div : forall a b (b_RANGE : (9223372036854775808 <= b < 18446744073709551616)%Z) @@ -2117,81 +2069,6 @@ Proof. ring. Qed. -Lemma full_div_longu_correct : - forall a b m - (b_NOT0 : (0 < Int64.unsigned b)%Z), - full_div_longu (Vlong a) (Vlong b) m = Vlong (Int64.repr (Int64.unsigned a / Int64.unsigned b))%Z. -Proof. - - Local Opaque anystep_div_longu. - intros. - unfold full_div_longu. - cbn. - unfold Int64.lt, Int64.ltu. - pose proof (Int64.unsigned_range a). - pose proof (Int64.unsigned_range b). - set (a' := Int64.unsigned a) in *. - set (b' := Int64.unsigned b) in *. - rewrite Int64.signed_zero. - rewrite Int64.unsigned_one. - destruct zlt. - { replace (Val.cmp_bool Cne - (Val.or Vtrue - (if negb (if zlt 1 b' then true else false) then Vtrue else Vfalse)) - (Vint Int.zero)) with (Some true) ; cycle 1. - { destruct zlt; reflexivity. - } - cbn. - destruct zlt; cbn. - { rewrite Zdiv_small by lia. - reflexivity. - } - rewrite one_bigb_div. - reflexivity. - { - change Int64.modulus with 18446744073709551616%Z in *. - split. 2: lia. - unfold b'. - rewrite Int64.unsigned_signed. - unfold Int64.lt. - rewrite Int64.signed_zero. - rewrite zlt_true by lia. - pose proof (Int64.signed_range b). - change Int64.min_signed with (-9223372036854775808)%Z in *. - change Int64.max_signed with (9223372036854775807)%Z in *. - change Int64.modulus with 18446744073709551616%Z in *. - lia. - } - change Int64.modulus with 18446744073709551616%Z in *. - lia. - } - destruct zlt; cbn. - { change (negb (Int.eq (Int.or Int.zero Int.zero) Int.zero)) with false. - cbn. - rewrite anystep_div_longu_correct. - reflexivity. - - change Int64.modulus with 18446744073709551616%Z in *. - split. lia. - rewrite Int64.unsigned_signed. - unfold Int64.lt. - rewrite Int64.signed_zero. - rewrite zlt_false by lia. - pose proof (Int64.signed_range b). - change Int64.min_signed with (-9223372036854775808)%Z in *. - change Int64.max_signed with (9223372036854775807)%Z in *. - change Int64.modulus with 18446744073709551616%Z in *. - lia. - } - change (negb (Int.eq (Int.or Int.zero Int.one) Int.zero)) with true. - cbn. - replace b' with 1%Z by lia. - rewrite Z.div_1_r. - unfold a'. - rewrite Int64.repr_unsigned. - reflexivity. -Qed. - Lemma repr_unsigned_sub2: forall a b, (Int64.repr (a - Int64.unsigned (Int64.repr b))) = Int64.repr (a - b). -- cgit From df6918762b58dcda4913d7f15b61bd64673af567 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 17:56:31 +0100 Subject: moved stuff to appropriate places and removed irrelevant content --- kvx/FPDivision64.v | 258 ---------------------------------------------------- lib/IEEE754_extra.v | 18 ++++ 2 files changed, 18 insertions(+), 258 deletions(-) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 8cab70b8..0d584283 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -30,25 +30,6 @@ Require Import Memory. From Gappa Require Import Gappa_tactic. - -Lemma Znearest_lub : - forall choice (n : Z) (x : R), (IZR n <= x)%R -> (n <= Znearest choice x)%Z. -Proof. - intros until x. intro BND. - pose proof (Zfloor_lub n x BND). - pose proof (Znearest_ge_floor choice x). - lia. -Qed. - -Lemma Znearest_glb : - forall choice (n : Z) (x : R), (x <= IZR n)%R -> (Znearest choice x <= n)%Z. -Proof. - intros until x. intro BND. - pose proof (Zceil_glb n x BND). - pose proof (Znearest_le_ceil choice x). - lia. -Qed. - Definition approx_inv_longu b := let invb_s := ExtValues.invfs (Val.singleoffloat (Val.maketotal (Val.floatoflongu b))) in let invb_d := Val.floatofsingle invb_s in @@ -66,245 +47,6 @@ Proof. lra. Qed. -Definition approx_inv_thresh := (25/2251799813685248)%R. -(* 1.11022302462516e-14 *) - -(* -Theorem approx_inv_longu_correct_abs : - forall b, - (0 < Int64.unsigned b)%Z -> - exists (f : float), - (approx_inv_longu (Vlong b)) = Vfloat f /\ - is_finite _ _ f = true /\ (Rabs((B2R _ _ f) - (1 / IZR (Int64.unsigned b))) <= approx_inv_thresh)%R. -Proof. - intros b NONZ. - unfold approx_inv_longu. - cbn. - econstructor. - split. - reflexivity. - Local Transparent Float.neg Float.of_single Float32.of_longu Float32.div Float.of_longu Float32.of_int Float.of_int. - unfold Float.fma, Float.neg, Float.of_single, Float32.of_longu, ExtFloat32.inv, Float32.div, Float.of_longu, ExtFloat32.one, Float32.of_int, ExtFloat.one, Float.of_int. - set (re := (@eq_refl Datatypes.comparison Lt)). - change (Int.signed (Int.repr 1)) with 1%Z. - set (b' := Int64.unsigned b) in *. - pose proof (Int64.unsigned_range b) as RANGE. - change Int64.modulus with 18446744073709551616%Z in RANGE. - assert(1 <= IZR b' <= 18446744073709551616)%R as RANGE'. - { split; apply IZR_le; lia. - } - - assert (-16777216 <= 1 <= 16777216)%Z as SILLY by lia. - destruct (BofZ_exact 24 128 re re 1 SILLY) as (C0R & C0F & _). - clear SILLY. - set (one_s := (BofZ 24 128 re re 1)) in *. - - pose proof (BofZ_correct 24 128 re re b') as C1. - cbn in C1. - rewrite Rlt_bool_true in C1; cycle 1. - { clear C1. - eapply (Rabs_relax (IZR 18446744073709551616)). - lra. - set (b'' := IZR b') in *. - gappa. - } - destruct C1 as (C1R & C1F & _). - set (b_s := (BofZ 24 128 re re b')) in *. - - assert(1 <= B2R 24 128 b_s <= 18446744073709551616)%R as b_s_RANGE. - { rewrite C1R. - gappa. - } - assert(B2R 24 128 b_s <> 0)%R as b_s_NONZ by lra. - - pose proof (Bdiv_correct 24 128 re re Float32.binop_nan mode_NE one_s b_s b_s_NONZ) as C2. - rewrite Rlt_bool_true in C2; cycle 1. - { clear C2. - apply Rabs_relax with (b := 1%R). - { cbn; lra. } - rewrite C0R. - set (r_b_s := B2R 24 128 b_s) in *. - cbn. - gappa. - } - - destruct C2 as (C2R & C2F & _). - set (invb_s := (Bdiv 24 128 re re Float32.binop_nan mode_NE one_s b_s)) in *. - rewrite C0F in C2F. - - assert ((1/18446744073709551616 <= B2R 24 128 invb_s <= 1)%R) as invb_s_RANGE. - { rewrite C2R. - set (r_b_s := B2R 24 128 b_s) in *. - rewrite C0R. - cbn. - gappa. - } - - pose proof (Bconv_correct 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s C2F) as C3. - rewrite Rlt_bool_true in C3; cycle 1. - { clear C3. - set (r_invb_s := (B2R 24 128 invb_s)) in *. - apply Rabs_relax with (b := 1%R). - { replace 1%R with (bpow radix2 0)%R by reflexivity. - apply bpow_lt. - lia. - } - cbn. - gappa. - } - - destruct C3 as (C3R & C3F & _). - set (invb_d := (Bconv 24 128 53 1024 re re Float.of_single_nan mode_NE invb_s)) in *. - assert ((1/18446744073709551616 <= B2R 53 1024 invb_d <= 1)%R) as invb_d_RANGE. - { - rewrite C3R. - set (r_invb_s := B2R 24 128 invb_s) in *. - cbn. - gappa. - } - - pose proof (is_finite_Bopp 53 1024 Float.neg_nan invb_d) as opp_finite. - rewrite C3F in opp_finite. - - pose proof (BofZ_correct 53 1024 re re 1) as C4. - rewrite Rlt_bool_true in C4; cycle 1. - { clear C4. - cbn. - eapply (Rabs_relax (IZR 18446744073709551616)). - lra. - set (b'' := IZR b') in *. - gappa. - } - destruct C4 as (C4R & C4F & _). - - pose proof (BofZ_correct 53 1024 re re b') as C5. - cbn in C5. - rewrite Rlt_bool_true in C5; cycle 1. - { clear C5. - eapply (Rabs_relax (IZR 18446744073709551616)). - lra. - set (b'' := IZR b') in *. - gappa. - } - destruct C5 as (C5R & C5F & _). - set (b_d := (BofZ 53 1024 re re b')) in *. - - assert(1 <= B2R 53 1024 b_d <= 18446744073709551616)%R as b_d_RANGE. - { rewrite C5R. - gappa. - } - - pose proof (Bfma_correct 53 1024 re re Float.fma_nan mode_NE - (Bopp 53 1024 Float.neg_nan invb_d) (BofZ 53 1024 re re b') - (BofZ 53 1024 re re 1) opp_finite C5F C4F) as C6. - rewrite Rlt_bool_true in C6; cycle 1. - { clear C6. - rewrite C4R. - rewrite B2R_Bopp. - cbn. - eapply (Rabs_relax (IZR 18446744073709551616)). - { lra. } - fold invb_d. - fold b_d. - set (r_invb_d := B2R 53 1024 invb_d) in *. - set (r_b_d := B2R 53 1024 b_d) in *. - gappa. - } - fold b_d in C6. - destruct C6 as (C6R & C6F & _). - - pose proof (Bfma_correct 53 1024 re re Float.fma_nan mode_NE - (Bfma 53 1024 re re Float.fma_nan mode_NE - (Bopp 53 1024 Float.neg_nan invb_d) b_d (BofZ 53 1024 re re 1)) - invb_d invb_d C6F C3F C3F) as C7. - rewrite Rlt_bool_true in C7; cycle 1. - { clear C7. - rewrite C6R. - rewrite B2R_Bopp. - eapply (Rabs_relax (bpow radix2 64)). - { apply bpow_lt. lia. } - rewrite C4R. - cbn. - set (r_invb_d := B2R 53 1024 invb_d) in *. - set (r_b_d := B2R 53 1024 b_d) in *. - gappa. - } - destruct C7 as (C7R & C7F & _). - - split. assumption. - rewrite C7R. - rewrite C6R. - rewrite C5R. - rewrite C4R. - rewrite B2R_Bopp. - rewrite C3R. - rewrite C2R. - rewrite C1R. - rewrite C0R. - cbn. - set(b1 := IZR b') in *. - replace (round radix2 (FLT_exp (-1074) 53) ZnearestE 1) with 1%R by gappa. - set (bd := round radix2 (FLT_exp (-1074) 53) ZnearestE b1). - set (x0 := round radix2 (FLT_exp (-1074) 53) ZnearestE - (round radix2 (FLT_exp (-149) 24) ZnearestE - (1 / round radix2 (FLT_exp (-149) 24) ZnearestE b1))). - set (alpha0 := (- x0 * bd + 1)%R). - set (y1 := (round radix2 (FLT_exp (-1074) 53) ZnearestE alpha0 * x0 + x0)%R). - set (x1 := round radix2 (FLT_exp (-1074) 53) ZnearestE y1). - replace (x1 - 1/b1)%R with ((y1-1/b1)+(x1-y1))%R by ring. - - assert(alpha0 = -((x0-1/bd)/(1/bd)))%R as alpha0_EQ. - { unfold alpha0. - field. - unfold bd. - gappa. - } - assert(y1-1/b1 = ((round radix2 (FLT_exp (-1074) 53) ZnearestE alpha0) - - alpha0) * x0 - + alpha0*(x0-1/b1) - ((bd-b1)/b1) * x0)%R as y1_EQ. - { unfold y1, alpha0. - field. - lra. - } - assert(Rabs alpha0 <= 257/2147483648)%R as alpha0_ABS. - { rewrite alpha0_EQ. - unfold x0, bd. - gappa. - } - assert (Rabs (x0 - 1 / b1) <= 3/33554432)%R as x0_delta_ABS. - { unfold x0. - gappa. - } - set (x0_delta := (x0 - 1 / b1)%R) in *. - assert (Rabs ((bd - b1) / b1) <= 1/9007199254740992)%R as bd_delta_ABS. - { unfold bd. - gappa. - } - set (bd_delta := ((bd - b1) / b1)%R) in *. - set (rnd_alpha0_delta := (round radix2 (FLT_exp (-1074) 53) ZnearestE alpha0 - alpha0)%R) in *. - assert (Rabs rnd_alpha0_delta <= 1/75557863725914323419136)%R as rnd_alpha0_delta_ABS. - { unfold rnd_alpha0_delta. - gappa. - } - assert (1/18446744073709551616 <= x0 <= 1)%R as x0_RANGE. - { unfold x0. - gappa. - } - assert (Rabs (y1 - 1 / b1) <= 49/4503599627370496)%R as y1_delta_ABS. - { rewrite y1_EQ. - gappa. - } - assert (Rabs(x1 - y1) <= 1/9007199254740992)%R as x1_delta_ABS. - { unfold x1. - gappa. - } - set (y1_delta := (y1 - 1 / b1)%R) in *. - set (x1_delta := (x1-y1)%R) in *. - unfold approx_inv_thresh. - gappa. -Qed. - *) - Local Notation "'rd'" := (round radix2 (FLT_exp (-1074) 53) ZnearestE). Local Notation "'rs'" := (round radix2 (FLT_exp (-149) 24) ZnearestE). diff --git a/lib/IEEE754_extra.v b/lib/IEEE754_extra.v index 636ea1ff..35feb29d 100644 --- a/lib/IEEE754_extra.v +++ b/lib/IEEE754_extra.v @@ -29,6 +29,24 @@ Require Import Coq.Logic.FunctionalExtensionality. Local Open Scope Z_scope. +Lemma Znearest_lub : + forall choice (n : Z) (x : R), (IZR n <= x)%R -> (n <= Znearest choice x)%Z. +Proof. + intros until x. intro BND. + pose proof (Zfloor_lub n x BND). + pose proof (Znearest_ge_floor choice x). + lia. +Qed. + +Lemma Znearest_glb : + forall choice (n : Z) (x : R), (x <= IZR n)%R -> (Znearest choice x <= n)%Z. +Proof. + intros until x. intro BND. + pose proof (Zceil_glb n x BND). + pose proof (Znearest_le_ceil choice x). + lia. +Qed. + Lemma Znearest_IZR : forall choice n, (Znearest choice (IZR n)) = n. Proof. -- cgit From 884cbaab11e294c9c5224b2fc98e68074d6a1584 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 20:42:32 +0100 Subject: twostep_mod_longu_mostb_correct --- kvx/FPDivision64.v | 94 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 94 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 0d584283..73dba8eb 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -1971,6 +1971,100 @@ Proof. reflexivity. Qed. +Definition step2_mod_long a b := + let q := step2_div_long' a b in + let r := Val.subl a (Val.mull q b) in + Val.select (Val.cmpl_bool Clt r (Vlong Int64.zero)) + (Val.addl r b) r Tlong. + +Definition twostep_mod_longu a b := + let q1 := step1_div_longu a b in + step2_mod_long (Val.subl a (Val.mull b q1)) b. + +Lemma vlong_eq: forall a b, (Vlong a) = (Vlong b) -> a = b. +Proof. + intros a b EQ. + congruence. +Qed. + +Lemma move_repr_in_mod : + forall a b c, + Int64.repr (a - b * c)%Z = + Int64.repr (a - b * Int64.unsigned (Int64.repr c))%Z. +Proof. + intros. + apply Int64.eqm_samerepr. + auto 10 with ints. +Qed. + +Lemma twostep_mod_longu_mostb_correct : + forall a b + (b_RANGE : (1 < Int64.unsigned b <= Int64.max_signed)%Z), + (twostep_mod_longu (Vlong a) (Vlong b)) = + (Val.maketotal (Val.modlu (Vlong a) (Vlong b))). +Proof. + intros. + Local Transparent twostep_div_longu. + pose proof (twostep_div_longu_mostb_correct a b b_RANGE) as div_correct. + unfold twostep_div_longu, twostep_mod_longu, step2_div_long, step2_mod_long in *. + set (q1 := (step1_div_longu (Vlong a) (Vlong b))) in *. + set (q2 := (step2_div_long' (Val.subl (Vlong a) (Val.mull (Vlong b) q1)) (Vlong b))) in *. + cbn in div_correct. + cbn. + unfold Int64.eq in *. + change (Int64.unsigned Int64.zero) with 0%Z in *. + rewrite zeq_false by lia. + rewrite zeq_false in div_correct by lia. + cbn in div_correct. + cbn. + destruct q1 as [ | | q1l | | | ] ; cbn in *; try discriminate. + destruct q2 as [ | | q2l | | | ] ; cbn in *; try discriminate. + rewrite <- (function_ite Vlong). + rewrite <- (function_ite Vlong) in div_correct. + cbn. cbn in div_correct. + unfold Int64.lt, Int64.sub, Int64.mul, Int64.add, Int64.divu, Int64.modu in *. + set (a' := Int64.unsigned a) in *. + set (b' := Int64.unsigned b) in *. + set (q1' := Int64.unsigned q1l) in *. + set (q2' := Int64.unsigned q2l) in *. + change (Int64.signed Int64.zero) with 0%Z in *. + replace (Int64.repr + (Int64.unsigned + (Int64.repr (a' - Int64.unsigned (Int64.repr (b' * q1')))) - + Int64.unsigned (Int64.repr (q2' * b')))) + with (Int64.repr (a' - (b' * q1') - (q2' * b')))%Z in * ; cycle 1. + { + apply Int64.eqm_samerepr. + auto 10 with ints. + } + replace (a' - b' * q1' - q2' * b')%Z with (a' - b' * (q1' + q2'))%Z in * by ring. + f_equal. + apply vlong_eq in div_correct. + rewrite Z.mod_eq by lia. + rewrite (move_repr_in_mod a' b' (a' / b'))%Z. + rewrite <- div_correct. + clear div_correct. + rewrite <- (move_repr_in_mod a' b')%Z. + + destruct zlt as [NEG | POS]. + 2: reflexivity. + rewrite repr_unsigned_add. + replace (a' - b' * (q1' + q2') + b')%Z with (a' - b' * (q1' + q2' - 1))%Z by ring. + apply Int64.eqm_samerepr. + assert (Int64.eqm (Int64.unsigned (Int64.repr (q2' + Int64.unsigned Int64.mone))) + (q2' -1))%Z as EQM. + { apply Int64.eqm_trans with (y := (q2' + Int64.unsigned Int64.mone)%Z). + apply Int64.eqm_sym. + apply Int64.eqm_unsigned_repr. + apply Int64.eqm_add. + apply Int64.eqm_refl. + exists (1)%Z. + reflexivity. + } + replace (q1' + q2' - 1)%Z with (q1' + (q2' - 1))%Z by ring. + auto with ints. +Qed. + Open Scope cminorsel_scope. Definition e_invfs a := Eop Oinvfs (a ::: Enil). Definition e_float_of_longu a := Eop Ofloatoflongu (a ::: Enil). -- cgit From 87bbbd23d6681bfdc29ac5f8bd24b28ba89da4aa Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 21:01:17 +0100 Subject: full2 mod correct --- kvx/FPDivision64.v | 98 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 98 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 73dba8eb..56170419 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2065,6 +2065,104 @@ Proof. auto with ints. Qed. +Definition full2_mod_longu a b m := + let is_big := Val.cmpl_bool Clt b (Vlong Int64.zero) in + let is_one := Val.cmplu_bool (Mem.valid_pointer m) Cle b (Vlong Int64.one) in + let is_special := Val.or (Val.of_optbool is_big) (Val.of_optbool is_one) in + let if_big := Val.select (Val.cmplu_bool (Mem.valid_pointer m) Cge a b) (Val.subl a b) a Tlong in + let if_special := Val.select is_big if_big (Vlong Int64.zero) Tlong in + Val.select (Val.cmp_bool Cne is_special (Vint Int.zero)) + if_special + (twostep_mod_longu a b) Tlong. + +Lemma full2_mod_longu_correct : + forall a b m + (b_NOT0 : (0 < Int64.unsigned b)%Z), + full2_mod_longu (Vlong a) (Vlong b) m = Vlong (Int64.repr ((Int64.unsigned a) mod (Int64.unsigned b)))%Z. +Proof. + + Local Opaque twostep_mod_longu. + intros. + unfold full2_mod_longu. + cbn. + unfold Int64.lt, Int64.ltu. + pose proof (Int64.unsigned_range a). + pose proof (Int64.unsigned_range b). + set (a' := Int64.unsigned a) in *. + set (b' := Int64.unsigned b) in *. + rewrite Int64.signed_zero. + rewrite Int64.unsigned_one. + destruct zlt. + { replace (Val.cmp_bool Cne + (Val.or Vtrue + (if negb (if zlt 1 b' then true else false) then Vtrue else Vfalse)) + (Vint Int.zero)) with (Some true) ; cycle 1. + { destruct zlt; reflexivity. + } + cbn. + rewrite Z.mod_eq by lia. + + destruct zlt; cbn. + { rewrite Zdiv_small by lia. + replace (a' - b' * 0)%Z with a' by ring. + unfold a'. + rewrite Int64.repr_unsigned. + reflexivity. + } + rewrite one_bigb_div. + { unfold Int64.sub. + fold a' b'. + repeat f_equal. ring. + } + { + change Int64.modulus with 18446744073709551616%Z in *. + split. 2: lia. + unfold b'. + rewrite Int64.unsigned_signed. + unfold Int64.lt. + rewrite Int64.signed_zero. + rewrite zlt_true by lia. + pose proof (Int64.signed_range b). + change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + destruct zlt; cbn. + { change (negb (Int.eq (Int.or Int.zero Int.zero) Int.zero)) with false. + cbn. + rewrite twostep_mod_longu_mostb_correct. + { + cbn. + unfold Int64.eq. + fold b'. + rewrite Int64.unsigned_zero. + rewrite zeq_false by lia. + reflexivity. + } + + change Int64.modulus with 18446744073709551616%Z in *. + split. lia. + rewrite Int64.unsigned_signed. + unfold Int64.lt. + rewrite Int64.signed_zero. + rewrite zlt_false by lia. + pose proof (Int64.signed_range b). + change Int64.min_signed with (-9223372036854775808)%Z in *. + change Int64.max_signed with (9223372036854775807)%Z in *. + change Int64.modulus with 18446744073709551616%Z in *. + lia. + } + change (negb (Int.eq (Int.or Int.zero Int.one) Int.zero)) with true. + cbn. + replace b' with 1%Z by lia. + rewrite Z.mod_1_r. + reflexivity. +Qed. + Open Scope cminorsel_scope. Definition e_invfs a := Eop Oinvfs (a ::: Enil). Definition e_float_of_longu a := Eop Ofloatoflongu (a ::: Enil). -- cgit From f651db91393a3df1b3c65e4249423a5c744761e5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 21:19:55 +0100 Subject: fp_modu64_correct --- kvx/FPDivision64.v | 30 ++++++++++++++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 56170419..1d0fd2ee 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2376,3 +2376,33 @@ Proof. rewrite <- full2_div_longu_correct with (m := memenv) by lia. apply e_step7_correct; assumption. Qed. + +Definition fp_modu64 a b := Elet a (Elet (lift b) (e_msubl (Eletvar 1%nat) (Eletvar 0%nat) + (fp_divu64 (Eletvar 1%nat) (Eletvar 0%nat)))). + +Theorem fp_modu64_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int64) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)) + (b_nz : (Int64.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_modu64 expr_a expr_b) + (Vlong (Int64.modu a b)). +Proof. + intros. + rewrite Int64.modu_divu; cycle 1. + { intro Z. + subst. + rewrite Int64.unsigned_zero in b_nz. + lia. + } + unfold fp_modu64. + Local Opaque fp_divu64. + repeat (econstructor + apply eval_lift + eassumption). + { apply fp_divu64_correct; + repeat (econstructor + apply eval_lift + eassumption). + } + cbn. + rewrite Int64.mul_commut. + reflexivity. +Qed. -- cgit From bc33b3536cc9985a6dddacadfaa9c2f5a111f4ac Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 21:25:09 +0100 Subject: fp_modu32 --- kvx/FPDivision32.v | 32 +++++++++++++++++++++++++++++++- 1 file changed, 31 insertions(+), 1 deletion(-) diff --git a/kvx/FPDivision32.v b/kvx/FPDivision32.v index 5cda1077..c920d22a 100644 --- a/kvx/FPDivision32.v +++ b/kvx/FPDivision32.v @@ -585,4 +585,34 @@ Proof. rewrite Int64.unsigned_repr by (cbn; lia). reflexivity. Qed. - + +Definition e_msubl a b c := Eop Omsub (a ::: b ::: c ::: Enil). +Definition fp_modu32 a b := Elet a (Elet (lift b) (e_msubl (Eletvar 1%nat) (Eletvar 0%nat) + (fp_divu32 (Eletvar 1%nat) (Eletvar 0%nat)))). + +Theorem fp_modu32_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : (Int.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_modu32 expr_a expr_b) + (Vint (Int.modu a b)). +Proof. + intros. + rewrite Int.modu_divu; cycle 1. + { intro Z. + subst. + rewrite Int.unsigned_zero in b_nz. + lia. + } + unfold fp_modu32. + Local Opaque fp_divu32. + repeat (econstructor + apply eval_lift + eassumption). + { apply fp_divu32_correct; + repeat (econstructor + apply eval_lift + eassumption). + } + cbn. + rewrite Int.mul_commut. + reflexivity. +Qed. -- cgit From 16715e5efd6ce899eb3d544fd71751f367eaa370 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 11 Feb 2022 21:36:57 +0100 Subject: modulos --- kvx/Builtins1.v | 18 +++++++++++++++++- kvx/CBuiltins.ml | 4 ++++ kvx/SelectOp.vp | 8 ++++++++ kvx/SelectOpproof.v | 44 ++++++++++++++++++++++++++++++++++++++++++++ 4 files changed, 73 insertions(+), 1 deletion(-) diff --git a/kvx/Builtins1.v b/kvx/Builtins1.v index 9a04815b..a0473d07 100644 --- a/kvx/Builtins1.v +++ b/kvx/Builtins1.v @@ -29,7 +29,9 @@ Inductive platform_builtin : Type := | BI_lround_ne | BI_luround_ne | BI_fp_udiv32 -| BI_fp_udiv64. +| BI_fp_udiv64 +| BI_fp_umod32 +| BI_fp_umod64. Local Open Scope string_scope. @@ -44,6 +46,8 @@ Definition platform_builtin_table : list (string * platform_builtin) := :: ("__builtin_luround_ne", BI_luround_ne) :: ("__builtin_fp_udiv32", BI_fp_udiv32) :: ("__builtin_fp_udiv64", BI_fp_udiv64) + :: ("__builtin_fp_umod32", BI_fp_umod32) + :: ("__builtin_fp_umod64", BI_fp_umod64) :: nil. Definition platform_builtin_sig (b: platform_builtin) : signature := @@ -62,6 +66,10 @@ Definition platform_builtin_sig (b: platform_builtin) : signature := mksignature (Tint :: Tint :: nil) Tint cc_default | BI_fp_udiv64 => mksignature (Tlong :: Tlong :: nil) Tlong cc_default + | BI_fp_umod32 => + mksignature (Tint :: Tint :: nil) Tint cc_default + | BI_fp_umod64 => + mksignature (Tlong :: Tlong :: nil) Tlong cc_default end. Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (platform_builtin_sig b)) := @@ -82,4 +90,12 @@ Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (pl (fun n1 n2 => if Int64.eq n2 Int64.zero then None else Some (Int64.divu n1 n2)) + | BI_fp_umod32 => mkbuiltin_n2p Tint Tint Tint + (fun n1 n2 => if Int.eq n2 Int.zero + then None + else Some (Int.modu n1 n2)) + | BI_fp_umod64 => mkbuiltin_n2p Tlong Tlong Tlong + (fun n1 n2 => if Int64.eq n2 Int64.zero + then None + else Some (Int64.modu n1 n2)) end. diff --git a/kvx/CBuiltins.ml b/kvx/CBuiltins.ml index 49534867..5ca5ddba 100644 --- a/kvx/CBuiltins.ml +++ b/kvx/CBuiltins.ml @@ -141,6 +141,10 @@ let builtins = { (TInt(IUInt, []), [TInt(IUInt, []); TInt(IUInt, [])], false); "__builtin_fp_udiv64", (TInt(IULong, []), [TInt(IULong, []); TInt(IULong, [])], false); + "__builtin_fp_umod32", + (TInt(IUInt, []), [TInt(IUInt, []); TInt(IUInt, [])], false); + "__builtin_fp_umod64", + (TInt(IULong, []), [TInt(IULong, []); TInt(IULong, [])], false); ] } diff --git a/kvx/SelectOp.vp b/kvx/SelectOp.vp index 72a6e4b3..089afe1d 100644 --- a/kvx/SelectOp.vp +++ b/kvx/SelectOp.vp @@ -762,6 +762,14 @@ Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr | a:::b:::Enil => Some (FPDivision64.fp_divu64 a b) | _ => None end) + | BI_fp_umod32 => (match args with + | a:::b:::Enil => Some (FPDivision32.fp_modu32 a b) + | _ => None + end) + | BI_fp_umod64 => (match args with + | a:::b:::Enil => Some (FPDivision64.fp_modu64 a b) + | _ => None + end) end. End SELECT. diff --git a/kvx/SelectOpproof.v b/kvx/SelectOpproof.v index e6cce0fa..cffa2583 100644 --- a/kvx/SelectOpproof.v +++ b/kvx/SelectOpproof.v @@ -1951,6 +1951,50 @@ Proof. } inv EVAL. constructor. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int.eq in EVAL. + change (Int.unsigned Int.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vint (Int.modu i i0)). split. + { + apply FPDivision32.fp_modu32_correct; auto. + pose proof (Int.unsigned_range i0). + lia. + } + inv EVAL. + constructor. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int64.eq in EVAL. + change (Int64.unsigned Int64.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int64.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vlong (Int64.modu i i0)). split. + { + apply FPDivision64.fp_modu64_correct; auto. + pose proof (Int64.unsigned_range i0). + lia. + } + inv EVAL. + constructor. Qed. End CMCONSTR. -- cgit From 98a115dece106f5036452b0c0bac04ad4a6e047e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 12 Feb 2022 13:15:16 +0100 Subject: fix bad reservation table for finvw --- kvx/PostpassSchedulingOracle.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/kvx/PostpassSchedulingOracle.ml b/kvx/PostpassSchedulingOracle.ml index f53da14b..e752624c 100644 --- a/kvx/PostpassSchedulingOracle.ml +++ b/kvx/PostpassSchedulingOracle.ml @@ -657,8 +657,8 @@ let rec_to_usage r = | Get -> bcu_tiny_tiny_mau_xnop | Fnegd | Fnegw | Fabsd | Fabsw | Fwidenlwd | Fmind | Fmaxd | Fminw | Fmaxw -> alu_lite - | Fnarrowdw -> alu_full - | Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw | Finvw + | Finvw | Fnarrowdw -> alu_full + | Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw | Ffmad | Ffmaw | Ffmsd | Ffmsw -> mau -- cgit From 14d1a795f3ab054374c4fbe8f25c2a388047f2bc Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 14 Feb 2022 20:40:21 +0100 Subject: fp_divs32_correct --- kvx/FPDivision32.v | 137 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 137 insertions(+) diff --git a/kvx/FPDivision32.v b/kvx/FPDivision32.v index c920d22a..edbcbe12 100644 --- a/kvx/FPDivision32.v +++ b/kvx/FPDivision32.v @@ -616,3 +616,140 @@ Proof. rewrite Int.mul_commut. reflexivity. Qed. + +Definition e_is_neg a := Eop (Ocmp (Ccompimm Clt Int.zero)) (a ::: Enil). +Definition e_xorw a b := Eop Oxor (a ::: b ::: Enil). +Definition e_ite ty c vc v1 v2 := Eop (Osel c ty) (v1 ::: v2 ::: vc ::: Enil). +Definition e_neg a := Eop Oneg (a ::: Enil). +Definition e_abs a := Eop (Oabsdiffimm Int.zero) (a ::: Enil). + +Definition fp_divs32 a b := + Elet a (Elet (lift b) + (Elet (fp_divu32 (e_abs (Eletvar (1%nat))) (e_abs (Eletvar (0%nat)))) + (e_ite Tint (Ccompu0 Cne) (e_xorw (e_is_neg (Eletvar 2%nat)) + (e_is_neg (Eletvar 1%nat))) + (e_neg (Eletvar 0%nat)) (Eletvar 0%nat)))). + +Lemma nonneg_signed_unsigned: + forall x (x_NONNEG : Int.signed x >= 0), + (Int.signed x) = (Int.unsigned x). +Proof. + intros. + pose proof (Int.unsigned_range x). + unfold Int.signed in *. + destruct zlt. reflexivity. + change Int.modulus with 4294967296%Z in *. + change Int.half_modulus with 2147483648%Z in *. + lia. +Qed. + +Lemma int_min_signed_unsigned : + (- Int.min_signed < Int.max_unsigned)%Z. +Proof. + reflexivity. +Qed. + +Lemma int_divs_divu : + forall a b + (b_NOT0 : Int.signed b <> 0), + Int.divs a b = if xorb (Int.lt a Int.zero) + (Int.lt b Int.zero) + then Int.neg (Int.divu (ExtValues.int_abs a) + (ExtValues.int_abs b)) + else Int.divu (ExtValues.int_abs a) (ExtValues.int_abs b). +Proof. + intros. + unfold Int.divs, Int.divu, Int.lt, ExtValues.int_abs. + pose proof (Int.signed_range a) as a_RANGE. + pose proof (Int.signed_range b) as b_RANGE. + change (Int.signed Int.zero) with 0%Z. + destruct zlt. + - cbn. rewrite (Z.abs_neq (Int.signed a)) by lia. + rewrite (Int.unsigned_repr (- Int.signed a)); cycle 1. + { pose proof int_min_signed_unsigned. lia. } + + destruct zlt. + + rewrite (Z.abs_neq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof int_min_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int.signed b)) at 1. + rewrite Z.quot_opp_r by lia. + rewrite <- (Z.opp_involutive (Int.signed a)) at 1. + rewrite Z.quot_opp_l by lia. + rewrite Z.quot_div_nonneg by lia. + rewrite Z.opp_involutive. + reflexivity. + + + rewrite (Z.abs_eq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof Int.max_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int.signed a)) at 1. + rewrite Z.quot_opp_l by lia. + rewrite Z.quot_div_nonneg by lia. + rewrite Int.neg_repr. + reflexivity. + + - cbn. rewrite (Z.abs_eq (Int.signed a)) by lia. + rewrite (Int.unsigned_repr (Int.signed a)); cycle 1. + { pose proof Int.max_signed_unsigned. lia. } + destruct zlt. + + rewrite (Z.abs_neq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof int_min_signed_unsigned. lia. } + rewrite Int.neg_repr. + rewrite <- (Z.opp_involutive (Int.signed b)) at 1. + rewrite Z.quot_opp_r by lia. + rewrite Z.quot_div_nonneg by lia. + reflexivity. + + + rewrite (Z.abs_eq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof Int.max_signed_unsigned. lia. } + rewrite Z.quot_div_nonneg by lia. + reflexivity. +Qed. + +Lemma nonzero_unsigned_signed : + forall b, Int.unsigned b > 0 -> Int.signed b <> 0. +Proof. + intros b GT EQ. + rewrite Int.unsigned_signed in GT. + unfold Int.lt in GT. + rewrite Int.signed_zero in GT. + destruct zlt in GT; lia. +Qed. + +Theorem fp_divs32_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : (Int.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_divs32 expr_a expr_b) + (Vint (Int.divs a b)). +Proof. + intros. + unfold fp_divs32. + Local Opaque fp_divu32. + repeat (econstructor + apply eval_lift + eassumption). + apply fp_divu32_correct. + all: repeat (econstructor + apply eval_lift + eassumption). + { unfold ExtValues.int_absdiff, ExtValues.Z_abs_diff. + rewrite Int.signed_zero. rewrite Z.sub_0_r. + rewrite Int.unsigned_repr. + { pose proof (nonzero_unsigned_signed b b_nz). + lia. + } + pose proof Int.max_signed_unsigned. + pose proof int_min_signed_unsigned. + pose proof (Int.signed_range b). + lia. + } + cbn. + rewrite int_divs_divu ; cycle 1. + { apply nonzero_unsigned_signed. assumption. } + unfold Int.lt, ExtValues.int_abs, ExtValues.int_absdiff, ExtValues.Z_abs_diff. + change (Int.signed Int.zero) with 0%Z. + repeat rewrite Z.sub_0_r. + destruct zlt; destruct zlt; reflexivity. +Qed. -- cgit From 3abc81d25a5216037e51eca5a820d6d9fa4649d8 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 15 Feb 2022 19:39:36 +0100 Subject: int_mods_modu --- kvx/FPDivision32.v | 58 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 58 insertions(+) diff --git a/kvx/FPDivision32.v b/kvx/FPDivision32.v index edbcbe12..4e9b2fb2 100644 --- a/kvx/FPDivision32.v +++ b/kvx/FPDivision32.v @@ -753,3 +753,61 @@ Proof. repeat rewrite Z.sub_0_r. destruct zlt; destruct zlt; reflexivity. Qed. + +Lemma int_mods_modu : + forall a b + (b_NOT0 : Int.signed b <> 0), + Int.mods a b = if Int.lt a Int.zero + then Int.neg (Int.modu (ExtValues.int_abs a) + (ExtValues.int_abs b)) + else Int.modu (ExtValues.int_abs a) (ExtValues.int_abs b). +Proof. + intros. + unfold Int.mods, Int.modu, Int.lt, ExtValues.int_abs. + pose proof (Int.signed_range a) as a_RANGE. + pose proof (Int.signed_range b) as b_RANGE. + change (Int.signed Int.zero) with 0%Z. + destruct zlt. + - cbn. rewrite (Z.abs_neq (Int.signed a)) by lia. + rewrite (Int.unsigned_repr (- Int.signed a)); cycle 1. + { pose proof int_min_signed_unsigned. lia. } + + destruct (zlt (Int.signed b) 0%Z). + + rewrite (Z.abs_neq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof int_min_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int.signed b)) at 1. + rewrite Z.rem_opp_r by lia. + rewrite <- (Z.opp_involutive (Int.signed a)) at 1. + rewrite Z.rem_opp_l by lia. + rewrite Z.rem_mod_nonneg by lia. + rewrite Int.neg_repr. + reflexivity. + + + rewrite (Z.abs_eq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof Int.max_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int.signed a)) at 1. + rewrite Z.rem_opp_l by lia. + rewrite Z.rem_mod_nonneg by lia. + rewrite Int.neg_repr. + reflexivity. + + - cbn. rewrite (Z.abs_eq (Int.signed a)) by lia. + rewrite (Int.unsigned_repr (Int.signed a)); cycle 1. + { pose proof Int.max_signed_unsigned. lia. } + destruct (zlt (Int.signed b) 0%Z). + + rewrite (Z.abs_neq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof int_min_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int.signed b)) at 1. + rewrite Z.rem_opp_r by lia. + rewrite Z.rem_mod_nonneg by lia. + reflexivity. + + + rewrite (Z.abs_eq (Int.signed b)) by lia. + rewrite Int.unsigned_repr ; cycle 1. + { pose proof Int.max_signed_unsigned. lia. } + rewrite Z.rem_mod_nonneg by lia. + reflexivity. +Qed. -- cgit From 759e546e570d82792e4611918cc365dd1000be61 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 15 Feb 2022 19:47:21 +0100 Subject: fp_divs32_correct --- kvx/FPDivision32.v | 41 +++++++++++++++++++++++++++++++++++++++++ 1 file changed, 41 insertions(+) diff --git a/kvx/FPDivision32.v b/kvx/FPDivision32.v index 4e9b2fb2..48e4bbbc 100644 --- a/kvx/FPDivision32.v +++ b/kvx/FPDivision32.v @@ -811,3 +811,44 @@ Proof. rewrite Z.rem_mod_nonneg by lia. reflexivity. Qed. + +Definition fp_mods32 a b := + Elet a (Elet (lift b) + (Elet (fp_modu32 (e_abs (Eletvar (1%nat))) (e_abs (Eletvar (0%nat)))) + (e_ite Tint (Ccomp0 Clt) (Eletvar 2%nat) + (e_neg (Eletvar 0%nat)) (Eletvar 0%nat)))). + +Theorem fp_mods32_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : (Int.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_mods32 expr_a expr_b) + (Vint (Int.mods a b)). +Proof. + intros. + unfold fp_mods32. + Local Opaque fp_modu32. + repeat (econstructor + apply eval_lift + eassumption). + apply fp_modu32_correct. + all: repeat (econstructor + apply eval_lift + eassumption). + { unfold ExtValues.int_absdiff, ExtValues.Z_abs_diff. + rewrite Int.signed_zero. rewrite Z.sub_0_r. + rewrite Int.unsigned_repr. + { pose proof (nonzero_unsigned_signed b b_nz). + lia. + } + pose proof Int.max_signed_unsigned. + pose proof int_min_signed_unsigned. + pose proof (Int.signed_range b). + lia. + } + cbn. + rewrite int_mods_modu ; cycle 1. + { apply nonzero_unsigned_signed. assumption. } + unfold Int.lt, ExtValues.int_abs, ExtValues.int_absdiff, ExtValues.Z_abs_diff. + change (Int.signed Int.zero) with 0%Z. + repeat rewrite Z.sub_0_r. + destruct zlt; reflexivity. +Qed. -- cgit From 191af63b4e3db42f203a5a2a62c5f924d7f37d81 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 15 Feb 2022 20:01:24 +0100 Subject: mods 64 --- kvx/FPDivision64.v | 235 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 235 insertions(+) diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 1d0fd2ee..113bf07a 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2406,3 +2406,238 @@ Proof. rewrite Int64.mul_commut. reflexivity. Qed. + +Definition e_is_negl a := Eop (Ocmp (Ccomplimm Clt Int64.zero)) (a ::: Enil). +Definition e_xorw a b := Eop Oxor (a ::: b ::: Enil). +Definition e_negl a := Eop Onegl (a ::: Enil). +Definition e_absl a := Eop (Oabsdifflimm Int64.zero) (a ::: Enil). + +Definition fp_divs64 a b := + Elet a (Elet (lift b) + (Elet (fp_divu64 (e_absl (Eletvar (1%nat))) (e_absl (Eletvar (0%nat)))) + (e_ite Tlong (Ccompu0 Cne) (e_xorw (e_is_negl (Eletvar 2%nat)) + (e_is_negl (Eletvar 1%nat))) + (e_negl (Eletvar 0%nat)) (Eletvar 0%nat)))). + +Lemma nonneg_signed_unsigned: + forall x (x_NONNEG : (Int64.signed x >= 0)%Z), + (Int64.signed x) = (Int64.unsigned x). +Proof. + intros. + pose proof (Int64.unsigned_range x). + unfold Int64.signed in *. + destruct zlt. reflexivity. + change Int64.modulus with 18446744073709551616%Z in *. + change Int64.half_modulus with 9223372036854775808%Z in *. + lia. +Qed. + +Lemma long_min_signed_unsigned : + (- Int64.min_signed < Int64.max_unsigned)%Z. +Proof. + reflexivity. +Qed. + +Lemma long_divs_divu : + forall a b + (b_NOT0 : (Int64.signed b <> 0)%Z), + Int64.divs a b = if xorb (Int64.lt a Int64.zero) + (Int64.lt b Int64.zero) + then Int64.neg (Int64.divu (ExtValues.long_abs a) + (ExtValues.long_abs b)) + else Int64.divu (ExtValues.long_abs a) (ExtValues.long_abs b). +Proof. + intros. + unfold Int64.divs, Int64.divu, Int64.lt, ExtValues.long_abs. + pose proof (Int64.signed_range a) as a_RANGE. + pose proof (Int64.signed_range b) as b_RANGE. + change (Int64.signed Int64.zero) with 0%Z. + destruct zlt. + - cbn. rewrite (Z.abs_neq (Int64.signed a)) by lia. + rewrite (Int64.unsigned_repr (- Int64.signed a)); cycle 1. + { pose proof long_min_signed_unsigned. lia. } + + destruct zlt. + + rewrite (Z.abs_neq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof long_min_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int64.signed b)) at 1. + rewrite Z.quot_opp_r by lia. + rewrite <- (Z.opp_involutive (Int64.signed a)) at 1. + rewrite Z.quot_opp_l by lia. + rewrite Z.quot_div_nonneg by lia. + rewrite Z.opp_involutive. + reflexivity. + + + rewrite (Z.abs_eq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof Int64.max_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int64.signed a)) at 1. + rewrite Z.quot_opp_l by lia. + rewrite Z.quot_div_nonneg by lia. + rewrite Int64.neg_repr. + reflexivity. + + - cbn. rewrite (Z.abs_eq (Int64.signed a)) by lia. + rewrite (Int64.unsigned_repr (Int64.signed a)); cycle 1. + { pose proof Int64.max_signed_unsigned. lia. } + destruct zlt. + + rewrite (Z.abs_neq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof long_min_signed_unsigned. lia. } + rewrite Int64.neg_repr. + rewrite <- (Z.opp_involutive (Int64.signed b)) at 1. + rewrite Z.quot_opp_r by lia. + rewrite Z.quot_div_nonneg by lia. + reflexivity. + + + rewrite (Z.abs_eq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof Int64.max_signed_unsigned. lia. } + rewrite Z.quot_div_nonneg by lia. + reflexivity. +Qed. + +Lemma nonzero_unsigned_signed : + forall b, (Int64.unsigned b > 0 -> Int64.signed b <> 0)%Z. +Proof. + intros b GT EQ. + rewrite Int64.unsigned_signed in GT. + unfold Int64.lt in GT. + rewrite Int64.signed_zero in GT. + destruct zlt in GT; lia. +Qed. + +Theorem fp_divs64_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int64) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)) + (b_nz : (Int64.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_divs64 expr_a expr_b) + (Vlong (Int64.divs a b)). +Proof. + intros. + unfold fp_divs64. + Local Opaque fp_divu64. + repeat (econstructor + apply eval_lift + eassumption). + apply fp_divu64_correct. + all: repeat (econstructor + apply eval_lift + eassumption). + { unfold ExtValues.long_absdiff, ExtValues.Z_abs_diff. + rewrite Int64.signed_zero. rewrite Z.sub_0_r. + rewrite Int64.unsigned_repr. + { pose proof (nonzero_unsigned_signed b b_nz). + lia. + } + pose proof Int64.max_signed_unsigned. + pose proof long_min_signed_unsigned. + pose proof (Int64.signed_range b). + lia. + } + cbn. + rewrite long_divs_divu ; cycle 1. + { apply nonzero_unsigned_signed. assumption. } + unfold Int64.lt, ExtValues.long_abs, ExtValues.long_absdiff, ExtValues.Z_abs_diff. + change (Int64.signed Int64.zero) with 0%Z. + repeat rewrite Z.sub_0_r. + destruct zlt; destruct zlt; reflexivity. +Qed. + +Lemma long_mods_modu : + forall a b + (b_NOT0 : (Int64.signed b <> 0)%Z), + Int64.mods a b = if Int64.lt a Int64.zero + then Int64.neg (Int64.modu (ExtValues.long_abs a) + (ExtValues.long_abs b)) + else Int64.modu (ExtValues.long_abs a) (ExtValues.long_abs b). +Proof. + intros. + unfold Int64.mods, Int64.modu, Int64.lt, ExtValues.long_abs. + pose proof (Int64.signed_range a) as a_RANGE. + pose proof (Int64.signed_range b) as b_RANGE. + change (Int64.signed Int64.zero) with 0%Z. + destruct zlt. + - cbn. rewrite (Z.abs_neq (Int64.signed a)) by lia. + rewrite (Int64.unsigned_repr (- Int64.signed a)); cycle 1. + { pose proof long_min_signed_unsigned. lia. } + + destruct (zlt (Int64.signed b) 0%Z). + + rewrite (Z.abs_neq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof long_min_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int64.signed b)) at 1. + rewrite Z.rem_opp_r by lia. + rewrite <- (Z.opp_involutive (Int64.signed a)) at 1. + rewrite Z.rem_opp_l by lia. + rewrite Z.rem_mod_nonneg by lia. + rewrite Int64.neg_repr. + reflexivity. + + + rewrite (Z.abs_eq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof Int64.max_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int64.signed a)) at 1. + rewrite Z.rem_opp_l by lia. + rewrite Z.rem_mod_nonneg by lia. + rewrite Int64.neg_repr. + reflexivity. + + - cbn. rewrite (Z.abs_eq (Int64.signed a)) by lia. + rewrite (Int64.unsigned_repr (Int64.signed a)); cycle 1. + { pose proof Int64.max_signed_unsigned. lia. } + destruct (zlt (Int64.signed b) 0%Z). + + rewrite (Z.abs_neq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof long_min_signed_unsigned. lia. } + rewrite <- (Z.opp_involutive (Int64.signed b)) at 1. + rewrite Z.rem_opp_r by lia. + rewrite Z.rem_mod_nonneg by lia. + reflexivity. + + + rewrite (Z.abs_eq (Int64.signed b)) by lia. + rewrite Int64.unsigned_repr ; cycle 1. + { pose proof Int64.max_signed_unsigned. lia. } + rewrite Z.rem_mod_nonneg by lia. + reflexivity. +Qed. + +Definition fp_mods64 a b := + Elet a (Elet (lift b) + (Elet (fp_modu64 (e_absl (Eletvar (1%nat))) (e_absl (Eletvar (0%nat)))) + (e_ite Tlong (Ccompl0 Clt) (Eletvar 2%nat) + (e_negl (Eletvar 0%nat)) (Eletvar 0%nat)))). + +Theorem fp_mods64_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int64) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)) + (b_nz : (Int64.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_mods64 expr_a expr_b) + (Vlong (Int64.mods a b)). +Proof. + intros. + unfold fp_mods64. + Local Opaque fp_modu64. + repeat (econstructor + apply eval_lift + eassumption). + apply fp_modu64_correct. + all: repeat (econstructor + apply eval_lift + eassumption). + { unfold ExtValues.long_absdiff, ExtValues.Z_abs_diff. + rewrite Int64.signed_zero. rewrite Z.sub_0_r. + rewrite Int64.unsigned_repr. + { pose proof (nonzero_unsigned_signed b b_nz). + lia. + } + pose proof Int64.max_signed_unsigned. + pose proof long_min_signed_unsigned. + pose proof (Int64.signed_range b). + lia. + } + cbn. + rewrite long_mods_modu ; cycle 1. + { apply nonzero_unsigned_signed. assumption. } + unfold Int64.lt, ExtValues.long_abs, ExtValues.long_absdiff, ExtValues.Z_abs_diff. + change (Int64.signed Int64.zero) with 0%Z. + repeat rewrite Z.sub_0_r. + destruct zlt; reflexivity. +Qed. -- cgit From 867b74d57fc2b834cc19176b650dc62a1e4e0fd2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 16 Feb 2022 09:14:01 +0100 Subject: simpler mod --- kvx/FPDivision32.v | 37 +++++++++++++++++++++++++++++++++---- kvx/FPDivision64.v | 35 +++++++++++++++++++++++++++++++---- 2 files changed, 64 insertions(+), 8 deletions(-) diff --git a/kvx/FPDivision32.v b/kvx/FPDivision32.v index 48e4bbbc..5a7b536f 100644 --- a/kvx/FPDivision32.v +++ b/kvx/FPDivision32.v @@ -812,23 +812,23 @@ Proof. reflexivity. Qed. -Definition fp_mods32 a b := +Definition fp_mods32z a b := Elet a (Elet (lift b) (Elet (fp_modu32 (e_abs (Eletvar (1%nat))) (e_abs (Eletvar (0%nat)))) (e_ite Tint (Ccomp0 Clt) (Eletvar 2%nat) (e_neg (Eletvar 0%nat)) (Eletvar 0%nat)))). -Theorem fp_mods32_correct : +Theorem fp_mods32z_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_a expr_b : expr) (a b : int) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) (b_nz : (Int.unsigned b > 0)%Z), - eval_expr ge sp cmenv memenv le (fp_mods32 expr_a expr_b) + eval_expr ge sp cmenv memenv le (fp_mods32z expr_a expr_b) (Vint (Int.mods a b)). Proof. intros. - unfold fp_mods32. + unfold fp_mods32z. Local Opaque fp_modu32. repeat (econstructor + apply eval_lift + eassumption). apply fp_modu32_correct. @@ -852,3 +852,32 @@ Proof. repeat rewrite Z.sub_0_r. destruct zlt; reflexivity. Qed. + +Definition e_msub a b c := Eop Omsub (a ::: b ::: c ::: Enil). + +Definition fp_mods32 a b := + Elet a (Elet (lift b) + (Elet (fp_divs32 (Eletvar (1%nat)) (Eletvar (0%nat))) + (e_msub (Eletvar 2%nat) (Eletvar 1%nat) (Eletvar 0%nat)))). + +Theorem fp_mods32_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vint a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vint b)) + (b_nz : (Int.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_mods32 expr_a expr_b) + (Vint (Int.mods a b)). +Proof. + intros. + rewrite Int.mods_divs. + unfold fp_mods32. + Local Opaque fp_divs32. + repeat (econstructor + apply eval_lift + eassumption). + { apply fp_divs32_correct; + repeat (econstructor + apply eval_lift + eassumption). + } + cbn. + rewrite Int.mul_commut. + reflexivity. +Qed. diff --git a/kvx/FPDivision64.v b/kvx/FPDivision64.v index 113bf07a..298831a0 100644 --- a/kvx/FPDivision64.v +++ b/kvx/FPDivision64.v @@ -2601,23 +2601,23 @@ Proof. reflexivity. Qed. -Definition fp_mods64 a b := +Definition fp_mods64z a b := Elet a (Elet (lift b) (Elet (fp_modu64 (e_absl (Eletvar (1%nat))) (e_absl (Eletvar (0%nat)))) (e_ite Tlong (Ccompl0 Clt) (Eletvar 2%nat) (e_negl (Eletvar 0%nat)) (Eletvar 0%nat)))). -Theorem fp_mods64_correct : +Theorem fp_mods64z_correct : forall (ge : genv) (sp: val) cmenv memenv (le : letenv) (expr_a expr_b : expr) (a b : int64) (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)) (b_nz : (Int64.unsigned b > 0)%Z), - eval_expr ge sp cmenv memenv le (fp_mods64 expr_a expr_b) + eval_expr ge sp cmenv memenv le (fp_mods64z expr_a expr_b) (Vlong (Int64.mods a b)). Proof. intros. - unfold fp_mods64. + unfold fp_mods64z. Local Opaque fp_modu64. repeat (econstructor + apply eval_lift + eassumption). apply fp_modu64_correct. @@ -2641,3 +2641,30 @@ Proof. repeat rewrite Z.sub_0_r. destruct zlt; reflexivity. Qed. + +Definition fp_mods64 a b := + Elet a (Elet (lift b) + (Elet (fp_divs64 (Eletvar (1%nat)) (Eletvar (0%nat))) + (e_msubl (Eletvar 2%nat) (Eletvar 1%nat) (Eletvar 0%nat)))). + +Theorem fp_mods64_correct : + forall (ge : genv) (sp: val) cmenv memenv + (le : letenv) (expr_a expr_b : expr) (a b : int64) + (EVAL_a : eval_expr ge sp cmenv memenv le expr_a (Vlong a)) + (EVAL_b : eval_expr ge sp cmenv memenv le expr_b (Vlong b)) + (b_nz : (Int64.unsigned b > 0)%Z), + eval_expr ge sp cmenv memenv le (fp_mods64 expr_a expr_b) + (Vlong (Int64.mods a b)). +Proof. + intros. + rewrite Int64.mods_divs. + unfold fp_mods64. + Local Opaque fp_divs64. + repeat (econstructor + apply eval_lift + eassumption). + { apply fp_divs64_correct; + repeat (econstructor + apply eval_lift + eassumption). + } + cbn. + rewrite Int64.mul_commut. + reflexivity. +Qed. -- cgit From eaae75e487a82021cc615856b31f86bd05853b1e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 16 Feb 2022 09:32:24 +0100 Subject: builtins pour signed div/mod --- kvx/Builtins1.v | 32 +++++++++++++++++++ kvx/CBuiltins.ml | 8 +++++ kvx/SelectOp.vp | 16 ++++++++++ kvx/SelectOpproof.v | 88 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 4 files changed, 144 insertions(+) diff --git a/kvx/Builtins1.v b/kvx/Builtins1.v index e3e4f95d..5536e58c 100644 --- a/kvx/Builtins1.v +++ b/kvx/Builtins1.v @@ -32,6 +32,10 @@ Inductive platform_builtin : Type := | BI_fp_udiv64 | BI_fp_umod32 | BI_fp_umod64 +| BI_fp_sdiv32 +| BI_fp_sdiv64 +| BI_fp_smod32 +| BI_fp_smod64 | BI_abs | BI_absl. @@ -50,6 +54,10 @@ Definition platform_builtin_table : list (string * platform_builtin) := :: ("__builtin_fp_udiv64", BI_fp_udiv64) :: ("__builtin_fp_umod32", BI_fp_umod32) :: ("__builtin_fp_umod64", BI_fp_umod64) + :: ("__builtin_fp_sdiv32", BI_fp_sdiv32) + :: ("__builtin_fp_sdiv64", BI_fp_sdiv64) + :: ("__builtin_fp_smod32", BI_fp_smod32) + :: ("__builtin_fp_smod64", BI_fp_smod64) :: ("__builtin_abs", BI_abs) :: ("__builtin_absl", BI_absl) :: nil. @@ -74,6 +82,14 @@ Definition platform_builtin_sig (b: platform_builtin) : signature := mksignature (Tint :: Tint :: nil) Tint cc_default | BI_fp_umod64 => mksignature (Tlong :: Tlong :: nil) Tlong cc_default + | BI_fp_sdiv32 => + mksignature (Tint :: Tint :: nil) Tint cc_default + | BI_fp_sdiv64 => + mksignature (Tlong :: Tlong :: nil) Tlong cc_default + | BI_fp_smod32 => + mksignature (Tint :: Tint :: nil) Tint cc_default + | BI_fp_smod64 => + mksignature (Tlong :: Tlong :: nil) Tlong cc_default | BI_abs => mksignature (Tint :: nil) Tint cc_default | BI_absl => @@ -106,6 +122,22 @@ Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (pl (fun n1 n2 => if Int64.eq n2 Int64.zero then None else Some (Int64.modu n1 n2)) + | BI_fp_sdiv32 => mkbuiltin_n2p Tint Tint Tint + (fun n1 n2 => if Int.eq n2 Int.zero + then None + else Some (Int.divs n1 n2)) + | BI_fp_sdiv64 => mkbuiltin_n2p Tlong Tlong Tlong + (fun n1 n2 => if Int64.eq n2 Int64.zero + then None + else Some (Int64.divs n1 n2)) + | BI_fp_smod32 => mkbuiltin_n2p Tint Tint Tint + (fun n1 n2 => if Int.eq n2 Int.zero + then None + else Some (Int.mods n1 n2)) + | BI_fp_smod64 => mkbuiltin_n2p Tlong Tlong Tlong + (fun n1 n2 => if Int64.eq n2 Int64.zero + then None + else Some (Int64.mods n1 n2)) | BI_abs => mkbuiltin_n1t Tint Tint ExtValues.int_abs | BI_absl => mkbuiltin_n1t Tlong Tlong ExtValues.long_abs end. diff --git a/kvx/CBuiltins.ml b/kvx/CBuiltins.ml index d712eecd..c0b69adf 100644 --- a/kvx/CBuiltins.ml +++ b/kvx/CBuiltins.ml @@ -145,6 +145,14 @@ let builtins = { (TInt(IUInt, []), [TInt(IUInt, []); TInt(IUInt, [])], false); "__builtin_fp_umod64", (TInt(IULong, []), [TInt(IULong, []); TInt(IULong, [])], false); + "__builtin_fp_sdiv32", + (TInt(IInt, []), [TInt(IInt, []); TInt(IInt, [])], false); + "__builtin_fp_sdiv64", + (TInt(ILong, []), [TInt(ILong, []); TInt(ILong, [])], false); + "__builtin_fp_smod32", + (TInt(IInt, []), [TInt(IInt, []); TInt(IInt, [])], false); + "__builtin_fp_smod64", + (TInt(ILong, []), [TInt(ILong, []); TInt(ILong, [])], false); "__builtin_abs", (TInt(IInt, []), [TInt(IInt, [])], false); "__builtin_absl", diff --git a/kvx/SelectOp.vp b/kvx/SelectOp.vp index 63c7d73b..5225a71c 100644 --- a/kvx/SelectOp.vp +++ b/kvx/SelectOp.vp @@ -788,6 +788,22 @@ Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr | a:::b:::Enil => Some (FPDivision64.fp_modu64 a b) | _ => None end) + | BI_fp_sdiv32 => (match args with + | a:::b:::Enil => Some (FPDivision32.fp_divs32 a b) + | _ => None + end) + | BI_fp_sdiv64 => (match args with + | a:::b:::Enil => Some (FPDivision64.fp_divs64 a b) + | _ => None + end) + | BI_fp_smod32 => (match args with + | a:::b:::Enil => Some (FPDivision32.fp_mods32 a b) + | _ => None + end) + | BI_fp_smod64 => (match args with + | a:::b:::Enil => Some (FPDivision64.fp_mods64 a b) + | _ => None + end) | BI_abs => gen_abs args | BI_absl => gen_absl args end. diff --git a/kvx/SelectOpproof.v b/kvx/SelectOpproof.v index a45d367f..19393091 100644 --- a/kvx/SelectOpproof.v +++ b/kvx/SelectOpproof.v @@ -2046,6 +2046,94 @@ Proof. } inv EVAL. constructor. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int.eq in EVAL. + change (Int.unsigned Int.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vint (Int.divs i i0)). split. + { + apply FPDivision32.fp_divs32_correct; auto. + pose proof (Int.unsigned_range i0). + lia. + } + inv EVAL. + constructor. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int64.eq in EVAL. + change (Int64.unsigned Int64.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int64.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vlong (Int64.divs i i0)). split. + { + apply FPDivision64.fp_divs64_correct; auto. + pose proof (Int64.unsigned_range i0). + lia. + } + inv EVAL. + constructor. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int.eq in EVAL. + change (Int.unsigned Int.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vint (Int.mods i i0)). split. + { + apply FPDivision32.fp_mods32_correct; auto. + pose proof (Int.unsigned_range i0). + lia. + } + inv EVAL. + constructor. + - cbn in *. + intro EVAL. + inv EVAL. discriminate. + inv H0. discriminate. + inv H2. 2: discriminate. + inv Heval. + intro EVAL. + destruct v1; try discriminate. + destruct v0; try discriminate. + unfold Int64.eq in EVAL. + change (Int64.unsigned Int64.zero) with 0 in EVAL. + unfold zeq in EVAL. + destruct (Z.eq_dec (Int64.unsigned i0) 0) as [EQ | NEQ]. + { discriminate. } + exists (Vlong (Int64.mods i i0)). split. + { + apply FPDivision64.fp_mods64_correct; auto. + pose proof (Int64.unsigned_range i0). + lia. + } + inv EVAL. + constructor. - apply eval_abs; assumption. - apply eval_absl; assumption. Qed. -- cgit From 2d9138547d93c32c0ec5ae54b4afc022f5c434ff Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 7 Mar 2022 16:13:58 +0100 Subject: installation --- INSTALL.md | 16 ++++++++++++++-- 1 file changed, 14 insertions(+), 2 deletions(-) diff --git a/INSTALL.md b/INSTALL.md index 5e2e800d..755ba690 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -34,11 +34,23 @@ Install dependencies available through opam ``` opam install coq menhir ``` - Note: it may happen that a newer version of Coq is not supported yet. You may downgrade to solve the problem: ``` -opam pin add coq 8.12.2 # example of Coq version +opam pin add coq 8.13.2 # example of Coq version +``` + +### For Kalray KVX +On this platform, we also need Gappa installed. +You may need to first install some development packages. On Ubuntu: +``` +apt install bison flex libmpfr-dev libboost-dev +``` + +This install Gappa and Flocq: +``` +opam pin add coq-flocq 3.4.0 --no-action +opam install gappa coq-gappa coq-flocq ``` ## Compilation -- cgit