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 (limited to '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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(-) (limited to 'kvx/FPDivision64.v') 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/FPDivision64.v | 65 +++++++++++++++++++++++++++++------------------------- 1 file changed, 35 insertions(+), 30 deletions(-) (limited to 'kvx/FPDivision64.v') 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). -- 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(-) (limited to 'kvx/FPDivision64.v') 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 ----------------------------------------------------- 1 file changed, 258 deletions(-) (limited to 'kvx/FPDivision64.v') 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). -- 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(+) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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(+) (limited to 'kvx/FPDivision64.v') 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 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(+) (limited to 'kvx/FPDivision64.v') 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/FPDivision64.v | 35 +++++++++++++++++++++++++++++++---- 1 file changed, 31 insertions(+), 4 deletions(-) (limited to 'kvx/FPDivision64.v') 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