Skip to content

Commit 8483ac7

Browse files
algebraic-devtydeuTwoFX
authored
fix: adjustments to the datetime library (#6431)
This PR fixes the `Repr` instance of the `Timestamp` type and changes the `PlainTime` type so that it always represents a clock time that may be a leap second. - Fix timestamp `Repr`. - The `PlainTime` type now always represents a clock time that may be a leap second. - Changed `readlink -f` to `IO.FS.realPath` --------- Co-authored-by: Mac Malone <tydeu@hatpress.net> Co-authored-by: Markus Himmel <markus@himmel-villmar.de>
1 parent 5f41cc7 commit 8483ac7

File tree

13 files changed

+54
-57
lines changed

13 files changed

+54
-57
lines changed

src/Std/Time/DateTime/PlainDateTime.lean

Lines changed: 5 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,6 @@ structure PlainDateTime where
2929
The `Time` component of a `PlainTime`
3030
-/
3131
time : PlainTime
32-
3332
deriving Inhabited, BEq, Repr
3433

3534
namespace PlainDateTime
@@ -123,7 +122,7 @@ def ofTimestampAssumingUTC (stamp : Timestamp) : PlainDateTime := Id.run do
123122

124123
return {
125124
date := PlainDate.ofYearMonthDayClip year hmon (Day.Ordinal.ofFin (Fin.succ mday))
126-
time := PlainTime.ofHourMinuteSecondsNano (leap := false) (hour.expandTop (by decide)) minute second nano
125+
time := PlainTime.ofHourMinuteSecondsNano (hour.expandTop (by decide)) minute (second.expandTop (by decide)) nano
127126
}
128127

129128
/--
@@ -199,7 +198,7 @@ Creates a new `PlainDateTime` by adjusting the `hour` component of its `time` to
199198
-/
200199
@[inline]
201200
def withHours (dt : PlainDateTime) (hour : Hour.Ordinal) : PlainDateTime :=
202-
{ dt with time := { dt.time with hour := hour } }
201+
{ dt with time := { dt.time with hour } }
203202

204203
/--
205204
Creates a new `PlainDateTime` by adjusting the `minute` component of its `time` to the given value.
@@ -212,7 +211,7 @@ def withMinutes (dt : PlainDateTime) (minute : Minute.Ordinal) : PlainDateTime :
212211
Creates a new `PlainDateTime` by adjusting the `second` component of its `time` to the given value.
213212
-/
214213
@[inline]
215-
def withSeconds (dt : PlainDateTime) (second : Sigma Second.Ordinal) : PlainDateTime :=
214+
def withSeconds (dt : PlainDateTime) (second : Second.Ordinal true) : PlainDateTime :=
216215
{ dt with time := { dt.time with second := second } }
217216

218217
/--
@@ -457,8 +456,8 @@ def millisecond (dt : PlainDateTime) : Millisecond.Ordinal :=
457456
Getter for the `Second` inside of a `PlainDateTime`.
458457
-/
459458
@[inline]
460-
def second (dt : PlainDateTime) : Second.Ordinal dt.time.second.fst :=
461-
dt.time.second.snd
459+
def second (dt : PlainDateTime) : Second.Ordinal true :=
460+
dt.time.second
462461

463462
/--
464463
Getter for the `Nanosecond.Ordinal` inside of a `PlainDateTime`.

src/Std/Time/DateTime/Timestamp.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -37,10 +37,10 @@ instance : OfNat Timestamp n where
3737
ofNat := ⟨OfNat.ofNat n⟩
3838

3939
instance : ToString Timestamp where
40-
toString s := toString s.val.toMilliseconds
40+
toString s := toString s.val.toSeconds
4141

4242
instance : Repr Timestamp where
43-
reprPrec s := reprPrec (toString s)
43+
reprPrec s := Repr.addAppParen ("Timestamp.ofNanosecondsSinceUnixEpoch " ++ repr s.val.toNanoseconds)
4444

4545
namespace Timestamp
4646

src/Std/Time/Format.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -280,7 +280,7 @@ def format (time : PlainTime) (format : String) : String :=
280280
Parses a time string in the 24-hour format (`HH:mm:ss`) and returns a `PlainTime`.
281281
-/
282282
def fromTime24Hour (input : String) : Except String PlainTime :=
283-
Formats.time24Hour.parseBuilder (fun h m s => some (PlainTime.ofHourMinuteSeconds h m s.snd)) input
283+
Formats.time24Hour.parseBuilder (fun h m s => some (PlainTime.ofHourMinuteSeconds h m s)) input
284284

285285
/--
286286
Formats a `PlainTime` value into a 24-hour format string (`HH:mm:ss`).
@@ -292,8 +292,8 @@ def toTime24Hour (input : PlainTime) : String :=
292292
Parses a time string in the lean 24-hour format (`HH:mm:ss.SSSSSSSSS` or `HH:mm:ss`) and returns a `PlainTime`.
293293
-/
294294
def fromLeanTime24Hour (input : String) : Except String PlainTime :=
295-
Formats.leanTime24Hour.parseBuilder (fun h m s n => some (PlainTime.ofHourMinuteSecondsNano h m s.snd n)) input
296-
<|> Formats.leanTime24HourNoNanos.parseBuilder (fun h m s => some (PlainTime.ofHourMinuteSecondsNano h m s.snd 0)) input
295+
Formats.leanTime24Hour.parseBuilder (fun h m s n => some <| PlainTime.ofHourMinuteSecondsNano h m s n) input
296+
<|> Formats.leanTime24HourNoNanos.parseBuilder (fun h m s => some <| PlainTime.ofHourMinuteSecondsNano h m s 0) input
297297

298298
/--
299299
Formats a `PlainTime` value into a 24-hour format string (`HH:mm:ss.SSSSSSSSS`).
@@ -307,7 +307,7 @@ Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `PlainT
307307
def fromTime12Hour (input : String) : Except String PlainTime := do
308308
let builder h m s a : Option PlainTime := do
309309
let value ← Internal.Bounded.ofInt? h.val
310-
some <| PlainTime.ofHourMinuteSeconds (HourMarker.toAbsolute a value) m s.snd
310+
some <| PlainTime.ofHourMinuteSeconds (HourMarker.toAbsolute a value) m s
311311

312312
Formats.time12Hour.parseBuilder builder input
313313

src/Std/Time/Format/Basic.lean

Lines changed: 10 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -741,7 +741,7 @@ private def toIsoString (offset : Offset) (withMinutes : Bool) (withSeconds : Bo
741741

742742
let data := s!"{sign}{pad time.hour.val}"
743743
let data := if withMinutes then s!"{data}{if colon then ":" else ""}{pad time.minute.val}" else data
744-
let data := if withSeconds ∧ time.second.snd.val ≠ 0 then s!"{data}{if colon then ":" else ""}{pad time.second.snd.val}" else data
744+
let data := if withSeconds ∧ time.second.val ≠ 0 then s!"{data}{if colon then ":" else ""}{pad time.second.val}" else data
745745

746746
data
747747

@@ -764,7 +764,7 @@ private def TypeFormat : Modifier → Type
764764
| .k _ => Bounded.LE 1 24
765765
| .H _ => Hour.Ordinal
766766
| .m _ => Minute.Ordinal
767-
| .s _ => Sigma Second.Ordinal
767+
| .s _ => Second.Ordinal true
768768
| .S _ => Nanosecond.Ordinal
769769
| .A _ => Millisecond.Offset
770770
| .n _ => Nanosecond.Ordinal
@@ -835,10 +835,10 @@ private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : Strin
835835
| .narrow => formatMarkerNarrow data
836836
| .h format => pad format.padding (data.val % 12)
837837
| .K format => pad format.padding (data.val % 12)
838-
| .k format => pad format.padding (data.val)
839-
| .H format => pad format.padding (data.val)
840-
| .m format => pad format.padding (data.val)
841-
| .s format => pad format.padding (data.snd.val)
838+
| .k format => pad format.padding data.val
839+
| .H format => pad format.padding data.val
840+
| .m format => pad format.padding data.val
841+
| .s format => pad format.padding data.val
842842
| .S format =>
843843
match format with
844844
| .nano => pad 9 data.val
@@ -1167,7 +1167,7 @@ private def parseWith : (mod : Modifier) → Parser (TypeFormat mod)
11671167
| .k format => parseNatToBounded (parseAtLeastNum format.padding)
11681168
| .H format => parseNatToBounded (parseAtLeastNum format.padding)
11691169
| .m format => parseNatToBounded (parseAtLeastNum format.padding)
1170-
| .s format => Sigma.mk true <$> (parseNatToBounded (parseAtLeastNum format.padding))
1170+
| .s format => parseNatToBounded (parseAtLeastNum format.padding)
11711171
| .S format =>
11721172
match format with
11731173
| .nano => parseNatToBounded (parseAtLeastNum 9)
@@ -1249,7 +1249,7 @@ private structure DateBuilder where
12491249
k : Option (Bounded.LE 1 24) := none
12501250
H : Option Hour.Ordinal := none
12511251
m : Option Minute.Ordinal := none
1252-
s : Option (Sigma Second.Ordinal) := none
1252+
s : Option (Second.Ordinal true) := none
12531253
S : Option Nanosecond.Ordinal := none
12541254
A : Option Millisecond.Offset := none
12551255
n : Option Nanosecond.Ordinal := none
@@ -1335,10 +1335,10 @@ private def build (builder : DateBuilder) (aw : Awareness) : Option aw.type :=
13351335
|>.getD ⟨0, by decide⟩
13361336

13371337
let minute := builder.m |>.getD 0
1338-
let second := builder.s |>.getD false, 0
1338+
let second := builder.s |>.getD 0
13391339
let nano := (builder.n <|> builder.S) |>.getD 0
13401340

1341-
let time : PlainTime
1341+
let time
13421342
:= PlainTime.ofNanoseconds <$> builder.N
13431343
<|> PlainTime.ofMilliseconds <$> builder.A
13441344
|>.getD (PlainTime.mk hour minute second nano)

src/Std/Time/Notation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -126,7 +126,7 @@ private def convertPlainDate (d : Std.Time.PlainDate) : MacroM (TSyntax `term) :
126126
`(Std.Time.PlainDate.ofYearMonthDayClip $(← syntaxInt d.year) $(← syntaxBounded d.month.val) $(← syntaxBounded d.day.val))
127127

128128
private def convertPlainTime (d : Std.Time.PlainTime) : MacroM (TSyntax `term) := do
129-
`(Std.Time.PlainTime.mk $(← syntaxBounded d.hour.val) $(← syntaxBounded d.minute.val) true, $(← syntaxBounded d.second.snd.val) $(← syntaxBounded d.nanosecond.val))
129+
`(Std.Time.PlainTime.mk $(← syntaxBounded d.hour.val) $(← syntaxBounded d.minute.val) $(← syntaxBounded d.second.val) $(← syntaxBounded d.nanosecond.val))
130130

131131
private def convertPlainDateTime (d : Std.Time.PlainDateTime) : MacroM (TSyntax `term) := do
132132
`(Std.Time.PlainDateTime.mk $(← convertPlainDate d.date) $(← convertPlainTime d.time))

src/Std/Time/Time/PlainTime.lean

Lines changed: 16 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -30,7 +30,7 @@ structure PlainTime where
3030
/--
3131
`Second` component of the `PlainTime`
3232
-/
33-
second : Sigma Second.Ordinal
33+
second : Second.Ordinal true
3434

3535
/--
3636
`Nanoseconds` component of the `PlainTime`
@@ -39,32 +39,32 @@ structure PlainTime where
3939
deriving Repr
4040

4141
instance : Inhabited PlainTime where
42-
default := ⟨0, 0, Sigma.mk false 0, 0, by decide⟩
42+
default := ⟨0, 0, 0, 0, by decide⟩
4343

4444
instance : BEq PlainTime where
4545
beq x y := x.hour.val == y.hour.val && x.minute == y.minute
46-
&& x.second.snd.val == y.second.snd.val && x.nanosecond == y.nanosecond
46+
&& x.second.val == y.second.val && x.nanosecond == y.nanosecond
4747

4848
namespace PlainTime
4949

5050
/--
5151
Creates a `PlainTime` value representing midnight (00:00:00.000000000).
5252
-/
5353
def midnight : PlainTime :=
54-
0, 0, true, 0, 0
54+
0, 0, 0, 0
5555

5656
/--
5757
Creates a `PlainTime` value from the provided hours, minutes, seconds and nanoseconds components.
5858
-/
5959
@[inline]
60-
def ofHourMinuteSecondsNano (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (nano : Nanosecond.Ordinal) : PlainTime :=
61-
⟨hour, minute, Sigma.mk leap second, nano⟩
60+
def ofHourMinuteSecondsNano (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal true) (nano : Nanosecond.Ordinal) : PlainTime :=
61+
⟨hour, minute, second, nano⟩
6262

6363
/--
6464
Creates a `PlainTime` value from the provided hours, minutes, and seconds.
6565
-/
6666
@[inline]
67-
def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : PlainTime :=
67+
def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal true) : PlainTime :=
6868
ofHourMinuteSecondsNano hour minute second 0
6969

7070
/--
@@ -73,7 +73,7 @@ Converts a `PlainTime` value to the total number of milliseconds.
7373
def toMilliseconds (time : PlainTime) : Millisecond.Offset :=
7474
time.hour.toOffset.toMilliseconds +
7575
time.minute.toOffset.toMilliseconds +
76-
time.second.snd.toOffset.toMilliseconds +
76+
time.second.toOffset.toMilliseconds +
7777
time.nanosecond.toOffset.toMilliseconds
7878

7979
/--
@@ -82,7 +82,7 @@ Converts a `PlainTime` value to the total number of nanoseconds.
8282
def toNanoseconds (time : PlainTime) : Nanosecond.Offset :=
8383
time.hour.toOffset.toNanoseconds +
8484
time.minute.toOffset.toNanoseconds +
85-
time.second.snd.toOffset.toNanoseconds +
85+
time.second.toOffset.toNanoseconds +
8686
time.nanosecond.toOffset
8787

8888
/--
@@ -91,15 +91,15 @@ Converts a `PlainTime` value to the total number of seconds.
9191
def toSeconds (time : PlainTime) : Second.Offset :=
9292
time.hour.toOffset.toSeconds +
9393
time.minute.toOffset.toSeconds +
94-
time.second.snd.toOffset
94+
time.second.toOffset
9595

9696
/--
9797
Converts a `PlainTime` value to the total number of minutes.
9898
-/
9999
def toMinutes (time : PlainTime) : Minute.Offset :=
100100
time.hour.toOffset.toMinutes +
101101
time.minute.toOffset +
102-
time.second.snd.toOffset.toMinutes
102+
time.second.toOffset.toMinutes
103103

104104
/--
105105
Converts a `PlainTime` value to the total number of hours.
@@ -115,9 +115,12 @@ def ofNanoseconds (nanos : Nanosecond.Offset) : PlainTime :=
115115
have remainingNanos := Bounded.LE.byEmod nanos.val 1000000000 (by decide)
116116
have hours := Bounded.LE.byEmod (totalSeconds.val / 3600) 24 (by decide)
117117
have minutes := (Bounded.LE.byEmod totalSeconds.val 3600 (by decide)).ediv 60 (by decide)
118+
118119
have seconds := Bounded.LE.byEmod totalSeconds.val 60 (by decide)
120+
have seconds := seconds.expandTop (by decide)
121+
119122
let nanos := Bounded.LE.byEmod nanos.val 1000000000 (by decide)
120-
PlainTime.mk hours minutes (Sigma.mk false seconds) nanos
123+
PlainTime.mk hours minutes seconds nanos
121124

122125
/--
123126
Creates a `PlainTime` value from a total number of millisecond.
@@ -222,7 +225,7 @@ def subMilliseconds (time : PlainTime) (millisToSub : Millisecond.Offset) : Plai
222225
Creates a new `PlainTime` by adjusting the `second` component to the given value.
223226
-/
224227
@[inline]
225-
def withSeconds (pt : PlainTime) (second : Sigma Second.Ordinal) : PlainTime :=
228+
def withSeconds (pt : PlainTime) (second : Second.Ordinal true) : PlainTime :=
226229
{ pt with second := second }
227230

228231
/--

src/Std/Time/Time/Unit/Second.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -30,7 +30,7 @@ instance : LE (Ordinal leap) where
3030
instance : LT (Ordinal leap) where
3131
lt x y := LT.lt x.val y.val
3232

33-
instance : Repr (Ordinal l) where
33+
instance : Repr (Ordinal leap) where
3434
reprPrec r := reprPrec r.val
3535

3636
instance : OfNat (Ordinal leap) n := by
@@ -39,10 +39,10 @@ instance : OfNat (Ordinal leap) n := by
3939
· exact inst
4040
· exact ⟨inst.ofNat.expandTop (by decide)⟩
4141

42-
instance {x y : Ordinal l} : Decidable (x ≤ y) :=
42+
instance {x y : Ordinal leap} : Decidable (x ≤ y) :=
4343
inferInstanceAs (Decidable (x.val ≤ y.val))
4444

45-
instance {x y : Ordinal l} : Decidable (x < y) :=
45+
instance {x y : Ordinal leap} : Decidable (x < y) :=
4646
inferInstanceAs (Decidable (x.val < y.val))
4747

4848
/--

src/Std/Time/Zoned/Database/TZdb.lean

Lines changed: 1 addition & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -71,12 +71,7 @@ def idFromPath (path : System.FilePath) : Option String := do
7171
Retrieves the timezone rules from the local timezone data file.
7272
-/
7373
def localRules (path : System.FilePath) : IO ZoneRules := do
74-
let localTimePath ←
75-
try
76-
IO.Process.run { cmd := "readlink", args := #["-f", path.toString] }
77-
catch _ =>
78-
throw <| IO.userError "cannot find the local timezone database"
79-
74+
let localTimePath ← IO.FS.realPath path
8075
if let some id := idFromPath localTimePath
8176
then parseTZIfFromDisk path id
8277
else throw (IO.userError "cannot read the id of the path.")

src/Std/Time/Zoned/DateTime.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -305,7 +305,7 @@ def withMinutes (dt : DateTime tz) (minute : Minute.Ordinal) : DateTime tz :=
305305
Creates a new `DateTime tz` by adjusting the `second` component.
306306
-/
307307
@[inline]
308-
def withSeconds (dt : DateTime tz) (second : Sigma Second.Ordinal) : DateTime tz :=
308+
def withSeconds (dt : DateTime tz) (second : Second.Ordinal true) : DateTime tz :=
309309
ofPlainDateTime (dt.date.get.withSeconds second) tz
310310

311311
/--
@@ -368,7 +368,7 @@ def minute (dt : DateTime tz) : Minute.Ordinal :=
368368
Getter for the `Second` inside of a `DateTime`
369369
-/
370370
@[inline]
371-
def second (dt : DateTime tz) : Second.Ordinal dt.date.get.time.second.fst :=
371+
def second (dt : DateTime tz) : Second.Ordinal true :=
372372
dt.date.get.second
373373

374374
/--

src/Std/Time/Zoned/ZonedDateTime.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -179,8 +179,8 @@ def minute (zdt : ZonedDateTime) : Minute.Ordinal :=
179179
Getter for the `Second` inside of a `ZonedDateTime`
180180
-/
181181
@[inline]
182-
def second (zdt : ZonedDateTime) : Second.Ordinal zdt.date.get.time.second.fst :=
183-
zdt.date.get.time.second.snd
182+
def second (zdt : ZonedDateTime) : Second.Ordinal true :=
183+
zdt.date.get.time.second
184184

185185
/--
186186
Getter for the `Millisecond` inside of a `ZonedDateTime`.
@@ -491,7 +491,7 @@ def withMinutes (dt : ZonedDateTime) (minute : Minute.Ordinal) : ZonedDateTime :
491491
Creates a new `ZonedDateTime` by adjusting the `second` component.
492492
-/
493493
@[inline]
494-
def withSeconds (dt : ZonedDateTime) (second : Sigma Second.Ordinal) : ZonedDateTime :=
494+
def withSeconds (dt : ZonedDateTime) (second : Second.Ordinal true) : ZonedDateTime :=
495495
let date := dt.date.get
496496
ZonedDateTime.ofPlainDateTime (date.withSeconds second) dt.rules
497497

tests/lean/run/timeAPI.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -587,7 +587,7 @@ Std.Time.Weekday.friday
587587
2023-06-09
588588
2023-06-09
589589
19517
590-
1686268800000
590+
1686268800
591591
1970-01-02
592592
593593
-/
@@ -667,7 +667,7 @@ Std.Time.Weekday.tuesday
667667
12
668668
3
669669
9938
670-
858650584000
670+
858650584
671671
1970-01-02T00:00:00.000000000
672672
673673
-/
@@ -741,7 +741,7 @@ Std.Time.Weekday.thursday
741741
37
742742
2
743743
19978
744-
1726117262000
744+
1726117262
745745
1970-01-02T00:00:00.000000000Z
746746
747747
-/
@@ -816,7 +816,7 @@ Std.Time.Weekday.tuesday
816816
12
817817
3
818818
9938
819-
858661384000
819+
858661384
820820
821821
-/
822822
#guard_msgs in

tests/lean/run/timeFormats.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -277,7 +277,7 @@ Format
277277

278278
def time₄ := time("23:13:12.324354679")
279279
def date₄ := date("2002-07-14")
280-
def datetime₅ := PlainDateTime.mk (PlainDate.ofYearMonthDayClip (-2000) 3 4) (PlainTime.mk 12 23 false, 12 0)
280+
def datetime₅ := PlainDateTime.mk (PlainDate.ofYearMonthDayClip (-2000) 3 4) (PlainTime.mk 12 23 12 0)
281281
def datetime₄ := datetime("2002-07-14T23:13:12.324354679")
282282
def zoned₄ := zoned("2002-07-14T23:13:12.324354679+09:00")
283283
def zoned₅ := zoned("2002-07-14T23:13:12.324354679+00:00")
@@ -806,7 +806,7 @@ info: ("19343232432-01-04T01:04:03.000000000",
806806
-/
807807
#guard_msgs in
808808
#eval
809-
let r := (PlainDateTime.mk (PlainDate.ofYearMonthDayClip 19343232432 1 4) (PlainTime.mk 25 64 true, 30))
809+
let r := PlainDateTime.mk (PlainDate.ofYearMonthDayClip 19343232432 1 4) (PlainTime.mk 25 64 3 0)
810810
let s := r.toLeanDateTimeString
811811
let r := PlainDateTime.parse s
812812
(s, r, datetime("1932-01-02T05:04:03.000000000"))

0 commit comments

Comments
 (0)