Skip to content

Commit 0fe672c

Browse files
theryproux01
authored andcommitted
Add Zfloor/Zceil to Reals
1 parent 83d0fad commit 0fe672c

File tree

3 files changed

+140
-0
lines changed

3 files changed

+140
-0
lines changed

doc/changelog/01-added/89-Zfloor.rst

+10
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,10 @@
1+
- in `Reals`
2+
3+
+ new file `Zfloor.v` with definitions of `Zfloor` and `Zceil`
4+
and corresponding lemmas `up_Zfloor`, `IZR_up_Zfloor`,
5+
`Zfloor_bound`, `Zfloor_lub`, `Zfloor_eq`, `Zfloor_le`,
6+
`Zfloor_addz`, `ZfloorZ`, `ZfloorNZ`, `ZfloorD_cond`, `Zceil_eq`,
7+
`Zceil_le`, `Zceil_addz`, `ZceilD_cond`, `ZfloorB_cond`,
8+
`ZceilB_cond`
9+
(`#89 <https://github.com/coq/stdlib/pull/89>`_,
10+
by Laurent Théry).

theories/Reals/Reals.v

+1
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,7 @@
2626

2727
Require Export Rbase.
2828
Require Export Rfunctions.
29+
Require Export Zfloor.
2930
Require Export SeqSeries.
3031
Require Export Rtrigo.
3132
Require Export Ranalysis.

theories/Reals/Zfloor.v

+129
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,129 @@
1+
(******************************************************************************)
2+
(* *)
3+
(* Zfloor x == floor function : R -> Z *)
4+
(* Zceil x == ceil function : R -> Z *)
5+
(* *)
6+
(******************************************************************************)
7+
8+
Require Import Rbase Rfunctions Lra Lia.
9+
10+
Local Open Scope R_scope.
11+
12+
Definition Zfloor (x : R) := (up x - 1)%Z.
13+
14+
Lemma up_Zfloor x : up x = (Zfloor x + 1)%Z.
15+
Proof. unfold Zfloor; lia. Qed.
16+
17+
Lemma IZR_up_Zfloor x : IZR (up x) = IZR (Zfloor x) + 1.
18+
Proof. now rewrite up_Zfloor, plus_IZR. Qed.
19+
20+
Lemma Zfloor_bound x : IZR (Zfloor x) <= x < IZR (Zfloor x) + 1.
21+
Proof.
22+
unfold Zfloor; rewrite minus_IZR.
23+
generalize (archimed x); lra.
24+
Qed.
25+
26+
Lemma Zfloor_lub (z : Z) x : IZR z <= x -> (z <= Zfloor x)%Z.
27+
Proof.
28+
intro H.
29+
assert (H1 : (z < Zfloor x + 1)%Z);[| lia].
30+
apply lt_IZR; rewrite plus_IZR.
31+
now generalize (Zfloor_bound x); lra.
32+
Qed.
33+
34+
Lemma Zfloor_eq (z : Z) x : IZR z <= x < IZR z + 1 -> Zfloor x = z.
35+
Proof.
36+
intro xB.
37+
assert (ZxB := Zfloor_bound x).
38+
assert (B : (Zfloor x < z + 1 /\ z <= Zfloor x)%Z) ; [| lia].
39+
split; [|apply Zfloor_lub; lra].
40+
apply lt_IZR; rewrite plus_IZR; lra.
41+
Qed.
42+
43+
Lemma Zfloor_le x y : x <= y -> (Zfloor x <= Zfloor y)%Z.
44+
Proof.
45+
intro H; apply Zfloor_lub; generalize (Zfloor_bound x); lra.
46+
Qed.
47+
48+
Lemma Zfloor_addz (z: Z) x : Zfloor (x + IZR z) = (Zfloor x + z)%Z.
49+
Proof.
50+
assert (ZB := Zfloor_bound x).
51+
now apply Zfloor_eq; rewrite plus_IZR; lra.
52+
Qed.
53+
54+
Lemma ZfloorZ (z : Z) : Zfloor (IZR z) = z.
55+
Proof. now apply Zfloor_eq; lra. Qed.
56+
57+
58+
Lemma ZfloorNZ (z : Z) : Zfloor (- IZR z) = (- z)%Z.
59+
Proof. now rewrite <- opp_IZR, ZfloorZ. Qed.
60+
61+
Lemma ZfloorD_cond r1 r2 :
62+
if Rle_dec (IZR (Zfloor r1) + IZR (Zfloor r2) + 1) (r1 + r2)
63+
then Zfloor (r1 + r2) = (Zfloor r1 + Zfloor r2 + 1)%Z
64+
else Zfloor (r1 + r2) = (Zfloor r1 + Zfloor r2)%Z.
65+
Proof.
66+
destruct (Zfloor_bound r1, Zfloor_bound r2) as [H1 H2].
67+
case Rle_dec; intro H.
68+
now apply Zfloor_eq; rewrite plus_IZR, plus_IZR; lra.
69+
now apply Zfloor_eq; rewrite plus_IZR; lra.
70+
Qed.
71+
72+
Definition Zceil (x : R) := (- Zfloor (- x))%Z.
73+
74+
Theorem Zceil_bound x : (IZR (Zceil x) - 1 < x <= IZR (Zceil x))%R.
75+
Proof.
76+
now unfold Zceil; generalize (Zfloor_bound (- x)); rewrite !opp_IZR; lra.
77+
Qed.
78+
79+
Theorem Zfloor_ceil_bound x : (IZR (Zfloor x) <= x <= IZR (Zceil x))%R.
80+
Proof. now generalize (Zfloor_bound x) (Zceil_bound x); lra. Qed.
81+
82+
Theorem ZceilN x : (Zceil (- x) = - Zfloor x)%Z.
83+
Proof. now unfold Zceil; rewrite Ropp_involutive. Qed.
84+
85+
Theorem ZfloorN x : (Zfloor (- x) = - Zceil x)%Z.
86+
Proof. unfold Zceil; lia. Qed.
87+
88+
Lemma Zceil_eq (z : Z) x : IZR z - 1 < x <= IZR z -> Zceil x = z.
89+
Proof.
90+
intro xB; assert (H : Zfloor (- x) = (- z)%Z); [|unfold Zceil; lia].
91+
now apply Zfloor_eq; rewrite opp_IZR; lra.
92+
Qed.
93+
94+
Lemma Zceil_le x y : x <= y -> (Zceil x <= Zceil y)%Z.
95+
Proof.
96+
intro xLy; apply Z.opp_le_mono; unfold Zceil; rewrite !Z.opp_involutive.
97+
now apply Zfloor_le; lra.
98+
Qed.
99+
100+
Lemma Zceil_addz (z: Z) x : Zceil (x + IZR z) = (Zceil x + z)%Z.
101+
Proof.
102+
now unfold Zceil; rewrite Ropp_plus_distr, <- opp_IZR, Zfloor_addz; lia.
103+
Qed.
104+
105+
Lemma ZceilD_cond r1 r2 :
106+
if Rle_dec (r1 + r2) (IZR (Zceil r1) + IZR (Zceil r2) - 1)
107+
then Zceil (r1 + r2) = (Zceil r1 + Zceil r2 - 1)%Z
108+
else Zceil (r1 + r2) = (Zceil r1 + Zceil r2)%Z.
109+
Proof.
110+
generalize (ZfloorD_cond (- r1) (- r2)).
111+
now unfold Zceil; rewrite !opp_IZR; do 2 case Rle_dec; try lra;
112+
rewrite Ropp_plus_distr; lia.
113+
Qed.
114+
115+
Lemma ZfloorB_cond r1 r2 :
116+
if Rle_dec (IZR (Zfloor r1) - IZR (Zceil r2) + 1) (r1 - r2)
117+
then Zfloor (r1 - r2) = (Zfloor r1 - Zceil r2 + 1)%Z
118+
else Zfloor (r1 - r2) = (Zfloor r1 - Zceil r2)%Z.
119+
Proof.
120+
now generalize (ZfloorD_cond r1 (- r2)); rewrite !ZfloorN, !opp_IZR.
121+
Qed.
122+
123+
Lemma ZceilB_cond r1 r2 :
124+
if Rle_dec (r1 - r2) (IZR (Zceil r1) - IZR (Zfloor r2) - 1)
125+
then Zceil (r1 - r2) = (Zceil r1 - Zfloor r2 - 1)%Z
126+
else Zceil (r1 - r2) = (Zceil r1 - Zfloor r2)%Z.
127+
Proof.
128+
now generalize (ZceilD_cond r1 (- r2)); rewrite !ZceilN, !opp_IZR.
129+
Qed.

0 commit comments

Comments
 (0)