From 9c1be6cdf377f6c82cf044dbbd708ebc1ba2c1e1 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 2 Aug 2024 18:28:34 -0300 Subject: [PATCH 001/118] feat: add datetime library --- src/Lean/Time/Bounded.lean | 361 +++++++++++++++++++ src/Lean/Time/Date.lean | 9 + src/Lean/Time/Date/Basic.lean | 7 + src/Lean/Time/Date/Date.lean | 139 ++++++++ src/Lean/Time/Date/Scalar.lean | 65 ++++ src/Lean/Time/Date/Unit/Basic.lean | 21 ++ src/Lean/Time/Date/Unit/Day.lean | 88 +++++ src/Lean/Time/Date/Unit/Lemmas.lean | 7 + src/Lean/Time/Date/Unit/Month.lean | 186 ++++++++++ src/Lean/Time/Date/Unit/WeekOfYear.lean | 63 ++++ src/Lean/Time/Date/Unit/Weekday.lean | 109 ++++++ src/Lean/Time/Date/Unit/Year.lean | 63 ++++ src/Lean/Time/Date/WeekDate.lean | 42 +++ src/Lean/Time/DateTime.lean | 25 ++ src/Lean/Time/DateTime/Basic.lean | 8 + src/Lean/Time/DateTime/NaiveDateTime.lean | 151 ++++++++ src/Lean/Time/DateTime/Timestamp.lean | 25 ++ src/Lean/Time/Duration.lean | 73 ++++ src/Lean/Time/Format.lean | 412 ++++++++++++++++++++++ src/Lean/Time/LessEq.lean | 27 ++ src/Lean/Time/Sign.lean | 30 ++ src/Lean/Time/Time.lean | 10 + src/Lean/Time/Time/Basic.lean | 7 + src/Lean/Time/Time/HourMarker.lean | 34 ++ src/Lean/Time/Time/Scalar.lean | 28 ++ src/Lean/Time/Time/Time.lean | 29 ++ src/Lean/Time/Time/Unit/Basic.lean | 40 +++ src/Lean/Time/Time/Unit/Hour.lean | 81 +++++ src/Lean/Time/Time/Unit/Minute.lean | 73 ++++ src/Lean/Time/Time/Unit/Nanosecond.lean | 45 +++ src/Lean/Time/Time/Unit/Second.lean | 72 ++++ src/Lean/Time/TimeZone.lean | 9 + src/Lean/Time/TimeZone/Basic.lean | 6 + src/Lean/Time/TimeZone/DateTime.lean | 88 +++++ src/Lean/Time/TimeZone/Offset.lean | 52 +++ src/Lean/Time/TimeZone/TimeZone.lean | 51 +++ src/Lean/Time/TimeZone/ZonedDateTime.lean | 88 +++++ src/Lean/Time/UnitVal.lean | 91 +++++ 38 files changed, 2715 insertions(+) create mode 100644 src/Lean/Time/Bounded.lean create mode 100644 src/Lean/Time/Date.lean create mode 100644 src/Lean/Time/Date/Basic.lean create mode 100644 src/Lean/Time/Date/Date.lean create mode 100644 src/Lean/Time/Date/Scalar.lean create mode 100644 src/Lean/Time/Date/Unit/Basic.lean create mode 100644 src/Lean/Time/Date/Unit/Day.lean create mode 100644 src/Lean/Time/Date/Unit/Lemmas.lean create mode 100644 src/Lean/Time/Date/Unit/Month.lean create mode 100644 src/Lean/Time/Date/Unit/WeekOfYear.lean create mode 100644 src/Lean/Time/Date/Unit/Weekday.lean create mode 100644 src/Lean/Time/Date/Unit/Year.lean create mode 100644 src/Lean/Time/Date/WeekDate.lean create mode 100644 src/Lean/Time/DateTime.lean create mode 100644 src/Lean/Time/DateTime/Basic.lean create mode 100644 src/Lean/Time/DateTime/NaiveDateTime.lean create mode 100644 src/Lean/Time/DateTime/Timestamp.lean create mode 100644 src/Lean/Time/Duration.lean create mode 100644 src/Lean/Time/Format.lean create mode 100644 src/Lean/Time/LessEq.lean create mode 100644 src/Lean/Time/Sign.lean create mode 100644 src/Lean/Time/Time.lean create mode 100644 src/Lean/Time/Time/Basic.lean create mode 100644 src/Lean/Time/Time/HourMarker.lean create mode 100644 src/Lean/Time/Time/Scalar.lean create mode 100644 src/Lean/Time/Time/Time.lean create mode 100644 src/Lean/Time/Time/Unit/Basic.lean create mode 100644 src/Lean/Time/Time/Unit/Hour.lean create mode 100644 src/Lean/Time/Time/Unit/Minute.lean create mode 100644 src/Lean/Time/Time/Unit/Nanosecond.lean create mode 100644 src/Lean/Time/Time/Unit/Second.lean create mode 100644 src/Lean/Time/TimeZone.lean create mode 100644 src/Lean/Time/TimeZone/Basic.lean create mode 100644 src/Lean/Time/TimeZone/DateTime.lean create mode 100644 src/Lean/Time/TimeZone/Offset.lean create mode 100644 src/Lean/Time/TimeZone/TimeZone.lean create mode 100644 src/Lean/Time/TimeZone/ZonedDateTime.lean create mode 100644 src/Lean/Time/UnitVal.lean diff --git a/src/Lean/Time/Bounded.lean b/src/Lean/Time/Bounded.lean new file mode 100644 index 000000000000..80c3d9dcc4d4 --- /dev/null +++ b/src/Lean/Time/Bounded.lean @@ -0,0 +1,361 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Init.Data.Int +import Lean.Time.LessEq + +namespace Lean +set_option linter.all true in + +/-- +A `Bounded` is represented by an `Int` that is contrained by a lower and higher bounded using some +relation `rel`. It includes all the integers that `rel lo val ∧ rel val hi`. +-/ +def Bounded (rel : Int → Int → Prop) (lo : Int) (hi : Int) := { val : Int // rel lo val ∧ rel val hi} + +namespace Bounded + +@[always_inline] +instance : LE (Bounded rel n m) where + le l r := l.val ≤ r.val + +@[always_inline] +instance : LT (Bounded rel n m) where + lt l r := l.val < r.val + +@[always_inline] +instance : Repr (Bounded rel m n) where + reprPrec n := reprPrec n.val + +@[always_inline] +instance : BEq (Bounded rel n m) where + beq x y := (x.val = y.val) + +@[always_inline] +instance {x y : Bounded rel a b} : Decidable (x ≤ y) := + dite (x.val ≤ y.val) isTrue isFalse + +/-- +A `Bounded` integer that the relation used is the the less-equal relation so, it includes all +integers that `lo ≤ val ≤ hi`. +-/ +abbrev LE := @Bounded LE.le + +instance [Le lo n] [Le n hi] : OfNat (Bounded.LE lo hi) n where + ofNat := ⟨n, And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)⟩ + +instance [Le lo hi] : Inhabited (Bounded.LE lo hi) where + default := ⟨lo, And.intro (Int.le_refl lo) (Int.ofNat_le.mpr Le.p)⟩ + +/-- +A `Bounded` integer that the relation used is the the less-than relation so, it includes all +integers that `lo < val < hi`. +-/ +abbrev LT := @Bounded LT.lt + +/-- +Creates a new `Bounded` Integer. +-/ +@[inline] +def mk {rel : Int → Int → Prop} (val : Int) (proof : rel lo val ∧ rel val hi) : @Bounded rel lo hi := + ⟨val, proof⟩ + +namespace LE + +/-- +Creates a new `Bounded` integer that the relation is less-equal. +-/ +@[inline] +def mk (val : Int) (proof : lo ≤ val ∧ val ≤ hi) : Bounded.LE lo hi := + ⟨val, proof⟩ + +/-- +Creates a new `Bounded` integer. +-/ +@[inline] +def ofInt { lo hi : Int } (val : Int) : Option (Bounded.LE lo hi) := + if h : lo ≤ val ∧ val ≤ hi + then some ⟨val, h⟩ + else none + +/-- +Convert a `Nat` to a `Bounded.LE`. +-/ +@[inline] +def ofNat (val : Nat) (h : val ≤ hi) : Bounded.LE 0 hi := + Bounded.mk val (And.intro (Int.ofNat_zero_le val) (Int.ofNat_le.mpr h)) + +/-- +Convert a `Nat` to a `Bounded.LE` if it checks. +-/ +@[inline] +def ofNat? { hi : Nat } (val : Nat) : Option (Bounded.LE 0 hi) := + if h : val ≤ hi then + ofNat val h + else + none + +/-- +Convert a `Nat` to a `Bounded.LE` using the lower boundary too. +-/ +@[inline] +def ofNat' (val : Nat) (h : lo ≤ val ∧ val ≤ hi) : Bounded.LE lo hi := + Bounded.mk val (And.intro (Int.ofNat_le.mpr h.left) (Int.ofNat_le.mpr h.right)) + +/-- +Convert a `Nat` to a `Bounded.LE` using the lower boundary too. +-/ +@[inline] +def force (val : Int) (h : lo ≤ hi) : Bounded.LE lo hi := + if h₀ : lo ≤ val then + if h₁ : val ≤ hi + then ⟨val, And.intro h₀ h₁⟩ + else ⟨hi, And.intro h (Int.le_refl hi)⟩ + else ⟨lo, And.intro (Int.le_refl lo) h⟩ + +/-- +Convert a `Nat` to a `Bounded.LE` using the lower boundary too. +-/ +@[inline] +def force! [Le lo hi] (val : Int) : Bounded.LE lo hi := + if h₀ : lo ≤ val then + if h₁ : val ≤ hi + then ⟨val, And.intro h₀ h₁⟩ + else panic! "greater than hi" + else panic! "lower than lo" + +/-- +Convert a `Bounded.LE` to a Nat. +-/ +@[inline] +def toNat (n : Bounded.LE lo hi) : Nat := + n.val.toNat + +/-- +Convert a `Bounded.LE` to a Nat. +-/ +@[inline] +def toNat' (n : Bounded.LE lo hi) (h : lo ≥ 0) : Nat := + let h₁ := (Int.le_trans h n.property.left) + match n.val, h₁ with + | .ofNat n, _ => n + | .negSucc _, h => nomatch h + +/-- +Convert a `Bounded.LE` to an Int. +-/ +@[inline] +def toInt (n : Bounded.LE lo hi) : Int := + n.val + +/-- +Convert a `Bounded.LE` to a `Fin`. +-/ +@[inline] +def toFin (n : Bounded.LE lo hi) (h₀ : 0 ≤ lo) (h₁ : lo < hi) : Fin (hi + 1).toNat := by + let h := n.property.right + let h₁ := Int.le_trans h₀ n.property.left + refine ⟨n.val.toNat, (Int.toNat_lt h₁).mpr ?_⟩ + rw [Int.toNat_of_nonneg (by omega)] + exact Int.lt_add_one_of_le h + +/-- +Convert a `Fin` to a `Bounded.LE`. +-/ +@[inline] +def ofFin (fin : Fin (Nat.succ hi)) : Bounded.LE 0 hi := + ofNat fin.val (Nat.le_of_lt_succ fin.isLt) + +/-- +Convert a `Fin` to a `Bounded.LE`. +-/ +@[inline] +def ofFin' {lo : Nat} (fin : Fin (Nat.succ hi)) (h : lo ≤ hi) : Bounded.LE lo hi := + if h₁ : fin.val ≥ lo + then ofNat' fin.val (And.intro h₁ ((Nat.le_of_lt_succ fin.isLt))) + else ofNat' lo (And.intro (Nat.le_refl lo) h) + +/-- +Creates a new `Bounded.LE` using a the modulus of a number. +-/ +@[inline] +def byEmod (b : Int) (i : Int) (hi : i > 0) : Bounded.LE 0 (i - 1) := by + refine ⟨b % i, And.intro ?_ ?_⟩ + · apply Int.emod_nonneg b + intro a + simp_all [Int.lt_irrefl] + · apply Int.le_of_lt_add_one + simp [Int.add_sub_assoc] + exact Int.emod_lt_of_pos b hi + +/-- +Creates a new `Bounded.LE` using a the Truncating modulus of a number. +-/ +@[inline] +def byMod (b : Int) (i : Int) (hi : 0 < i) : Bounded.LE (- (i - 1)) (i - 1) := by + refine ⟨b.mod i, And.intro ?_ ?_⟩ + · simp [Int.mod] + split <;> try contradiction + next m n => + let h := Int.emod_nonneg (a := m) (b := n) (Int.ne_of_gt hi) + apply (Int.le_trans · h) + apply Int.le_of_neg_le_neg + simp_all + exact (Int.le_sub_one_of_lt hi) + next m n => + apply Int.neg_le_neg + have h := Int.mod_lt_of_pos (m + 1) hi + exact Int.le_sub_one_of_lt h + · exact Int.le_sub_one_of_lt (Int.mod_lt_of_pos b hi) + +/-- +Adjust the bounds of a `Bounded` by setting the lower bound to zero and the maximum value to (m - n). +-/ +@[inline] +def truncate (bounded : Bounded.LE n m) : Bounded.LE 0 (m - n) := by + let ⟨left, right⟩ := bounded.property + refine ⟨bounded.val - n, And.intro ?_ ?_⟩ + all_goals omega + +/-- +Adjust the bounds of a `Bounded` by changing the higher bound if another value `j` satisfies the same +constraint. +-/ +@[inline] +def truncateTop (bounded : Bounded.LE n m) (h : bounded.val ≤ j) : Bounded.LE n j := by + refine ⟨bounded.val, And.intro ?_ ?_⟩ + · exact bounded.property.left + · exact h + +/-- +Adjust the bounds of a `Bounded` by changing the lower bound if another value `j` satisfies the same +constraint. +-/ +@[inline] +def truncateBottom (bounded : Bounded.LE n m) (h : bounded.val ≥ j) : Bounded.LE j m := by + refine ⟨bounded.val, And.intro ?_ ?_⟩ + · exact h + · exact bounded.property.right + +/-- +Adjust the bounds of a `Bounded` by adding a constant value to both the lower and upper bounds. +-/ +@[inline] +def neg (bounded : Bounded.LE n m) : Bounded.LE (-m) (-n) := by + refine ⟨-bounded.val, And.intro ?_ ?_⟩ + · exact Int.neg_le_neg bounded.property.right + · exact Int.neg_le_neg bounded.property.left + +/-- +Adjust the bounds of a `Bounded` by adding a constant value to both the lower and upper bounds. +-/ +@[inline] +def add (bounded : Bounded.LE n m) (num : Int) : Bounded.LE (n + num) (m + num) := by + refine ⟨bounded.val + num, And.intro ?_ ?_⟩ + all_goals apply (Int.add_le_add · (Int.le_refl num)) + · exact bounded.property.left + · exact bounded.property.right + +/-- +Adjust the bounds of a `Bounded` by adding a constant value to the upper bounds. +-/ +@[inline] +def addTop (bounded : Bounded.LE n m) (num : Nat) : Bounded.LE n (m + num) := by + refine ⟨bounded.val + num, And.intro ?_ ?_⟩ + · let h := Int.add_le_add bounded.property.left (Int.ofNat_zero_le num) + simp at h + exact h + · exact Int.add_le_add bounded.property.right (Int.le_refl num) + +/-- +Adjust the bounds of a `Bounded` by adding a constant value to the lower bounds. +-/ +@[inline] +def subBottom (bounded : Bounded.LE n m) (num : Nat) : Bounded.LE (n - num) m := by + refine ⟨bounded.val - num, And.intro ?_ ?_⟩ + · exact Int.add_le_add bounded.property.left (Int.le_refl (-num)) + · let h := Int.sub_le_sub bounded.property.right (Int.ofNat_zero_le num) + simp at h + exact h + +/-- +Adds two `Bounded` and adjust the boundaries. +-/ +@[inline] +def addBounds (bounded : Bounded.LE n m) (bounded₂ : Bounded.LE n m) : Bounded.LE (n + n) (m + m) := by + refine ⟨bounded.val + bounded₂.val, And.intro ?_ ?_⟩ + · exact Int.add_le_add bounded.property.left bounded₂.property.left + · exact Int.add_le_add bounded.property.right bounded₂.property.right + +/-- +Adjust the bounds of a `Bounded` by subtracting a constant value to both the lower and upper bounds. +-/ +@[inline] +def sub (bounded : Bounded.LE n m) (num : Int) : Bounded.LE (n - num) (m - num) := + add bounded (-num) + +/-- +Adds two `Bounded` and adjust the boundaries. +-/ +@[inline] +def subBounds (bounded : Bounded.LE n m) (bounded₂ : Bounded.LE n m) : Bounded.LE (n - m) (m - n) := by + refine ⟨bounded.val - bounded₂.val, And.intro ?_ ?_⟩ + · exact Int.sub_le_sub bounded.property.left bounded₂.property.right + · exact Int.sub_le_sub bounded.property.right bounded₂.property.left + +/-- +Adjust the bounds of a `Bounded` by applying the emod operation constraining the lower bound to 0 and +the upper bound to the value. +-/ +@[inline] +def emod (bounded : Bounded.LE n num) (num : Int) (hi : 0 < num) : Bounded.LE 0 (num - 1) := + byEmod bounded.val num hi + +/-- +Adjust the bounds of a `Bounded` by applying the mod operation. +-/ +@[inline] +def mod (bounded : Bounded.LE n num) (num : Int) (hi : 0 < num) : Bounded.LE (- (num - 1)) (num - 1) := + byMod bounded.val num hi + +/-- +Adjust the bounds of a `Bounded` by applying the div operation. +-/ +@[inline] +def div (bounded : Bounded.LE n m) (num : Int) (h : num > 0) : Bounded.LE (n / num) (m / num) := by + let ⟨left, right⟩ := bounded.property + refine ⟨bounded.val / num, And.intro ?_ ?_⟩ + apply Int.ediv_le_ediv + · exact h + · exact left + · apply Int.ediv_le_ediv + · exact h + · exact right + +/-- +Expand the bottom of the bounded to a number `nhi` is `hi` is less or equal to the previous higher bound. +-/ +@[inline] +def expandTop (bounded : Bounded.LE lo hi) (h : hi ≤ nhi) : Bounded.LE lo nhi := + ⟨bounded.val, And.intro bounded.property.left (Int.le_trans bounded.property.right h)⟩ + +/-- +Expand the bottom of the bounded to a number `nlo` if `lo` is greater or equal to the previous lower bound. +-/ +@[inline] +def expandBottom (bounded : Bounded.LE lo hi) (h : nlo ≤ lo) : Bounded.LE nlo hi := + ⟨bounded.val, And.intro (Int.le_trans h bounded.property.left) bounded.property.right⟩ + +/-- +Adds one to the value of the bounded if the value is less than the higher bound of the bounded number. +-/ +@[inline] +def succ (bounded : Bounded.LE lo hi) (h : bounded.val < hi) : Bounded.LE lo hi := + let left := bounded.property.left + ⟨bounded.val + 1, And.intro (by omega) (by omega)⟩ + +end LE +end Bounded diff --git a/src/Lean/Time/Date.lean b/src/Lean/Time/Date.lean new file mode 100644 index 000000000000..e298cfdb95c6 --- /dev/null +++ b/src/Lean/Time/Date.lean @@ -0,0 +1,9 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Date.Basic +import Lean.Time.Date.Scalar +import Lean.Time.Date.Date diff --git a/src/Lean/Time/Date/Basic.lean b/src/Lean/Time/Date/Basic.lean new file mode 100644 index 000000000000..4f10e9df36ae --- /dev/null +++ b/src/Lean/Time/Date/Basic.lean @@ -0,0 +1,7 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Date.Unit.Basic diff --git a/src/Lean/Time/Date/Date.lean b/src/Lean/Time/Date/Date.lean new file mode 100644 index 000000000000..4449e466128c --- /dev/null +++ b/src/Lean/Time/Date/Date.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Date.Basic +import Lean.Time.Date.Scalar + +namespace Lean +namespace Date + +/-- +Date in YMD format. +-/ +structure Date where + year : Year.Offset + month : Month.Ordinal + day : Day.Ordinal + valid : year.valid month day + deriving Repr + +namespace Date + +/-- +Force the date to be valid. +-/ +def force (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Date := + let ⟨day, valid⟩ := month.forceDay year.isLeap day + Date.mk year month day valid + +/-- +Creates a new `Date` using YMD. +-/ +def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option Date := + if valid : year.valid month day + then some (Date.mk year month day valid) + else none + +/-- +Creates a new `Date` using YO. +-/ +def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : Date := + let ⟨⟨month, day⟩, valid⟩ := ordinal.toMonthAndDay + Date.mk year month day valid + +/-- +Creates a new `Date` using the `Day.Offset` which `0` corresponds the UNIX Epoch. +-/ +def ofDaysSinceUNIXEpoch (day : Day.Offset) : Date := + let z := day.toInt + 719468 + let era := (if z ≥ 0 then z else z - 146096) / 146097 + let doe := z - era * 146097 + let yoe := (doe - doe / 1460 + doe / 36524 - doe / 146096) / 365 + let y := yoe + era * 400 + let doy := doe - (365 * yoe + yoe / 4 - yoe / 100) + let mp : Int := (5 * doy + 2) / 153 + let d := doy - (153 * mp + 2) / 5 + let m := mp + (if mp < 10 then 3 else -9) + let y := y + (if m <= 2 then 1 else 0) + + .force y (.force m (by decide)) (.force (d + 1) (by decide)) + +/-- +Returns the `Weekday` of a `Date`. +-/ +def weekday (date : Date) : Weekday := + let q := date.day.toInt + let m := date.month.toInt + let y := date.year.toInt + + let y := if m < 2 then y - 1 else y + let m := if m < 2 then m + 12 else m + + let k := y % 100 + let j := y / 100 + let part := q + (13 * (m + 1)) / 5 + k + (k / 4) + let h := if y ≥ 1582 then part + (j/4) - 2*j else part + 5 - j + let d := (h + 5) % 7 + + .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ + +/-- +Returns the `Weekday` of a `Date` using Zeller's Congruence for the Julian calendar. +-/ +def weekdayJulian (date : Date) : Weekday := + let month := date.month.toInt + let year := date.year.toInt + + let q := date.day.toInt + let m := if month < 3 then month + 12 else month + let y := if month < 3 then year - 1 else year + + let k := y % 100 + let j := y / 100 + + let h := (q + (13 * (m + 1)) / 5 + k + (k / 4) + 5 - j) % 7 + let d := (h + 5 - 1) % 7 + + .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ + +/-- +Convert `Date` to `Day.Offset` since the UNIX Epoch. +-/ +def toDaysSinceUNIXEpoch (date : Date) : Day.Offset := + let y : Int := if date.month.toInt > 2 then date.year else date.year.toInt - 1 + let era : Int := (if y ≥ 0 then y else y - 399) / 400 + let yoe : Int := y - era * 400 + let m : Int := date.month.toInt + let d : Int := date.day.toInt + let doy := (153 * (m + (if m > 2 then -3 else 9)) + 2) / 5 + d - 1 + let doe := yoe * 365 + yoe / 4 - yoe / 100 + doy + + .ofInt (era * 146097 + doe - 719468) + +/-- +Returns the `Scalar` starting from the UNIX epoch. +-/ +def toScalar (date : Date) : Scalar := + ⟨toDaysSinceUNIXEpoch date⟩ + +/-- +Creates a new `Date` from a `Scalar` that begins on the epoch. +-/ +def ofScalar (scalar : Scalar) : Date := + ofDaysSinceUNIXEpoch scalar.day + +/-- +Calculate the Year.Offset from a Date +-/ +def yearsSince (date : Date) (year : Year.Offset) : Year.Offset := + date.year - year + +instance : HAdd Date Day.Offset Date where + hAdd date day := ofScalar (toScalar date + ⟨day⟩) + +instance : HAdd Date Scalar Date where + hAdd date day := ofScalar (toScalar date + day) diff --git a/src/Lean/Time/Date/Scalar.lean b/src/Lean/Time/Date/Scalar.lean new file mode 100644 index 000000000000..1701e31565e3 --- /dev/null +++ b/src/Lean/Time/Date/Scalar.lean @@ -0,0 +1,65 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Date.Basic + +namespace Lean +namespace Date + +/-- +`Scalar` represents a date offset, using the number of day as the underlying unit. +-/ +structure Scalar where + day : Day.Offset + deriving Repr, BEq, Inhabited + +instance : Add Scalar where add x y := ⟨x.day + y.day⟩ +instance : Sub Scalar where sub x y := ⟨x.day - y.day⟩ +instance : Mul Scalar where mul x y := ⟨x.day * y.day⟩ +instance : Div Scalar where div x y := ⟨x.day / y.day⟩ +instance : Neg Scalar where neg x := ⟨-x.day⟩ + +namespace Scalar + +/-- +Creates a `Scalar` from a given number of day. +-/ +def ofDays (day : Int) : Scalar := + ⟨UnitVal.ofInt day⟩ + +/-- +Retrieves the number of day from a `Scalar`. +-/ +def toDays (scalar : Scalar) : Int := + scalar.day.val + +/-- +Adds a specified number of day to the `Scalar`, returning a new `Scalar`. +-/ +def addDays (scalar : Scalar) (day : Day.Offset) : Scalar := + ⟨scalar.day + day⟩ + +/-- +Subtracts a specified number of day from the `Scalar`, returning a new `Scalar`. +-/ +def subDays (scalar : Scalar) (day : Day.Offset) : Scalar := + ⟨scalar.day - day⟩ + +/-- +Converts a `Scalar` to a `Day.Offset`. +-/ +def toOffset (scalar : Scalar) : Day.Offset := + scalar.day + +/-- +Creates a `Scalar` from a `Day.Offset`. +-/ +def ofOffset (offset : Day.Offset) : Scalar := + ⟨offset⟩ + +end Scalar +end Date diff --git a/src/Lean/Time/Date/Unit/Basic.lean b/src/Lean/Time/Date/Unit/Basic.lean new file mode 100644 index 000000000000..a80b6b75eaf8 --- /dev/null +++ b/src/Lean/Time/Date/Unit/Basic.lean @@ -0,0 +1,21 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Date.Unit.Day +import Lean.Time.Date.Unit.Month +import Lean.Time.Date.Unit.Year +import Lean.Time.Date.Unit.Weekday +import Lean.Time.Date.Unit.WeekOfYear + +namespace Lean +namespace Date.Day.Ordinal.OfYear + +@[inline] +def toMonthAndDay (year : Year.Offset) (ordinal : OfYear year.isLeap) : { val : Month.Ordinal × Ordinal // Year.Offset.valid year (Prod.fst val) (Prod.snd val) } := + Month.Ordinal.ofOrdinal ordinal + +end Date.Day.Ordinal.OfYear diff --git a/src/Lean/Time/Date/Unit/Day.lean b/src/Lean/Time/Date/Unit/Day.lean new file mode 100644 index 000000000000..a93a186a654e --- /dev/null +++ b/src/Lean/Time/Date/Unit/Day.lean @@ -0,0 +1,88 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat + +namespace Lean +namespace Date +open Lean Time + +set_option linter.all true + +namespace Day + +/-- +`Ordinal` represents a bounded value for days, which ranges between 0 and 31. +-/ +def Ordinal := Bounded.LE 1 31 + deriving Repr, BEq, LE, LT + +instance [Le 1 n] [Le n 31] : OfNat Ordinal n where + ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) + +instance { x y : Ordinal } : Decidable (x ≤ y) := + dite (x.val ≤ y.val) isTrue isFalse + +instance : Inhabited Ordinal where default := 1 + +/-- +`Ordinal.OfYear` represents a bounded value for days, which ranges between 0 and 31. +-/ +def Ordinal.OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) + +instance [Le 1 n] [Le n (if leap then 366 else 365)] : OfNat (Ordinal.OfYear leap) n where + ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) + +instance : Inhabited (Ordinal.OfYear leap) where + default := by + refine ⟨1, And.intro (by decide) ?_⟩ + split <;> simp + +/-- +`Offset` represents an offset in days. It is defined as an `Int`. +-/ +def Offset : Type := UnitVal 86400 + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ + +namespace Ordinal + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 31 := by decide) : Ordinal := + Bounded.LE.ofNat' data h + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds, if its 0 then its converted +to 1. +-/ +def ofFin (data : Fin 32) : Ordinal := + Bounded.LE.ofFin' data (by decide) + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal + +namespace Offset + +/-- +Convert `Day.Offset` into `Second.Offset`. +-/ +def toSeconds (days : Offset) : Second.Offset := + days.mul 86400 + +end Offset +end Day diff --git a/src/Lean/Time/Date/Unit/Lemmas.lean b/src/Lean/Time/Date/Unit/Lemmas.lean new file mode 100644 index 000000000000..46e50091b717 --- /dev/null +++ b/src/Lean/Time/Date/Unit/Lemmas.lean @@ -0,0 +1,7 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Init.Data.Int diff --git a/src/Lean/Time/Date/Unit/Month.lean b/src/Lean/Time/Date/Unit/Month.lean new file mode 100644 index 000000000000..47256755d364 --- /dev/null +++ b/src/Lean/Time/Date/Unit/Month.lean @@ -0,0 +1,186 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat +import Lean.Time.Time.Basic +import Lean.Time.Date.Unit.Day + +namespace Lean +namespace Date +open Lean Time + +namespace Month +/-- +`Ordinal` represents a bounded value for months, which ranges between 1 and 12. +-/ +def Ordinal := Bounded.LE 1 12 + deriving Repr, BEq, LE + +instance [Le 1 n] [Le n 12] : OfNat Ordinal n where + ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) + +instance : Inhabited Ordinal where default := 1 + +/-- +`Offset` represents an offset in months. It is defined as an `Int`. +-/ +def Offset : Type := Int + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + +instance : OfNat Offset n := ⟨Int.ofNat n⟩ + +namespace Ordinal + +@[inline] def january : Ordinal := 1 +@[inline] def february : Ordinal := 2 +@[inline] def march : Ordinal := 3 +@[inline] def april : Ordinal := 4 +@[inline] def may : Ordinal := 5 +@[inline] def june : Ordinal := 6 +@[inline] def july : Ordinal := 7 +@[inline] def august : Ordinal := 8 +@[inline] def september : Ordinal := 9 +@[inline] def october : Ordinal := 10 +@[inline] def november : Ordinal := 11 +@[inline] def december : Ordinal := 12 + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 12 := by decide) : Ordinal := + Bounded.LE.ofNat' data h + +/-- +Converts a `Ordinal` into a `Nat`. +-/ +@[inline] +def toNat (month : Ordinal) : Nat := + Bounded.LE.toNat month + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds, if its 0 then its converted +to 1. +-/ +@[inline] +def ofFin (data : Fin 13) : Ordinal := + Bounded.LE.ofFin' data (by decide) + +/-- +Transforms `Month.Ordinal` into `Second.Offset`. +-/ +def toSeconds (leap : Bool) (month : Ordinal) : Second.Offset := + let daysAcc := #[0, 31, 59, 90, 120, 151, 181, 212, 243, 273, 304, 334] + let time := daysAcc[month.toNat]! * 86400 + if leap && month.toNat ≥ 2 + then time + 86400 + else time + +/-- +Transforms `Month.Ordinal` into `Minute.Offset`. +-/ +@[inline] +def toMinute (leap : Bool) (month : Ordinal) : Minute.Offset := + toSeconds leap month + |>.div 60 + +/-- +Transforms `Month.Ordinal` into `Hour.Offset`. +-/ +@[inline] +def toHours (leap : Bool) (month : Ordinal) : Hour.Offset := + toMinute leap month + |>.div 60 + +/-- +Transforms `Month.Ordinal` into `Day.Offset`. +-/ +@[inline] +def toDays (leap : Bool) (month : Ordinal) : Day.Offset := + toSeconds leap month + |>.convert + +/-- +Size in days of each month if the year is not leap. +-/ +@[inline] +def monthSizesNonLeap : { val : Array Day.Ordinal // val.size = 12 } := + ⟨#[ 31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31 ], by simp⟩ + +/-- +Gets the number of days in a month. +-/ +def days' (leap : Bool) (month : Ordinal) : Day.Ordinal := + if month.val = 2 then + if leap then 29 else 28 + else by + let ⟨months, p⟩ := monthSizesNonLeap + let r : Fin 12 := (month.sub 1).toFin (by decide) (by decide) + rw [← p] at r + exact months.get r + +/-- +Check if the day is valid in a Month and a leap Year. +-/ +def valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := + day ≤ days' leap month + +instance : Decidable (valid leap month day) := + dite (day ≤ days' leap month) isTrue isFalse + +/-- +Gets the number of days in a month. +-/ +def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // valid leap month day } := + ⟨days' leap month, Int.le_refl ((days' leap month).val)⟩ + +/-- +Forces the day to be on the valid range. +-/ +@[inline] +def forceDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : Day.Ordinal //valid leap month day } := + let max : Day.Ordinal := month.days leap + if h : day.val > max.val + then ⟨max, Int.le_refl max.val⟩ + else ⟨⟨day.val, day.property⟩, Int.not_lt.mp h⟩ + +/-- +Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. +-/ +@[inline] +def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day.Ordinal // valid leap (Prod.fst val) (Prod.snd val) } := Id.run do + let rec go (idx : Fin 12) (cumulative : Fin 366) := + let month := Month.Ordinal.ofFin idx.succ + let ⟨days, valid⟩ := days leap month + + if h : cumulative.val < ordinal.val ∧ ordinal.val ≤ cumulative.val + days.val then + let bounded := + Bounded.LE.mk ordinal.val h |>.sub ↑↑cumulative + let bounded : Bounded.LE 1 days.val := by + simp [← Int.add_comm, Int.sub_self] at bounded + rw [← Int.add_comm 1 (↑↑cumulative), Int.add_sub_assoc, Int.sub_self] at bounded + exact bounded + let p₁ := bounded.property.right + let p := And.intro bounded.property.left (Int.le_trans bounded.property.right days.property.right) + let days₁ : Day.Ordinal := ⟨bounded.val, p⟩ + let h1 : Month.Ordinal.valid leap month days₁ := Int.le_trans p₁ valid + ⟨⟨month, days₁⟩, h1⟩ + else + if h : idx.val ≥ 11 then + -- Need to remove this in the future. + let ⟨day, valid⟩ := forceDay leap 1 1 + ⟨⟨1, day⟩, valid⟩ + else + go ⟨idx.val + 1, Nat.succ_le_succ (Nat.not_le.mp h)⟩ cumulative + termination_by 12 - idx.val + + go 1 0 + +end Ordinal +end Month diff --git a/src/Lean/Time/Date/Unit/WeekOfYear.lean b/src/Lean/Time/Date/Unit/WeekOfYear.lean new file mode 100644 index 000000000000..57b077534e35 --- /dev/null +++ b/src/Lean/Time/Date/Unit/WeekOfYear.lean @@ -0,0 +1,63 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat +import Lean.Time.Date.Unit.Day +import Lean.Time.Date.Unit.Month + +namespace Lean +namespace Date + +set_option linter.all true + +namespace WeekOfYear + +/-- +`Ordinal` represents a bounded value for weeks, which ranges between 1 and 53. +-/ +def Ordinal := Bounded.LE 1 53 + deriving Repr, BEq, LE, LT + +instance [Le 1 n] [Le n 53] : OfNat Ordinal n where + ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) + +instance : Inhabited Ordinal where default := 1 + +/-- +`Offset` represents an offset in weeks. It is defined as an `Int`. +-/ +def Offset : Type := UnitVal (86400 * 7) + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ + +namespace Ordinal + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 53 := by decide) : Ordinal := + Bounded.LE.ofNat' data h + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +-/ +def ofFin (data : Fin 54) : Ordinal := + Bounded.LE.ofFin' data (by decide) + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal + +end WeekOfYear +end Date diff --git a/src/Lean/Time/Date/Unit/Weekday.lean b/src/Lean/Time/Date/Unit/Weekday.lean new file mode 100644 index 000000000000..87c27f5ea3c0 --- /dev/null +++ b/src/Lean/Time/Date/Unit/Weekday.lean @@ -0,0 +1,109 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat +import Lean.Time.Date.Unit.Day +import Lean.Time.Date.Unit.Month + +namespace Lean +namespace Date + +set_option linter.all true + +/-- Defines the enumeration for days of the week. Each variant corresponds to a day of the week, +from Monday to Sunday. -/ +inductive Weekday + /-- Monday. -/ + | mon + /-- Tuesday. -/ + | tue + /-- Wednesday. -/ + | wed + /-- Thursday. -/ + | thu + /-- Friday. -/ + | fri + /-- Saturday. -/ + | sat + /-- Sunday. -/ + | sun + deriving Repr, Inhabited, BEq + +namespace Weekday + +/-- +Converts a `Fin 7` representing a day index into a corresponding `Weekday`. This function is +useful for mapping numerical representations to days of the week. +-/ +def ofFin : Fin 7 → Weekday + | 0 => .mon + | 1 => .tue + | 2 => .wed + | 3 => .thu + | 4 => .fri + | 5 => .sat + | 6 => .sun + +/-- +Converts a `Weekday` to a `Nat`. +-/ +def toNat : Weekday → Nat + | .mon => 0 + | .tue => 1 + | .wed => 2 + | .thu => 3 + | .fri => 4 + | .sat => 5 + | .sun => 6 + +/-- +Converts a `Weekday` to a `Fin`. +-/ +def toFin : Weekday → Nat + | .mon => 0 + | .tue => 1 + | .wed => 2 + | .thu => 3 + | .fri => 4 + | .sat => 5 + | .sun => 6 + +/-- +Converts a `Nat` to a `Option Weekday`. +-/ +def ofNat? : Nat → Option Weekday + | 0 => some .mon + | 1 => some .tue + | 2 => some .wed + | 3 => some .thu + | 4 => some .fri + | 5 => some .sat + | 6 => some .sun + | _ => none + +/-- +Converts a `Nat` to a `Weekday`. Panics if the value provided is invalid. +-/ +@[inline] +def ofNat! (n : Nat) : Weekday := + match ofNat? n with + | some res => res + | none => panic! "invalid weekday" + +/-- +Gets the next `Weekday`. +-/ +def next : Weekday → Weekday + | .mon => .sun + | .tue => .mon + | .wed => .tue + | .thu => .wed + | .fri => .thu + | .sat => .fri + | .sun => .sat diff --git a/src/Lean/Time/Date/Unit/Year.lean b/src/Lean/Time/Date/Unit/Year.lean new file mode 100644 index 000000000000..3ebcef013f87 --- /dev/null +++ b/src/Lean/Time/Date/Unit/Year.lean @@ -0,0 +1,63 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat +import Lean.Time.Date.Unit.Day +import Lean.Time.Date.Unit.Month + +namespace Lean +namespace Date + +set_option linter.all true + +namespace Year + +/-- +`Offset` represents an offset in years. It is defined as an `Int`. +-/ +def Offset : Type := Int + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + +instance : OfNat Offset n := ⟨Int.ofNat n⟩ + +namespace Offset + +/-- +Convert the `Year` offset to an `Int`. +-/ +@[inline] +def toInt (offset : Offset) : Int := + offset + +/-- +Convert the `Year` offset to a `Month` Offset. +-/ +@[inline] +def toMonths (val : Offset) : Month.Offset := + val.mul 12 + +/-- +Checks if the `Year` is a Gregorian Leap Year. +-/ +@[inline] +def isLeap (y : Offset) : Bool := + y.toInt.mod 4 = 0 ∧ (y.toInt.mod 100 ≠ 0 ∨ y.toInt.mod 400 = 0) + +/-- +Forces the day to be on the valid range. +-/ +@[inline] +def valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := + month.valid year.isLeap day + +instance : Decidable (valid year month day) := + dite (month.valid year.isLeap day) isTrue isFalse + +end Offset +end Year diff --git a/src/Lean/Time/Date/WeekDate.lean b/src/Lean/Time/Date/WeekDate.lean new file mode 100644 index 000000000000..6bdb97d8b38c --- /dev/null +++ b/src/Lean/Time/Date/WeekDate.lean @@ -0,0 +1,42 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Date.Unit.Year +import Lean.Time.Date.Unit.WeekOfYear +import Lean.Time.Date.Scalar +import Lean.Time.Date.Date + +namespace Lean +namespace Date + +/-- +`WeekDate` represents a date using a combination of a week of the year and the year. +-/ +structure WeekDate where + year : Year.Offset + week : WeekOfYear.Ordinal + deriving Repr, BEq, Inhabited + +namespace WeekDate + +/-- +Converts a `WeekDate` to a `Scalar`. +-/ +def toScalar (wd : WeekDate) : Scalar := + let days := wd.year.toInt * 365 + wd.week.val * 7 + Scalar.ofDays days + +/-- +Creates a `WeekDate` from a `Scalar`. +-/ +def fromScalar (scalar : Scalar) : WeekDate := + let totalDays := scalar.toDays + let year := totalDays / 365 + let week := + Bounded.LE.byEmod totalDays 365 (by decide) + |>.div 7 (by decide) + |>.add 1 + { year := year, week := week } diff --git a/src/Lean/Time/DateTime.lean b/src/Lean/Time/DateTime.lean new file mode 100644 index 000000000000..7acc71628232 --- /dev/null +++ b/src/Lean/Time/DateTime.lean @@ -0,0 +1,25 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.DateTime.NaiveDateTime +import Lean.Time.DateTime.Timestamp + +namespace Lean +namespace DateTime.Timestamp + +/-- +Converts a `NaiveDateTime` to a `Timestamp` +-/ +@[inline] +def ofNaiveDateTime (naive : NaiveDateTime) : Timestamp := + naive.toTimestamp + +/-- +Converts a `Timestamp` to a `NaiveDateTime` +-/ +@[inline] +def toNaiveDateTime (timestamp : Timestamp) : NaiveDateTime := + NaiveDateTime.ofTimestamp timestamp diff --git a/src/Lean/Time/DateTime/Basic.lean b/src/Lean/Time/DateTime/Basic.lean new file mode 100644 index 000000000000..61458d98a6fe --- /dev/null +++ b/src/Lean/Time/DateTime/Basic.lean @@ -0,0 +1,8 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Init.Data.Int +import Lean.Time.LessEq diff --git a/src/Lean/Time/DateTime/NaiveDateTime.lean b/src/Lean/Time/DateTime/NaiveDateTime.lean new file mode 100644 index 000000000000..f230be45f1e3 --- /dev/null +++ b/src/Lean/Time/DateTime/NaiveDateTime.lean @@ -0,0 +1,151 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Init.Data.Int +import Lean.Time.LessEq +import Lean.Time.Date +import Lean.Time.Time +import Lean.Time.DateTime.Timestamp + +namespace Lean +namespace DateTime +open Date Time + +/-- +Date time format with Year, Month, Day, Hour, Minute, Seconds and Nanoseconds. +-/ +structure NaiveDateTime where + date : Date.Date + time : Time.Time + deriving Repr + +namespace NaiveDateTime + +/-- +Converts a `NaiveDateTime` into a `Timestamp` +-/ +def toTimestamp (dt : NaiveDateTime) : Timestamp := + let days := dt.date.toScalar.day + let second := dt.time.toSeconds + days.toSeconds + second + +/-- +Converts a UNIX `Timestamp` into a `NaiveDateTime`. +-/ +def ofTimestamp (stamp : Timestamp) : NaiveDateTime := Id.run do + let leapYearEpoch := 11017 + let daysPer400Y := 365 * 400 + 97 + let daysPer100Y := 365 * 100 + 24 + let daysPer4Y := 365 * 4 + 1 + + let secs := stamp.toInt + let daysSinceEpoch := secs / 86400 + let boundedDaysSinceEpoch := daysSinceEpoch + + let mut rawDays := boundedDaysSinceEpoch - leapYearEpoch + let mut rem := Bounded.LE.byMod secs 86400 (by decide) + + let ⟨remSecs, days⟩ := + if h : rem.val ≤ -1 then + let remSecs := rem.truncateTop h + let remSecs : Bounded.LE 1 86399 := remSecs.add 86400 + let rawDays := rawDays - 1 + (remSecs.expandBottom (by decide), rawDays) + else + let h := rem.truncateBottom (Int.not_le.mp h) + (h, rawDays) + + let mut quadracentennialCycles := days / daysPer400Y; + let mut remDays := days % daysPer400Y; + + if remDays < 0 then + remDays := remDays + daysPer400Y + quadracentennialCycles := quadracentennialCycles - 1 + + let mut centenialCycles := remDays / daysPer100Y; + + if centenialCycles = 4 then + centenialCycles := centenialCycles - 1 + + remDays := remDays - centenialCycles * daysPer100Y + let mut quadrennialCycles := remDays / daysPer4Y; + + if quadrennialCycles = 25 then + quadrennialCycles := quadrennialCycles - 1 + + remDays := remDays - quadrennialCycles * daysPer4Y + let mut remYears := remDays / 365; + + if remYears = 4 then + remYears := remYears - 1 + + remDays := remDays - remYears * 365 + + let mut year := 2000 + remYears + 4 * quadrennialCycles + 100 * centenialCycles + 400 * quadracentennialCycles + let months := [31, 30, 31, 30, 31, 31, 30, 31, 30, 31, 31, 29]; + let mut mon : Fin 13 := 0; + + for monLen in months do + mon := mon + 1; + if remDays < monLen then + break + remDays := remDays - monLen + + let mday : Fin 31 := Fin.ofNat (Int.toNat remDays) + + let hmon ← + if h₁ : mon.val > 10 + then do + year := year + 1 + pure (Month.Ordinal.ofNat (mon.val - 10) (by omega)) + else + pure (Month.Ordinal.ofNat (mon.val + 2) (by omega)) + + let second : Bounded.LE 0 59 := remSecs.emod 60 (by decide) + let minute : Bounded.LE 0 59 := (remSecs.div 60 (by decide)).emod 60 (by decide) + let hour : Bounded.LE 0 23 := remSecs.div 3600 (by decide) + + return { + date := Date.force year hmon (Day.Ordinal.ofFin (Fin.succ mday)) + time := Time.mk (hour.expandTop (by decide)) minute (second.expandTop (by decide)) 0 + } + +/-- +Getter for the `Year` inside of a `NaiveDateTime` +-/ +@[inline] +def year (dt : NaiveDateTime) : Year.Offset := + dt.date.year +/-- +Getter for the `Month` inside of a `NaiveDateTime` +-/ +@[inline] +def month (dt : NaiveDateTime) : Month.Ordinal := + dt.date.month +/-- +Getter for the `Day` inside of a `NaiveDateTime` +-/ +@[inline] +def day (dt : NaiveDateTime) : Day.Ordinal := + dt.date.day +/-- +Getter for the `Hour` inside of a `NaiveDateTime` +-/ +@[inline] +def hour (dt : NaiveDateTime) : Hour.Ordinal := + dt.time.hour +/-- +Getter for the `Minute` inside of a `NaiveDateTime` +-/ +@[inline] +def minute (dt : NaiveDateTime) : Minute.Ordinal := + dt.time.minute +/-- +Getter for the `Second` inside of a `NaiveDateTime` +-/ +@[inline] +def second (dt : NaiveDateTime) : Second.Ordinal := + dt.time.second diff --git a/src/Lean/Time/DateTime/Timestamp.lean b/src/Lean/Time/DateTime/Timestamp.lean new file mode 100644 index 000000000000..f17e6fed14af --- /dev/null +++ b/src/Lean/Time/DateTime/Timestamp.lean @@ -0,0 +1,25 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Init.Data.Int +import Lean.Time.LessEq +import Lean.Time.Time + +namespace Lean +namespace DateTime +open Time + +/-- +Seconds since the UNIX Epoch. +-/ +def Timestamp := Second.Offset + deriving Repr, BEq + +instance : OfNat Timestamp n where + ofNat := UnitVal.ofNat n + +instance : HAdd Timestamp Second.Offset Timestamp where + hAdd x y := UnitVal.add x y diff --git a/src/Lean/Time/Duration.lean b/src/Lean/Time/Duration.lean new file mode 100644 index 000000000000..63db8fd707af --- /dev/null +++ b/src/Lean/Time/Duration.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Bounded +import Lean.Time.Time + +namespace Lean + +open Time + +/-- +`Instant` represents a place in time with second and nanoseconds precision. +-/ +structure Instant where + second : Second.Offset + nano : Nanosecond.Ordinal + valid : second.val ≥ 0 ∧ nano.val ≥ 0 + deriving Repr + +/-- +Time duration with nanosecond precision. This type allows negative duration. +-/ +structure Duration where + second : Second.Offset + nano : Nanosecond.Span + +namespace Instant + +/-- +Gets the difference of two `Instant` in a `Duration`. +-/ +def sub (t₁ t₂ : Instant) : Duration := + let nsec_diff := (t₁.nano).subBounds (t₂.nano) + let sec_diff := (t₁.second) - (t₂.second) + if h : sec_diff.val > 0 ∧ nsec_diff.val ≤ -1 then + let truncated := nsec_diff.truncateTop h.right + { second := (sec_diff - 1), nano := truncated.addTop 1000000000 } + else if h₁ : sec_diff.val < 0 ∧ nsec_diff.val ≥ 1 then + let truncated := nsec_diff.truncateBottom h₁.right + { second := (sec_diff + 1), nano := (truncated.subBottom 1000000000) } + else + { second := sec_diff, nano := nsec_diff } + +instance : HSub Instant Instant Duration where + hSub := Instant.sub + +/-- +Gets how much time elapsed since another `Instant` and returns a `Duration`. +-/ +@[inline] +def since (instant : Instant) (since : Instant) : Duration := + Instant.sub since instant + +end Instant + +namespace Duration + +/-- Returns a `Duration` representing the given number of second. -/ +def ofSeconds (second : Second.Offset) : Duration := + { second := second, nano := Bounded.LE.mk 0 (by decide) } + +/-- Returns a `Duration` representing the given number of minute. -/ +def ofMinutes (minute : Minute.Offset) : Duration := + { second := minute.toSeconds * 60, nano := Bounded.LE.mk 0 (by decide) } + +/-- Returns a `Duration` representing the given number of hour. -/ +def ofHours (hour : Hour.Offset) : Duration := + { second := hour.toSeconds * 3600, nano := Bounded.LE.mk 0 (by decide) } + +end Duration diff --git a/src/Lean/Time/Format.lean b/src/Lean/Time/Format.lean new file mode 100644 index 000000000000..4732a827cac4 --- /dev/null +++ b/src/Lean/Time/Format.lean @@ -0,0 +1,412 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Date +import Lean.Time.Time +import Lean.Time.DateTime +import Lean.Time.TimeZone +import Lean.Data.Parsec + +namespace Lean +namespace Format +open Lean.Parsec Time Date TimeZone DateTime + +private inductive Modifier + | YYYY + | YY + | MMMM + | MMM + | MM + | M + | DD + | D + | EEEE + | EEE + | hh + | h + | HH + | H + | AA + | aa + | mm + | m + | ss + | s + | ZZZZZ + | ZZZZ + | ZZZ + | Z + deriving Repr + +/-- +The part of a formatting string. a string is just a text and a modifier is in the format `%0T` where +0 is the quantity of left pad and `T` the `Modifier`. +-/ +private inductive FormatPart + | string (val : String) + | modifier (left_pad: Nat) (modifier : Modifier) + deriving Repr + +/-- +The format of date and time string. +-/ +abbrev Format := List FormatPart + +private def isLetter (c : Char) : Bool := + (c >= 'a' && c <= 'z') || (c >= 'A' && c <= 'Z') + +private def isNonLetter (c: Char) : Bool := ¬isLetter c + +private def parseModifier : Lean.Parsec Modifier + := pstring "YYYY" *> pure .YYYY + <|> pstring "YY" *> pure .YY + <|> pstring "MMMM" *> pure .MMMM + <|> pstring "MMM" *> pure .MMM + <|> pstring "MM" *> pure .MM + <|> pstring "M" *> pure .M + <|> pstring "DD" *> pure .DD + <|> pstring "D" *> pure .D + <|> pstring "EEEE" *> pure .EEEE + <|> pstring "EEE" *> pure .EEE + <|> pstring "hh" *> pure .hh + <|> pstring "h" *> pure .h + <|> pstring "HH" *> pure .HH + <|> pstring "H" *> pure .H + <|> pstring "AA" *> pure .AA + <|> pstring "aa" *> pure .aa + <|> pstring "mm" *> pure .mm + <|> pstring "m" *> pure .m + <|> pstring "ss" *> pure .ss + <|> pstring "s" *> pure .s + <|> pstring "ZZZZZ" *> pure .ZZZZZ + <|> pstring "ZZZZ" *> pure .ZZZZ + <|> pstring "ZZZ" *> pure .ZZZ + <|> pstring "Z" *> pure .Z + +private def pnumber : Lean.Parsec Nat := do + let numbers ← manyChars digit + return numbers.foldl (λacc char => acc * 10 + (char.toNat - 48)) 0 + +private def parseFormatPart : Lean.Parsec FormatPart + := (.modifier <$> (pchar '%' *> pnumber) <*> parseModifier) + <|> (pchar '\\') *> anyChar <&> (.string ∘ toString) + <|> (pchar '\"' *> many1Chars (satisfy (· ≠ '\"')) <* pchar '\"') <&> .string + <|> (pchar '\'' *> many1Chars (satisfy (· ≠ '\'')) <* pchar '\'') <&> .string + <|> many1Chars (satisfy (fun x => x ≠ '%' ∧ x ≠ '\'' ∧ x ≠ '\"')) <&> .string + +private def specParser : Lean.Parsec Format := + (Array.toList <$> Lean.Parsec.many parseFormatPart) <* eof + +private def specParse (s: String) : Except String Format := + specParser.run s + +def Format.spec! (s: String) : Format := + match specParser.run s with + | .ok s => s + | .error s => panic! s + +-- Pretty printer + +private def unabbrevMonth (month: Month.Ordinal) : String := + match month.val, month.property with + | 1, _ => "January" + | 2, _ => "February" + | 3, _ => "March" + | 4, _ => "April" + | 5, _ => "May" + | 6, _ => "June" + | 7, _ => "July" + | 8, _ => "August" + | 9, _ => "September" + | 10, _ => "October" + | 11, _ => "November" + | 12, _ => "December" + +private def abbrevMonth (month: Month.Ordinal) : String := + match month.val, month.property with + | 1, _ => "Jan" + | 2, _ => "Feb" + | 3, _ => "Mar" + | 4, _ => "Apr" + | 5, _ => "May" + | 6, _ => "Jun" + | 7, _ => "Jul" + | 8, _ => "Aug" + | 9, _ => "Sep" + | 10, _ => "Oct" + | 11, _ => "Nov" + | 12, _ => "Dec" + +private def abbrevDayOfWeek (day: Weekday) : String := + match day with + | .sun => "Sun" + | .mon => "Mon" + | .tue => "Tue" + | .wed => "Wed" + | .thu => "Thu" + | .fri => "Fri" + | .sat => "Sat" + +private def dayOfWeek (day: Weekday) : String := + match day with + | .sun => "Sunday" + | .mon => "Monday" + | .tue => "Tuesday" + | .wed => "Wednesday" + | .thu => "Thusday" + | .fri => "Friday" + | .sat => "Saturday" + +private def leftPad (n : Nat) (a : Char) (s : String) : String := + "".pushn a (n - s.length) ++ s + +private def formatWithDate (date : ZonedDateTime) : Modifier → String + | .YYYY => s!"{leftPad 4 '0' (toString date.year)}" + | .YY => s!"{leftPad 2 '0' (toString $ date.year.toNat % 100)}" + | .MMMM => unabbrevMonth date.month + | .MMM => abbrevMonth date.month + | .MM => s!"{leftPad 2 '0' (toString $ date.month.toNat)}" + | .M => s!"{date.month.toNat}" + | .DD => s!"{leftPad 2 '0' (toString $ date.day.toNat)}" + | .D => s!"{date.day.toNat}" + | .EEEE => dayOfWeek date.weekday + | .EEE => abbrevDayOfWeek date.weekday + | .hh => s!"{leftPad 2 '0' (toString date.hour.toNat)}" + | .h => s!"{date.hour.toNat}" + | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' $ toString hour}" + | .H => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{hour}" + | .AA => if date.hour.toNat < 12 then "AM" else "PM" + | .aa => if date.hour.toNat < 12 then "am" else "pm" + | .mm => s!"{leftPad 2 '0' $ toString date.minute.toNat}" + | .m => s!"{date.minute.toNat}" + | .ss => s!"{leftPad 2 '0' $ toString date.second.toNat}" + | .s => s!"{date.second.toNat}" + | .ZZZZZ => date.offset.toIsoString true + | .ZZZZ => date.offset.toIsoString false + | .ZZZ => if date.offset.second.val = 0 then "UTC" else date.offset.toIsoString false + | .Z => if date.offset.second.val = 0 then "Z" else date.offset.toIsoString true + +private def formatPartWithDate (date : ZonedDateTime) : FormatPart → String + | .string s => s + | .modifier p t => leftPad p ' ' (formatWithDate date t) + +-- Parser + +@[simp] +private def SingleFormatType : Modifier → Type + | .YYYY => Year.Offset + | .YY => Year.Offset + | .MMMM => Month.Ordinal + | .MMM => Month.Ordinal + | .MM => Month.Ordinal + | .M => Month.Ordinal + | .DD => Day.Ordinal + | .D => Day.Ordinal + | .EEEE => Weekday + | .EEE => Weekday + | .hh => Hour.Ordinal + | .h => Hour.Ordinal + | .HH => Hour.Ordinal + | .H => Hour.Ordinal + | .AA => HourMarker + | .aa => HourMarker + | .mm => Minute.Ordinal + | .m => Minute.Ordinal + | .ss => Second.Ordinal + | .s => Second.Ordinal + | .ZZZZZ => Offset + | .ZZZZ => Offset + | .ZZZ => Offset + | .Z => Offset + +private def transform (n: β → Option α) (p: Lean.Parsec β) : Lean.Parsec α := do + let res ← p + match n res with + | some n => pure n + | none => fail "cannot parse" + +private def parseMonth : Lean.Parsec Month.Ordinal + := (pstring "Jan" *> pure 1) + <|> (pstring "Feb" *> pure 2) + <|> (pstring "Mar" *> pure 3) + <|> (pstring "Apr" *> pure 4) + <|> (pstring "May" *> pure 5) + <|> (pstring "Jun" *> pure 6) + <|> (pstring "Jul" *> pure 7) + <|> (pstring "Aug" *> pure 8) + <|> (pstring "Sep" *> pure 9) + <|> (pstring "Oct" *> pure 10) + <|> (pstring "Nov" *> pure 11) + <|> (pstring "Dec" *> pure 12) + +private def parseMonthUnabbrev : Lean.Parsec Month.Ordinal + := (pstring "January" *> pure 1) + <|> (pstring "February" *> pure 2) + <|> (pstring "March" *> pure 3) + <|> (pstring "April" *> pure 4) + <|> (pstring "May" *> pure 5) + <|> (pstring "June" *> pure 6) + <|> (pstring "July" *> pure 7) + <|> (pstring "August" *> pure 8) + <|> (pstring "September" *> pure 9) + <|> (pstring "October" *> pure 10) + <|> (pstring "November" *> pure 11) + <|> (pstring "December" *> pure 12) + +private def parseWeekday : Lean.Parsec Weekday + := (pstring "Mon" *> pure Weekday.mon) + <|> (pstring "Tue" *> pure Weekday.tue) + <|> (pstring "Wed" *> pure Weekday.wed) + <|> (pstring "Thu" *> pure Weekday.thu) + <|> (pstring "Fri" *> pure Weekday.fri) + <|> (pstring "Sat" *> pure Weekday.sat) + <|> (pstring "Sun" *> pure Weekday.sun) + +private def parseWeekdayUnnabrev : Lean.Parsec Weekday + := (pstring "Monday" *> pure Weekday.mon) + <|> (pstring "Tuesday" *> pure Weekday.tue) + <|> (pstring "Wednesday" *> pure Weekday.wed) + <|> (pstring "Thursday" *> pure Weekday.thu) + <|> (pstring "Friday" *> pure Weekday.fri) + <|> (pstring "Saturday" *> pure Weekday.sat) + <|> (pstring "Sunday" *> pure Weekday.sun) + +private def parserUpperHourMarker : Lean.Parsec HourMarker + := (pstring "AM" *> pure HourMarker.am) + <|> (pstring "PM" *> pure HourMarker.pm) + +private def parserLowerHourMarker : Lean.Parsec HourMarker + := (pstring "am" *> pure HourMarker.am) + <|> (pstring "pm" *> pure HourMarker.pm) + +private def twoDigit : Lean.Parsec Int := do + let digit1 ← Lean.Parsec.digit + let digit2 ← Lean.Parsec.digit + return String.toNat! s!"{digit1}{digit2}" + +private def parseYearTwo : Lean.Parsec Int :=do + let year ← twoDigit + return if year < 70 then 2000 + year else 1900 + year + +private def timeOffset (colon: Bool) : Lean.Parsec Offset := do + let sign : Int ← (pstring "-" *> pure (-1)) <|> (pstring "+" *> pure 1) + let hour ← twoDigit + if colon then let _ ← pstring ":" + let minutes ← twoDigit + let res := (hour * 3600 + minutes * 60) * sign + pure (Offset.ofSeconds (UnitVal.ofInt res)) + +private def timeOrUTC (utcString: String) (colon: Bool) : Lean.Parsec Offset := + (pstring utcString *> pure Offset.zero) <|> timeOffset colon + +private def number : Lean.Parsec Nat := do + String.toNat! <$> Lean.Parsec.many1Chars Lean.Parsec.digit + +private def singleDigit : Lean.Parsec Nat := do + let digit1 ← Lean.Parsec.digit + return String.toNat! s!"{digit1}" + +private def fourDigit : Lean.Parsec Int := do + let digit1 ← Lean.Parsec.digit + let digit2 ← Lean.Parsec.digit + let digit3 ← Lean.Parsec.digit + let digit4 ← Lean.Parsec.digit + return String.toNat! s!"{digit1}{digit2}{digit3}{digit4}" + +private def parserWithFormat : (typ: Modifier) → Lean.Parsec (SingleFormatType typ) + | .YYYY => fourDigit + | .YY => parseYearTwo + | .MMMM => parseMonthUnabbrev + | .MMM => parseMonth + | .MM => transform Bounded.LE.ofInt twoDigit + | .M => transform Bounded.LE.ofInt number + | .DD => transform Bounded.LE.ofInt twoDigit + | .D => transform Bounded.LE.ofInt number + | .EEEE => parseWeekdayUnnabrev + | .EEE => parseWeekday + | .hh => transform Bounded.LE.ofInt twoDigit + | .h => transform Bounded.LE.ofInt number + | .HH => transform Bounded.LE.ofInt twoDigit + | .H => transform Bounded.LE.ofInt number + | .AA => parserUpperHourMarker + | .aa => parserLowerHourMarker + | .mm => transform Bounded.LE.ofInt twoDigit + | .m => transform Bounded.LE.ofInt number + | .ss => transform Bounded.LE.ofInt twoDigit + | .s => transform Bounded.LE.ofInt number + | .ZZZZZ => timeOffset true + | .ZZZZ => timeOffset false + | .ZZZ => timeOrUTC "UTC" false + | .Z => timeOrUTC "Z" true + +structure DateBuilder where + tz : Offset := Offset.zero + year : Year.Offset := 0 + month : Month.Ordinal := 1 + day : Day.Ordinal := 1 + hour : Hour.Ordinal := 0 + minute : Minute.Ordinal := 0 + second : Second.Ordinal := 0 + +def DateBuilder.build (builder : DateBuilder) : ZonedDateTime := + { + fst := TimeZone.mk builder.tz "GMT" + snd := DateTime.ofNaiveDateTime { + date := Date.force builder.year builder.month builder.day + time := Time.mk builder.hour builder.minute builder.second 0 + } (TimeZone.mk builder.tz "GMT") + } + +private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : SingleFormatType typ) : DateBuilder := + match typ with + | .YYYY => { data with year := value } + | .YY => { data with year := value } + | .MMMM => { data with month := value } + | .MMM => { data with month := value } + | .MM => { data with month := value } + | .M => { data with month := value } + | .DD => { data with day := value } + | .D => { data with day := value } + | .EEEE => data + | .EEE => data + | .hh => { data with hour := value } + | .h => { data with hour := value } + | .HH => { data with hour := value } + | .H => { data with hour := value } + | .AA => { data with hour := HourMarker.toAbsolute! value data.hour } + | .aa => { data with hour := HourMarker.toAbsolute! value data.hour } + | .mm => { data with minute := value } + | .m => { data with minute := value } + | .ss => { data with second := value } + | .s => { data with second := value } + | .ZZZZZ => { data with tz := value } + | .ZZZZ => { data with tz := value } + | .ZZZ => { data with tz := value } + | .Z => { data with tz := value } + +private def addData (data : DateBuilder) : FormatPart → DateBuilder + | .string s => data + | .modifier _ m => addDataInDateTime data m sorry + +-- Internals + +/- +def Format.format (x: Format) (date : ZonedDateTime) : String := + x.map (formatPartWithDate date) + |> String.join + +def Formats.ISO8601 := Format.spec! "%YYYY-%MM-%DD'T'%hh:%mm:%ss'Z'" + +-- Tests + +def x : Timestamp := 1722631000 + +def UTCM3 := (TimeZone.mk (Offset.ofSeconds (-10800)) "Brasilia") + +def date := ZonedDateTime.ofTimestamp x UTCM3 + +#eval Formats.ISO8601.format date +-/ diff --git a/src/Lean/Time/LessEq.lean b/src/Lean/Time/LessEq.lean new file mode 100644 index 000000000000..7704a4997a6f --- /dev/null +++ b/src/Lean/Time/LessEq.lean @@ -0,0 +1,27 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Init +import Lean.Data.Rat + +namespace Lean +set_option linter.all true + +/-- +Type class to check if a number `n` is less or equal than `m` in compile time and give me a proof. +-/ +class Le (n : Nat) (m : Nat) where + /-- The proof that n ≤ m. -/ + p : n ≤ m + +instance : Le n n where + p := Nat.le_refl n + +instance (m : Nat) : Le 0 m where + p := Nat.zero_le m + +instance {n m : Nat} [Le n m] : Le (Nat.succ n) (Nat.succ m) where + p := Nat.succ_le_succ (Le.p) diff --git a/src/Lean/Time/Sign.lean b/src/Lean/Time/Sign.lean new file mode 100644 index 000000000000..b5a385455a81 --- /dev/null +++ b/src/Lean/Time/Sign.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Bounded + +namespace Lean +set_option linter.all true + +/-- +A `Sign` is a type that can have three values -1, 0 and 1 that represents the sign of a value. +-/ +def Sign := Bounded.LE (-1) 1 + +namespace Sign + +/-- +Gets the `Sign` out of a integer. +-/ +@[inline] +def ofInt (val : Int) : Sign := by + refine ⟨val.sign, ?_⟩ + simp [Int.sign] + split <;> simp + +@[inline] +def apply (sign : Sign) (val : Int) : Int := + val * sign.val diff --git a/src/Lean/Time/Time.lean b/src/Lean/Time/Time.lean new file mode 100644 index 000000000000..a9b295ed8d36 --- /dev/null +++ b/src/Lean/Time/Time.lean @@ -0,0 +1,10 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time.HourMarker +import Lean.Time.Time.Basic +import Lean.Time.Time.Scalar +import Lean.Time.Time.Time diff --git a/src/Lean/Time/Time/Basic.lean b/src/Lean/Time/Time/Basic.lean new file mode 100644 index 000000000000..af3f62132446 --- /dev/null +++ b/src/Lean/Time/Time/Basic.lean @@ -0,0 +1,7 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time.Unit.Basic diff --git a/src/Lean/Time/Time/HourMarker.lean b/src/Lean/Time/Time/HourMarker.lean new file mode 100644 index 000000000000..2c415707b2f7 --- /dev/null +++ b/src/Lean/Time/Time/HourMarker.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time.Basic + +namespace Lean +namespace Time + +/-- +`HourMarker` represents the two 12-hour periods of the day: `am` for hour between 12:00 AM and +11:59 AM, and `pm` for hour between 12:00 PM and 11:59 PM. +-/ +inductive HourMarker + | am + | pm + +/-- +Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker`. +-/ +def HourMarker.toAbsolute (marker : HourMarker) (time : Hour.Ordinal) (h₀ : time.toInt ≤ 12) : Hour.Ordinal := + match marker with + | .am => time + | .pm => Bounded.LE.mk (time.toInt + 12) (And.intro (Int.add_le_add (c := 0) time.property.left (by decide)) (Int.add_le_add_right h₀ 12)) + +/-- +Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker` without a proof. +-/ +def HourMarker.toAbsolute! (marker : HourMarker) (time : Hour.Ordinal) : Hour.Ordinal := + match marker with + | .am => time + | .pm => Bounded.LE.ofFin (Fin.ofNat (time.val.toNat + 12)) diff --git a/src/Lean/Time/Time/Scalar.lean b/src/Lean/Time/Time/Scalar.lean new file mode 100644 index 000000000000..a6a62d557e05 --- /dev/null +++ b/src/Lean/Time/Time/Scalar.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time.Basic + +namespace Lean +namespace Time + +structure Scalar where + second : Second.Offset + nano : Nanosecond.Ordinal + +namespace Scalar + +def toSeconds (time : Scalar) : Second.Offset := + time.second + +def toMinutes (time : Scalar) : Minute.Offset := + time.second.toMinutes + +def toHours (time : Scalar) : Hour.Offset := + time.second.toHours + +end Scalar +end Time diff --git a/src/Lean/Time/Time/Time.lean b/src/Lean/Time/Time/Time.lean new file mode 100644 index 000000000000..e692fbd227ed --- /dev/null +++ b/src/Lean/Time/Time/Time.lean @@ -0,0 +1,29 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time.Basic + +namespace Lean +namespace Time + +structure Time where + hour : Hour.Ordinal + minute : Minute.Ordinal + second : Second.Ordinal + nano : Nanosecond.Ordinal + deriving Repr, Inhabited, BEq + +namespace Time + +def toSeconds (time : Time) : Second.Offset := + time.hour.toOffset.toSeconds + + time.minute.toOffset.toSeconds + + time.second.toOffset + +def toMinutes (time : Time) : Minute.Offset := + time.hour.toOffset.toMinutes + + time.minute.toOffset + + time.second.toOffset.toMinutes diff --git a/src/Lean/Time/Time/Unit/Basic.lean b/src/Lean/Time/Time/Unit/Basic.lean new file mode 100644 index 000000000000..cab3ad927554 --- /dev/null +++ b/src/Lean/Time/Time/Unit/Basic.lean @@ -0,0 +1,40 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Time.Unit.Hour +import Lean.Time.Time.Unit.Minute +import Lean.Time.Time.Unit.Second +import Lean.Time.Time.Unit.Nanosecond + +set_option linter.all true + +namespace Lean +namespace Time +namespace Second.Offset + +/-- Convert `Second.Offset` to `Minute.Offset` -/ +@[inline] +def toMinutes (offset : Second.Offset) : Minute.Offset := + offset.div 60 + +/-- Convert `Second.Offset` to `Hour.Offset` -/ +@[inline] +def toHours (offset : Second.Offset) : Hour.Offset := + offset.div 3600 + +end Second.Offset + +namespace Minute.Offset + +/-- Convert `Minute.Offset` to `Hour.Offset` -/ +@[inline] +def toHours (offset : Minute.Offset) : Hour.Offset := + offset.div 60 + +end Minute.Offset + +end Time diff --git a/src/Lean/Time/Time/Unit/Hour.lean b/src/Lean/Time/Time/Unit/Hour.lean new file mode 100644 index 000000000000..947587593f58 --- /dev/null +++ b/src/Lean/Time/Time/Unit/Hour.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat +import Lean.Time.Time.Unit.Minute +import Lean.Time.Time.Unit.Second + +namespace Lean +namespace Time + +set_option linter.all true + +namespace Hour + +/-- +`Ordinal` represents a bounded value for hour, which ranges between 0 and 23. +-/ +def Ordinal := Bounded.LE 0 24 + deriving Repr, BEq, LE, LT + +instance [Le n 24] : OfNat Ordinal n where ofNat := Bounded.LE.ofNat n Le.p + +instance : Inhabited Ordinal where default := 0 + +/-- +`Offset` represents an offset in hour. It is defined as an `Int`. +-/ +def Offset : Type := UnitVal 3600 + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ + +namespace Ordinal + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≤ 24 := by decide) : Ordinal := + Bounded.LE.ofNat data h + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +-/ +@[inline] +def ofFin (data : Fin 25) : Ordinal := + Bounded.LE.ofFin data + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +@[inline] +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal + +namespace Offset + +/-- +Convert the `Hour` offset to a `Second` Offset. +-/ +@[inline] +def toSeconds (val : Offset) : Second.Offset := + val.mul 3600 + +/-- +Convert the `Hour` offset to a `Minute` Offset. +-/ +@[inline] +def toMinutes (val : Offset) : Minute.Offset := + val.mul 60 + +end Offset +end Hour diff --git a/src/Lean/Time/Time/Unit/Minute.lean b/src/Lean/Time/Time/Unit/Minute.lean new file mode 100644 index 000000000000..24410e572a4c --- /dev/null +++ b/src/Lean/Time/Time/Unit/Minute.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat +import Lean.Time.Time.Unit.Second + +namespace Lean +namespace Time + +set_option linter.all true + +namespace Minute + +/-- +`Ordinal` represents a bounded value for minute, which ranges between 0 and 59. +-/ +def Ordinal := Bounded.LE 0 59 + deriving Repr, BEq, LE + +instance [Le n 59] : OfNat Ordinal n where ofNat := Bounded.LE.ofNat n Le.p + +instance : Inhabited Ordinal where default := 0 + +/-- +`Offset` represents an offset in minute. It is defined as an `Int`. +-/ +def Offset : Type := UnitVal 60 + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + +instance : OfNat Offset n := ⟨UnitVal.ofInt <| Int.ofNat n⟩ + +namespace Ordinal + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≤ 59 := by decide) : Ordinal := + Bounded.LE.ofNat data h + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +-/ +@[inline] +def ofFin (data : Fin 60) : Ordinal := + Bounded.LE.ofFin data + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +@[inline] +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal + +namespace Offset + +/-- +Converts an `Offset` to `Second.Offset`. +-/ +@[inline] +def toSeconds (val : Offset) : Second.Offset := + val.mul 60 + +end Offset +end Minute diff --git a/src/Lean/Time/Time/Unit/Nanosecond.lean b/src/Lean/Time/Time/Unit/Nanosecond.lean new file mode 100644 index 000000000000..1f793df2449c --- /dev/null +++ b/src/Lean/Time/Time/Unit/Nanosecond.lean @@ -0,0 +1,45 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat + +namespace Lean +namespace Time + +set_option linter.all true + +namespace Nanosecond + +/-- +`Ordinal` represents a bounded value for second, which ranges between 0 and 60. +This accounts for potential leap second. +-/ +def Ordinal := Bounded.LE 0 999999999 + deriving Repr, BEq, LE, LT + +instance : OfNat Ordinal n where ofNat := Bounded.LE.ofFin (Fin.ofNat n) + +instance : Inhabited Ordinal where default := 0 + +/-- +`Offset` represents an offset in nanoseconds. It is defined as an `Int`. +-/ +def Offset : Type := UnitVal (1 / 1000000000) + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ + +/-- +`Span` represents a bounded value for nanoseconds, ranging between -999999999 and 999999999. +This can be used for operations that involve differences or adjustments within this range. +-/ +def Span := Bounded.LE (-999999999) 999999999 + deriving Repr, BEq, LE, LT + +instance : Inhabited Span where default := Bounded.LE.mk 0 (by decide) diff --git a/src/Lean/Time/Time/Unit/Second.lean b/src/Lean/Time/Time/Unit/Second.lean new file mode 100644 index 000000000000..82a090b7d7c3 --- /dev/null +++ b/src/Lean/Time/Time/Unit/Second.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.UnitVal +import Lean.Time.Bounded +import Lean.Time.LessEq +import Lean.Data.Rat +import Lean.Time.Time.Unit.Nanosecond + +namespace Lean +namespace Time + +set_option linter.all true + +namespace Second + +/-- +`Ordinal` represents a bounded value for second, which ranges between 0 and 60. +This accounts for potential leap second. +-/ +def Ordinal := Bounded.LE 0 60 + deriving Repr, BEq, LE + +instance [Le n 60] : OfNat Ordinal n where ofNat := Bounded.LE.ofNat n Le.p + +instance : Inhabited Ordinal where default := 0 + +/-- +`Offset` represents an offset in second. It is defined as an `Int`. +-/ +def Offset : Type := UnitVal 1 + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ + +namespace Ordinal + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≤ 60 := by decide) : Ordinal := + Bounded.LE.ofNat data h + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +-/ +@[inline] +def ofFin (data : Fin 61) : Ordinal := + Bounded.LE.ofFin data + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +@[inline] +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal + +namespace Offset + +def ofNanosecond (offset : Nanosecond.Offset) : Second.Offset × Nanosecond.Span := + let second := offset.val.div 1000000000 + let nano := Bounded.LE.byMod offset.val 1000000000 (by decide) + ⟨UnitVal.ofInt second, nano⟩ + +end Offset +end Second diff --git a/src/Lean/Time/TimeZone.lean b/src/Lean/Time/TimeZone.lean new file mode 100644 index 000000000000..6a7fe1250b11 --- /dev/null +++ b/src/Lean/Time/TimeZone.lean @@ -0,0 +1,9 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.TimeZone.Basic +import Lean.Time.TimeZone.DateTime +import Lean.Time.TimeZone.ZonedDateTime diff --git a/src/Lean/Time/TimeZone/Basic.lean b/src/Lean/Time/TimeZone/Basic.lean new file mode 100644 index 000000000000..41ea2ee86f7e --- /dev/null +++ b/src/Lean/Time/TimeZone/Basic.lean @@ -0,0 +1,6 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +import Lean.Time.TimeZone.Offset diff --git a/src/Lean/Time/TimeZone/DateTime.lean b/src/Lean/Time/TimeZone/DateTime.lean new file mode 100644 index 000000000000..450096602673 --- /dev/null +++ b/src/Lean/Time/TimeZone/DateTime.lean @@ -0,0 +1,88 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time +import Lean.Time.Date +import Lean.Time.DateTime +import Lean.Time.TimeZone.TimeZone + +namespace Lean +namespace TimeZone +open Time Date DateTime + +/-- +It stores a `Timestamp`, a `NaiveDateTime` and a `TimeZone` +-/ +structure DateTime (tz : TimeZone) where + private mk :: + timestamp : Timestamp + date : NaiveDateTime + deriving Repr + +namespace DateTime + +/-- +Creates a new DateTime out of a `Timestamp` +-/ +def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := + let date := (tm + tz.toSeconds).toNaiveDateTime + DateTime.mk tm date + +/-- +Creates a new DateTime out of a `NaiveDateTime` +-/ +def ofNaiveDateTime (date : NaiveDateTime) (tz : TimeZone) : DateTime tz := + let tm := date.toTimestamp + DateTime.mk tm date + +/-- +Getter for the `Year` inside of a `DateTime` +-/ +@[inline] +def year (dt : DateTime tz) : Year.Offset := + dt.date.year + +/-- +Getter for the `Month` inside of a `DateTime` +-/ +@[inline] +def month (dt : DateTime tz) : Month.Ordinal := + dt.date.month + +/-- +Getter for the `Day` inside of a `DateTime` +-/ +@[inline] +def day (dt : DateTime tz) : Day.Ordinal := + dt.date.day + +/-- +Getter for the `Hour` inside of a `DateTime` +-/ +@[inline] +def hour (dt : DateTime tz) : Hour.Ordinal := + dt.date.hour + +/-- +Getter for the `Minute` inside of a `DateTime` +-/ +@[inline] +def minute (dt : DateTime tz) : Minute.Ordinal := + dt.date.minute + +/-- +Getter for the `Second` inside of a `DateTime` +-/ +@[inline] +def second (dt : DateTime tz) : Second.Ordinal := + dt.date.second + +/-- +Gets the `Weekday` of a DateTime. +-/ +@[inline] +def weekday (dt : DateTime tz) : Weekday := + dt.date.date.weekday diff --git a/src/Lean/Time/TimeZone/Offset.lean b/src/Lean/Time/TimeZone/Offset.lean new file mode 100644 index 000000000000..a8ffa3ec6e0a --- /dev/null +++ b/src/Lean/Time/TimeZone/Offset.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time.Unit.Basic + +namespace Lean +namespace TimeZone +open Time + +/-- +Represents a timezone offset with an hour and second component. +-/ +structure Offset where + hour: Hour.Offset + second: Second.Offset + deriving Repr, Inhabited + +namespace Offset + +/-- +Converts an `Offset` to a string in ISO 8601 format. The `colon` parameter determines if the hour +and minute components are separated by a colon (e.g., "+01:00" or "+0100"). +-/ +def toIsoString (offset : Offset) (colon : Bool) : String := + let (sign, time) := if offset.second.val ≥ 0 then ("+", offset.second) else ("-", -offset.second) + let hour : Hour.Offset := time.div 3600 + let minute := Int.div (Int.mod time.val 3600) 60 + let hourStr := if hour.val < 10 then s!"0{hour.val}" else toString hour.val + let minuteStr := if minute < 10 then s!"0{minute}" else toString minute + if colon then s!"{sign}{hourStr}:{minuteStr}" + else s!"{sign}{hourStr}{minuteStr}" + +/-- +A zero `Offset` representing UTC (no offset). +-/ +def zero : Offset := + { hour := 0, second := 0 } + +/-- +Creates an `Offset` from a given number of hour. +-/ +def ofHours (n: Hour.Offset) : Offset := + mk n n.toSeconds + +/-- +Creates an `Offset` from a given number of second. +-/ +def ofSeconds (n: Second.Offset) : Offset := + mk n.toHours n diff --git a/src/Lean/Time/TimeZone/TimeZone.lean b/src/Lean/Time/TimeZone/TimeZone.lean new file mode 100644 index 000000000000..092a7702d10b --- /dev/null +++ b/src/Lean/Time/TimeZone/TimeZone.lean @@ -0,0 +1,51 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time.Unit.Basic +import Lean.Time.TimeZone.Basic + +namespace Lean +namespace TimeZone +open Time + +/-- +An enumeration representing different time zones. +-/ +structure TimeZone where + offset : Offset + name : String + +namespace TimeZone + +/-- +A zeroed `Timezone` representing UTC (no offset). +-/ +def UTC (name : String) : TimeZone := + TimeZone.mk (Offset.zero) name + +/-- +A zeroed `Timezone` representing GMT (no offset). +-/ +def GMT (name : String) : TimeZone := + TimeZone.mk (Offset.zero) name + +/-- +Creates a `Timestamp` from a given number of hour. +-/ +def ofHours (name : String) (n: Hour.Offset) : TimeZone := + TimeZone.mk (Offset.ofHours n) name + +/-- +Creates a `Timestamp` from a given number of second. +-/ +def ofSeconds (name : String) (n: Second.Offset) : TimeZone := + TimeZone.mk (Offset.ofSeconds n) name + +/-- +Gets the number of seconds in a timezone offset. +-/ +def toSeconds (tz : TimeZone) : Second.Offset := + tz.offset.second diff --git a/src/Lean/Time/TimeZone/ZonedDateTime.lean b/src/Lean/Time/TimeZone/ZonedDateTime.lean new file mode 100644 index 000000000000..9f73d44d9432 --- /dev/null +++ b/src/Lean/Time/TimeZone/ZonedDateTime.lean @@ -0,0 +1,88 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Lean.Time.Time +import Lean.Time.Date +import Lean.Time.DateTime +import Lean.Time.TimeZone.TimeZone +import Lean.Time.TimeZone.DateTime + +namespace Lean +namespace TimeZone +open Time Date DateTime + +def ZonedDateTime := Sigma DateTime + +namespace ZonedDateTime +open DateTime + +/-- +Creates a new `ZonedDateTime` out of a `Timestamp` +-/ +def ofTimestamp (tm : DateTime.Timestamp) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofTimestamp tm tz⟩ + +/-- +Creates a new `ZonedDateTime` out of a `NaiveDateTime` +-/ +def ofNaiveDateTime (date : DateTime.NaiveDateTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofNaiveDateTime date tz⟩ + +/-- +Getter for the `Year` inside of a `ZonedDateTime` +-/ +@[inline] +def year (zdt : ZonedDateTime) : Year.Offset := + zdt.snd.year + +/-- +Getter for the `Month` inside of a `ZonedDateTime` +-/ +@[inline] +def month (zdt : ZonedDateTime) : Month.Ordinal := + zdt.snd.month + +/-- +Getter for the `Day` inside of a `ZonedDateTime` +-/ +@[inline] +def day (zdt : ZonedDateTime) : Day.Ordinal := + zdt.snd.day + +/-- +Getter for the `Hour` inside of a `ZonedDateTime` +-/ +@[inline] +def hour (zdt : ZonedDateTime) : Hour.Ordinal := + zdt.snd.hour + +/-- +Getter for the `Minute` inside of a `ZonedDateTime` +-/ +@[inline] +def minute (zdt : ZonedDateTime) : Minute.Ordinal := + zdt.snd.minute + +/-- +Getter for the `Second` inside of a `ZonedDateTime` +-/ +@[inline] +def second (zdt : ZonedDateTime) : Second.Ordinal := + zdt.snd.second + +/-- +Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` +-/ +@[inline] +def offset (zdt : ZonedDateTime) : Offset := + zdt.fst.offset + +/-- +Returns the weekday. +-/ +@[inline] +def weekday (zdt : ZonedDateTime) : Weekday := + zdt.snd.weekday diff --git a/src/Lean/Time/UnitVal.lean b/src/Lean/Time/UnitVal.lean new file mode 100644 index 000000000000..bc7c0221d0f3 --- /dev/null +++ b/src/Lean/Time/UnitVal.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Init +import Lean.Data.Rat + +namespace Lean +set_option linter.all true + +/-- +A structure representing a unit of a given ratio type `α`. +-/ +structure UnitVal (α : Rat) where + private mk :: + /-- + Value inside the UnitVal Value + -/ + val : Int + deriving Inhabited, BEq + +namespace UnitVal + +/-- +Creates a `UnitVal` from an `Int`. +-/ +@[inline] +def ofInt (value : Int) : UnitVal α := + ⟨value⟩ + +/-- +Creates a `UnitVal` from a `Nat`. +-/ +@[inline] +def ofNat (value : Nat) : UnitVal α := + ⟨value⟩ + +/-- +Converts a `UnitVal` to an `Int`. +-/ +@[inline] +def toInt (unit : UnitVal α) : Int := + unit.val + +/-- +Multiplies the `UnitVal` by an `Int`, resulting in a new `UnitVal` with an adjusted ratio. +-/ +@[inline] +def mul (unit : UnitVal a) (factor : Int) : UnitVal (a / factor) := + ⟨unit.val * factor⟩ + +/-- +Divides the `UnitVal` by an `Int`, resulting in a new `UnitVal` with an adjusted ratio. +-/ +@[inline] +def div (unit : UnitVal a) (divisor : Int) : UnitVal (a * divisor) := + ⟨unit.val / divisor⟩ + +/-- +Adds two `UnitVal` values of the same ratio. +-/ +@[inline] +def add (u1 : UnitVal α) (u2 : UnitVal α) : UnitVal α := + ⟨u1.val + u2.val⟩ + +/-- +Subtracts one `UnitVal` value from another of the same ratio. +-/ +@[inline] +def sub (u1 : UnitVal α) (u2 : UnitVal α) : UnitVal α := + ⟨u1.val - u2.val⟩ + +/-- +Converts an `Offset` to another unit type. +-/ +def convert (val : UnitVal a) : UnitVal b := + let ratio := a.div b + ofInt <| val.toInt * ratio.num / ratio.den + +instance : OfNat (UnitVal α) n where ofNat := ⟨Int.ofNat n⟩ +instance : Repr (UnitVal α) where reprPrec x p := reprPrec x.val p +instance : LE (UnitVal α) where le x y := x.val ≤ y.val +instance : LT (UnitVal α) where lt x y := x.val < y.val +instance : Add (UnitVal α) where add x y := ⟨x.val + y.val⟩ +instance : Sub (UnitVal α) where sub x y := ⟨x.val - y.val⟩ +instance : Mul (UnitVal α) where mul x y := ⟨x.val * y.val⟩ +instance : Div (UnitVal α) where div x y := ⟨x.val / y.val⟩ +instance : Neg (UnitVal α) where neg x := ⟨-x.val⟩ +instance : ToString (UnitVal n) where toString n := toString n.val From 856f569aae73d87d917104cc07f8e315beb4ccdb Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 2 Aug 2024 18:56:38 -0300 Subject: [PATCH 002/118] chore: remove useless piece of code --- src/Lean/Time/Format.lean | 20 -------------------- 1 file changed, 20 deletions(-) diff --git a/src/Lean/Time/Format.lean b/src/Lean/Time/Format.lean index 4732a827cac4..65f9aa7685d2 100644 --- a/src/Lean/Time/Format.lean +++ b/src/Lean/Time/Format.lean @@ -390,23 +390,3 @@ private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : Sin private def addData (data : DateBuilder) : FormatPart → DateBuilder | .string s => data | .modifier _ m => addDataInDateTime data m sorry - --- Internals - -/- -def Format.format (x: Format) (date : ZonedDateTime) : String := - x.map (formatPartWithDate date) - |> String.join - -def Formats.ISO8601 := Format.spec! "%YYYY-%MM-%DD'T'%hh:%mm:%ss'Z'" - --- Tests - -def x : Timestamp := 1722631000 - -def UTCM3 := (TimeZone.mk (Offset.ofSeconds (-10800)) "Brasilia") - -def date := ZonedDateTime.ofTimestamp x UTCM3 - -#eval Formats.ISO8601.format date --/ From a9bf653bf368534ad3535f8f780efb848a63fccc Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 5 Aug 2024 15:13:35 -0300 Subject: [PATCH 003/118] feat: format and moved the folder to the right place --- src/Lean/Time/Date/Unit/Basic.lean | 21 - src/Lean/Time/Format.lean | 392 ----------- src/Lean/Time/Time.lean | 10 - src/Lean/Time/TimeZone.lean | 9 - src/Std/Time.lean | 10 + src/{Lean => Std}/Time/Bounded.lean | 23 +- .../Basic.lean => Std/Time/Date.lean} | 6 +- .../Time/Time => Std/Time/Date}/Basic.lean | 2 +- src/{Lean => Std}/Time/Date/Date.lean | 21 +- src/{Lean => Std}/Time/Date/Scalar.lean | 8 +- src/Std/Time/Date/Unit/Basic.lean | 25 + src/{Lean => Std}/Time/Date/Unit/Day.lean | 16 +- src/{Lean => Std}/Time/Date/Unit/Month.lean | 18 +- .../Time/Date/Unit/WeekOfYear.lean | 17 +- src/{Lean => Std}/Time/Date/Unit/Weekday.lean | 18 +- src/{Lean => Std}/Time/Date/Unit/Year.lean | 16 +- src/{Lean => Std}/Time/Date/WeekDate.lean | 22 +- src/{Lean => Std}/Time/DateTime.lean | 11 +- .../Time/DateTime/Basic.lean} | 2 +- .../Time/DateTime/NaiveDateTime.lean | 23 +- .../Time/DateTime/Timestamp.lean | 17 +- src/{Lean => Std}/Time/Duration.lean | 21 +- src/Std/Time/Format.lean | 200 ++++++ src/Std/Time/Format/Basic.lean | 648 ++++++++++++++++++ src/{Lean => Std}/Time/LessEq.lean | 3 +- src/{Lean => Std}/Time/Sign.lean | 18 +- src/Std/Time/Time.lean | 10 + .../Time/Date => Std/Time/Time}/Basic.lean | 2 +- src/{Lean => Std}/Time/Time/HourMarker.lean | 8 +- src/{Lean => Std}/Time/Time/Scalar.lean | 26 +- src/{Lean => Std}/Time/Time/Time.lean | 23 +- src/{Lean => Std}/Time/Time/Unit/Basic.lean | 17 +- src/{Lean => Std}/Time/Time/Unit/Hour.lean | 17 +- src/Std/Time/Time/Unit/Millisecond.lean | 36 + src/{Lean => Std}/Time/Time/Unit/Minute.lean | 15 +- .../Time/Time/Unit/Nanosecond.lean | 30 +- src/{Lean => Std}/Time/Time/Unit/Second.lean | 24 +- src/{Lean => Std}/Time/UnitVal.lean | 19 +- .../Time/Date.lean => Std/Time/Zoned.lean} | 6 +- .../TimeZone => Std/Time/Zoned}/Basic.lean | 2 +- .../TimeZone => Std/Time/Zoned}/DateTime.lean | 43 +- .../TimeZone => Std/Time/Zoned}/Offset.lean | 19 +- .../TimeZone => Std/Time/Zoned}/TimeZone.lean | 20 +- .../Time/Zoned}/ZonedDateTime.lean | 54 +- 44 files changed, 1333 insertions(+), 615 deletions(-) delete mode 100644 src/Lean/Time/Date/Unit/Basic.lean delete mode 100644 src/Lean/Time/Format.lean delete mode 100644 src/Lean/Time/Time.lean delete mode 100644 src/Lean/Time/TimeZone.lean create mode 100644 src/Std/Time.lean rename src/{Lean => Std}/Time/Bounded.lean (92%) rename src/{Lean/Time/DateTime/Basic.lean => Std/Time/Date.lean} (59%) rename src/{Lean/Time/Time => Std/Time/Date}/Basic.lean (83%) rename src/{Lean => Std}/Time/Date/Date.lean (92%) rename src/{Lean => Std}/Time/Date/Scalar.lean (94%) create mode 100644 src/Std/Time/Date/Unit/Basic.lean rename src/{Lean => Std}/Time/Date/Unit/Day.lean (93%) rename src/{Lean => Std}/Time/Date/Unit/Month.lean (96%) rename src/{Lean => Std}/Time/Date/Unit/WeekOfYear.lean (88%) rename src/{Lean => Std}/Time/Date/Unit/Weekday.lean (90%) rename src/{Lean => Std}/Time/Date/Unit/Year.lean (86%) rename src/{Lean => Std}/Time/Date/WeekDate.lean (70%) rename src/{Lean => Std}/Time/DateTime.lean (75%) rename src/{Lean/Time/Date/Unit/Lemmas.lean => Std/Time/DateTime/Basic.lean} (87%) rename src/{Lean => Std}/Time/DateTime/NaiveDateTime.lean (94%) rename src/{Lean => Std}/Time/DateTime/Timestamp.lean (65%) rename src/{Lean => Std}/Time/Duration.lean (87%) create mode 100644 src/Std/Time/Format.lean create mode 100644 src/Std/Time/Format/Basic.lean rename src/{Lean => Std}/Time/LessEq.lean (97%) rename src/{Lean => Std}/Time/Sign.lean (73%) create mode 100644 src/Std/Time/Time.lean rename src/{Lean/Time/Date => Std/Time/Time}/Basic.lean (83%) rename src/{Lean => Std}/Time/Time/HourMarker.lean (89%) rename src/{Lean => Std}/Time/Time/Scalar.lean (58%) rename src/{Lean => Std}/Time/Time/Time.lean (53%) rename src/{Lean => Std}/Time/Time/Unit/Basic.lean (79%) rename src/{Lean => Std}/Time/Time/Unit/Hour.lean (90%) create mode 100644 src/Std/Time/Time/Unit/Millisecond.lean rename src/{Lean => Std}/Time/Time/Unit/Minute.lean (91%) rename src/{Lean => Std}/Time/Time/Unit/Nanosecond.lean (70%) rename src/{Lean => Std}/Time/Time/Unit/Second.lean (77%) rename src/{Lean => Std}/Time/UnitVal.lean (91%) rename src/{Lean/Time/Date.lean => Std/Time/Zoned.lean} (63%) rename src/{Lean/Time/TimeZone => Std/Time/Zoned}/Basic.lean (82%) rename src/{Lean/Time/TimeZone => Std/Time/Zoned}/DateTime.lean (68%) rename src/{Lean/Time/TimeZone => Std/Time/Zoned}/Offset.lean (84%) rename src/{Lean/Time/TimeZone => Std/Time/Zoned}/TimeZone.lean (74%) rename src/{Lean/Time/TimeZone => Std/Time/Zoned}/ZonedDateTime.lean (58%) diff --git a/src/Lean/Time/Date/Unit/Basic.lean b/src/Lean/Time/Date/Unit/Basic.lean deleted file mode 100644 index a80b6b75eaf8..000000000000 --- a/src/Lean/Time/Date/Unit/Basic.lean +++ /dev/null @@ -1,21 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Lean.Time.UnitVal -import Lean.Time.Date.Unit.Day -import Lean.Time.Date.Unit.Month -import Lean.Time.Date.Unit.Year -import Lean.Time.Date.Unit.Weekday -import Lean.Time.Date.Unit.WeekOfYear - -namespace Lean -namespace Date.Day.Ordinal.OfYear - -@[inline] -def toMonthAndDay (year : Year.Offset) (ordinal : OfYear year.isLeap) : { val : Month.Ordinal × Ordinal // Year.Offset.valid year (Prod.fst val) (Prod.snd val) } := - Month.Ordinal.ofOrdinal ordinal - -end Date.Day.Ordinal.OfYear diff --git a/src/Lean/Time/Format.lean b/src/Lean/Time/Format.lean deleted file mode 100644 index 65f9aa7685d2..000000000000 --- a/src/Lean/Time/Format.lean +++ /dev/null @@ -1,392 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Lean.Time.Date -import Lean.Time.Time -import Lean.Time.DateTime -import Lean.Time.TimeZone -import Lean.Data.Parsec - -namespace Lean -namespace Format -open Lean.Parsec Time Date TimeZone DateTime - -private inductive Modifier - | YYYY - | YY - | MMMM - | MMM - | MM - | M - | DD - | D - | EEEE - | EEE - | hh - | h - | HH - | H - | AA - | aa - | mm - | m - | ss - | s - | ZZZZZ - | ZZZZ - | ZZZ - | Z - deriving Repr - -/-- -The part of a formatting string. a string is just a text and a modifier is in the format `%0T` where -0 is the quantity of left pad and `T` the `Modifier`. --/ -private inductive FormatPart - | string (val : String) - | modifier (left_pad: Nat) (modifier : Modifier) - deriving Repr - -/-- -The format of date and time string. --/ -abbrev Format := List FormatPart - -private def isLetter (c : Char) : Bool := - (c >= 'a' && c <= 'z') || (c >= 'A' && c <= 'Z') - -private def isNonLetter (c: Char) : Bool := ¬isLetter c - -private def parseModifier : Lean.Parsec Modifier - := pstring "YYYY" *> pure .YYYY - <|> pstring "YY" *> pure .YY - <|> pstring "MMMM" *> pure .MMMM - <|> pstring "MMM" *> pure .MMM - <|> pstring "MM" *> pure .MM - <|> pstring "M" *> pure .M - <|> pstring "DD" *> pure .DD - <|> pstring "D" *> pure .D - <|> pstring "EEEE" *> pure .EEEE - <|> pstring "EEE" *> pure .EEE - <|> pstring "hh" *> pure .hh - <|> pstring "h" *> pure .h - <|> pstring "HH" *> pure .HH - <|> pstring "H" *> pure .H - <|> pstring "AA" *> pure .AA - <|> pstring "aa" *> pure .aa - <|> pstring "mm" *> pure .mm - <|> pstring "m" *> pure .m - <|> pstring "ss" *> pure .ss - <|> pstring "s" *> pure .s - <|> pstring "ZZZZZ" *> pure .ZZZZZ - <|> pstring "ZZZZ" *> pure .ZZZZ - <|> pstring "ZZZ" *> pure .ZZZ - <|> pstring "Z" *> pure .Z - -private def pnumber : Lean.Parsec Nat := do - let numbers ← manyChars digit - return numbers.foldl (λacc char => acc * 10 + (char.toNat - 48)) 0 - -private def parseFormatPart : Lean.Parsec FormatPart - := (.modifier <$> (pchar '%' *> pnumber) <*> parseModifier) - <|> (pchar '\\') *> anyChar <&> (.string ∘ toString) - <|> (pchar '\"' *> many1Chars (satisfy (· ≠ '\"')) <* pchar '\"') <&> .string - <|> (pchar '\'' *> many1Chars (satisfy (· ≠ '\'')) <* pchar '\'') <&> .string - <|> many1Chars (satisfy (fun x => x ≠ '%' ∧ x ≠ '\'' ∧ x ≠ '\"')) <&> .string - -private def specParser : Lean.Parsec Format := - (Array.toList <$> Lean.Parsec.many parseFormatPart) <* eof - -private def specParse (s: String) : Except String Format := - specParser.run s - -def Format.spec! (s: String) : Format := - match specParser.run s with - | .ok s => s - | .error s => panic! s - --- Pretty printer - -private def unabbrevMonth (month: Month.Ordinal) : String := - match month.val, month.property with - | 1, _ => "January" - | 2, _ => "February" - | 3, _ => "March" - | 4, _ => "April" - | 5, _ => "May" - | 6, _ => "June" - | 7, _ => "July" - | 8, _ => "August" - | 9, _ => "September" - | 10, _ => "October" - | 11, _ => "November" - | 12, _ => "December" - -private def abbrevMonth (month: Month.Ordinal) : String := - match month.val, month.property with - | 1, _ => "Jan" - | 2, _ => "Feb" - | 3, _ => "Mar" - | 4, _ => "Apr" - | 5, _ => "May" - | 6, _ => "Jun" - | 7, _ => "Jul" - | 8, _ => "Aug" - | 9, _ => "Sep" - | 10, _ => "Oct" - | 11, _ => "Nov" - | 12, _ => "Dec" - -private def abbrevDayOfWeek (day: Weekday) : String := - match day with - | .sun => "Sun" - | .mon => "Mon" - | .tue => "Tue" - | .wed => "Wed" - | .thu => "Thu" - | .fri => "Fri" - | .sat => "Sat" - -private def dayOfWeek (day: Weekday) : String := - match day with - | .sun => "Sunday" - | .mon => "Monday" - | .tue => "Tuesday" - | .wed => "Wednesday" - | .thu => "Thusday" - | .fri => "Friday" - | .sat => "Saturday" - -private def leftPad (n : Nat) (a : Char) (s : String) : String := - "".pushn a (n - s.length) ++ s - -private def formatWithDate (date : ZonedDateTime) : Modifier → String - | .YYYY => s!"{leftPad 4 '0' (toString date.year)}" - | .YY => s!"{leftPad 2 '0' (toString $ date.year.toNat % 100)}" - | .MMMM => unabbrevMonth date.month - | .MMM => abbrevMonth date.month - | .MM => s!"{leftPad 2 '0' (toString $ date.month.toNat)}" - | .M => s!"{date.month.toNat}" - | .DD => s!"{leftPad 2 '0' (toString $ date.day.toNat)}" - | .D => s!"{date.day.toNat}" - | .EEEE => dayOfWeek date.weekday - | .EEE => abbrevDayOfWeek date.weekday - | .hh => s!"{leftPad 2 '0' (toString date.hour.toNat)}" - | .h => s!"{date.hour.toNat}" - | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' $ toString hour}" - | .H => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{hour}" - | .AA => if date.hour.toNat < 12 then "AM" else "PM" - | .aa => if date.hour.toNat < 12 then "am" else "pm" - | .mm => s!"{leftPad 2 '0' $ toString date.minute.toNat}" - | .m => s!"{date.minute.toNat}" - | .ss => s!"{leftPad 2 '0' $ toString date.second.toNat}" - | .s => s!"{date.second.toNat}" - | .ZZZZZ => date.offset.toIsoString true - | .ZZZZ => date.offset.toIsoString false - | .ZZZ => if date.offset.second.val = 0 then "UTC" else date.offset.toIsoString false - | .Z => if date.offset.second.val = 0 then "Z" else date.offset.toIsoString true - -private def formatPartWithDate (date : ZonedDateTime) : FormatPart → String - | .string s => s - | .modifier p t => leftPad p ' ' (formatWithDate date t) - --- Parser - -@[simp] -private def SingleFormatType : Modifier → Type - | .YYYY => Year.Offset - | .YY => Year.Offset - | .MMMM => Month.Ordinal - | .MMM => Month.Ordinal - | .MM => Month.Ordinal - | .M => Month.Ordinal - | .DD => Day.Ordinal - | .D => Day.Ordinal - | .EEEE => Weekday - | .EEE => Weekday - | .hh => Hour.Ordinal - | .h => Hour.Ordinal - | .HH => Hour.Ordinal - | .H => Hour.Ordinal - | .AA => HourMarker - | .aa => HourMarker - | .mm => Minute.Ordinal - | .m => Minute.Ordinal - | .ss => Second.Ordinal - | .s => Second.Ordinal - | .ZZZZZ => Offset - | .ZZZZ => Offset - | .ZZZ => Offset - | .Z => Offset - -private def transform (n: β → Option α) (p: Lean.Parsec β) : Lean.Parsec α := do - let res ← p - match n res with - | some n => pure n - | none => fail "cannot parse" - -private def parseMonth : Lean.Parsec Month.Ordinal - := (pstring "Jan" *> pure 1) - <|> (pstring "Feb" *> pure 2) - <|> (pstring "Mar" *> pure 3) - <|> (pstring "Apr" *> pure 4) - <|> (pstring "May" *> pure 5) - <|> (pstring "Jun" *> pure 6) - <|> (pstring "Jul" *> pure 7) - <|> (pstring "Aug" *> pure 8) - <|> (pstring "Sep" *> pure 9) - <|> (pstring "Oct" *> pure 10) - <|> (pstring "Nov" *> pure 11) - <|> (pstring "Dec" *> pure 12) - -private def parseMonthUnabbrev : Lean.Parsec Month.Ordinal - := (pstring "January" *> pure 1) - <|> (pstring "February" *> pure 2) - <|> (pstring "March" *> pure 3) - <|> (pstring "April" *> pure 4) - <|> (pstring "May" *> pure 5) - <|> (pstring "June" *> pure 6) - <|> (pstring "July" *> pure 7) - <|> (pstring "August" *> pure 8) - <|> (pstring "September" *> pure 9) - <|> (pstring "October" *> pure 10) - <|> (pstring "November" *> pure 11) - <|> (pstring "December" *> pure 12) - -private def parseWeekday : Lean.Parsec Weekday - := (pstring "Mon" *> pure Weekday.mon) - <|> (pstring "Tue" *> pure Weekday.tue) - <|> (pstring "Wed" *> pure Weekday.wed) - <|> (pstring "Thu" *> pure Weekday.thu) - <|> (pstring "Fri" *> pure Weekday.fri) - <|> (pstring "Sat" *> pure Weekday.sat) - <|> (pstring "Sun" *> pure Weekday.sun) - -private def parseWeekdayUnnabrev : Lean.Parsec Weekday - := (pstring "Monday" *> pure Weekday.mon) - <|> (pstring "Tuesday" *> pure Weekday.tue) - <|> (pstring "Wednesday" *> pure Weekday.wed) - <|> (pstring "Thursday" *> pure Weekday.thu) - <|> (pstring "Friday" *> pure Weekday.fri) - <|> (pstring "Saturday" *> pure Weekday.sat) - <|> (pstring "Sunday" *> pure Weekday.sun) - -private def parserUpperHourMarker : Lean.Parsec HourMarker - := (pstring "AM" *> pure HourMarker.am) - <|> (pstring "PM" *> pure HourMarker.pm) - -private def parserLowerHourMarker : Lean.Parsec HourMarker - := (pstring "am" *> pure HourMarker.am) - <|> (pstring "pm" *> pure HourMarker.pm) - -private def twoDigit : Lean.Parsec Int := do - let digit1 ← Lean.Parsec.digit - let digit2 ← Lean.Parsec.digit - return String.toNat! s!"{digit1}{digit2}" - -private def parseYearTwo : Lean.Parsec Int :=do - let year ← twoDigit - return if year < 70 then 2000 + year else 1900 + year - -private def timeOffset (colon: Bool) : Lean.Parsec Offset := do - let sign : Int ← (pstring "-" *> pure (-1)) <|> (pstring "+" *> pure 1) - let hour ← twoDigit - if colon then let _ ← pstring ":" - let minutes ← twoDigit - let res := (hour * 3600 + minutes * 60) * sign - pure (Offset.ofSeconds (UnitVal.ofInt res)) - -private def timeOrUTC (utcString: String) (colon: Bool) : Lean.Parsec Offset := - (pstring utcString *> pure Offset.zero) <|> timeOffset colon - -private def number : Lean.Parsec Nat := do - String.toNat! <$> Lean.Parsec.many1Chars Lean.Parsec.digit - -private def singleDigit : Lean.Parsec Nat := do - let digit1 ← Lean.Parsec.digit - return String.toNat! s!"{digit1}" - -private def fourDigit : Lean.Parsec Int := do - let digit1 ← Lean.Parsec.digit - let digit2 ← Lean.Parsec.digit - let digit3 ← Lean.Parsec.digit - let digit4 ← Lean.Parsec.digit - return String.toNat! s!"{digit1}{digit2}{digit3}{digit4}" - -private def parserWithFormat : (typ: Modifier) → Lean.Parsec (SingleFormatType typ) - | .YYYY => fourDigit - | .YY => parseYearTwo - | .MMMM => parseMonthUnabbrev - | .MMM => parseMonth - | .MM => transform Bounded.LE.ofInt twoDigit - | .M => transform Bounded.LE.ofInt number - | .DD => transform Bounded.LE.ofInt twoDigit - | .D => transform Bounded.LE.ofInt number - | .EEEE => parseWeekdayUnnabrev - | .EEE => parseWeekday - | .hh => transform Bounded.LE.ofInt twoDigit - | .h => transform Bounded.LE.ofInt number - | .HH => transform Bounded.LE.ofInt twoDigit - | .H => transform Bounded.LE.ofInt number - | .AA => parserUpperHourMarker - | .aa => parserLowerHourMarker - | .mm => transform Bounded.LE.ofInt twoDigit - | .m => transform Bounded.LE.ofInt number - | .ss => transform Bounded.LE.ofInt twoDigit - | .s => transform Bounded.LE.ofInt number - | .ZZZZZ => timeOffset true - | .ZZZZ => timeOffset false - | .ZZZ => timeOrUTC "UTC" false - | .Z => timeOrUTC "Z" true - -structure DateBuilder where - tz : Offset := Offset.zero - year : Year.Offset := 0 - month : Month.Ordinal := 1 - day : Day.Ordinal := 1 - hour : Hour.Ordinal := 0 - minute : Minute.Ordinal := 0 - second : Second.Ordinal := 0 - -def DateBuilder.build (builder : DateBuilder) : ZonedDateTime := - { - fst := TimeZone.mk builder.tz "GMT" - snd := DateTime.ofNaiveDateTime { - date := Date.force builder.year builder.month builder.day - time := Time.mk builder.hour builder.minute builder.second 0 - } (TimeZone.mk builder.tz "GMT") - } - -private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : SingleFormatType typ) : DateBuilder := - match typ with - | .YYYY => { data with year := value } - | .YY => { data with year := value } - | .MMMM => { data with month := value } - | .MMM => { data with month := value } - | .MM => { data with month := value } - | .M => { data with month := value } - | .DD => { data with day := value } - | .D => { data with day := value } - | .EEEE => data - | .EEE => data - | .hh => { data with hour := value } - | .h => { data with hour := value } - | .HH => { data with hour := value } - | .H => { data with hour := value } - | .AA => { data with hour := HourMarker.toAbsolute! value data.hour } - | .aa => { data with hour := HourMarker.toAbsolute! value data.hour } - | .mm => { data with minute := value } - | .m => { data with minute := value } - | .ss => { data with second := value } - | .s => { data with second := value } - | .ZZZZZ => { data with tz := value } - | .ZZZZ => { data with tz := value } - | .ZZZ => { data with tz := value } - | .Z => { data with tz := value } - -private def addData (data : DateBuilder) : FormatPart → DateBuilder - | .string s => data - | .modifier _ m => addDataInDateTime data m sorry diff --git a/src/Lean/Time/Time.lean b/src/Lean/Time/Time.lean deleted file mode 100644 index a9b295ed8d36..000000000000 --- a/src/Lean/Time/Time.lean +++ /dev/null @@ -1,10 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Lean.Time.Time.HourMarker -import Lean.Time.Time.Basic -import Lean.Time.Time.Scalar -import Lean.Time.Time.Time diff --git a/src/Lean/Time/TimeZone.lean b/src/Lean/Time/TimeZone.lean deleted file mode 100644 index 6a7fe1250b11..000000000000 --- a/src/Lean/Time/TimeZone.lean +++ /dev/null @@ -1,9 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Lean.Time.TimeZone.Basic -import Lean.Time.TimeZone.DateTime -import Lean.Time.TimeZone.ZonedDateTime diff --git a/src/Std/Time.lean b/src/Std/Time.lean new file mode 100644 index 000000000000..a5f786b1806f --- /dev/null +++ b/src/Std/Time.lean @@ -0,0 +1,10 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +import Std.Time.Time +import Std.Time.Date +import Std.Time.DateTime +import Std.Time.Zoned +import Std.Time.Format diff --git a/src/Lean/Time/Bounded.lean b/src/Std/Time/Bounded.lean similarity index 92% rename from src/Lean/Time/Bounded.lean rename to src/Std/Time/Bounded.lean index 80c3d9dcc4d4..43e0edd24514 100644 --- a/src/Lean/Time/Bounded.lean +++ b/src/Std/Time/Bounded.lean @@ -5,9 +5,11 @@ Authors: Sofia Rodrigues -/ prelude import Init.Data.Int -import Lean.Time.LessEq +import Std.Time.LessEq + +namespace Std +namespace Time -namespace Lean set_option linter.all true in /-- @@ -322,6 +324,23 @@ def mod (bounded : Bounded.LE n num) (num : Int) (hi : 0 < num) : Bounded.LE (- byMod bounded.val num hi /-- +Adjust the bounds of a `Bounded` by applying the multiplication operation with a positive number. +-/ +@[inline] +def mul_pos (bounded : Bounded.LE n m) (num : Int) (h : num ≥ 0) : Bounded.LE (n * num) (m * num) := by + refine ⟨bounded.val * num, And.intro ?_ ?_⟩ + · exact Int.mul_le_mul_of_nonneg_right bounded.property.left h + · exact Int.mul_le_mul_of_nonneg_right bounded.property.right h + +/-- +Adjust the bounds of a `Bounded` by applying the multiplication operation with a positive number. +-/ +@[inline] +def mul_neg (bounded : Bounded.LE n m) (num : Int) (h : num ≤ 0) : Bounded.LE (m * num) (n * num) := by + refine ⟨bounded.val * num, And.intro ?_ ?_⟩ + · exact Int.mul_le_mul_of_nonpos_right bounded.property.right h + · exact Int.mul_le_mul_of_nonpos_right bounded.property.left h +/-- Adjust the bounds of a `Bounded` by applying the div operation. -/ @[inline] diff --git a/src/Lean/Time/DateTime/Basic.lean b/src/Std/Time/Date.lean similarity index 59% rename from src/Lean/Time/DateTime/Basic.lean rename to src/Std/Time/Date.lean index 61458d98a6fe..80dbf3ad7c46 100644 --- a/src/Lean/Time/DateTime/Basic.lean +++ b/src/Std/Time/Date.lean @@ -4,5 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Init.Data.Int -import Lean.Time.LessEq +import Std.Time.Date.Basic +import Std.Time.Date.Scalar +import Std.Time.Date.Date +import Std.Time.Date.WeekDate diff --git a/src/Lean/Time/Time/Basic.lean b/src/Std/Time/Date/Basic.lean similarity index 83% rename from src/Lean/Time/Time/Basic.lean rename to src/Std/Time/Date/Basic.lean index af3f62132446..e50fbc17d028 100644 --- a/src/Lean/Time/Time/Basic.lean +++ b/src/Std/Time/Date/Basic.lean @@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time.Unit.Basic +import Std.Time.Date.Unit.Basic diff --git a/src/Lean/Time/Date/Date.lean b/src/Std/Time/Date/Date.lean similarity index 92% rename from src/Lean/Time/Date/Date.lean rename to src/Std/Time/Date/Date.lean index 4449e466128c..3e1e4273dcf0 100644 --- a/src/Lean/Time/Date/Date.lean +++ b/src/Std/Time/Date/Date.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Date.Basic -import Lean.Time.Date.Scalar +import Std.Time.UnitVal +import Std.Time.Date.Basic +import Std.Time.Date.Scalar -namespace Lean -namespace Date +namespace Std +namespace Time /-- Date in YMD format. @@ -30,6 +30,9 @@ def force (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Dat let ⟨day, valid⟩ := month.forceDay year.isLeap day Date.mk year month day valid +instance : Inhabited Date where + default := force 0 1 1 + /-- Creates a new `Date` using YMD. -/ @@ -117,13 +120,13 @@ def toDaysSinceUNIXEpoch (date : Date) : Day.Offset := /-- Returns the `Scalar` starting from the UNIX epoch. -/ -def toScalar (date : Date) : Scalar := +def toScalar (date : Date) : Date.Scalar := ⟨toDaysSinceUNIXEpoch date⟩ /-- Creates a new `Date` from a `Scalar` that begins on the epoch. -/ -def ofScalar (scalar : Scalar) : Date := +def ofScalar (scalar : Date.Scalar) : Date := ofDaysSinceUNIXEpoch scalar.day /-- @@ -137,3 +140,7 @@ instance : HAdd Date Day.Offset Date where instance : HAdd Date Scalar Date where hAdd date day := ofScalar (toScalar date + day) + +end Date +end Time +end Std diff --git a/src/Lean/Time/Date/Scalar.lean b/src/Std/Time/Date/Scalar.lean similarity index 94% rename from src/Lean/Time/Date/Scalar.lean rename to src/Std/Time/Date/Scalar.lean index 1701e31565e3..982135b13085 100644 --- a/src/Lean/Time/Date/Scalar.lean +++ b/src/Std/Time/Date/Scalar.lean @@ -4,10 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Date.Basic +import Std.Time.UnitVal +import Std.Time.Date.Basic -namespace Lean +namespace Std +namespace Time namespace Date /-- @@ -63,3 +64,4 @@ def ofOffset (offset : Day.Offset) : Scalar := end Scalar end Date +end Time diff --git a/src/Std/Time/Date/Unit/Basic.lean b/src/Std/Time/Date/Unit/Basic.lean new file mode 100644 index 000000000000..7fca03623866 --- /dev/null +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -0,0 +1,25 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.UnitVal +import Std.Time.Date.Unit.Day +import Std.Time.Date.Unit.Month +import Std.Time.Date.Unit.Year +import Std.Time.Date.Unit.Weekday +import Std.Time.Date.Unit.WeekOfYear + +namespace Std +namespace Time.Day.Ordinal.OfYear + +/-- +Conevrts a `Year` and a `Ordinal.OfYear` to a valid day and month. +-/ +@[inline] +def toMonthAndDay (year : Year.Offset) (ordinal : OfYear year.isLeap) : { val : Month.Ordinal × Day.Ordinal // Year.Offset.valid year (Prod.fst val) (Prod.snd val) } := + Month.Ordinal.ofOrdinal ordinal + +end Time.Day.Ordinal.OfYear +end Std diff --git a/src/Lean/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean similarity index 93% rename from src/Lean/Time/Date/Unit/Day.lean rename to src/Std/Time/Date/Unit/Day.lean index a93a186a654e..1741e1ad2e17 100644 --- a/src/Lean/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -4,15 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.Time +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -namespace Lean -namespace Date -open Lean Time +namespace Std +namespace Time +open Lean set_option linter.all true @@ -86,3 +86,5 @@ def toSeconds (days : Offset) : Second.Offset := end Offset end Day +end Time +end Std diff --git a/src/Lean/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean similarity index 96% rename from src/Lean/Time/Date/Unit/Month.lean rename to src/Std/Time/Date/Unit/Month.lean index 47256755d364..d72df64db255 100644 --- a/src/Lean/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -4,15 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -import Lean.Time.Time.Basic -import Lean.Time.Date.Unit.Day +import Std.Time.Time.Basic +import Std.Time.Date.Unit.Day -namespace Lean -namespace Date +namespace Std +namespace Time open Lean Time namespace Month @@ -128,6 +128,7 @@ def days' (leap : Bool) (month : Ordinal) : Day.Ordinal := /-- Check if the day is valid in a Month and a leap Year. -/ +@[inline] def valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := day ≤ days' leap month @@ -137,6 +138,7 @@ instance : Decidable (valid leap month day) := /-- Gets the number of days in a month. -/ +@[inline] def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // valid leap month day } := ⟨days' leap month, Int.le_refl ((days' leap month).val)⟩ @@ -184,3 +186,5 @@ def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day end Ordinal end Month +end Time +end Std diff --git a/src/Lean/Time/Date/Unit/WeekOfYear.lean b/src/Std/Time/Date/Unit/WeekOfYear.lean similarity index 88% rename from src/Lean/Time/Date/Unit/WeekOfYear.lean rename to src/Std/Time/Date/Unit/WeekOfYear.lean index 57b077534e35..e38597b6121f 100644 --- a/src/Lean/Time/Date/Unit/WeekOfYear.lean +++ b/src/Std/Time/Date/Unit/WeekOfYear.lean @@ -4,15 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -import Lean.Time.Date.Unit.Day -import Lean.Time.Date.Unit.Month +import Std.Time.Date.Unit.Day +import Std.Time.Date.Unit.Month -namespace Lean -namespace Date +namespace Std +namespace Time set_option linter.all true @@ -60,4 +60,5 @@ def toOffset (ordinal : Ordinal) : Offset := end Ordinal end WeekOfYear -end Date +end Time +end Std diff --git a/src/Lean/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean similarity index 90% rename from src/Lean/Time/Date/Unit/Weekday.lean rename to src/Std/Time/Date/Unit/Weekday.lean index 87c27f5ea3c0..cb001a21d845 100644 --- a/src/Lean/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -4,15 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -import Lean.Time.Date.Unit.Day -import Lean.Time.Date.Unit.Month +import Std.Time.Date.Unit.Day +import Std.Time.Date.Unit.Month -namespace Lean -namespace Date +namespace Std +namespace Time set_option linter.all true @@ -107,3 +107,7 @@ def next : Weekday → Weekday | .fri => .thu | .sat => .fri | .sun => .sat + +end Weekday +end Time +end Std diff --git a/src/Lean/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean similarity index 86% rename from src/Lean/Time/Date/Unit/Year.lean rename to src/Std/Time/Date/Unit/Year.lean index 3ebcef013f87..ca2968677606 100644 --- a/src/Lean/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -4,15 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -import Lean.Time.Date.Unit.Day -import Lean.Time.Date.Unit.Month +import Std.Time.Date.Unit.Day +import Std.Time.Date.Unit.Month -namespace Lean -namespace Date +namespace Std +namespace Time set_option linter.all true @@ -61,3 +61,5 @@ instance : Decidable (valid year month day) := end Offset end Year +end Time +end Std diff --git a/src/Lean/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean similarity index 70% rename from src/Lean/Time/Date/WeekDate.lean rename to src/Std/Time/Date/WeekDate.lean index 6bdb97d8b38c..2ef7fc11d765 100644 --- a/src/Lean/Time/Date/WeekDate.lean +++ b/src/Std/Time/Date/WeekDate.lean @@ -4,13 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Date.Unit.Year -import Lean.Time.Date.Unit.WeekOfYear -import Lean.Time.Date.Scalar -import Lean.Time.Date.Date +import Std.Time.Date.Unit.Year +import Std.Time.Date.Unit.WeekOfYear +import Std.Time.Date.Scalar +import Std.Time.Date.Date -namespace Lean -namespace Date +namespace Std +namespace Time /-- `WeekDate` represents a date using a combination of a week of the year and the year. @@ -25,14 +25,14 @@ namespace WeekDate /-- Converts a `WeekDate` to a `Scalar`. -/ -def toScalar (wd : WeekDate) : Scalar := +def toScalar (wd : WeekDate) : Date.Scalar := let days := wd.year.toInt * 365 + wd.week.val * 7 - Scalar.ofDays days + Date.Scalar.ofDays days /-- Creates a `WeekDate` from a `Scalar`. -/ -def fromScalar (scalar : Scalar) : WeekDate := +def fromScalar (scalar : Date.Scalar) : WeekDate := let totalDays := scalar.toDays let year := totalDays / 365 let week := @@ -40,3 +40,7 @@ def fromScalar (scalar : Scalar) : WeekDate := |>.div 7 (by decide) |>.add 1 { year := year, week := week } + +end WeekDate +end Time +end Std diff --git a/src/Lean/Time/DateTime.lean b/src/Std/Time/DateTime.lean similarity index 75% rename from src/Lean/Time/DateTime.lean rename to src/Std/Time/DateTime.lean index 7acc71628232..cbcc86dd9fa2 100644 --- a/src/Lean/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -4,11 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.DateTime.NaiveDateTime -import Lean.Time.DateTime.Timestamp +import Std.Time.DateTime.NaiveDateTime +import Std.Time.DateTime.Timestamp -namespace Lean -namespace DateTime.Timestamp +namespace Std +namespace Time +namespace Timestamp + +set_option linter.all true /-- Converts a `NaiveDateTime` to a `Timestamp` diff --git a/src/Lean/Time/Date/Unit/Lemmas.lean b/src/Std/Time/DateTime/Basic.lean similarity index 87% rename from src/Lean/Time/Date/Unit/Lemmas.lean rename to src/Std/Time/DateTime/Basic.lean index 46e50091b717..45d10e698ca4 100644 --- a/src/Lean/Time/Date/Unit/Lemmas.lean +++ b/src/Std/Time/DateTime/Basic.lean @@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Init.Data.Int +import Std.Time.LessEq diff --git a/src/Lean/Time/DateTime/NaiveDateTime.lean b/src/Std/Time/DateTime/NaiveDateTime.lean similarity index 94% rename from src/Lean/Time/DateTime/NaiveDateTime.lean rename to src/Std/Time/DateTime/NaiveDateTime.lean index f230be45f1e3..18b7f46cdccd 100644 --- a/src/Lean/Time/DateTime/NaiveDateTime.lean +++ b/src/Std/Time/DateTime/NaiveDateTime.lean @@ -5,22 +5,21 @@ Authors: Sofia Rodrigues -/ prelude import Init.Data.Int -import Lean.Time.LessEq -import Lean.Time.Date -import Lean.Time.Time -import Lean.Time.DateTime.Timestamp +import Std.Time.LessEq +import Std.Time.Date +import Std.Time.Time +import Std.Time.DateTime.Timestamp -namespace Lean -namespace DateTime -open Date Time +namespace Std +namespace Time /-- Date time format with Year, Month, Day, Hour, Minute, Seconds and Nanoseconds. -/ structure NaiveDateTime where - date : Date.Date - time : Time.Time - deriving Repr + date : Date + time : Time + deriving Repr, Inhabited namespace NaiveDateTime @@ -149,3 +148,7 @@ Getter for the `Second` inside of a `NaiveDateTime` @[inline] def second (dt : NaiveDateTime) : Second.Ordinal := dt.time.second + +end NaiveDateTime +end Time +end Std diff --git a/src/Lean/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean similarity index 65% rename from src/Lean/Time/DateTime/Timestamp.lean rename to src/Std/Time/DateTime/Timestamp.lean index f17e6fed14af..df76ddb413d0 100644 --- a/src/Lean/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -5,21 +5,26 @@ Authors: Sofia Rodrigues -/ prelude import Init.Data.Int -import Lean.Time.LessEq -import Lean.Time.Time +import Std.Time.LessEq +import Std.Time.Time -namespace Lean -namespace DateTime -open Time +namespace Std +namespace Time /-- Seconds since the UNIX Epoch. -/ def Timestamp := Second.Offset - deriving Repr, BEq + deriving Repr, BEq, Inhabited instance : OfNat Timestamp n where ofNat := UnitVal.ofNat n instance : HAdd Timestamp Second.Offset Timestamp where hAdd x y := UnitVal.add x y + +instance : HSub Timestamp Second.Offset Timestamp where + hSub x y := UnitVal.sub x y + +end Time +end Std diff --git a/src/Lean/Time/Duration.lean b/src/Std/Time/Duration.lean similarity index 87% rename from src/Lean/Time/Duration.lean rename to src/Std/Time/Duration.lean index 63db8fd707af..2d3b3ee0f707 100644 --- a/src/Lean/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -4,12 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Bounded -import Lean.Time.Time +import Std.Time.Bounded +import Std.Time.Time -namespace Lean - -open Time +namespace Std +namespace Time /-- `Instant` represents a place in time with second and nanoseconds precision. @@ -58,15 +57,21 @@ end Instant namespace Duration -/-- Returns a `Duration` representing the given number of second. -/ +/-- +Returns a `Duration` representing the given number of second. +-/ def ofSeconds (second : Second.Offset) : Duration := { second := second, nano := Bounded.LE.mk 0 (by decide) } -/-- Returns a `Duration` representing the given number of minute. -/ +/-- +Returns a `Duration` representing the given number of minute. +-/ def ofMinutes (minute : Minute.Offset) : Duration := { second := minute.toSeconds * 60, nano := Bounded.LE.mk 0 (by decide) } -/-- Returns a `Duration` representing the given number of hour. -/ +/-- +Returns a `Duration` representing the given number of hour. +-/ def ofHours (hour : Hour.Offset) : Duration := { second := hour.toSeconds * 3600, nano := Bounded.LE.mk 0 (by decide) } diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean new file mode 100644 index 000000000000..49ed77f68c71 --- /dev/null +++ b/src/Std/Time/Format.lean @@ -0,0 +1,200 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.Format.Basic + +open Std.Time + +namespace Formats + +set_option linter.all true + +/-- +The ISO8601 format, which is always 24 or 27 characters long, used for representing date and time in +a standardized format. The format follows the pattern `YYYY-MM-DD'T'hh:mm:ss'Z'`. +-/ +def ISO8601 : Format .any := Format.spec! "YYYY-MM-DD'T'hh:mm:ss.sssZ" + +/-- +The AmericanDate format, which follows the pattern `MM/DD/YYYY`. +-/ +def AmericanDate : Format .any := ⟨[Modifier.MM, "/", Modifier.DD, "/", Modifier.YYYY]⟩ + +/-- +The EuropeanDate format, which follows the pattern `DD/MM/YYYY`. +-/ +def EuropeanDate : Format .any := ⟨[Modifier.DD, "/", Modifier.MM, "/", Modifier.YYYY]⟩ + +/-- +The Time12Hour format, which follows the pattern `hh:mm:ss aa` for representing time +in a 12-hour clock format with an AM/PM marker. +-/ +def Time12Hour : Format .any := ⟨[Modifier.HH, ":", Modifier.mm, ":", Modifier.ss, " ", Modifier.aa]⟩ + +/-- +The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time +in a 24-hour clock format. +-/ +def Time24Hour : Format .any := ⟨[Modifier.hh, ":", Modifier.mm, ":", Modifier.ss]⟩ + +/-- +The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used +in SQL databases to represent dates. +-/ +def SQLDate : Format .any := ⟨[Modifier.YYYY, "-", Modifier.MM, "-", Modifier.DD]⟩ + +/-- +The LongDateFormat, which follows the pattern `EEEE, MMMM D, YYYY hh:mm:ss` for +representing a full date and time with the day of the week and month name. +-/ +def LongDateFormat : Format (.only .GMT) := Format.spec! "EEEE, MMMM D, YYYY hh:mm:ss" + +/-- +The AscTime format, which follows the pattern `EEE MMM d hh:mm:ss YYYY`. This format +is often used in older systems for logging and time-stamping events. +-/ +def AscTime : Format (.only .GMT) := Format.spec! "EEE MMM d hh:mm:ss YYYY" + +/-- +The RFC822 format, which follows the pattern `EEE, DD MMM YYYY hh:mm:ss ZZZ`. +This format is used in email headers and HTTP headers. +-/ +def RFC822 : Format .any := Format.spec! "EEE, DD MMM YYYY hh:mm:ss ZZZ" + +/-- +The RFC850 format, which follows the pattern `EEEE, DD-MMM-YY hh:mm:ss ZZZ`. +This format is an older standard for representing date and time in headers. +-/ +def RFC850 : Format .any := Format.spec! "EEEE, DD-MMM-YY hh:mm:ss ZZZ" + +end Formats + +namespace Date + +/-- +Parses a date string in the American format (`MM/DD/YYYY`) and returns a `Date`. +-/ +def fromAmericanDateString (input : String) : Except String Date := do + let res ← Formats.AmericanDate.parseBuilder (λm d y => Date.ofYearMonthDay y m d) input + match res with + | .some res => pure res + | none => .error "Invalid date." + +end Date + +namespace Time + +/-- +Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `Time`. +-/ +def fromTime24Hour (input : String) : Except String Time := + Formats.Time24Hour.parseBuilder (Time.ofHourMinuteSeconds) input + +/-- +Formats a `Time` value into a 24-hour format string (`hh:mm:ss`). +-/ +def toTime24Hour (input : Time) : String := + Formats.Time24Hour.formatBuilder input.hour input.minute input.second + +/-- +Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `Time`. +-/ +def fromTime12Hour (input : String) : Except String Time := + Formats.Time12Hour.parseBuilder (λh m s a => Time.ofHourMinuteSeconds (HourMarker.toAbsolute! a h) m s) input + +/-- +Formats a `Time` value into a 12-hour format string (`hh:mm:ss aa`). +-/ +def toTime12Hour (input : Time) : String := + Formats.Time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.val ≥ 12 then HourMarker.pm else HourMarker.am) + +end Time + +namespace ZonedDateTime + +/-- +Parses a `String` in the `ISO8601` format and returns a `DateTime`. +-/ +def fromISO8601String (input : String) : Except String ZonedDateTime := + Formats.ISO8601.parse input + +/-- +Formats a `ZonedDateTime` value into an ISO8601 string. +-/ +def toISO8601String (date : ZonedDateTime) : String := + Formats.ISO8601.format date.snd + +/-- +Parses a `String` in the `RFC822` format and returns a `DateTime`. +-/ +def fromRFC822String (input : String) : Except String ZonedDateTime := + Formats.RFC822.parse input + +/-- +Formats a `ZonedDateTime` value into an RFC822 format string. +-/ +def toRFC822String (date : ZonedDateTime) : String := + Formats.RFC822.format date.snd + +/-- +Parses a `String` in the `RFC850` format and returns a `DateTime`. +-/ +def fromRFC850String (input : String) : Except String ZonedDateTime := + Formats.RFC850.parse input + +/-- +Formats a `ZonedDateTime` value into an RFC850 format string. +-/ +def toRFC850String (date : ZonedDateTime) : String := + Formats.RFC850.format date.snd + +end ZonedDateTime + +namespace DateTime + +/-- +Parses a `String` in the `AscTime` format and returns a `DateTime` object in the GMT time zone. +-/ +def fromAscTimeString (input : String) : Except String (DateTime .GMT) := + Formats.AscTime.parse input + +/-- +Formats a `DateTime` value into an AscTime format string. +-/ +def toAscTimeString (datetime : DateTime .GMT) : String := + Formats.AscTime.format datetime + +/-- +Parses a `String` in the `LongDateFormat` and returns a `DateTime` object in the GMT time zone. +-/ +def fromLongDateFormatString (input : String) : Except String (DateTime .GMT) := + Formats.LongDateFormat.parse input + +/-- +Formats a `DateTime` value into a LongDateFormat string. +-/ +def toLongDateFormatString (datetime : DateTime .GMT) : String := + Formats.LongDateFormat.format datetime + +/-- +Formats a `ZonedDateTime` value into an ISO8601 string. +-/ +def toISO8601String (date : ZonedDateTime) : String := + Formats.ISO8601.format date.snd + +/-- +Formats a `ZonedDateTime` value into an RFC822 format string. +-/ +def toRFC822String (date : ZonedDateTime) : String := + Formats.RFC822.format date.snd + +/-- +Formats a `ZonedDateTime` value into an RFC850 format string. +-/ +def toRFC850String (date : ZonedDateTime) : String := + Formats.RFC850.format date.snd + +end DateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean new file mode 100644 index 000000000000..c510efd0e77d --- /dev/null +++ b/src/Std/Time/Format/Basic.lean @@ -0,0 +1,648 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.Date +import Std.Time.Time +import Std.Time.DateTime +import Std.Time.Zoned +import Lean.Data.Parsec + +namespace Std +namespace Time + +open Lean.Parsec Lean Time Date TimeZone DateTime + +/-- +The `Modifier` inductive type represents various formatting options for date and time components. +These modifiers are typically used in formatting functions to generate human-readable date and time strings. + +- `YYYY`: Four-digit year (e.g., 2024). +- `YY`: Two-digit year (e.g., 24 for 2024). +- `MMMM`: Full month name (e.g., January, February). +- `MMM`: Abbreviated month name (e.g., Jan, Feb). +- `MM`: Two-digit month (e.g., 01 for January). +- `M`: One or two-digit month (e.g., 1 for January, 10 for October). +- `DD`: Two-digit day of the month (e.g., 01, 02). +- `D`: One or two-digit day of the month (e.g., 1, 2). +- `d`: One or two digit day of the month with space padding at the beggining (e.g. 1, 12). +- `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday). +- `EEE`: Abbreviated day of the week (e.g., Mon, Tue). +- `hh`: Two-digit hour in 24-hour format (e.g., 01, 02). +- `h`: One or two-digit hour in 24-hour format (e.g., 1, 2). +- `HH`: Two-digit hour in 12-hour format (e.g., 13, 14). +- `H`: One or two-digit hour in 12-hour format (e.g., 1, 2). +- `AA`: Uppercase AM/PM indicator (e.g., AM, PM). +- `aa`: Lowercase am/pm indicator (e.g., am, pm). +- `mm`: Two-digit minute (e.g., 01, 02). +- `m`: One or two-digit minute (e.g., 1, 2). +- `sss`: Three-digit milliseconds (e.g., 001, 202). +- `ss`: Two-digit second (e.g., 01, 02). +- `s`: One or two-digit second (e.g., 1, 2). +- `ZZZZZ`: Full timezone offset including hours and minutes (e.g., +03:00). +- `ZZZZ`: Timezone offset including hours and minutes without the colon (e.g., +0300). +- `ZZZ`: Like ZZZZ but with a special case "UTC" for UTC. +- `Z`: Like ZZZZZ but with a special case "Z" for UTC. +- `z`: Name of the time-zone like (Brasilia Standard Time). +-/ +inductive Modifier + | YYYY + | YY + | MMMM + | MMM + | MM + | M + | DD + | D + | d + | EEEE + | EEE + | hh + | h + | HH + | H + | AA + | aa + | mm + | m + | sss + | ss + | s + | ZZZZZ + | ZZZZ + | ZZZ + | Z + | z + deriving Repr + +/-- +The part of a formatting string. a string is just a text and a modifier is in the format `%0T` where +0 is the quantity of left pad and `T` the `Modifier`. +-/ +inductive FormatPart + | string (val : String) + | modifier (modifier : Modifier) + deriving Repr + +instance : Coe String FormatPart where + coe := .string + +instance : Coe Modifier FormatPart where + coe := .modifier + +/-- +The format of date and time string. +-/ +abbrev FormatString := List FormatPart + +inductive Awareness + | only : TimeZone → Awareness + | any + +namespace Awareness + +instance : Coe TimeZone Awareness where + coe := .only + +@[simp] +def type (x : Awareness) : Type := + match x with + | .any => ZonedDateTime + | .only tz => DateTime tz + +instance : Inhabited (type aw) where + default := by + simp [type] + split <;> exact Inhabited.default + +def getD (x : Awareness) (default : TimeZone) : TimeZone := + match x with + | .any => default + | .only tz => tz + +end Awareness + +structure Format (awareness : Awareness) where + string : FormatString + deriving Inhabited, Repr + +private def isNonLetter : Char → Bool := not ∘ Char.isAlpha + +private def parseModifier : Lean.Parsec Modifier + := pstring "YYYY" *> pure .YYYY + <|> pstring "YY" *> pure .YY + <|> pstring "MMMM" *> pure .MMMM + <|> pstring "MMM" *> pure .MMM + <|> pstring "MM" *> pure .MM + <|> pstring "M" *> pure .M + <|> pstring "DD" *> pure .DD + <|> pstring "D" *> pure .D + <|> pstring "d" *> pure .d + <|> pstring "EEEE" *> pure .EEEE + <|> pstring "EEE" *> pure .EEE + <|> pstring "hh" *> pure .hh + <|> pstring "h" *> pure .h + <|> pstring "HH" *> pure .HH + <|> pstring "H" *> pure .H + <|> pstring "AA" *> pure .AA + <|> pstring "aa" *> pure .aa + <|> pstring "mm" *> pure .mm + <|> pstring "m" *> pure .m + <|> pstring "sss" *> pure .sss + <|> pstring "ss" *> pure .ss + <|> pstring "s" *> pure .s + <|> pstring "ZZZZZ" *> pure .ZZZZZ + <|> pstring "ZZZZ" *> pure .ZZZZ + <|> pstring "ZZZ" *> pure .ZZZ + <|> pstring "Z" *> pure .Z + <|> pstring "z" *> pure .z + +def isFormatStart : Char → Bool := Char.isAlpha + +private def pnumber : Lean.Parsec Nat := do + let numbers ← manyChars digit + return numbers.foldl (λacc char => acc * 10 + (char.toNat - 48)) 0 + +private def parseFormatPart : Lean.Parsec FormatPart + := (.modifier <$> parseModifier) + <|> (pchar '\\') *> anyChar <&> (.string ∘ toString) + <|> (pchar '\"' *> many1Chars (satisfy (· ≠ '\"')) <* pchar '\"') <&> .string + <|> (pchar '\'' *> many1Chars (satisfy (· ≠ '\'')) <* pchar '\'') <&> .string + <|> many1Chars (satisfy (fun x => ¬isFormatStart x ∧ x ≠ '\'' ∧ x ≠ '\"')) <&> .string + +private def specParser : Lean.Parsec FormatString := + (Array.toList <$> Lean.Parsec.many parseFormatPart) <* eof + +private def specParse (s: String) : Except String FormatString := + specParser.run s + +-- Pretty printer + +private def unabbrevMonth (month: Month.Ordinal) : String := + match month.val, month.property with + | 1, _ => "January" + | 2, _ => "February" + | 3, _ => "March" + | 4, _ => "April" + | 5, _ => "May" + | 6, _ => "June" + | 7, _ => "July" + | 8, _ => "August" + | 9, _ => "September" + | 10, _ => "October" + | 11, _ => "November" + | 12, _ => "December" + +private def abbrevMonth (month: Month.Ordinal) : String := + match month.val, month.property with + | 1, _ => "Jan" + | 2, _ => "Feb" + | 3, _ => "Mar" + | 4, _ => "Apr" + | 5, _ => "May" + | 6, _ => "Jun" + | 7, _ => "Jul" + | 8, _ => "Aug" + | 9, _ => "Sep" + | 10, _ => "Oct" + | 11, _ => "Nov" + | 12, _ => "Dec" + +private def abbrevDayOfWeek (day: Weekday) : String := + match day with + | .sun => "Sun" + | .mon => "Mon" + | .tue => "Tue" + | .wed => "Wed" + | .thu => "Thu" + | .fri => "Fri" + | .sat => "Sat" + +private def dayOfWeek (day: Weekday) : String := + match day with + | .sun => "Sunday" + | .mon => "Monday" + | .tue => "Tuesday" + | .wed => "Wednesday" + | .thu => "Thusday" + | .fri => "Friday" + | .sat => "Saturday" + +private def leftPad (n : Nat) (a : Char) (s : String) : String := + "".pushn a (n - s.length) ++ s + +private def formatWithDate (date : DateTime tz) : Modifier → String + | .YYYY => s!"{leftPad 4 '0' (toString date.year)}" + | .YY => s!"{leftPad 2 '0' (toString $ date.year.toNat % 100)}" + | .MMMM => unabbrevMonth date.month + | .MMM => abbrevMonth date.month + | .MM => s!"{leftPad 2 '0' (toString $ date.month.toNat)}" + | .M => s!"{date.month.toNat}" + | .DD => s!"{leftPad 2 '0' (toString $ date.day.toNat)}" + | .D => s!"{date.day.toNat}" + | .d => s!"{leftPad 2 ' ' $ toString date.day.toNat}" + | .EEEE => dayOfWeek date.weekday + | .EEE => abbrevDayOfWeek date.weekday + | .hh => s!"{leftPad 2 '0' (toString date.hour.toNat)}" + | .h => s!"{date.hour.toNat}" + | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' $ toString hour}" + | .H => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{hour}" + | .AA => if date.hour.toNat < 12 then "AM" else "PM" + | .aa => if date.hour.toNat < 12 then "am" else "pm" + | .mm => s!"{leftPad 2 '0' $ toString date.minute.toNat}" + | .m => s!"{date.minute.toNat}" + | .sss => s!"{leftPad 3 '0' $ toString date.milliseconds.toNat}" + | .ss => s!"{leftPad 2 '0' $ toString date.second.toNat}" + | .s => s!"{date.second.toNat}" + | .ZZZZZ => tz.offset.toIsoString true + | .ZZZZ => tz.offset.toIsoString false + | .ZZZ => if tz.offset.second.val = 0 then "UTC" else tz.offset.toIsoString false + | .Z => if tz.offset.second.val = 0 then "Z" else tz.offset.toIsoString true + | .z => tz.name + +private def formatPartWithDate (date : DateTime z) : FormatPart → String + | .string s => s + | .modifier t => formatWithDate date t + +-- Parser + +@[simp] +private def SingleFormatType : Modifier → Type + | .YYYY => Year.Offset + | .YY => Year.Offset + | .MMMM => Month.Ordinal + | .MMM => Month.Ordinal + | .MM => Month.Ordinal + | .M => Month.Ordinal + | .DD => Day.Ordinal + | .D => Day.Ordinal + | .d => Day.Ordinal + | .EEEE => Weekday + | .EEE => Weekday + | .hh => Hour.Ordinal + | .h => Hour.Ordinal + | .HH => Hour.Ordinal + | .H => Hour.Ordinal + | .AA => HourMarker + | .aa => HourMarker + | .mm => Minute.Ordinal + | .m => Minute.Ordinal + | .sss => Millisecond.Ordinal + | .ss => Second.Ordinal + | .s => Second.Ordinal + | .ZZZZZ => Offset + | .ZZZZ => Offset + | .ZZZ => Offset + | .Z => Offset + | .z => String + +private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) : String := + match modifier with + | .YYYY => s!"{leftPad 4 '0' (toString data.toNat)}" + | .YY => s!"{leftPad 2 '0' (toString $ data.toNat % 100)}" + | .MMMM => unabbrevMonth data + | .MMM => abbrevMonth data + | .MM => s!"{leftPad 2 '0' (toString $ data.toNat)}" + | .M => s!"{data.toNat}" + | .DD => s!"{leftPad 2 '0' (toString $ data.toNat)}" + | .D => s!"{data.toNat}" + | .d => s!"{leftPad 2 ' ' $ toString data.toNat}" + | .EEEE => dayOfWeek data + | .EEE => abbrevDayOfWeek data + | .hh => s!"{leftPad 2 '0' (toString data.toNat)}" + | .h => s!"{data.toNat}" + | .HH => let hour := data.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' $ toString hour}" + | .H => let hour := data.val % 12; if hour == 0 then "12" else s!"{hour}" + | .AA => match data with | .am => "AM" | .pm => "PM" + | .aa => match data with | .am => "am" | .pm => "pm" + | .mm => s!"{leftPad 2 '0' $ toString data.toNat}" + | .m => s!"{data.toNat}" + | .sss => s!"{leftPad 3 '0' $ toString data.toNat}" + | .ss => s!"{leftPad 2 '0' $ toString data.toNat}" + | .s => s!"{data.toNat}" + | .ZZZZZ => data.toIsoString true + | .ZZZZ => data.toIsoString false + | .ZZZ => if data.second.val = 0 then "UTC" else data.toIsoString false + | .Z => if data.second.val = 0 then "Z" else data.toIsoString true + | .z => data + +@[simp] +def FormatType (result: Type) : FormatString → Type + | .modifier entry :: xs => (SingleFormatType entry) → (FormatType result xs) + | .string _ :: xs => (FormatType result xs) + | [] => result + +private def position : Parsec Nat := λs => (ParseResult.success s (s.pos.byteIdx)) + +private def size (data : Parsec α) : Parsec (α × Nat) := do + let st ← position + let res ← data + let en ← position + pure (res, en-st) + +private def transform (n: β → Option α) (p: Lean.Parsec β) : Lean.Parsec α := do + let res ← p + match n res with + | some n => pure n + | none => fail "cannot parse" + +private def parseMonth : Lean.Parsec Month.Ordinal + := (pstring "Jan" *> pure 1) + <|> (pstring "Feb" *> pure 2) + <|> (pstring "Mar" *> pure 3) + <|> (pstring "Apr" *> pure 4) + <|> (pstring "May" *> pure 5) + <|> (pstring "Jun" *> pure 6) + <|> (pstring "Jul" *> pure 7) + <|> (pstring "Aug" *> pure 8) + <|> (pstring "Sep" *> pure 9) + <|> (pstring "Oct" *> pure 10) + <|> (pstring "Nov" *> pure 11) + <|> (pstring "Dec" *> pure 12) + +private def parseMonthUnabbrev : Lean.Parsec Month.Ordinal + := (pstring "January" *> pure 1) + <|> (pstring "February" *> pure 2) + <|> (pstring "March" *> pure 3) + <|> (pstring "April" *> pure 4) + <|> (pstring "May" *> pure 5) + <|> (pstring "June" *> pure 6) + <|> (pstring "July" *> pure 7) + <|> (pstring "August" *> pure 8) + <|> (pstring "September" *> pure 9) + <|> (pstring "October" *> pure 10) + <|> (pstring "November" *> pure 11) + <|> (pstring "December" *> pure 12) + +private def parseWeekday : Lean.Parsec Weekday + := (pstring "Mon" *> pure Weekday.mon) + <|> (pstring "Tue" *> pure Weekday.tue) + <|> (pstring "Wed" *> pure Weekday.wed) + <|> (pstring "Thu" *> pure Weekday.thu) + <|> (pstring "Fri" *> pure Weekday.fri) + <|> (pstring "Sat" *> pure Weekday.sat) + <|> (pstring "Sun" *> pure Weekday.sun) + +private def parseWeekdayUnnabrev : Lean.Parsec Weekday + := (pstring "Monday" *> pure Weekday.mon) + <|> (pstring "Tuesday" *> pure Weekday.tue) + <|> (pstring "Wednesday" *> pure Weekday.wed) + <|> (pstring "Thursday" *> pure Weekday.thu) + <|> (pstring "Friday" *> pure Weekday.fri) + <|> (pstring "Saturday" *> pure Weekday.sat) + <|> (pstring "Sunday" *> pure Weekday.sun) + +private def parserUpperHourMarker : Lean.Parsec HourMarker + := (pstring "AM" *> pure HourMarker.am) + <|> (pstring "PM" *> pure HourMarker.pm) + +private def parserLowerHourMarker : Lean.Parsec HourMarker + := (pstring "am" *> pure HourMarker.am) + <|> (pstring "pm" *> pure HourMarker.pm) + +private def threeDigit : Lean.Parsec Int := do + let digit1 ← Lean.Parsec.digit + let digit2 ← Lean.Parsec.digit + let digit3 ← Lean.Parsec.digit + return String.toNat! s!"{digit1}{digit2}{digit3}" + +private def twoDigit : Lean.Parsec Int := do + let digit1 ← Lean.Parsec.digit + let digit2 ← Lean.Parsec.digit + return String.toNat! s!"{digit1}{digit2}" + +private def parseYearTwo : Lean.Parsec Int :=do + let year ← twoDigit + return if year < 70 then 2000 + year else 1900 + year + +private def timeOffset (colon: Bool) : Lean.Parsec Offset := do + let sign : Int ← (pstring "-" *> pure (-1)) <|> (pstring "+" *> pure 1) + let hour ← twoDigit + if colon then let _ ← pstring ":" + let minutes ← twoDigit + let res := (hour * 3600 + minutes * 60) * sign + pure (Offset.ofSeconds (UnitVal.ofInt res)) + +private def timeOrUTC (utcString: String) (colon: Bool) : Lean.Parsec Offset := + (pstring utcString *> pure Offset.zero) <|> timeOffset colon + +private def number : Lean.Parsec Nat := do + String.toNat! <$> Lean.Parsec.many1Chars Lean.Parsec.digit + +private def singleDigit : Lean.Parsec Nat := do + let digit1 ← Lean.Parsec.digit + return String.toNat! s!"{digit1}" + +private def fourDigit : Lean.Parsec Int := do + let digit1 ← Lean.Parsec.digit + let digit2 ← Lean.Parsec.digit + let digit3 ← Lean.Parsec.digit + let digit4 ← Lean.Parsec.digit + return String.toNat! s!"{digit1}{digit2}{digit3}{digit4}" + +private def parserWithFormat : (typ: Modifier) → Lean.Parsec (SingleFormatType typ) + | .YYYY => fourDigit + | .YY => parseYearTwo + | .MMMM => parseMonthUnabbrev + | .MMM => parseMonth + | .MM => transform Bounded.LE.ofInt twoDigit + | .M => transform Bounded.LE.ofInt number + | .DD => transform Bounded.LE.ofInt twoDigit + | .D => transform Bounded.LE.ofInt number + | .d => transform Bounded.LE.ofInt (Lean.Parsec.orElse twoDigit (λ_ => pchar ' ' *> (singleDigit))) + | .EEEE => parseWeekdayUnnabrev + | .EEE => parseWeekday + | .hh => transform Bounded.LE.ofInt twoDigit + | .h => transform Bounded.LE.ofInt number + | .HH => do + let res : Bounded.LE 1 13 ← transform Bounded.LE.ofInt twoDigit + return (res.sub 1).expandTop (by decide) + | .H => do + let res : Bounded.LE 1 13 ← transform Bounded.LE.ofInt number + return (res.sub 1).expandTop (by decide) + | .AA => parserUpperHourMarker + | .aa => parserLowerHourMarker + | .mm => transform Bounded.LE.ofInt twoDigit + | .m => transform Bounded.LE.ofInt number + | .sss => transform Bounded.LE.ofInt threeDigit + | .ss => transform Bounded.LE.ofInt twoDigit + | .s => transform Bounded.LE.ofInt number + | .ZZZZZ => timeOffset true + | .ZZZZ => timeOffset false + | .ZZZ => timeOrUTC "UTC" false + | .Z => timeOrUTC "Z" true + | .z => many1Chars (satisfy (λc => c == ' ' || c.isAlpha)) + +private structure DateBuilder where + tzName : String := "Greenwich Mean Time" + tz : Offset := Offset.zero + year : Year.Offset := 0 + month : Month.Ordinal := 1 + day : Day.Ordinal := 1 + hour : Hour.Ordinal := 0 + minute : Minute.Ordinal := 0 + second : Second.Ordinal := 0 + millisecond : Millisecond.Ordinal := 0 + +private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : aw.type := + let build := DateTime.ofNaiveDateTime { + date := Date.force builder.year builder.month builder.day + time := Time.mk builder.hour builder.minute builder.second (.ofMillisecond builder.millisecond) + } + + match aw with + | .any => + let tz₁ := TimeZone.mk builder.tz builder.tzName + ⟨tz₁, build tz₁⟩ + | .only tz => + build tz + +private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : SingleFormatType typ) : DateBuilder := + match typ with + | .YYYY | .YY => { data with year := value } + | .MMMM | .MMM | .MM | .M => { data with month := value } + | .DD | .D | .d => { data with day := value } + | .EEEE | .EEE => data + | .hh | .h | .HH | .H => { data with hour := value } + | .AA | .aa => { data with hour := HourMarker.toAbsolute! value data.hour } + | .mm | .m => { data with minute := value } + | .sss => { data with millisecond := value } + | .ss | .s => { data with second := value } + | .ZZZZZ | .ZZZZ | .ZZZ + | .Z => { data with tz := value } + | .z => { data with tzName := value } + +private def formatParser (date : DateBuilder) : FormatPart → Lean.Parsec DateBuilder + | .modifier mod => addDataInDateTime date mod <$> parserWithFormat mod + | .string s => skipString s *> pure date + +-- API + +namespace Format + +/-- +Constructs a new `Format` specification for a date-time string. Modifiers can be combined to create +custom formats, such as %YYYY, MMMM, D". + +### Supported Modifiers: +- `YYYY`: Four-digit year (e.g., 2024). +- `YY`: Two-digit year (e.g., 24 for 2024). +- `MMMM`: Full month name (e.g., January, February). +- `MMM`: Abbreviated month name (e.g., Jan, Feb). +- `MM`: Two-digit month (e.g., 01 for January). +- `M`: One or two-digit month (e.g., 1 for January, 10 for October). +- `DD`: Two-digit day of the month (e.g., 01, 02). +- `D`: One or two-digit day of the month (e.g., 1, 2). +- `d`: One or two-digit day of the month with left padding with spaces. (e.g., 1, 2). +- `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday). +- `EEE`: Abbreviated day of the week (e.g., Mon, Tue). +- `hh`: Two-digit hour in 12-hour format (e.g., 01, 02). +- `h`: One or two-digit hour in 12-hour format (e.g., 1, 2). +- `HH`: Two-digit hour in 24-hour format (e.g., 13, 14). +- `H`: One or two-digit hour in 24-hour format (e.g., 1, 2). +- `AA`: Uppercase AM/PM indicator (e.g., AM, PM). +- `aa`: Lowercase am/pm indicator (e.g., am, pm). +- `mm`: Two-digit minute (e.g., 01, 02). +- `m`: One or two-digit minute (e.g., 1, 2). +- `sss`: Three-digit millisecond (e.g., 001, 002). +- `ss`: Two-digit second (e.g., 01, 02). +- `s`: One or two-digit second (e.g., 1, 2). +- `ZZZZZ`: Full timezone offset with hours and minutes (e.g., +03:00). +- `ZZZZ`: Timezone offset with hours and minutes, without the colon (e.g., +0300). +- `ZZZ`: Like `ZZZZ`, but displays "UTC" for UTC time. +- `Z`: Like `ZZZZZ`, but displays "Z" for UTC time. +- `z`: Timezone name (e.g., Brasilia Standard Time). + +Example usage: +- `"YYYY-MM-DD HH:mm:ss ZZZZ"` → "2024-08-04 14:23:45 +0300" +-/ +def spec (input : String) : Except String (Format tz) := do + let string ← specParser.run input + return ⟨string⟩ + +/-- +Builds a new `Format` specification for a Date-time, panics if the input string results in an error. +-/ +def spec! (input : String) : Format tz := + match specParser.run input with + | .ok res => ⟨res⟩ + | .error res => panic! res + +/-- +Formats the date using the format into a String. +-/ +def format (format : Format aw) (date : DateTime tz) : String := + let func := + match aw with + | .any => formatPartWithDate date + | .only tz => formatPartWithDate (date.convertTimeZone tz) + + format.string.map func + |> String.join + +/-- +Formats the date using the format into a String. +-/ +def formatBuilder (format : Format aw) : FormatType String format.string := + let rec go (data : String) : (format : FormatString) → FormatType String format + | .modifier x :: xs => λres => go (data ++ formatPart x res) xs + | .string x :: xs => go (data ++ x) xs + | [] => data + go "" format.string + +/-- +Parser for a ZonedDateTime. +-/ +def parser (format : FormatString) (aw : Awareness) : Parsec (aw.type) := + let rec go (date : DateBuilder) (x : FormatString) : Parsec aw.type := + match x with + | x :: xs => formatParser date x >>= (go · xs) + | [] => pure (date.build aw) + go {} format + +/-- +Parser for a format with a builder. +-/ +def builderParser (format: FormatString) (func: FormatType α format) : Lean.Parsec α := + let rec go (format : FormatString) (func: FormatType α format) : Parsec α := + match format with + | .modifier x :: xs => do + let res ← parserWithFormat x + go xs (func res) + | .string s :: xs => skipString s *> (go xs func) + | [] => pure func + go format func + +/-- +Parses the input string into a `ZoneDateTime` +-/ +def parse (format : Format aw) (input : String) : Except String aw.type := + (parser format.string aw).run input + +/- +Parses the input string into a `ZoneDateTime`, panics if its wrong +-/ +def parse! (format : Format aw) (input : String) : aw.type := + match parse format input with + | .ok res => res + | .error err => panic! err + +/-- +Parses and instead of using a builder to build a date, it uses a builder function instead. +-/ +def parseBuilder (format : Format aw) (builder : FormatType α format.string) (input : String) : Except String α := + (builderParser format.string builder).run input + +/-- +Parses and instead of using a builder to build a date, it uses a builder function instead. +-/ +def parseBuilder! [Inhabited α] (format : Format aw) (builder : FormatType α format.string) (input : String) : α := + match parseBuilder format builder input with + | .ok res => res + | .error err => panic! err + +end Format +end Time +end Std diff --git a/src/Lean/Time/LessEq.lean b/src/Std/Time/LessEq.lean similarity index 97% rename from src/Lean/Time/LessEq.lean rename to src/Std/Time/LessEq.lean index 7704a4997a6f..b248b0f9b795 100644 --- a/src/Lean/Time/LessEq.lean +++ b/src/Std/Time/LessEq.lean @@ -7,7 +7,8 @@ prelude import Init import Lean.Data.Rat -namespace Lean +namespace Std + set_option linter.all true /-- diff --git a/src/Lean/Time/Sign.lean b/src/Std/Time/Sign.lean similarity index 73% rename from src/Lean/Time/Sign.lean rename to src/Std/Time/Sign.lean index b5a385455a81..90dc41dc66b2 100644 --- a/src/Lean/Time/Sign.lean +++ b/src/Std/Time/Sign.lean @@ -4,9 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Bounded +import Std.Time.Bounded + +namespace Std +namespace Time -namespace Lean set_option linter.all true /-- @@ -14,6 +16,11 @@ A `Sign` is a type that can have three values -1, 0 and 1 that represents the si -/ def Sign := Bounded.LE (-1) 1 +instance : ToString Sign where + toString + | ⟨-1, _⟩ => "-" + | _ => "+" + namespace Sign /-- @@ -25,6 +32,13 @@ def ofInt (val : Int) : Sign := by simp [Int.sign] split <;> simp +/-- +Applies the sign to a value. +-/ @[inline] def apply (sign : Sign) (val : Int) : Int := val * sign.val + +end Sign +end Time +end Std diff --git a/src/Std/Time/Time.lean b/src/Std/Time/Time.lean new file mode 100644 index 000000000000..bac27a2e8850 --- /dev/null +++ b/src/Std/Time/Time.lean @@ -0,0 +1,10 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.Time.HourMarker +import Std.Time.Time.Basic +import Std.Time.Time.Scalar +import Std.Time.Time.Time diff --git a/src/Lean/Time/Date/Basic.lean b/src/Std/Time/Time/Basic.lean similarity index 83% rename from src/Lean/Time/Date/Basic.lean rename to src/Std/Time/Time/Basic.lean index 4f10e9df36ae..754d0bc6bea7 100644 --- a/src/Lean/Time/Date/Basic.lean +++ b/src/Std/Time/Time/Basic.lean @@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Date.Unit.Basic +import Std.Time.Time.Unit.Basic diff --git a/src/Lean/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean similarity index 89% rename from src/Lean/Time/Time/HourMarker.lean rename to src/Std/Time/Time/HourMarker.lean index 2c415707b2f7..1d79edeb31a3 100644 --- a/src/Lean/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -4,17 +4,21 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time.Basic +import Std.Time.Time.Basic -namespace Lean +namespace Std namespace Time +set_option linter.all true + /-- `HourMarker` represents the two 12-hour periods of the day: `am` for hour between 12:00 AM and 11:59 AM, and `pm` for hour between 12:00 PM and 11:59 PM. -/ inductive HourMarker + /-- Ante meridiem. -/ | am + /-- Post meridiem. -/ | pm /-- diff --git a/src/Lean/Time/Time/Scalar.lean b/src/Std/Time/Time/Scalar.lean similarity index 58% rename from src/Lean/Time/Time/Scalar.lean rename to src/Std/Time/Time/Scalar.lean index a6a62d557e05..1b080b3a242e 100644 --- a/src/Lean/Time/Time/Scalar.lean +++ b/src/Std/Time/Time/Scalar.lean @@ -4,25 +4,47 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time.Basic +import Std.Time.Time.Basic -namespace Lean +namespace Std namespace Time +namespace Time + +set_option linter.all true +/-- +Represents a scalar value of time. +-/ structure Scalar where + /-- + The quantity of seconds. + -/ second : Second.Offset + /-- + The quantity of nanoseconds. + -/ nano : Nanosecond.Ordinal namespace Scalar +/-- +Converts the scalar value to seconds. +-/ def toSeconds (time : Scalar) : Second.Offset := time.second +/-- +Converts the scalar value to minutes. +-/ def toMinutes (time : Scalar) : Minute.Offset := time.second.toMinutes +/-- +Converts the scalar value to hours. +-/ def toHours (time : Scalar) : Hour.Offset := time.second.toHours end Scalar end Time +end Time diff --git a/src/Lean/Time/Time/Time.lean b/src/Std/Time/Time/Time.lean similarity index 53% rename from src/Lean/Time/Time/Time.lean rename to src/Std/Time/Time/Time.lean index e692fbd227ed..45a6ef53637e 100644 --- a/src/Lean/Time/Time/Time.lean +++ b/src/Std/Time/Time/Time.lean @@ -4,11 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time.Basic +import Std.Time.Time.Basic -namespace Lean +namespace Std namespace Time +/-- +Represents a specific point in time, including hours, minutes, seconds, and nanoseconds. +-/ structure Time where hour : Hour.Ordinal minute : Minute.Ordinal @@ -18,12 +21,28 @@ structure Time where namespace Time +/-- +Creates a `Time` value from hours, minutes, and seconds, setting nanoseconds to zero. +-/ +def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal) : Time := + ⟨hour, minute, second, 0⟩ + +/-- +Converts a `Time` value to the total number of seconds since midnight. +-/ def toSeconds (time : Time) : Second.Offset := time.hour.toOffset.toSeconds + time.minute.toOffset.toSeconds + time.second.toOffset +/-- +Converts a `Time` value to the total number of minutes since midnight. +-/ def toMinutes (time : Time) : Minute.Offset := time.hour.toOffset.toMinutes + time.minute.toOffset + time.second.toOffset.toMinutes + +end Time +end Time +end Std diff --git a/src/Lean/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean similarity index 79% rename from src/Lean/Time/Time/Unit/Basic.lean rename to src/Std/Time/Time/Unit/Basic.lean index cab3ad927554..8c7c2fe0f289 100644 --- a/src/Lean/Time/Time/Unit/Basic.lean +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -4,18 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Time.Unit.Hour -import Lean.Time.Time.Unit.Minute -import Lean.Time.Time.Unit.Second -import Lean.Time.Time.Unit.Nanosecond +import Std.Time.UnitVal +import Std.Time.Time.Unit.Hour +import Std.Time.Time.Unit.Minute +import Std.Time.Time.Unit.Second +import Std.Time.Time.Unit.Nanosecond -set_option linter.all true - -namespace Lean +namespace Std namespace Time namespace Second.Offset +set_option linter.all true + /-- Convert `Second.Offset` to `Minute.Offset` -/ @[inline] def toMinutes (offset : Second.Offset) : Minute.Offset := @@ -38,3 +38,4 @@ def toHours (offset : Minute.Offset) : Hour.Offset := end Minute.Offset end Time +end Std diff --git a/src/Lean/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean similarity index 90% rename from src/Lean/Time/Time/Unit/Hour.lean rename to src/Std/Time/Time/Unit/Hour.lean index 947587593f58..62e5a85a4f54 100644 --- a/src/Lean/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -4,20 +4,19 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -import Lean.Time.Time.Unit.Minute -import Lean.Time.Time.Unit.Second +import Std.Time.Time.Unit.Minute +import Std.Time.Time.Unit.Second -namespace Lean +namespace Std namespace Time +namespace Hour set_option linter.all true -namespace Hour - /-- `Ordinal` represents a bounded value for hour, which ranges between 0 and 23. -/ @@ -79,3 +78,5 @@ def toMinutes (val : Offset) : Minute.Offset := end Offset end Hour +end Time +end Std diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean new file mode 100644 index 000000000000..04b34b44f70d --- /dev/null +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -0,0 +1,36 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq +import Lean.Data.Rat + +namespace Std +namespace Time +namespace Millisecond + +/-- +`Ordinal` represents a bounded value for milliseconds, which ranges between 0 and 999. +-/ +def Ordinal := Bounded.LE 0 999 + deriving Repr, BEq, LE, LT + +instance : OfNat Ordinal n where ofNat := Bounded.LE.ofFin (Fin.ofNat n) + +instance : Inhabited Ordinal where default := 0 + +/-- +`Offset` represents an offset in milliseconds. It is defined as an `Int`. +-/ +def Offset : Type := UnitVal (1 / 1000) + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ + +end Millisecond +end Time +end Std diff --git a/src/Lean/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean similarity index 91% rename from src/Lean/Time/Time/Unit/Minute.lean rename to src/Std/Time/Time/Unit/Minute.lean index 24410e572a4c..cf5d944baa7a 100644 --- a/src/Lean/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -4,19 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -import Lean.Time.Time.Unit.Second +import Std.Time.Time.Unit.Second -namespace Lean +namespace Std namespace Time +namespace Minute set_option linter.all true -namespace Minute - /-- `Ordinal` represents a bounded value for minute, which ranges between 0 and 59. -/ @@ -71,3 +70,5 @@ def toSeconds (val : Offset) : Second.Offset := end Offset end Minute +end Time +end Std diff --git a/src/Lean/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean similarity index 70% rename from src/Lean/Time/Time/Unit/Nanosecond.lean rename to src/Std/Time/Time/Unit/Nanosecond.lean index 1f793df2449c..9ab9a7b295b5 100644 --- a/src/Lean/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat +import Std.Time.Time.Unit.Millisecond -namespace Lean +namespace Std namespace Time set_option linter.all true @@ -27,6 +28,23 @@ instance : OfNat Ordinal n where ofNat := Bounded.LE.ofFin (Fin.ofNat n) instance : Inhabited Ordinal where default := 0 +namespace Ordinal + +/-- +Convert to `Millisecond.Ordinal` +-/ +def toMillisecond (nano : Ordinal) : Millisecond.Ordinal := + nano.div 1000000 (by decide) + +/-- +Convert from `Millisecond.Ordinal` +-/ +def ofMillisecond (nano : Millisecond.Ordinal) : Nanosecond.Ordinal := + nano.mul_pos 1000000 (by decide) + |>.expandTop (by decide) + +end Ordinal + /-- `Offset` represents an offset in nanoseconds. It is defined as an `Int`. -/ @@ -43,3 +61,7 @@ def Span := Bounded.LE (-999999999) 999999999 deriving Repr, BEq, LE, LT instance : Inhabited Span where default := Bounded.LE.mk 0 (by decide) + +end Nanosecond +end Time +end Std diff --git a/src/Lean/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean similarity index 77% rename from src/Lean/Time/Time/Unit/Second.lean rename to src/Std/Time/Time/Unit/Second.lean index 82a090b7d7c3..dd412023010b 100644 --- a/src/Lean/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -4,19 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.UnitVal -import Lean.Time.Bounded -import Lean.Time.LessEq +import Std.Time.UnitVal +import Std.Time.Bounded +import Std.Time.LessEq import Lean.Data.Rat -import Lean.Time.Time.Unit.Nanosecond +import Std.Time.Time.Unit.Nanosecond -namespace Lean +namespace Std namespace Time +namespace Second set_option linter.all true -namespace Second - /-- `Ordinal` represents a bounded value for second, which ranges between 0 and 60. This accounts for potential leap second. @@ -60,13 +59,6 @@ def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val end Ordinal - -namespace Offset - -def ofNanosecond (offset : Nanosecond.Offset) : Second.Offset × Nanosecond.Span := - let second := offset.val.div 1000000000 - let nano := Bounded.LE.byMod offset.val 1000000000 (by decide) - ⟨UnitVal.ofInt second, nano⟩ - -end Offset end Second +end Time +end Std diff --git a/src/Lean/Time/UnitVal.lean b/src/Std/Time/UnitVal.lean similarity index 91% rename from src/Lean/Time/UnitVal.lean rename to src/Std/Time/UnitVal.lean index bc7c0221d0f3..29b8223100da 100644 --- a/src/Lean/Time/UnitVal.lean +++ b/src/Std/Time/UnitVal.lean @@ -6,8 +6,12 @@ Authors: Sofia Rodrigues prelude import Init import Lean.Data.Rat +import Std.Time.Sign + +namespace Std +namespace Time +open Lean -namespace Lean set_option linter.all true /-- @@ -72,6 +76,15 @@ Subtracts one `UnitVal` value from another of the same ratio. def sub (u1 : UnitVal α) (u2 : UnitVal α) : UnitVal α := ⟨u1.val - u2.val⟩ +/-- +Apply the `Sign` to a value. +-/ +@[inline] +def apply (u1 : UnitVal α) (sign : Sign) : UnitVal α := + ⟨u1.val * sign.val⟩ + + + /-- Converts an `Offset` to another unit type. -/ @@ -89,3 +102,7 @@ instance : Mul (UnitVal α) where mul x y := ⟨x.val * y.val⟩ instance : Div (UnitVal α) where div x y := ⟨x.val / y.val⟩ instance : Neg (UnitVal α) where neg x := ⟨-x.val⟩ instance : ToString (UnitVal n) where toString n := toString n.val + +end UnitVal +end Time +end Std diff --git a/src/Lean/Time/Date.lean b/src/Std/Time/Zoned.lean similarity index 63% rename from src/Lean/Time/Date.lean rename to src/Std/Time/Zoned.lean index e298cfdb95c6..224a5ff528c8 100644 --- a/src/Lean/Time/Date.lean +++ b/src/Std/Time/Zoned.lean @@ -4,6 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Date.Basic -import Lean.Time.Date.Scalar -import Lean.Time.Date.Date +import Std.Time.Zoned.Basic +import Std.Time.Zoned.DateTime +import Std.Time.Zoned.ZonedDateTime diff --git a/src/Lean/Time/TimeZone/Basic.lean b/src/Std/Time/Zoned/Basic.lean similarity index 82% rename from src/Lean/Time/TimeZone/Basic.lean rename to src/Std/Time/Zoned/Basic.lean index 41ea2ee86f7e..d6624ccc2884 100644 --- a/src/Lean/Time/TimeZone/Basic.lean +++ b/src/Std/Time/Zoned/Basic.lean @@ -3,4 +3,4 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ -import Lean.Time.TimeZone.Offset +import Std.Time.Zoned.Offset diff --git a/src/Lean/Time/TimeZone/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean similarity index 68% rename from src/Lean/Time/TimeZone/DateTime.lean rename to src/Std/Time/Zoned/DateTime.lean index 450096602673..94de5383d779 100644 --- a/src/Lean/Time/TimeZone/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -4,14 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time -import Lean.Time.Date -import Lean.Time.DateTime -import Lean.Time.TimeZone.TimeZone +import Std.Time.Time +import Std.Time.Date +import Std.Time.DateTime +import Std.Time.Zoned.TimeZone -namespace Lean -namespace TimeZone -open Time Date DateTime +namespace Std +namespace Time /-- It stores a `Timestamp`, a `NaiveDateTime` and a `TimeZone` @@ -20,20 +19,36 @@ structure DateTime (tz : TimeZone) where private mk :: timestamp : Timestamp date : NaiveDateTime - deriving Repr + deriving Repr, Inhabited namespace DateTime /-- -Creates a new DateTime out of a `Timestamp` +Creates a new `DateTime` out of a `Timestamp` -/ +@[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := let date := (tm + tz.toSeconds).toNaiveDateTime DateTime.mk tm date +/-- +Creates a new `Timestamp` out of a `DateTime` +-/ +@[inline] +def toTimestamp (date : DateTime tz) : Timestamp := + date.timestamp + +/-- +Changes the `TimeZone` to a new one. +-/ +@[inline] +def convertTimeZone (date : DateTime tz) (tz₁ : TimeZone) : DateTime tz₁ := + ofTimestamp (date.toTimestamp) tz₁ + /-- Creates a new DateTime out of a `NaiveDateTime` -/ +@[inline] def ofNaiveDateTime (date : NaiveDateTime) (tz : TimeZone) : DateTime tz := let tm := date.toTimestamp DateTime.mk tm date @@ -81,8 +96,18 @@ def second (dt : DateTime tz) : Second.Ordinal := dt.date.second /-- +Getter for the `Milliseconds` inside of a `DateTime` +-/ +@[inline] +def milliseconds (dt : DateTime tz) : Millisecond.Ordinal := + dt.date.time.nano.toMillisecond +/-- Gets the `Weekday` of a DateTime. -/ @[inline] def weekday (dt : DateTime tz) : Weekday := dt.date.date.weekday + +end DateTime +end Time +end Std diff --git a/src/Lean/Time/TimeZone/Offset.lean b/src/Std/Time/Zoned/Offset.lean similarity index 84% rename from src/Lean/Time/TimeZone/Offset.lean rename to src/Std/Time/Zoned/Offset.lean index a8ffa3ec6e0a..57c1b4549a30 100644 --- a/src/Lean/Time/TimeZone/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -4,20 +4,21 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time.Unit.Basic +import Std.Time.Time.Unit.Basic +import Std.Time.Sign -namespace Lean -namespace TimeZone -open Time +namespace Std +namespace Time /-- Represents a timezone offset with an hour and second component. -/ -structure Offset where +structure TimeZone.Offset where hour: Hour.Offset second: Second.Offset deriving Repr, Inhabited +namespace TimeZone namespace Offset /-- @@ -25,7 +26,8 @@ Converts an `Offset` to a string in ISO 8601 format. The `colon` parameter deter and minute components are separated by a colon (e.g., "+01:00" or "+0100"). -/ def toIsoString (offset : Offset) (colon : Bool) : String := - let (sign, time) := if offset.second.val ≥ 0 then ("+", offset.second) else ("-", -offset.second) + let sign := Sign.ofInt offset.second.val + let time := offset.second.apply sign let hour : Hour.Offset := time.div 3600 let minute := Int.div (Int.mod time.val 3600) 60 let hourStr := if hour.val < 10 then s!"0{hour.val}" else toString hour.val @@ -50,3 +52,8 @@ Creates an `Offset` from a given number of second. -/ def ofSeconds (n: Second.Offset) : Offset := mk n.toHours n + +end Offset +end TimeZone +end Time +end Std diff --git a/src/Lean/Time/TimeZone/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean similarity index 74% rename from src/Lean/Time/TimeZone/TimeZone.lean rename to src/Std/Time/Zoned/TimeZone.lean index 092a7702d10b..c46f04d317be 100644 --- a/src/Lean/Time/TimeZone/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -4,33 +4,33 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time.Unit.Basic -import Lean.Time.TimeZone.Basic +import Std.Time.Time.Unit.Basic +import Std.Time.Zoned.Basic -namespace Lean -namespace TimeZone -open Time +namespace Std +namespace Time /-- An enumeration representing different time zones. -/ structure TimeZone where - offset : Offset + offset : TimeZone.Offset name : String + deriving Inhabited namespace TimeZone /-- A zeroed `Timezone` representing UTC (no offset). -/ -def UTC (name : String) : TimeZone := - TimeZone.mk (Offset.zero) name +def UTC : TimeZone := + TimeZone.mk (Offset.zero) "Coordinated Universal Time" /-- A zeroed `Timezone` representing GMT (no offset). -/ -def GMT (name : String) : TimeZone := - TimeZone.mk (Offset.zero) name +def GMT : TimeZone := + TimeZone.mk (Offset.zero) "Greenwich Mean Time" /-- Creates a `Timestamp` from a given number of hour. diff --git a/src/Lean/Time/TimeZone/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean similarity index 58% rename from src/Lean/Time/TimeZone/ZonedDateTime.lean rename to src/Std/Time/Zoned/ZonedDateTime.lean index 9f73d44d9432..9c43211d9752 100644 --- a/src/Lean/Time/TimeZone/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -4,31 +4,53 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Time.Time -import Lean.Time.Date -import Lean.Time.DateTime -import Lean.Time.TimeZone.TimeZone -import Lean.Time.TimeZone.DateTime - -namespace Lean -namespace TimeZone +import Std.Time.Time +import Std.Time.Date +import Std.Time.DateTime +import Std.Time.Zoned.TimeZone +import Std.Time.Zoned.DateTime + +namespace Std +namespace Time open Time Date DateTime def ZonedDateTime := Sigma DateTime +instance : CoeDep ZonedDateTime d (DateTime d.fst) where + coe := d.snd + +instance : Inhabited ZonedDateTime where + default := ⟨Inhabited.default, Inhabited.default⟩ + namespace ZonedDateTime open DateTime /-- Creates a new `ZonedDateTime` out of a `Timestamp` -/ -def ofTimestamp (tm : DateTime.Timestamp) (tz : TimeZone) : ZonedDateTime := +@[inline] +def ofTimestamp (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := ⟨tz, DateTime.ofTimestamp tm tz⟩ +/-- +Creates a new `Timestamp` out of a `ZonedDateTime` +-/ +@[inline] +def toTimestamp (date : ZonedDateTime) : Timestamp := + date.snd.toTimestamp + +/-- +Changes the `TimeZone` to a new one. +-/ +@[inline] +def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := + ofTimestamp (date.toTimestamp) tz₁ + /-- Creates a new `ZonedDateTime` out of a `NaiveDateTime` -/ -def ofNaiveDateTime (date : DateTime.NaiveDateTime) (tz : TimeZone) : ZonedDateTime := +@[inline] +def ofNaiveDateTime (date : NaiveDateTime) (tz : TimeZone) : ZonedDateTime := ⟨tz, DateTime.ofNaiveDateTime date tz⟩ /-- @@ -77,8 +99,14 @@ def second (zdt : ZonedDateTime) : Second.Ordinal := Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` -/ @[inline] -def offset (zdt : ZonedDateTime) : Offset := +def offset (zdt : ZonedDateTime) : TimeZone.Offset := zdt.fst.offset +/-- +Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` +-/ +@[inline] +def timezone (zdt : ZonedDateTime) : TimeZone := + zdt.fst /-- Returns the weekday. @@ -86,3 +114,7 @@ Returns the weekday. @[inline] def weekday (zdt : ZonedDateTime) : Weekday := zdt.snd.weekday + +end ZonedDateTime +end Time +end Std From a2e3a4e604ff59fd7dd607206dc52e3e289b3b2d Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 7 Aug 2024 17:39:26 -0300 Subject: [PATCH 004/118] feat: api docs, parser changes, notation, renames and improve documentation --- src/Std.lean | 1 + src/Std/Time.lean | 166 ++++++++++- src/Std/Time/Date.lean | 3 +- .../Time/Date/{Date.lean => LocalDate.lean} | 75 +++-- src/Std/Time/Date/Scalar.lean | 7 +- src/Std/Time/Date/Unit/Basic.lean | 3 +- src/Std/Time/Date/Unit/Day.lean | 20 +- src/Std/Time/Date/Unit/Month.lean | 59 ++-- src/Std/Time/Date/Unit/WeekOfYear.lean | 5 +- src/Std/Time/Date/Unit/Weekday.lean | 5 +- src/Std/Time/Date/Unit/Year.lean | 13 +- src/Std/Time/Date/WeekDate.lean | 9 +- src/Std/Time/DateTime.lean | 14 +- ...{NaiveDateTime.lean => LocalDateTime.lean} | 52 ++-- src/Std/Time/DateTime/Timestamp.lean | 3 +- src/Std/Time/Duration.lean | 3 +- src/Std/Time/Format.lean | 61 ++-- src/Std/Time/Format/Basic.lean | 96 +++---- .../{DateTime/Basic.lean => Internal.lean} | 5 +- src/Std/Time/{ => Internal}/Bounded.lean | 16 +- src/Std/Time/{ => Internal}/LessEq.lean | 6 + src/Std/Time/{ => Internal}/Sign.lean | 4 +- src/Std/Time/{ => Internal}/UnitVal.lean | 5 +- src/Std/Time/Notation.lean | 263 ++++++++++++++++++ src/Std/Time/Time.lean | 5 +- src/Std/Time/Time/HourMarker.lean | 3 +- .../Time/Time/{Time.lean => LocalTime.lean} | 18 +- src/Std/Time/Time/Scalar.lean | 50 ---- src/Std/Time/Time/Unit/Basic.lean | 3 +- src/Std/Time/Time/Unit/Hour.lean | 8 +- src/Std/Time/Time/Unit/Millisecond.lean | 5 +- src/Std/Time/Time/Unit/Minute.lean | 5 +- src/Std/Time/Time/Unit/Nanosecond.lean | 17 +- src/Std/Time/Time/Unit/Second.lean | 5 +- src/Std/Time/Zoned.lean | 1 - src/Std/Time/Zoned/Basic.lean | 6 - src/Std/Time/Zoned/DateTime.lean | 14 +- src/Std/Time/Zoned/Offset.lean | 17 +- src/Std/Time/Zoned/TimeZone.lean | 6 +- src/Std/Time/Zoned/ZonedDateTime.lean | 10 +- 40 files changed, 730 insertions(+), 337 deletions(-) rename src/Std/Time/Date/{Date.lean => LocalDate.lean} (58%) rename src/Std/Time/DateTime/{NaiveDateTime.lean => LocalDateTime.lean} (70%) rename src/Std/Time/{DateTime/Basic.lean => Internal.lean} (55%) rename src/Std/Time/{ => Internal}/Bounded.lean (95%) rename src/Std/Time/{ => Internal}/LessEq.lean (90%) rename src/Std/Time/{ => Internal}/Sign.lean (91%) rename src/Std/Time/{ => Internal}/UnitVal.lean (97%) create mode 100644 src/Std/Time/Notation.lean rename src/Std/Time/Time/{Time.lean => LocalTime.lean} (62%) delete mode 100644 src/Std/Time/Time/Scalar.lean delete mode 100644 src/Std/Time/Zoned/Basic.lean diff --git a/src/Std.lean b/src/Std.lean index 47f874eb224c..a59bea450007 100644 --- a/src/Std.lean +++ b/src/Std.lean @@ -6,3 +6,4 @@ Authors: Sebastian Ullrich prelude import Std.Data import Std.Sat +import Std.Time diff --git a/src/Std/Time.lean b/src/Std/Time.lean index a5f786b1806f..2913fe3a13bc 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -5,6 +5,170 @@ Authors: Sofia Rodrigues -/ import Std.Time.Time import Std.Time.Date -import Std.Time.DateTime import Std.Time.Zoned import Std.Time.Format +import Std.Time.DateTime +import Std.Time.Notation + +namespace Time + +/-! +# Time + +The Lean4 API for date, time, and duration functionalities, following the ISO standards. + +# Overview + +This module of the standard library defines various concepts related to time, such as dates, times, +time zones and durations. These types are designed to be strongly-typed and to avoid problems with +conversion. It offers both unbounded and bounded variants to suit different use cases, like +adding days to a date or representing ordinal values. + +# Date and Time Components + +Date and time components are classified into different types based how you SHOULD use them. These +components are categorized as: + +## Offset + +Offsets represent unbounded shifts in specific date or time units. They are typically used in operations +like `date.addDays` where a `Day.Offset` is the parameter. Some offsets, such as `Month.Offset`, do not +correspond directly to a specific duration in seconds, as their value depends on the context (if +the year is leap or the size of the month). Offsets with a clear correspondent to seconds can be +converted because they use a internal type called `UnitVal`. + +- Types with a correspondence to seconds: + - `Day.Offset` + - `Hour.Offset` + - `WeekOfYear.Offset` + - `Millisecond.Offset` + - `Nanosecond.Offset` + - `Second.Offset` + +- Types without a correspondence to seconds: + - `Month.Offset` + - `Year.Offset` + +## Ordinal + +Ordinal types represent specific bounded values in reference to another unit, e.g, `Day.Ordinal` +represents a day in a month, ranging from 1 to 31. Some ordinal types like `Hour.Ordinal` and `Second.Ordinal`, +allow for values beyond the normal range (e.g, 24 hours and 61 seconds) to accomodate special cases +with leap seconds like `24:00:00` that is valid in ISO 8601. + +- Ordinal types: + - `Day.Ordinal`: Ranges from 1 to 31. + - `Day.Ordinal.OfYear`: Ranges from 1 to (365 or 366). + - `Month.Ordinal`: Ranges from 1 to 12. + - `WeekOfYear.Ordinal`: Ranges from 1 to 53. + - `Hour.Ordinal`: Ranges from 0 to 24. + - `Millisecond.Ordinal`: Ranges from 0 to 999. + - `Minute.Ordinal`: Ranges from 0 to 59. + - `Nanosecond.Ordinal`: Ranges from 0 to 999,999,999. + - `Second.Ordinal`: Ranges from 0 to 60. + - `Weekday`: That is a inductive type with all the seven days. + +## Span + +Span types are used as subcomponents of other types. They represent a range of values in the limits +of the parent type, e.g, `Nanosecond.Span` ranges from -999,999,999 to 999,999,999, as 1,000,000,000 +nanoseconds corresponds to one second. + +- Span types: + - `Nanosecond.Span`: Ranges from -999,999,999 to 999,999,999. + +# Date and Time Types + +Dates and times are composed of these components. Dates are "absolute" value in contrast with Offsets +that are just shifts in dates and times. Types like `Date` are made using of components such as `Year.Offset`, +`Month.Ordinal`, and `Day.Ordinal`, with a proof of the date's validity. Some types, like `Date.Scalar`, +are wrappers around offsets (e.g., `Day.Offset`) and represent absolute dates relative to the UNIX Epoch. + +## Date +These types provide precision down to the day level, useful for representing and manipulating dates. + +- **`LocalDate`:** Represents a calendar date in the format `YYYY-MM-DD`. +- **`LocalDate.Scalar`:** : Represents an absolute date with day-level precision in a compact format `DD`. + It encapsulates days since the UNIX Epoch. +- **`WeekDate`:** Uses the `YYYY-Www` format with week level precision. + +## Time +These types offer precision down to the nanosecond level, useful for representing and manipulating time of day. + +- **`LocalTime`**: Represents a time of day in the format `HH:mm:ss.SSSSSSSSS`. + +## Date and time +Combines date and time into a single representation, useful for precise timestamps and scheduling. + +- **`LocalDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. +- **`Timestamp`**: Represents a point in time with second-level precision. + +## Zoned date and times. +Combines date, time and time zones. + +- **`DateTime`**: Represents both date and time but with a time zone in the type constructor. +- **`ZonedDateTime`**: An existencial version of the `DateTime`. + +## Duration +Represents spans of time and the difference between two points in time. + +- **`Instant`**: Represents a non-negative instant in time with nanosecond precision, represents a + specific point on the time axis. +- **`Duration`**: Represents the time span or difference between two `Instant` values with nanosecond precision. + +# Formats + +Format strings are used to convert between `String` representations and date/time types, e.g, `YYYY-MM-DD'T'hh:mm:ss.sssZ`. +The supported formats include: + +- **Year:** + - `YYYY`: Four-digit year (e.g., 2024) + - `YY`: Two-digit year (e.g., 24 for 2024) +- **Month:** + - `MMMM`: Full month name (e.g., January, February) + - `MMM`: Abbreviated month name (e.g., Jan, Feb) + - `MM`: Two-digit month (e.g., 01 for January) + - `M`: One or two-digit month (e.g., 1 for January, 10 for October) +- **Day:** + - `DD`: Two-digit day of the month (e.g., 01, 02) + - `D`: One or two-digit day of the month (e.g., 1, 2) + - `d`: One or two-digit day of the month with space padding (e.g., " 1", "12") +- **Day of Week:** + - `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday) + - `EEE`: Abbreviated day of the week (e.g., Mon, Tue) +- **Hour:** + - `hh`: Two-digit hour in 24-hour format (e.g., 13, 14) + - `h`: One or two-digit hour in 24-hour format (e.g., 1, 2) + - `HH`: Two-digit hour in 12-hour format (e.g., 01, 02) + - `H`: One or two-digit hour in 12-hour format (e.g., 1, 2) +- **AM/PM Indicator:** + - `AA`: Uppercase AM/PM (e.g., AM, PM) + - `aa`: Lowercase am/pm (e.g., am, pm) +- **Minute:** + - `mm`: Two-digit minute (e.g., 01, 02) + - `m`: One or two-digit minute (e.g., 1, 2) +- **Second:** + - `sss`: Three-digit milliseconds (e.g., 001, 202) + - `ss`: Two-digit second (e.g., 01, 02) + - `s`: One or two-digit second (e.g., 1, 2) +- **Timezone:** + - `ZZZZZ`: Full timezone offset with hours and minutes (e.g., +03:00) + - `ZZZZ`: Timezone offset without the colon (e.g., +0300) + - `ZZZ`: Like `ZZZZ`, but with "UTC" for UTC + - `Z`: Like `ZZZZZ`, but with "Z" for UTC + - `z`: Timezone name (e.g., Brasilia Standard Time) + +# Macros + +In order to help the user build dates easily, there are a lot of macros available for creating dates. + +- **`date% DD-MM-YYYY`**: Defines a date in the `DD-MM-YYYY` format. +- **`date% HH:mm:ss`**: Defines a time in the `HH:mm:ss` format. +- **`date% HH:mm:ss.sssssss`**: Defines a time in the `HH:mm:ss.sssssss` format, including fractional seconds. +- **`date% YYYY-MM-DD:HH:mm:ss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss` format. +- **`date% YYYY-MM-DDTHH:mm:ssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ssZ` format. +- **`offset% +HH:mm`**: Defines a timezone offset in the format `+HH:mm`. +- **`timezone% NAME/ID offset% +HH:mm`**: Defines a timezone with a name and an offset. +- **`date-spec% "format"`**: Defines a date specification format at compile time using the provided format string. + +-/ diff --git a/src/Std/Time/Date.lean b/src/Std/Time/Date.lean index 80dbf3ad7c46..b7634f036bac 100644 --- a/src/Std/Time/Date.lean +++ b/src/Std/Time/Date.lean @@ -5,6 +5,5 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Date.Basic -import Std.Time.Date.Scalar -import Std.Time.Date.Date import Std.Time.Date.WeekDate +import Std.Time.Date.LocalDate diff --git a/src/Std/Time/Date/Date.lean b/src/Std/Time/Date/LocalDate.lean similarity index 58% rename from src/Std/Time/Date/Date.lean rename to src/Std/Time/Date/LocalDate.lean index 3e1e4273dcf0..9c805e4cb819 100644 --- a/src/Std/Time/Date/Date.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal +import Std.Time.Internal import Std.Time.Date.Basic import Std.Time.Date.Scalar @@ -14,44 +14,44 @@ namespace Time /-- Date in YMD format. -/ -structure Date where +structure LocalDate where year : Year.Offset month : Month.Ordinal day : Day.Ordinal - valid : year.valid month day + valid : year.Valid month day deriving Repr -namespace Date +namespace LocalDate /-- Force the date to be valid. -/ -def force (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Date := - let ⟨day, valid⟩ := month.forceDay year.isLeap day - Date.mk year month day valid +def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : LocalDate := + let ⟨day, valid⟩ := month.clipDay year.isLeap day + LocalDate.mk year month day valid -instance : Inhabited Date where - default := force 0 1 1 +instance : Inhabited LocalDate where + default := clip 0 1 1 /-- -Creates a new `Date` using YMD. +Creates a new `LocalDate` using YMD. -/ -def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option Date := - if valid : year.valid month day - then some (Date.mk year month day valid) +def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option LocalDate := + if valid : year.Valid month day + then some (LocalDate.mk year month day valid) else none /-- -Creates a new `Date` using YO. +Creates a new `LocalDate` using YO. -/ -def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : Date := +def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : LocalDate := let ⟨⟨month, day⟩, valid⟩ := ordinal.toMonthAndDay - Date.mk year month day valid + LocalDate.mk year month day valid /-- -Creates a new `Date` using the `Day.Offset` which `0` corresponds the UNIX Epoch. +Creates a new `LocalDate` using the `Day.Offset` which `0` corresponds the UNIX Epoch. -/ -def ofDaysSinceUNIXEpoch (day : Day.Offset) : Date := +def ofDaysSinceUNIXEpoch (day : Day.Offset) : LocalDate := let z := day.toInt + 719468 let era := (if z ≥ 0 then z else z - 146096) / 146097 let doe := z - era * 146097 @@ -63,12 +63,12 @@ def ofDaysSinceUNIXEpoch (day : Day.Offset) : Date := let m := mp + (if mp < 10 then 3 else -9) let y := y + (if m <= 2 then 1 else 0) - .force y (.force m (by decide)) (.force (d + 1) (by decide)) + .clip y (.clip m (by decide)) (.clip (d + 1) (by decide)) /-- -Returns the `Weekday` of a `Date`. +Returns the `Weekday` of a `LocalDate`. -/ -def weekday (date : Date) : Weekday := +def weekday (date : LocalDate) : Weekday := let q := date.day.toInt let m := date.month.toInt let y := date.year.toInt @@ -85,9 +85,9 @@ def weekday (date : Date) : Weekday := .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ /-- -Returns the `Weekday` of a `Date` using Zeller's Congruence for the Julian calendar. +Returns the `Weekday` of a `LocalDate` using Zeller's Congruence for the Julian calendar. -/ -def weekdayJulian (date : Date) : Weekday := +def weekdayJulian (date : LocalDate) : Weekday := let month := date.month.toInt let year := date.year.toInt @@ -104,9 +104,9 @@ def weekdayJulian (date : Date) : Weekday := .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ /-- -Convert `Date` to `Day.Offset` since the UNIX Epoch. +Convert `LocalDate` to `Day.Offset` since the UNIX Epoch. -/ -def toDaysSinceUNIXEpoch (date : Date) : Day.Offset := +def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := let y : Int := if date.month.toInt > 2 then date.year else date.year.toInt - 1 let era : Int := (if y ≥ 0 then y else y - 399) / 400 let yoe : Int := y - era * 400 @@ -118,29 +118,26 @@ def toDaysSinceUNIXEpoch (date : Date) : Day.Offset := .ofInt (era * 146097 + doe - 719468) /-- -Returns the `Scalar` starting from the UNIX epoch. +Convert a `Scalar` to a `LocalDate` since the UNIX Epoch. -/ -def toScalar (date : Date) : Date.Scalar := - ⟨toDaysSinceUNIXEpoch date⟩ +def ofScalar (day : Scalar) : LocalDate := + ofDaysSinceUNIXEpoch day.day /-- -Creates a new `Date` from a `Scalar` that begins on the epoch. +Convert a `LocalDate` to a `Scalar` since the UNIX Epoch. -/ -def ofScalar (scalar : Date.Scalar) : Date := - ofDaysSinceUNIXEpoch scalar.day +def toScalar (date : LocalDate) : Scalar := + ⟨toDaysSinceUNIXEpoch date⟩ /-- -Calculate the Year.Offset from a Date +Calculate the Year.Offset from a LocalDate -/ -def yearsSince (date : Date) (year : Year.Offset) : Year.Offset := +def yearsSince (date : LocalDate) (year : Year.Offset) : Year.Offset := date.year - year -instance : HAdd Date Day.Offset Date where - hAdd date day := ofScalar (toScalar date + ⟨day⟩) - -instance : HAdd Date Scalar Date where - hAdd date day := ofScalar (toScalar date + day) +instance : HAdd LocalDate Day.Offset LocalDate where + hAdd date day := ofDaysSinceUNIXEpoch (toDaysSinceUNIXEpoch date + day) -end Date +end LocalDate end Time end Std diff --git a/src/Std/Time/Date/Scalar.lean b/src/Std/Time/Date/Scalar.lean index 982135b13085..2ff31be8ab8f 100644 --- a/src/Std/Time/Date/Scalar.lean +++ b/src/Std/Time/Date/Scalar.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal +import Std.Time.Internal import Std.Time.Date.Basic namespace Std namespace Time -namespace Date +namespace LocalDate +open Internal /-- `Scalar` represents a date offset, using the number of day as the underlying unit. @@ -63,5 +64,5 @@ def ofOffset (offset : Day.Offset) : Scalar := ⟨offset⟩ end Scalar -end Date +end LocalDate end Time diff --git a/src/Std/Time/Date/Unit/Basic.lean b/src/Std/Time/Date/Unit/Basic.lean index 7fca03623866..27f5b563f887 100644 --- a/src/Std/Time/Date/Unit/Basic.lean +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal import Std.Time.Date.Unit.Day import Std.Time.Date.Unit.Month import Std.Time.Date.Unit.Year @@ -18,7 +17,7 @@ namespace Time.Day.Ordinal.OfYear Conevrts a `Year` and a `Ordinal.OfYear` to a valid day and month. -/ @[inline] -def toMonthAndDay (year : Year.Offset) (ordinal : OfYear year.isLeap) : { val : Month.Ordinal × Day.Ordinal // Year.Offset.valid year (Prod.fst val) (Prod.snd val) } := +def toMonthAndDay (year : Year.Offset) (ordinal : OfYear year.isLeap) : { val : Month.Ordinal × Day.Ordinal // Year.Offset.Valid year (Prod.fst val) (Prod.snd val) } := Month.Ordinal.ofOrdinal ordinal end Time.Day.Ordinal.OfYear diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 1741e1ad2e17..ccc1609d04a7 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -5,21 +5,19 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Time -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat namespace Std namespace Time -open Lean +open Lean Internal set_option linter.all true namespace Day /-- -`Ordinal` represents a bounded value for days, which ranges between 0 and 31. +`Ordinal` represents a bounded value for days, which ranges between 1 and 31. -/ def Ordinal := Bounded.LE 1 31 deriving Repr, BEq, LE, LT @@ -28,17 +26,21 @@ instance [Le 1 n] [Le n 31] : OfNat Ordinal n where ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) instance { x y : Ordinal } : Decidable (x ≤ y) := - dite (x.val ≤ y.val) isTrue isFalse + inferInstanceAs (Decidable (x.val ≤ y.val)) instance : Inhabited Ordinal where default := 1 /-- -`Ordinal.OfYear` represents a bounded value for days, which ranges between 0 and 31. +`Ordinal.OfYear` represents a bounded value for days, which ranges between 0 and 366 if the year +is a leap year or 365. -/ def Ordinal.OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) -instance [Le 1 n] [Le n (if leap then 366 else 365)] : OfNat (Ordinal.OfYear leap) n where - ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) +instance [Le 1 n] [Le n 365] : OfNat (Ordinal.OfYear leap) n where + ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr (by have := Le.p (n := n) (m := 365); split <;> omega))) + +instance : OfNat (Ordinal.OfYear true) 366 where + ofNat := Bounded.LE.mk (Int.ofNat 366) (by decide) instance : Inhabited (Ordinal.OfYear leap) where default := by diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index d72df64db255..6b8fd613ef2c 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -4,16 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Time.Basic import Std.Time.Date.Unit.Day namespace Std namespace Time -open Lean Time +open Internal namespace Month /-- @@ -111,12 +109,12 @@ Size in days of each month if the year is not leap. -/ @[inline] def monthSizesNonLeap : { val : Array Day.Ordinal // val.size = 12 } := - ⟨#[ 31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31 ], by simp⟩ + ⟨#[31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31], by simp⟩ /-- -Gets the number of days in a month. +Gets the number of days in a month without a proof of validity of the ordinal in a month and year. -/ -def days' (leap : Bool) (month : Ordinal) : Day.Ordinal := +def daysWithoutProof (leap : Bool) (month : Ordinal) : Day.Ordinal := if month.val = 2 then if leap then 29 else 28 else by @@ -126,27 +124,27 @@ def days' (leap : Bool) (month : Ordinal) : Day.Ordinal := exact months.get r /-- -Check if the day is valid in a Month and a leap Year. +Check if the day is valid in a month and a leap year. -/ @[inline] -def valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := - day ≤ days' leap month +def Valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := + day ≤ daysWithoutProof leap month -instance : Decidable (valid leap month day) := - dite (day ≤ days' leap month) isTrue isFalse +instance : Decidable (Valid leap month day) := + dite (day ≤ daysWithoutProof leap month) isTrue isFalse /-- -Gets the number of days in a month. +Gets the number of days in a month along side a proof of it's validity. -/ @[inline] -def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // valid leap month day } := - ⟨days' leap month, Int.le_refl ((days' leap month).val)⟩ +def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // Valid leap month day } := + ⟨daysWithoutProof leap month, Int.le_refl ((daysWithoutProof leap month).val)⟩ /-- -Forces the day to be on the valid range. +Clips the day to be on the valid range. -/ @[inline] -def forceDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : Day.Ordinal //valid leap month day } := +def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : Day.Ordinal // Valid leap month day } := let max : Day.Ordinal := month.days leap if h : day.val > max.val then ⟨max, Int.le_refl max.val⟩ @@ -156,33 +154,30 @@ def forceDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. -/ @[inline] -def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day.Ordinal // valid leap (Prod.fst val) (Prod.snd val) } := Id.run do +def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } := Id.run do let rec go (idx : Fin 12) (cumulative : Fin 366) := let month := Month.Ordinal.ofFin idx.succ let ⟨days, valid⟩ := days leap month if h : cumulative.val < ordinal.val ∧ ordinal.val ≤ cumulative.val + days.val then - let bounded := - Bounded.LE.mk ordinal.val h |>.sub ↑↑cumulative - let bounded : Bounded.LE 1 days.val := by - simp [← Int.add_comm, Int.sub_self] at bounded - rw [← Int.add_comm 1 (↑↑cumulative), Int.add_sub_assoc, Int.sub_self] at bounded - exact bounded - let p₁ := bounded.property.right - let p := And.intro bounded.property.left (Int.le_trans bounded.property.right days.property.right) - let days₁ : Day.Ordinal := ⟨bounded.val, p⟩ - let h1 : Month.Ordinal.valid leap month days₁ := Int.le_trans p₁ valid - ⟨⟨month, days₁⟩, h1⟩ + let bounded := Bounded.LE.mk ordinal.val h |>.sub cumulative + + let bounded : Bounded.LE 1 days.val := bounded.cast + (by simp [Int.add_comm _ 1, Int.add_assoc, ←Int.sub_eq_add_neg];) + (by simp [Int.add_comm _ days.val, Int.add_assoc, ←Int.sub_eq_add_neg];) + + let ⟨left, right⟩ := bounded.property + let days₁ : Day.Ordinal := ⟨bounded.val, And.intro left (Int.le_trans right days.property.right)⟩ + ⟨⟨month, days₁⟩, Int.le_trans right valid⟩ else if h : idx.val ≥ 11 then -- Need to remove this in the future. - let ⟨day, valid⟩ := forceDay leap 1 1 + let ⟨day, valid⟩ := clipDay leap 1 1 ⟨⟨1, day⟩, valid⟩ else go ⟨idx.val + 1, Nat.succ_le_succ (Nat.not_le.mp h)⟩ cumulative termination_by 12 - idx.val - - go 1 0 + go 0 0 end Ordinal end Month diff --git a/src/Std/Time/Date/Unit/WeekOfYear.lean b/src/Std/Time/Date/Unit/WeekOfYear.lean index e38597b6121f..65de02c2966a 100644 --- a/src/Std/Time/Date/Unit/WeekOfYear.lean +++ b/src/Std/Time/Date/Unit/WeekOfYear.lean @@ -4,15 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Date.Unit.Day import Std.Time.Date.Unit.Month namespace Std namespace Time +open Internal set_option linter.all true diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index cb001a21d845..74390715112b 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -4,15 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Date.Unit.Day import Std.Time.Date.Unit.Month namespace Std namespace Time +open Internal set_option linter.all true diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index ca2968677606..ee5332a2fe0a 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -4,15 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Date.Unit.Day import Std.Time.Date.Unit.Month namespace Std namespace Time +open Internal set_option linter.all true @@ -53,11 +52,11 @@ def isLeap (y : Offset) : Bool := Forces the day to be on the valid range. -/ @[inline] -def valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := - month.valid year.isLeap day +def Valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := + month.Valid year.isLeap day -instance : Decidable (valid year month day) := - dite (month.valid year.isLeap day) isTrue isFalse +instance : Decidable (Valid year month day) := + dite (month.Valid year.isLeap day) isTrue isFalse end Offset end Year diff --git a/src/Std/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean index 2ef7fc11d765..05aa35e0b0bc 100644 --- a/src/Std/Time/Date/WeekDate.lean +++ b/src/Std/Time/Date/WeekDate.lean @@ -7,10 +7,11 @@ prelude import Std.Time.Date.Unit.Year import Std.Time.Date.Unit.WeekOfYear import Std.Time.Date.Scalar -import Std.Time.Date.Date +import Std.Time.Date.LocalDate namespace Std namespace Time +open Internal /-- `WeekDate` represents a date using a combination of a week of the year and the year. @@ -25,14 +26,14 @@ namespace WeekDate /-- Converts a `WeekDate` to a `Scalar`. -/ -def toScalar (wd : WeekDate) : Date.Scalar := +def toScalar (wd : WeekDate) : LocalDate.Scalar := let days := wd.year.toInt * 365 + wd.week.val * 7 - Date.Scalar.ofDays days + LocalDate.Scalar.ofDays days /-- Creates a `WeekDate` from a `Scalar`. -/ -def fromScalar (scalar : Date.Scalar) : WeekDate := +def fromScalar (scalar : LocalDate.Scalar) : WeekDate := let totalDays := scalar.toDays let year := totalDays / 365 let week := diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index cbcc86dd9fa2..b1cea7f83849 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.DateTime.NaiveDateTime +import Std.Time.DateTime.LocalDateTime import Std.Time.DateTime.Timestamp namespace Std @@ -14,15 +14,15 @@ namespace Timestamp set_option linter.all true /-- -Converts a `NaiveDateTime` to a `Timestamp` +Converts a `LocalDateTime` to a `Timestamp` -/ @[inline] -def ofNaiveDateTime (naive : NaiveDateTime) : Timestamp := - naive.toTimestamp +def ofLocalDateTime (Local : LocalDateTime) : Timestamp := + Local.toTimestamp /-- -Converts a `Timestamp` to a `NaiveDateTime` +Converts a `Timestamp` to a `LocalDateTime` -/ @[inline] -def toNaiveDateTime (timestamp : Timestamp) : NaiveDateTime := - NaiveDateTime.ofTimestamp timestamp +def toLocalDateTime (timestamp : Timestamp) : LocalDateTime := + LocalDateTime.ofTimestamp timestamp diff --git a/src/Std/Time/DateTime/NaiveDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean similarity index 70% rename from src/Std/Time/DateTime/NaiveDateTime.lean rename to src/Std/Time/DateTime/LocalDateTime.lean index 18b7f46cdccd..43de1f68fc41 100644 --- a/src/Std/Time/DateTime/NaiveDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -4,37 +4,37 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Init.Data.Int -import Std.Time.LessEq import Std.Time.Date import Std.Time.Time +import Std.Time.Internal import Std.Time.DateTime.Timestamp namespace Std namespace Time +open Internal /-- Date time format with Year, Month, Day, Hour, Minute, Seconds and Nanoseconds. -/ -structure NaiveDateTime where - date : Date - time : Time +structure LocalDateTime where + date : LocalDate + time : LocalTime deriving Repr, Inhabited -namespace NaiveDateTime +namespace LocalDateTime /-- -Converts a `NaiveDateTime` into a `Timestamp` +Converts a `LocalDateTime` into a `Std.Time.Timestamp` -/ -def toTimestamp (dt : NaiveDateTime) : Timestamp := - let days := dt.date.toScalar.day +def toTimestamp (dt : LocalDateTime) : Timestamp := + let days := dt.date.toDaysSinceUNIXEpoch let second := dt.time.toSeconds days.toSeconds + second /-- -Converts a UNIX `Timestamp` into a `NaiveDateTime`. +Converts a UNIX `Timestamp` into a `LocalDateTime`. -/ -def ofTimestamp (stamp : Timestamp) : NaiveDateTime := Id.run do +def ofTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do let leapYearEpoch := 11017 let daysPer400Y := 365 * 400 + 97 let daysPer100Y := 365 * 100 + 24 @@ -108,47 +108,47 @@ def ofTimestamp (stamp : Timestamp) : NaiveDateTime := Id.run do let hour : Bounded.LE 0 23 := remSecs.div 3600 (by decide) return { - date := Date.force year hmon (Day.Ordinal.ofFin (Fin.succ mday)) - time := Time.mk (hour.expandTop (by decide)) minute (second.expandTop (by decide)) 0 + date := LocalDate.clip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) + time := LocalTime.mk (hour.expandTop (by decide)) minute (second.expandTop (by decide)) 0 } /-- -Getter for the `Year` inside of a `NaiveDateTime` +Getter for the `Year` inside of a `LocalDateTime` -/ @[inline] -def year (dt : NaiveDateTime) : Year.Offset := +def year (dt : LocalDateTime) : Year.Offset := dt.date.year /-- -Getter for the `Month` inside of a `NaiveDateTime` +Getter for the `Month` inside of a `LocalDateTime` -/ @[inline] -def month (dt : NaiveDateTime) : Month.Ordinal := +def month (dt : LocalDateTime) : Month.Ordinal := dt.date.month /-- -Getter for the `Day` inside of a `NaiveDateTime` +Getter for the `Day` inside of a `LocalDateTime` -/ @[inline] -def day (dt : NaiveDateTime) : Day.Ordinal := +def day (dt : LocalDateTime) : Day.Ordinal := dt.date.day /-- -Getter for the `Hour` inside of a `NaiveDateTime` +Getter for the `Hour` inside of a `LocalDateTime` -/ @[inline] -def hour (dt : NaiveDateTime) : Hour.Ordinal := +def hour (dt : LocalDateTime) : Hour.Ordinal := dt.time.hour /-- -Getter for the `Minute` inside of a `NaiveDateTime` +Getter for the `Minute` inside of a `LocalDateTime` -/ @[inline] -def minute (dt : NaiveDateTime) : Minute.Ordinal := +def minute (dt : LocalDateTime) : Minute.Ordinal := dt.time.minute /-- -Getter for the `Second` inside of a `NaiveDateTime` +Getter for the `Second` inside of a `LocalDateTime` -/ @[inline] -def second (dt : NaiveDateTime) : Second.Ordinal := +def second (dt : LocalDateTime) : Second.Ordinal := dt.time.second -end NaiveDateTime +end LocalDateTime end Time end Std diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index df76ddb413d0..1a78ae1244a9 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude +import Std.Time.Internal import Init.Data.Int -import Std.Time.LessEq import Std.Time.Time namespace Std namespace Time +open Internal /-- Seconds since the UNIX Epoch. diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 2d3b3ee0f707..e307524572d4 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -4,11 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Bounded +import Std.Time.Internal import Std.Time.Time namespace Std namespace Time +open Internal /-- `Instant` represents a place in time with second and nanoseconds precision. diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 49ed77f68c71..a96f0a97d76b 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -5,9 +5,10 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Format.Basic +import Std.Time.Notation -open Std.Time - +namespace Std +namespace Time namespace Formats set_option linter.all true @@ -16,102 +17,102 @@ set_option linter.all true The ISO8601 format, which is always 24 or 27 characters long, used for representing date and time in a standardized format. The format follows the pattern `YYYY-MM-DD'T'hh:mm:ss'Z'`. -/ -def ISO8601 : Format .any := Format.spec! "YYYY-MM-DD'T'hh:mm:ss.sssZ" +def ISO8601 : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" /-- The AmericanDate format, which follows the pattern `MM/DD/YYYY`. -/ -def AmericanDate : Format .any := ⟨[Modifier.MM, "/", Modifier.DD, "/", Modifier.YYYY]⟩ +def AmericanDate : Format .any := date-spec% "MM-DD-YYYY" /-- The EuropeanDate format, which follows the pattern `DD/MM/YYYY`. -/ -def EuropeanDate : Format .any := ⟨[Modifier.DD, "/", Modifier.MM, "/", Modifier.YYYY]⟩ +def EuropeanDate : Format .any := date-spec% "DD-MM-YYYY" /-- -The Time12Hour format, which follows the pattern `hh:mm:ss aa` for representing time +The Time12Hour format, which follows the pattern `HH:mm:ss aa` for representing time in a 12-hour clock format with an AM/PM marker. -/ -def Time12Hour : Format .any := ⟨[Modifier.HH, ":", Modifier.mm, ":", Modifier.ss, " ", Modifier.aa]⟩ +def Time12Hour : Format .any := date-spec% "HH:mm:ss aa" /-- The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time in a 24-hour clock format. -/ -def Time24Hour : Format .any := ⟨[Modifier.hh, ":", Modifier.mm, ":", Modifier.ss]⟩ +def Time24Hour : Format .any := date-spec% "hh:mm:ss" /-- The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used in SQL databases to represent dates. -/ -def SQLDate : Format .any := ⟨[Modifier.YYYY, "-", Modifier.MM, "-", Modifier.DD]⟩ +def SQLDate : Format .any := date-spec% "YYYY-MM-DD" /-- The LongDateFormat, which follows the pattern `EEEE, MMMM D, YYYY hh:mm:ss` for representing a full date and time with the day of the week and month name. -/ -def LongDateFormat : Format (.only .GMT) := Format.spec! "EEEE, MMMM D, YYYY hh:mm:ss" +def LongDateFormat : Format (.only .GMT) := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss" /-- The AscTime format, which follows the pattern `EEE MMM d hh:mm:ss YYYY`. This format is often used in older systems for logging and time-stamping events. -/ -def AscTime : Format (.only .GMT) := Format.spec! "EEE MMM d hh:mm:ss YYYY" +def AscTime : Format (.only .GMT) := date-spec% "EEE MMM d hh:mm:ss YYYY" /-- The RFC822 format, which follows the pattern `EEE, DD MMM YYYY hh:mm:ss ZZZ`. This format is used in email headers and HTTP headers. -/ -def RFC822 : Format .any := Format.spec! "EEE, DD MMM YYYY hh:mm:ss ZZZ" +def RFC822 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" /-- The RFC850 format, which follows the pattern `EEEE, DD-MMM-YY hh:mm:ss ZZZ`. This format is an older standard for representing date and time in headers. -/ -def RFC850 : Format .any := Format.spec! "EEEE, DD-MMM-YY hh:mm:ss ZZZ" +def RFC850 : Format .any := date-spec% "EEEE, DD-MMM-YY hh:mm:ss ZZZ" end Formats -namespace Date +namespace LocalDate /-- -Parses a date string in the American format (`MM/DD/YYYY`) and returns a `Date`. +Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. -/ -def fromAmericanDateString (input : String) : Except String Date := do - let res ← Formats.AmericanDate.parseBuilder (λm d y => Date.ofYearMonthDay y m d) input +def fromAmericanDateString (input : String) : Except String LocalDate := do + let res ← Formats.AmericanDate.parseBuilder (λm d y => LocalDate.ofYearMonthDay y m d) input match res with | .some res => pure res | none => .error "Invalid date." -end Date +end LocalDate -namespace Time +namespace LocalTime /-- -Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `Time`. +Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `LocalTime`. -/ -def fromTime24Hour (input : String) : Except String Time := - Formats.Time24Hour.parseBuilder (Time.ofHourMinuteSeconds) input +def fromTime24Hour (input : String) : Except String LocalTime := + Formats.Time24Hour.parseBuilder (LocalTime.ofHourMinuteSeconds) input /-- -Formats a `Time` value into a 24-hour format string (`hh:mm:ss`). +Formats a `LocalTime` value into a 24-hour format string (`hh:mm:ss`). -/ -def toTime24Hour (input : Time) : String := +def toTime24Hour (input : LocalTime) : String := Formats.Time24Hour.formatBuilder input.hour input.minute input.second /-- -Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `Time`. +Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `LocalTime`. -/ -def fromTime12Hour (input : String) : Except String Time := - Formats.Time12Hour.parseBuilder (λh m s a => Time.ofHourMinuteSeconds (HourMarker.toAbsolute! a h) m s) input +def fromTime12Hour (input : String) : Except String LocalTime := + Formats.Time12Hour.parseBuilder (λh m s a => LocalTime.ofHourMinuteSeconds (HourMarker.toAbsolute' a h) m s) input /-- -Formats a `Time` value into a 12-hour format string (`hh:mm:ss aa`). +Formats a `LocalTime` value into a 12-hour format string (`hh:mm:ss aa`). -/ -def toTime12Hour (input : Time) : String := +def toTime12Hour (input : LocalTime) : String := Formats.Time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.val ≥ 12 then HourMarker.pm else HourMarker.am) -end Time +end LocalTime namespace ZonedDateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index c510efd0e77d..1851e4cdc035 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -12,8 +12,10 @@ import Lean.Data.Parsec namespace Std namespace Time +open Internal -open Lean.Parsec Lean Time Date TimeZone DateTime +open Lean.Parsec.String +open Lean.Parsec Lean LocalTime LocalDate TimeZone DateTime /-- The `Modifier` inductive type represents various formatting options for date and time components. @@ -130,7 +132,7 @@ structure Format (awareness : Awareness) where private def isNonLetter : Char → Bool := not ∘ Char.isAlpha -private def parseModifier : Lean.Parsec Modifier +private def parseModifier : Parser Modifier := pstring "YYYY" *> pure .YYYY <|> pstring "YY" *> pure .YY <|> pstring "MMMM" *> pure .MMMM @@ -161,19 +163,19 @@ private def parseModifier : Lean.Parsec Modifier def isFormatStart : Char → Bool := Char.isAlpha -private def pnumber : Lean.Parsec Nat := do +private def pnumber : Parser Nat := do let numbers ← manyChars digit return numbers.foldl (λacc char => acc * 10 + (char.toNat - 48)) 0 -private def parseFormatPart : Lean.Parsec FormatPart +private def parseFormatPart : Parser FormatPart := (.modifier <$> parseModifier) - <|> (pchar '\\') *> anyChar <&> (.string ∘ toString) + <|> (pchar '\\') *> any <&> (.string ∘ toString) <|> (pchar '\"' *> many1Chars (satisfy (· ≠ '\"')) <* pchar '\"') <&> .string <|> (pchar '\'' *> many1Chars (satisfy (· ≠ '\'')) <* pchar '\'') <&> .string <|> many1Chars (satisfy (fun x => ¬isFormatStart x ∧ x ≠ '\'' ∧ x ≠ '\"')) <&> .string -private def specParser : Lean.Parsec FormatString := - (Array.toList <$> Lean.Parsec.many parseFormatPart) <* eof +private def specParser : Parser FormatString := + (Array.toList <$> many parseFormatPart) <* eof private def specParse (s: String) : Except String FormatString := specParser.run s @@ -334,21 +336,21 @@ def FormatType (result: Type) : FormatString → Type | .string _ :: xs => (FormatType result xs) | [] => result -private def position : Parsec Nat := λs => (ParseResult.success s (s.pos.byteIdx)) +private def position : Parser Nat := λs => (ParseResult.success s (s.pos.byteIdx)) -private def size (data : Parsec α) : Parsec (α × Nat) := do +private def size (data : Parser α) : Parser (α × Nat) := do let st ← position let res ← data let en ← position pure (res, en-st) -private def transform (n: β → Option α) (p: Lean.Parsec β) : Lean.Parsec α := do +private def transform (n: β → Option α) (p: Parser β) : Parser α := do let res ← p match n res with | some n => pure n | none => fail "cannot parse" -private def parseMonth : Lean.Parsec Month.Ordinal +private def parseMonth : Parser Month.Ordinal := (pstring "Jan" *> pure 1) <|> (pstring "Feb" *> pure 2) <|> (pstring "Mar" *> pure 3) @@ -362,7 +364,7 @@ private def parseMonth : Lean.Parsec Month.Ordinal <|> (pstring "Nov" *> pure 11) <|> (pstring "Dec" *> pure 12) -private def parseMonthUnabbrev : Lean.Parsec Month.Ordinal +private def parseMonthUnabbrev : Parser Month.Ordinal := (pstring "January" *> pure 1) <|> (pstring "February" *> pure 2) <|> (pstring "March" *> pure 3) @@ -376,7 +378,7 @@ private def parseMonthUnabbrev : Lean.Parsec Month.Ordinal <|> (pstring "November" *> pure 11) <|> (pstring "December" *> pure 12) -private def parseWeekday : Lean.Parsec Weekday +private def parseWeekday : Parser Weekday := (pstring "Mon" *> pure Weekday.mon) <|> (pstring "Tue" *> pure Weekday.tue) <|> (pstring "Wed" *> pure Weekday.wed) @@ -385,7 +387,7 @@ private def parseWeekday : Lean.Parsec Weekday <|> (pstring "Sat" *> pure Weekday.sat) <|> (pstring "Sun" *> pure Weekday.sun) -private def parseWeekdayUnnabrev : Lean.Parsec Weekday +private def parseWeekdayUnnabrev : Parser Weekday := (pstring "Monday" *> pure Weekday.mon) <|> (pstring "Tuesday" *> pure Weekday.tue) <|> (pstring "Wednesday" *> pure Weekday.wed) @@ -394,55 +396,55 @@ private def parseWeekdayUnnabrev : Lean.Parsec Weekday <|> (pstring "Saturday" *> pure Weekday.sat) <|> (pstring "Sunday" *> pure Weekday.sun) -private def parserUpperHourMarker : Lean.Parsec HourMarker +private def parserUpperHourMarker : Parser HourMarker := (pstring "AM" *> pure HourMarker.am) <|> (pstring "PM" *> pure HourMarker.pm) -private def parserLowerHourMarker : Lean.Parsec HourMarker +private def parserLowerHourMarker : Parser HourMarker := (pstring "am" *> pure HourMarker.am) <|> (pstring "pm" *> pure HourMarker.pm) -private def threeDigit : Lean.Parsec Int := do - let digit1 ← Lean.Parsec.digit - let digit2 ← Lean.Parsec.digit - let digit3 ← Lean.Parsec.digit +private def threeDigit : Parser Int := do + let digit1 ← digit + let digit2 ← digit + let digit3 ← digit return String.toNat! s!"{digit1}{digit2}{digit3}" -private def twoDigit : Lean.Parsec Int := do - let digit1 ← Lean.Parsec.digit - let digit2 ← Lean.Parsec.digit +private def twoDigit : Parser Int := do + let digit1 ← digit + let digit2 ← digit return String.toNat! s!"{digit1}{digit2}" -private def parseYearTwo : Lean.Parsec Int :=do +private def parseYearTwo : Parser Int :=do let year ← twoDigit return if year < 70 then 2000 + year else 1900 + year -private def timeOffset (colon: Bool) : Lean.Parsec Offset := do +private def timeOffset (colon: Bool) : Parser Offset := do let sign : Int ← (pstring "-" *> pure (-1)) <|> (pstring "+" *> pure 1) let hour ← twoDigit - if colon then let _ ← pstring ":" + if colon then discard <| pstring ":" let minutes ← twoDigit let res := (hour * 3600 + minutes * 60) * sign pure (Offset.ofSeconds (UnitVal.ofInt res)) -private def timeOrUTC (utcString: String) (colon: Bool) : Lean.Parsec Offset := +private def timeOrUTC (utcString: String) (colon: Bool) : Parser Offset := (pstring utcString *> pure Offset.zero) <|> timeOffset colon -private def number : Lean.Parsec Nat := do - String.toNat! <$> Lean.Parsec.many1Chars Lean.Parsec.digit +private def number : Parser Nat := do + String.toNat! <$> many1Chars digit -private def singleDigit : Lean.Parsec Nat := do - let digit1 ← Lean.Parsec.digit +private def singleDigit : Parser Nat := do + let digit1 ← digit return String.toNat! s!"{digit1}" -private def fourDigit : Lean.Parsec Int := do - let digit1 ← Lean.Parsec.digit - let digit2 ← Lean.Parsec.digit - let digit3 ← Lean.Parsec.digit - let digit4 ← Lean.Parsec.digit +private def fourDigit : Parser Int := do + let digit1 ← digit + let digit2 ← digit + let digit3 ← digit + let digit4 ← digit return String.toNat! s!"{digit1}{digit2}{digit3}{digit4}" -private def parserWithFormat : (typ: Modifier) → Lean.Parsec (SingleFormatType typ) +private def parserWithFormat : (typ: Modifier) → Parser (SingleFormatType typ) | .YYYY => fourDigit | .YY => parseYearTwo | .MMMM => parseMonthUnabbrev @@ -451,7 +453,7 @@ private def parserWithFormat : (typ: Modifier) → Lean.Parsec (SingleFormatType | .M => transform Bounded.LE.ofInt number | .DD => transform Bounded.LE.ofInt twoDigit | .D => transform Bounded.LE.ofInt number - | .d => transform Bounded.LE.ofInt (Lean.Parsec.orElse twoDigit (λ_ => pchar ' ' *> (singleDigit))) + | .d => transform Bounded.LE.ofInt (orElse twoDigit (λ_ => pchar ' ' *> (singleDigit))) | .EEEE => parseWeekdayUnnabrev | .EEE => parseWeekday | .hh => transform Bounded.LE.ofInt twoDigit @@ -487,9 +489,9 @@ private structure DateBuilder where millisecond : Millisecond.Ordinal := 0 private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : aw.type := - let build := DateTime.ofNaiveDateTime { - date := Date.force builder.year builder.month builder.day - time := Time.mk builder.hour builder.minute builder.second (.ofMillisecond builder.millisecond) + let build := DateTime.ofLocalDateTime { + date := LocalDate.clip builder.year builder.month builder.day + time := LocalTime.mk builder.hour builder.minute builder.second (.ofMillisecond builder.millisecond) } match aw with @@ -506,7 +508,7 @@ private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : Sin | .DD | .D | .d => { data with day := value } | .EEEE | .EEE => data | .hh | .h | .HH | .H => { data with hour := value } - | .AA | .aa => { data with hour := HourMarker.toAbsolute! value data.hour } + | .AA | .aa => { data with hour := HourMarker.toAbsolute' value data.hour } | .mm | .m => { data with minute := value } | .sss => { data with millisecond := value } | .ss | .s => { data with second := value } @@ -514,7 +516,7 @@ private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : Sin | .Z => { data with tz := value } | .z => { data with tzName := value } -private def formatParser (date : DateBuilder) : FormatPart → Lean.Parsec DateBuilder +private def formatParser (date : DateBuilder) : FormatPart → Parser DateBuilder | .modifier mod => addDataInDateTime date mod <$> parserWithFormat mod | .string s => skipString s *> pure date @@ -595,8 +597,8 @@ def formatBuilder (format : Format aw) : FormatType String format.string := /-- Parser for a ZonedDateTime. -/ -def parser (format : FormatString) (aw : Awareness) : Parsec (aw.type) := - let rec go (date : DateBuilder) (x : FormatString) : Parsec aw.type := +def parser (format : FormatString) (aw : Awareness) : Parser (aw.type) := + let rec go (date : DateBuilder) (x : FormatString) : Parser aw.type := match x with | x :: xs => formatParser date x >>= (go · xs) | [] => pure (date.build aw) @@ -605,8 +607,8 @@ def parser (format : FormatString) (aw : Awareness) : Parsec (aw.type) := /-- Parser for a format with a builder. -/ -def builderParser (format: FormatString) (func: FormatType α format) : Lean.Parsec α := - let rec go (format : FormatString) (func: FormatType α format) : Parsec α := +def builderParser (format: FormatString) (func: FormatType α format) : Parser α := + let rec go (format : FormatString) (func: FormatType α format) : Parser α := match format with | .modifier x :: xs => do let res ← parserWithFormat x diff --git a/src/Std/Time/DateTime/Basic.lean b/src/Std/Time/Internal.lean similarity index 55% rename from src/Std/Time/DateTime/Basic.lean rename to src/Std/Time/Internal.lean index 45d10e698ca4..1ce45c0915d8 100644 --- a/src/Std/Time/DateTime/Basic.lean +++ b/src/Std/Time/Internal.lean @@ -4,4 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.LessEq +import Std.Time.Internal.Bounded +import Std.Time.Internal.LessEq +import Std.Time.Internal.Sign +import Std.Time.Internal.UnitVal diff --git a/src/Std/Time/Bounded.lean b/src/Std/Time/Internal/Bounded.lean similarity index 95% rename from src/Std/Time/Bounded.lean rename to src/Std/Time/Internal/Bounded.lean index 43e0edd24514..bd09fca71144 100644 --- a/src/Std/Time/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -5,15 +5,16 @@ Authors: Sofia Rodrigues -/ prelude import Init.Data.Int -import Std.Time.LessEq +import Std.Time.Internal.LessEq namespace Std namespace Time +namespace Internal set_option linter.all true in /-- -A `Bounded` is represented by an `Int` that is contrained by a lower and higher bounded using some +A `Bounded` is represented by an `Int` that is constrained by a lower and higher bounded using some relation `rel`. It includes all the integers that `rel lo val ∧ rel val hi`. -/ def Bounded (rel : Int → Int → Prop) (lo : Int) (hi : Int) := { val : Int // rel lo val ∧ rel val hi} @@ -52,6 +53,10 @@ instance [Le lo n] [Le n hi] : OfNat (Bounded.LE lo hi) n where instance [Le lo hi] : Inhabited (Bounded.LE lo hi) where default := ⟨lo, And.intro (Int.le_refl lo) (Int.ofNat_le.mpr Le.p)⟩ +def cast {rel : Int → Int → Prop} {lo₁ lo₂ hi₁ hi₂ : Int} (h₁ : lo₁ = lo₂) (h₂ : hi₁ = hi₂) + (b : Bounded rel lo₁ hi₁) : Bounded rel lo₂ hi₂ := + .mk b.val ⟨h₁ ▸ b.property.1, h₂ ▸ b.property.2⟩ + /-- A `Bounded` integer that the relation used is the the less-than relation so, it includes all integers that `lo < val < hi`. @@ -111,7 +116,7 @@ def ofNat' (val : Nat) (h : lo ≤ val ∧ val ≤ hi) : Bounded.LE lo hi := Convert a `Nat` to a `Bounded.LE` using the lower boundary too. -/ @[inline] -def force (val : Int) (h : lo ≤ hi) : Bounded.LE lo hi := +def clip (val : Int) (h : lo ≤ hi) : Bounded.LE lo hi := if h₀ : lo ≤ val then if h₁ : val ≤ hi then ⟨val, And.intro h₀ h₁⟩ @@ -122,7 +127,7 @@ def force (val : Int) (h : lo ≤ hi) : Bounded.LE lo hi := Convert a `Nat` to a `Bounded.LE` using the lower boundary too. -/ @[inline] -def force! [Le lo hi] (val : Int) : Bounded.LE lo hi := +def clip! [Le lo hi] (val : Int) : Bounded.LE lo hi := if h₀ : lo ≤ val then if h₁ : val ≤ hi then ⟨val, And.intro h₀ h₁⟩ @@ -378,3 +383,6 @@ def succ (bounded : Bounded.LE lo hi) (h : bounded.val < hi) : Bounded.LE lo hi end LE end Bounded +end Internal +end Time +end Std diff --git a/src/Std/Time/LessEq.lean b/src/Std/Time/Internal/LessEq.lean similarity index 90% rename from src/Std/Time/LessEq.lean rename to src/Std/Time/Internal/LessEq.lean index b248b0f9b795..c74f7d921731 100644 --- a/src/Std/Time/LessEq.lean +++ b/src/Std/Time/Internal/LessEq.lean @@ -8,6 +8,8 @@ import Init import Lean.Data.Rat namespace Std +namespace Time +namespace Internal set_option linter.all true @@ -26,3 +28,7 @@ instance (m : Nat) : Le 0 m where instance {n m : Nat} [Le n m] : Le (Nat.succ n) (Nat.succ m) where p := Nat.succ_le_succ (Le.p) + +end Internal +end Time +end Std diff --git a/src/Std/Time/Sign.lean b/src/Std/Time/Internal/Sign.lean similarity index 91% rename from src/Std/Time/Sign.lean rename to src/Std/Time/Internal/Sign.lean index 90dc41dc66b2..d3127637409b 100644 --- a/src/Std/Time/Sign.lean +++ b/src/Std/Time/Internal/Sign.lean @@ -4,10 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Bounded +import Std.Time.Internal.Bounded namespace Std namespace Time +namespace Internal set_option linter.all true @@ -40,5 +41,6 @@ def apply (sign : Sign) (val : Int) : Int := val * sign.val end Sign +end Internal end Time end Std diff --git a/src/Std/Time/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean similarity index 97% rename from src/Std/Time/UnitVal.lean rename to src/Std/Time/Internal/UnitVal.lean index 29b8223100da..67aadb96e879 100644 --- a/src/Std/Time/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -6,10 +6,12 @@ Authors: Sofia Rodrigues prelude import Init import Lean.Data.Rat -import Std.Time.Sign +import Std.Time.Internal.Sign namespace Std namespace Time +namespace Internal + open Lean set_option linter.all true @@ -104,5 +106,6 @@ instance : Neg (UnitVal α) where neg x := ⟨-x.val⟩ instance : ToString (UnitVal n) where toString n := toString n.val end UnitVal +end Internal end Time end Std diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean new file mode 100644 index 000000000000..66ae77dfd35e --- /dev/null +++ b/src/Std/Time/Notation.lean @@ -0,0 +1,263 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.Date +import Std.Time.Time +import Std.Time.DateTime +import Std.Time.Zoned +import Std.Time.Format.Basic +import Lean.Elab.Command + +namespace Std +namespace Time +open Lean.Elab Term Lean Parser Command Std + +set_option linter.all true + +/-- +Category of units that are valid inside a date. +-/ +declare_syntax_cat date_component + +/-- +Raw numbers +-/ +syntax num : date_component + +/-- +Strings for dates like `"1970"-"1"-"1"` +-/ +syntax str : date_component + +/-- +Variables inside of the date. +-/ +syntax ident : date_component + +private def parseComponent : TSyntax `date_component -> MacroM (TSyntax `term) + | `(date_component| $num:num) => `($num) + | `(date_component| $str:str) => `($(Syntax.mkNumLit str.getString)) + | `(date_component| $ident:ident) => `($ident:ident) + | syn => Macro.throwErrorAt syn "unsupported syntax" + +/-- +Category of dates like `DD-MM-YYYY`. +-/ +declare_syntax_cat date + + +/-- +Date in `DD-MM-YYYY` format. +-/ +syntax date_component noWs "-" noWs date_component noWs "-" noWs date_component : date + +private def parseDate : TSyntax `date -> MacroM (TSyntax `term) + | `(date|$day:date_component-$month:date_component-$year:date_component) => do + `(Std.Time.LocalDate.mk $(← parseComponent year) $(← parseComponent month) $(← parseComponent day) (by decide)) + | syn => Macro.throwErrorAt syn "unsupported type" + +/-- +Category of times like `HH:mm:ss`. +-/ +declare_syntax_cat time + +/-- +Date in `HH-mm-ss` format. +-/ +syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component : time + +/-- +Date in `HH-mm-ss.sssssss` format. +-/ +syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component "." noWs date_component : time + +private def parseTime : TSyntax `time -> MacroM (TSyntax `term) + | `(time| $hour:date_component:$minute:date_component:$second:date_component) => do + `(Std.Time.LocalTime.mk $(← parseComponent hour) $(← parseComponent minute) $(← parseComponent second) 0) + | `(time| $hour:date_component:$minute:date_component:$second:date_component.$nanos:date_component) => do + `(Std.Time.LocalTime.mk $(← parseComponent hour) $(← parseComponent minute) $(← parseComponent second) $(← parseComponent nanos)) + | syn => Macro.throwErrorAt syn "unsupported syntax" + +/-- +Category of date and time together. +-/ +declare_syntax_cat datetime + +/-- +Date and time in `YYYY-MM-DD:HH:mm:ss` format. +-/ +syntax date noWs ":" noWs time : datetime + +/-- +Date but using timestamp. +-/ +syntax date_component : datetime + +private def parseDateTime : TSyntax `datetime -> MacroM (TSyntax `term) + | `(datetime| $date:date:$time:time) => do + `(Std.Time.LocalDateTime.mk $(← parseDate date) $(← parseTime time)) + | `(datetime|$tm:date_component) => do + `(Std.Time.LocalDateTime.ofTimestamp $(← parseComponent tm)) + | syn => Macro.throwErrorAt syn "unsupported syntax" + +/-- +Category of offsets like `+00:00`. +-/ +declare_syntax_cat offset + +/-- +Timezone in `+HH:mm`, or `-HH:mm` format. +-/ +syntax ("+" <|> "-") date_component ":" date_component : offset + +private def parseOffset : TSyntax `offset -> MacroM (TSyntax `term) + | `(offset| +$hour:date_component:$minutes:date_component) => do + `(Std.Time.TimeZone.Offset.ofHoursAndMinutes $(← parseComponent hour) $(← parseComponent minutes)) + | `(offset| -$hour:date_component:$minutes:date_component) => do + `(Std.Time.TimeZone.Offset.ofHoursAndMinutes (- $(← parseComponent hour)) (-$(← parseComponent minutes))) + | syn => Macro.throwErrorAt syn "unsupported syntax" + +/-- +Category of timezones like `Z`, `UTC`, `GMT` and `+HH:mm`. +-/ +declare_syntax_cat zone + +/-- +Timezone in `+HHmm`, or `-HHmm` format. +-/ +syntax (offset <|> "UTC" <|> "GMT" <|> "Z" <|> ("(" term ")")) : zone + +private def parseZone : TSyntax `zone -> MacroM (TSyntax `term) + | `(zone| ($term)) => do `($term) + | `(zone| UTC) => do `(Std.Time.TimeZone.UTC) + | `(zone| GMT) => do `(Std.Time.TimeZone.GMT) + | `(zone| Z) => do `(Std.Time.TimeZone.UTC) + | `(zone| $offset:offset) => do `(Std.Time.TimeZone.mk $(← parseOffset offset) "Offset") + | syn => Macro.throwErrorAt syn "unsupported syntax" + +/-- +Datetimes with zone. +-/ +declare_syntax_cat zoned + +/-- +Datetime with zone in `YYYY-MM-DDTHH:mm:ssZ` format. +-/ +syntax datetime zone : zoned + +private def parseZoned : TSyntax `zoned -> MacroM (TSyntax `term) + | `(zoned| $timestamp:num $zone) => do + `(Std.Time.DateTime.ofTimestamp $timestamp $(← parseZone zone)) + | `(zoned| $datetime:datetime $zone) => do + `(Std.Time.DateTime.ofLocalDateTime $(← parseDateTime datetime) $(← parseZone zone)) + | syn => Macro.throwErrorAt syn "unsupported syntax" + +/-- +Syntax for zones. +-/ +syntax "zone%" zone : term + +/-- +Zoned date times. +-/ +syntax "date%" zoned : term + +/-- +Normal datetimes. +-/ +syntax "date%" datetime : term + +/-- +Macro for defining dates. +-/ +syntax (name := dateDate) "date%" date : term + +/-- +Macro for defining times. +-/ +syntax "date%" time : term + +/-- +Macro for creating offsets +-/ +syntax "offset%" offset : term + +/-- +Macro for creating timezones +-/ +syntax "timezone%" ident noWs "/" noWs ident offset : term + +/-- +Macro for creating timezones with string +-/ +syntax "timezone%" str offset : term + +macro_rules + | `(date% $date:date) => parseDate date + | `(date% $time:time) => parseTime time + | `(date% $datetime:datetime) => parseDateTime datetime + | `(date% $zoned:zoned) => parseZoned zoned + | `(offset% $offset:offset) => parseOffset offset + | `(timezone% $str $offset:offset) => do + do `(Std.Time.TimeZone.mk $(← parseOffset offset) $str) + | `(timezone% $name/$sub $offset:offset) => do + let name := s!"{name.getId.toString}/{sub.getId.toString}" + let syn := Syntax.mkStrLit name + do `(Std.Time.TimeZone.mk $(← parseOffset offset) $syn) + +/-- +Macro for creating formats +-/ + +private def convertModifier : Modifier → MacroM (TSyntax `term) + | .YYYY => `(Std.Time.Modifier.YYYY) + | .YY => `(Std.Time.Modifier.YY) + | .MMMM => `(Std.Time.Modifier.MMMM) + | .MMM => `(Std.Time.Modifier.MMM) + | .MM => `(Std.Time.Modifier.MM) + | .M => `(Std.Time.Modifier.M) + | .DD => `(Std.Time.Modifier.DD) + | .D => `(Std.Time.Modifier.D) + | .d => `(Std.Time.Modifier.d) + | .EEEE => `(Std.Time.Modifier.EEEE) + | .EEE => `(Std.Time.Modifier.EEE) + | .hh => `(Std.Time.Modifier.hh) + | .h => `(Std.Time.Modifier.h) + | .HH => `(Std.Time.Modifier.HH) + | .H => `(Std.Time.Modifier.H) + | .AA => `(Std.Time.Modifier.AA) + | .aa => `(Std.Time.Modifier.aa) + | .mm => `(Std.Time.Modifier.mm) + | .m => `(Std.Time.Modifier.m) + | .sss => `(Std.Time.Modifier.sss) + | .ss => `(Std.Time.Modifier.ss) + | .s => `(Std.Time.Modifier.s) + | .ZZZZZ => `(Std.Time.Modifier.ZZZZZ) + | .ZZZZ => `(Std.Time.Modifier.ZZZZ) + | .ZZZ => `(Std.Time.Modifier.ZZZ) + | .Z => `(Std.Time.Modifier.Z) + | .z => `(Std.Time.Modifier.z) + +private def convertFormatPart : FormatPart → MacroM (TSyntax `term) + | .string s => `(.string $(Syntax.mkStrLit s)) + | .modifier mod => do `(.modifier $(← convertModifier mod)) + +/-- +Syntax for defining a date spec at compile time. +-/ +syntax "date-spec%" str : term + +macro_rules + | `(date-spec% $format_string) => do + let input := format_string.getString + let format : Except String (Format .any) := Format.spec input + match format with + | .ok res => + let alts ← res.string.mapM convertFormatPart + let alts := alts.foldl Syntax.TSepArray.push (Syntax.TSepArray.mk #[] (sep := ",")) + `(⟨[$alts,*]⟩) + | .error err => + Macro.throwErrorAt format_string s!"cannot compile spec: {err}" diff --git a/src/Std/Time/Time.lean b/src/Std/Time/Time.lean index bac27a2e8850..467a4fb35bb9 100644 --- a/src/Std/Time/Time.lean +++ b/src/Std/Time/Time.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time.HourMarker import Std.Time.Time.Basic -import Std.Time.Time.Scalar -import Std.Time.Time.Time +import Std.Time.Time.HourMarker +import Std.Time.Time.LocalTime diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 1d79edeb31a3..1e0f26072316 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -8,6 +8,7 @@ import Std.Time.Time.Basic namespace Std namespace Time +open Internal set_option linter.all true @@ -32,7 +33,7 @@ def HourMarker.toAbsolute (marker : HourMarker) (time : Hour.Ordinal) (h₀ : ti /-- Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker` without a proof. -/ -def HourMarker.toAbsolute! (marker : HourMarker) (time : Hour.Ordinal) : Hour.Ordinal := +def HourMarker.toAbsolute' (marker : HourMarker) (time : Hour.Ordinal) : Hour.Ordinal := match marker with | .am => time | .pm => Bounded.LE.ofFin (Fin.ofNat (time.val.toNat + 12)) diff --git a/src/Std/Time/Time/Time.lean b/src/Std/Time/Time/LocalTime.lean similarity index 62% rename from src/Std/Time/Time/Time.lean rename to src/Std/Time/Time/LocalTime.lean index 45a6ef53637e..84389d851989 100644 --- a/src/Std/Time/Time/Time.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -12,37 +12,37 @@ namespace Time /-- Represents a specific point in time, including hours, minutes, seconds, and nanoseconds. -/ -structure Time where +structure LocalTime where hour : Hour.Ordinal minute : Minute.Ordinal second : Second.Ordinal nano : Nanosecond.Ordinal deriving Repr, Inhabited, BEq -namespace Time +namespace LocalTime /-- -Creates a `Time` value from hours, minutes, and seconds, setting nanoseconds to zero. +Creates a `LocalTime` value from hours, minutes, and seconds, setting nanoseconds to zero. -/ -def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal) : Time := +def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal) : LocalTime := ⟨hour, minute, second, 0⟩ /-- -Converts a `Time` value to the total number of seconds since midnight. +Converts a `LocalTime` value to the total number of seconds since midnight. -/ -def toSeconds (time : Time) : Second.Offset := +def toSeconds (time : LocalTime) : Second.Offset := time.hour.toOffset.toSeconds + time.minute.toOffset.toSeconds + time.second.toOffset /-- -Converts a `Time` value to the total number of minutes since midnight. +Converts a `LocalTime` value to the total number of minutes since midnight. -/ -def toMinutes (time : Time) : Minute.Offset := +def toMinutes (time : LocalTime) : Minute.Offset := time.hour.toOffset.toMinutes + time.minute.toOffset + time.second.toOffset.toMinutes -end Time +end LocalTime end Time end Std diff --git a/src/Std/Time/Time/Scalar.lean b/src/Std/Time/Time/Scalar.lean deleted file mode 100644 index 1b080b3a242e..000000000000 --- a/src/Std/Time/Time/Scalar.lean +++ /dev/null @@ -1,50 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Std.Time.Time.Basic - -namespace Std -namespace Time -namespace Time - -set_option linter.all true - -/-- -Represents a scalar value of time. --/ -structure Scalar where - /-- - The quantity of seconds. - -/ - second : Second.Offset - /-- - The quantity of nanoseconds. - -/ - nano : Nanosecond.Ordinal - -namespace Scalar - -/-- -Converts the scalar value to seconds. --/ -def toSeconds (time : Scalar) : Second.Offset := - time.second - -/-- -Converts the scalar value to minutes. --/ -def toMinutes (time : Scalar) : Minute.Offset := - time.second.toMinutes - -/-- -Converts the scalar value to hours. --/ -def toHours (time : Scalar) : Hour.Offset := - time.second.toHours - -end Scalar -end Time -end Time diff --git a/src/Std/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean index 8c7c2fe0f289..e1213290bb23 100644 --- a/src/Std/Time/Time/Unit/Basic.lean +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal +import Std.Time.Internal import Std.Time.Time.Unit.Hour import Std.Time.Time.Unit.Minute import Std.Time.Time.Unit.Second @@ -13,6 +13,7 @@ import Std.Time.Time.Unit.Nanosecond namespace Std namespace Time namespace Second.Offset +open Internal set_option linter.all true diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 62e5a85a4f54..46834ed32358 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Time.Unit.Minute import Std.Time.Time.Unit.Second @@ -14,11 +12,13 @@ import Std.Time.Time.Unit.Second namespace Std namespace Time namespace Hour +open Internal set_option linter.all true /-- -`Ordinal` represents a bounded value for hour, which ranges between 0 and 23. +`Ordinal` represents a bounded value for hour, which ranges between 0 and 24. It is bounded to 24 +because 24:00:00 is a valid date with leap seconds. -/ def Ordinal := Bounded.LE 0 24 deriving Repr, BEq, LE, LT diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 04b34b44f70d..37ec556c6457 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -4,14 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat namespace Std namespace Time namespace Millisecond +open Internal /-- `Ordinal` represents a bounded value for milliseconds, which ranges between 0 and 999. diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index cf5d944baa7a..d889483c93dd 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -4,15 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Time.Unit.Second namespace Std namespace Time namespace Minute +open Internal set_option linter.all true diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index 9ab9a7b295b5..a970aebb4f91 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -4,22 +4,19 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Time.Unit.Millisecond namespace Std namespace Time +namespace Nanosecond +open Internal set_option linter.all true -namespace Nanosecond - /-- -`Ordinal` represents a bounded value for second, which ranges between 0 and 60. -This accounts for potential leap second. +`Ordinal` represents a bounded value for nanoseconds, which ranges between 0 and 999999999. -/ def Ordinal := Bounded.LE 0 999999999 deriving Repr, BEq, LE, LT @@ -30,6 +27,12 @@ instance : Inhabited Ordinal where default := 0 namespace Ordinal +/-- +`Ordinal` represents a bounded value for nanoseconds in a day, which ranges between 0 and 86400000000000. +-/ +def OfDay := Bounded.LE 0 86400000000000 + deriving Repr, BEq, LE, LT + /-- Convert to `Millisecond.Ordinal` -/ diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index dd412023010b..89a03a354825 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -4,15 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.UnitVal -import Std.Time.Bounded -import Std.Time.LessEq +import Std.Time.Internal import Lean.Data.Rat import Std.Time.Time.Unit.Nanosecond namespace Std namespace Time namespace Second +open Internal set_option linter.all true diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 224a5ff528c8..f327cae98b01 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -4,6 +4,5 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Zoned.Basic import Std.Time.Zoned.DateTime import Std.Time.Zoned.ZonedDateTime diff --git a/src/Std/Time/Zoned/Basic.lean b/src/Std/Time/Zoned/Basic.lean deleted file mode 100644 index d6624ccc2884..000000000000 --- a/src/Std/Time/Zoned/Basic.lean +++ /dev/null @@ -1,6 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -import Std.Time.Zoned.Offset diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 94de5383d779..dc263f7e046b 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time -import Std.Time.Date import Std.Time.DateTime import Std.Time.Zoned.TimeZone @@ -13,12 +11,12 @@ namespace Std namespace Time /-- -It stores a `Timestamp`, a `NaiveDateTime` and a `TimeZone` +It stores a `Timestamp`, a `LocalDateTime` and a `TimeZone` -/ structure DateTime (tz : TimeZone) where private mk :: timestamp : Timestamp - date : NaiveDateTime + date : LocalDateTime deriving Repr, Inhabited namespace DateTime @@ -28,7 +26,7 @@ Creates a new `DateTime` out of a `Timestamp` -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := - let date := (tm + tz.toSeconds).toNaiveDateTime + let date := (tm + tz.toSeconds).toLocalDateTime DateTime.mk tm date /-- @@ -46,11 +44,11 @@ def convertTimeZone (date : DateTime tz) (tz₁ : TimeZone) : DateTime tz₁ := ofTimestamp (date.toTimestamp) tz₁ /-- -Creates a new DateTime out of a `NaiveDateTime` +Creates a new DateTime out of a `LocalDateTime` -/ @[inline] -def ofNaiveDateTime (date : NaiveDateTime) (tz : TimeZone) : DateTime tz := - let tm := date.toTimestamp +def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : DateTime tz := + let tm := date.toTimestamp - tz.toSeconds DateTime.mk tm date /-- diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 57c1b4549a30..bb0ff04a952d 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -4,21 +4,23 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time.Unit.Basic -import Std.Time.Sign +import Std.Time.Time +import Std.Time.Internal namespace Std namespace Time +namespace TimeZone +open Internal /-- Represents a timezone offset with an hour and second component. -/ -structure TimeZone.Offset where +structure Offset where + private mk :: hour: Hour.Offset second: Second.Offset deriving Repr, Inhabited -namespace TimeZone namespace Offset /-- @@ -47,6 +49,13 @@ Creates an `Offset` from a given number of hour. def ofHours (n: Hour.Offset) : Offset := mk n n.toSeconds +/-- +Creates an `Offset` from a given number of hour and minuets. +-/ +def ofHoursAndMinutes (n: Hour.Offset) (m: Minute.Offset) : Offset := + let secs := n.toSeconds + m.toSeconds + mk secs.toHours secs + /-- Creates an `Offset` from a given number of second. -/ diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index c46f04d317be..cf6dcafcf190 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time.Unit.Basic -import Std.Time.Zoned.Basic +import Std.Time.Time +import Std.Time.Zoned.Offset namespace Std namespace Time @@ -16,7 +16,7 @@ An enumeration representing different time zones. structure TimeZone where offset : TimeZone.Offset name : String - deriving Inhabited + deriving Inhabited, Repr namespace TimeZone diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 9c43211d9752..bf392cad0b13 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -4,15 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time -import Std.Time.Date -import Std.Time.DateTime import Std.Time.Zoned.TimeZone import Std.Time.Zoned.DateTime namespace Std namespace Time -open Time Date DateTime def ZonedDateTime := Sigma DateTime @@ -47,11 +43,11 @@ def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := ofTimestamp (date.toTimestamp) tz₁ /-- -Creates a new `ZonedDateTime` out of a `NaiveDateTime` +Creates a new `ZonedDateTime` out of a `LocalDateTime` -/ @[inline] -def ofNaiveDateTime (date : NaiveDateTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofNaiveDateTime date tz⟩ +def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofLocalDateTime date tz⟩ /-- Getter for the `Year` inside of a `ZonedDateTime` From 21563e8501a9e172798f783fd7fd9b812a9c3849 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 7 Aug 2024 18:39:38 -0300 Subject: [PATCH 005/118] refactor: rename some functions --- src/Std/Time/Format.lean | 2 +- src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Internal/Bounded.lean | 2 +- src/Std/Time/Time/HourMarker.lean | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index a96f0a97d76b..2363f449d51d 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -104,7 +104,7 @@ def toTime24Hour (input : LocalTime) : String := Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `LocalTime`. -/ def fromTime12Hour (input : String) : Except String LocalTime := - Formats.Time12Hour.parseBuilder (λh m s a => LocalTime.ofHourMinuteSeconds (HourMarker.toAbsolute' a h) m s) input + Formats.Time12Hour.parseBuilder (λh m s a => LocalTime.ofHourMinuteSeconds (HourMarker.toAbsoluteShift a h) m s) input /-- Formats a `LocalTime` value into a 12-hour format string (`hh:mm:ss aa`). diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 1851e4cdc035..25ca3a5a335b 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -508,7 +508,7 @@ private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : Sin | .DD | .D | .d => { data with day := value } | .EEEE | .EEE => data | .hh | .h | .HH | .H => { data with hour := value } - | .AA | .aa => { data with hour := HourMarker.toAbsolute' value data.hour } + | .AA | .aa => { data with hour := HourMarker.toAbsoluteShift value data.hour } | .mm | .m => { data with minute := value } | .sss => { data with millisecond := value } | .ss | .s => { data with second := value } diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index bd09fca71144..6c3c24917010 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -39,7 +39,7 @@ instance : BEq (Bounded rel n m) where @[always_inline] instance {x y : Bounded rel a b} : Decidable (x ≤ y) := - dite (x.val ≤ y.val) isTrue isFalse + inferInstanceAs (Decidable (x.val ≤ y.val)) /-- A `Bounded` integer that the relation used is the the less-equal relation so, it includes all diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 1e0f26072316..9bee8ac07e3c 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -33,7 +33,7 @@ def HourMarker.toAbsolute (marker : HourMarker) (time : Hour.Ordinal) (h₀ : ti /-- Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker` without a proof. -/ -def HourMarker.toAbsolute' (marker : HourMarker) (time : Hour.Ordinal) : Hour.Ordinal := +def HourMarker.toAbsoluteShift (marker : HourMarker) (time : Hour.Ordinal) : Hour.Ordinal := match marker with | .am => time | .pm => Bounded.LE.ofFin (Fin.ofNat (time.val.toNat + 12)) From e0da8a6550361325b1c38c9685e594ddcdd1d24a Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 9 Aug 2024 02:46:52 -0300 Subject: [PATCH 006/118] feat: zone rules --- src/Std/Time.lean | 15 +- src/Std/Time/Date/LocalDate.lean | 13 -- src/Std/Time/Date/Scalar.lean | 68 ------- src/Std/Time/Date/WeekDate.lean | 13 +- src/Std/Time/DateTime/Timestamp.lean | 7 + src/Std/Time/Duration.lean | 71 ++++++-- src/Std/Time/Internal/UnitVal.lean | 1 - src/Std/Time/Zoned/Database/Basic.lean | 79 +++++++++ src/Std/Time/Zoned/Database/TzIf.lean | 235 +++++++++++++++++++++++++ src/Std/Time/Zoned/ZoneRules.lean | 88 +++++++++ 10 files changed, 479 insertions(+), 111 deletions(-) delete mode 100644 src/Std/Time/Date/Scalar.lean create mode 100644 src/Std/Time/Zoned/Database/Basic.lean create mode 100644 src/Std/Time/Zoned/Database/TzIf.lean create mode 100644 src/Std/Time/Zoned/ZoneRules.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 2913fe3a13bc..3f84a1da7b69 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -35,7 +35,7 @@ Offsets represent unbounded shifts in specific date or time units. They are typi like `date.addDays` where a `Day.Offset` is the parameter. Some offsets, such as `Month.Offset`, do not correspond directly to a specific duration in seconds, as their value depends on the context (if the year is leap or the size of the month). Offsets with a clear correspondent to seconds can be -converted because they use a internal type called `UnitVal`. +converted because they use an internal type called `UnitVal`. - Types with a correspondence to seconds: - `Day.Offset` @@ -51,7 +51,7 @@ converted because they use a internal type called `UnitVal`. ## Ordinal -Ordinal types represent specific bounded values in reference to another unit, e.g, `Day.Ordinal` +Ordinal types represent specific bounded values in reference to another unit, e.g., `Day.Ordinal` represents a day in a month, ranging from 1 to 31. Some ordinal types like `Hour.Ordinal` and `Second.Ordinal`, allow for values beyond the normal range (e.g, 24 hours and 61 seconds) to accomodate special cases with leap seconds like `24:00:00` that is valid in ISO 8601. @@ -81,15 +81,12 @@ nanoseconds corresponds to one second. Dates and times are composed of these components. Dates are "absolute" value in contrast with Offsets that are just shifts in dates and times. Types like `Date` are made using of components such as `Year.Offset`, -`Month.Ordinal`, and `Day.Ordinal`, with a proof of the date's validity. Some types, like `Date.Scalar`, -are wrappers around offsets (e.g., `Day.Offset`) and represent absolute dates relative to the UNIX Epoch. +`Month.Ordinal`, and `Day.Ordinal`, with a proof of the date's validity. ## Date These types provide precision down to the day level, useful for representing and manipulating dates. - **`LocalDate`:** Represents a calendar date in the format `YYYY-MM-DD`. -- **`LocalDate.Scalar`:** : Represents an absolute date with day-level precision in a compact format `DD`. - It encapsulates days since the UNIX Epoch. - **`WeekDate`:** Uses the `YYYY-Www` format with week level precision. ## Time @@ -101,13 +98,14 @@ These types offer precision down to the nanosecond level, useful for representin Combines date and time into a single representation, useful for precise timestamps and scheduling. - **`LocalDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. -- **`Timestamp`**: Represents a point in time with second-level precision. +- **`Timestamp`**: Represents a point in time with second-level precision. It starts on the UNIX +dpoch and it should be used when you receive or need to send timestamps to another systems. ## Zoned date and times. Combines date, time and time zones. - **`DateTime`**: Represents both date and time but with a time zone in the type constructor. -- **`ZonedDateTime`**: An existencial version of the `DateTime`. +- **`ZonedDateTime`**: An existential version of the `DateTime`. ## Duration Represents spans of time and the difference between two points in time. @@ -170,5 +168,4 @@ In order to help the user build dates easily, there are a lot of macros availabl - **`offset% +HH:mm`**: Defines a timezone offset in the format `+HH:mm`. - **`timezone% NAME/ID offset% +HH:mm`**: Defines a timezone with a name and an offset. - **`date-spec% "format"`**: Defines a date specification format at compile time using the provided format string. - -/ diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index 9c805e4cb819..2a2d01861592 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -6,7 +6,6 @@ Authors: Sofia Rodrigues prelude import Std.Time.Internal import Std.Time.Date.Basic -import Std.Time.Date.Scalar namespace Std namespace Time @@ -117,18 +116,6 @@ def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := .ofInt (era * 146097 + doe - 719468) -/-- -Convert a `Scalar` to a `LocalDate` since the UNIX Epoch. --/ -def ofScalar (day : Scalar) : LocalDate := - ofDaysSinceUNIXEpoch day.day - -/-- -Convert a `LocalDate` to a `Scalar` since the UNIX Epoch. --/ -def toScalar (date : LocalDate) : Scalar := - ⟨toDaysSinceUNIXEpoch date⟩ - /-- Calculate the Year.Offset from a LocalDate -/ diff --git a/src/Std/Time/Date/Scalar.lean b/src/Std/Time/Date/Scalar.lean deleted file mode 100644 index 2ff31be8ab8f..000000000000 --- a/src/Std/Time/Date/Scalar.lean +++ /dev/null @@ -1,68 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Std.Time.Internal -import Std.Time.Date.Basic - -namespace Std -namespace Time -namespace LocalDate -open Internal - -/-- -`Scalar` represents a date offset, using the number of day as the underlying unit. --/ -structure Scalar where - day : Day.Offset - deriving Repr, BEq, Inhabited - -instance : Add Scalar where add x y := ⟨x.day + y.day⟩ -instance : Sub Scalar where sub x y := ⟨x.day - y.day⟩ -instance : Mul Scalar where mul x y := ⟨x.day * y.day⟩ -instance : Div Scalar where div x y := ⟨x.day / y.day⟩ -instance : Neg Scalar where neg x := ⟨-x.day⟩ - -namespace Scalar - -/-- -Creates a `Scalar` from a given number of day. --/ -def ofDays (day : Int) : Scalar := - ⟨UnitVal.ofInt day⟩ - -/-- -Retrieves the number of day from a `Scalar`. --/ -def toDays (scalar : Scalar) : Int := - scalar.day.val - -/-- -Adds a specified number of day to the `Scalar`, returning a new `Scalar`. --/ -def addDays (scalar : Scalar) (day : Day.Offset) : Scalar := - ⟨scalar.day + day⟩ - -/-- -Subtracts a specified number of day from the `Scalar`, returning a new `Scalar`. --/ -def subDays (scalar : Scalar) (day : Day.Offset) : Scalar := - ⟨scalar.day - day⟩ - -/-- -Converts a `Scalar` to a `Day.Offset`. --/ -def toOffset (scalar : Scalar) : Day.Offset := - scalar.day - -/-- -Creates a `Scalar` from a `Day.Offset`. --/ -def ofOffset (offset : Day.Offset) : Scalar := - ⟨offset⟩ - -end Scalar -end LocalDate -end Time diff --git a/src/Std/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean index 05aa35e0b0bc..d546c5a6e320 100644 --- a/src/Std/Time/Date/WeekDate.lean +++ b/src/Std/Time/Date/WeekDate.lean @@ -6,7 +6,6 @@ Authors: Sofia Rodrigues prelude import Std.Time.Date.Unit.Year import Std.Time.Date.Unit.WeekOfYear -import Std.Time.Date.Scalar import Std.Time.Date.LocalDate namespace Std @@ -24,17 +23,17 @@ structure WeekDate where namespace WeekDate /-- -Converts a `WeekDate` to a `Scalar`. +Converts a `WeekDate` to a `Day.Offset`. -/ -def toScalar (wd : WeekDate) : LocalDate.Scalar := +def toDays (wd : WeekDate) : Day.Offset := let days := wd.year.toInt * 365 + wd.week.val * 7 - LocalDate.Scalar.ofDays days + UnitVal.mk days /-- -Creates a `WeekDate` from a `Scalar`. +Creates a `WeekDate` from a `Day.Offset`. -/ -def fromScalar (scalar : LocalDate.Scalar) : WeekDate := - let totalDays := scalar.toDays +def fromScalar (scalar : Day.Offset) : WeekDate := + let totalDays := scalar.val let year := totalDays / 365 let week := Bounded.LE.byEmod totalDays 365 (by decide) diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 1a78ae1244a9..e4c5fd7612fa 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -18,6 +18,12 @@ Seconds since the UNIX Epoch. def Timestamp := Second.Offset deriving Repr, BEq, Inhabited +namespace Timestamp + +@[inline] +def toSeconds (tm : Timestamp) : Second.Offset := + tm + instance : OfNat Timestamp n where ofNat := UnitVal.ofNat n @@ -27,5 +33,6 @@ instance : HAdd Timestamp Second.Offset Timestamp where instance : HSub Timestamp Second.Offset Timestamp where hSub x y := UnitVal.sub x y +end Timestamp end Time end Std diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index e307524572d4..4f538f952a40 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal -import Std.Time.Time +import Time.Internal +import Time.Time namespace Std namespace Time @@ -17,7 +17,7 @@ open Internal structure Instant where second : Second.Offset nano : Nanosecond.Ordinal - valid : second.val ≥ 0 ∧ nano.val ≥ 0 + valid : second.val ≥ 0 deriving Repr /-- @@ -26,6 +26,8 @@ Time duration with nanosecond precision. This type allows negative duration. structure Duration where second : Second.Offset nano : Nanosecond.Span + proof : second.val = 0 ∨ (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) + deriving Repr namespace Instant @@ -35,14 +37,39 @@ Gets the difference of two `Instant` in a `Duration`. def sub (t₁ t₂ : Instant) : Duration := let nsec_diff := (t₁.nano).subBounds (t₂.nano) let sec_diff := (t₁.second) - (t₂.second) - if h : sec_diff.val > 0 ∧ nsec_diff.val ≤ -1 then - let truncated := nsec_diff.truncateTop h.right - { second := (sec_diff - 1), nano := truncated.addTop 1000000000 } - else if h₁ : sec_diff.val < 0 ∧ nsec_diff.val ≥ 1 then + + if h₀ : sec_diff.val > 0 ∧ nsec_diff.val ≤ -1 then by + let truncated := nsec_diff.truncateTop h₀.right + let nano := truncated.addTop 1000000000 + let proof₁ : 0 ≤ sec_diff.val - 1 := Int.le_sub_one_of_lt h₀.left + refine { second := UnitVal.mk (sec_diff.val - 1), nano, proof := ?_ } + simp [nano, truncated, Bounded.LE.addTop, Bounded.LE.truncateTop] + refine Or.intro_right _ (Or.intro_left _ (And.intro proof₁ ?_)) + let h₃ := (Int.add_le_add_iff_left 1000000000).mpr nsec_diff.property.left + rw [Int.add_comm] + exact Int.le_trans (by decide) h₃ + else if h₁ : sec_diff.val < 0 ∧ nsec_diff.val ≥ 1 then by + let second := sec_diff.val + 1 let truncated := nsec_diff.truncateBottom h₁.right - { second := (sec_diff + 1), nano := (truncated.subBottom 1000000000) } - else - { second := sec_diff, nano := nsec_diff } + let nano := truncated.subBottom 1000000000 + refine { second := UnitVal.mk second, nano, proof := ?_ } + simp [nano, truncated, Bounded.LE.subBottom, Bounded.LE.truncateBottom] + refine Or.intro_right _ (Or.intro_right _ (And.intro ?_ ?_)) + · exact h₁.left + · let h₃ := Int.sub_le_sub_right nsec_diff.property.right 1000000000 + simp at h₃ + exact Int.le_trans h₃ (by decide) + else by + refine { second := sec_diff, nano := nsec_diff, proof := ?_ } + if h₄ : sec_diff.val > 0 then + let h₅ := Int.not_le.mp (not_and.mp h₀ h₄) + refine Or.intro_right _ (Or.intro_left _ (And.intro (Int.le_of_lt h₄) (Int.add_one_le_iff.mp h₅))) + else if h₅ : sec_diff.val < 0 then + let h₆ := Int.not_le.mp (not_and.mp h₁ h₅) + refine Or.intro_right _ (Or.intro_right _ (And.intro (Int.le_of_lt h₅) (Int.le_sub_one_of_lt h₆))) + else + let h₈ := Int.eq_iff_le_and_ge.mpr (And.intro (Int.not_lt.mp h₄) (Int.not_lt.mp h₅)) + exact Or.intro_left _ h₈ instance : HSub Instant Instant Duration where hSub := Instant.sub @@ -58,22 +85,40 @@ end Instant namespace Duration +/-- +Proof that for nano = 0 then second.val = 0 ∨ (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) +is valid. +-/ +theorem nano_zero + {second : Second.Offset} + {nano : Nanosecond.Span} + (h : nano.val = 0) + : second.val = 0 ∨ (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) + := by + if h₄ : second.val > 0 then + exact Or.intro_right _ (Or.intro_left _ (And.intro (Int.le_of_lt h₄) (Int.eq_iff_le_and_ge.mp h).right)) + else if h₅ : second.val < 0 then + exact Or.intro_right _ (Or.intro_right _ (And.intro (Int.le_of_lt h₅) (Int.eq_iff_le_and_ge.mp h).left)) + else + let h₈ := Int.eq_iff_le_and_ge.mpr (And.intro (Int.not_lt.mp h₄) (Int.not_lt.mp h₅)) + exact Or.intro_left _ h₈ + /-- Returns a `Duration` representing the given number of second. -/ def ofSeconds (second : Second.Offset) : Duration := - { second := second, nano := Bounded.LE.mk 0 (by decide) } + { second := second, nano := Bounded.LE.mk 0 (by decide), proof := nano_zero (by simp [Bounded.LE.mk]) } /-- Returns a `Duration` representing the given number of minute. -/ def ofMinutes (minute : Minute.Offset) : Duration := - { second := minute.toSeconds * 60, nano := Bounded.LE.mk 0 (by decide) } + { second := minute.toSeconds * 60, nano := Bounded.LE.mk 0 (by decide), proof := nano_zero (by simp [Bounded.LE.mk]) } /-- Returns a `Duration` representing the given number of hour. -/ def ofHours (hour : Hour.Offset) : Duration := - { second := hour.toSeconds * 3600, nano := Bounded.LE.mk 0 (by decide) } + { second := hour.toSeconds * 3600, nano := Bounded.LE.mk 0 (by decide), proof := nano_zero (by simp [Bounded.LE.mk]) } end Duration diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 67aadb96e879..4f1ef3e71655 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -20,7 +20,6 @@ set_option linter.all true A structure representing a unit of a given ratio type `α`. -/ structure UnitVal (α : Rat) where - private mk :: /-- Value inside the UnitVal Value -/ diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean new file mode 100644 index 000000000000..ef2331d1a9ea --- /dev/null +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +import Std.Time.Zoned.Database.TzIf +import Std.Time.Zoned.ZoneRules +import Std.Time.Zoned.TimeZone +import Std.Time.DateTime + +namespace Std +namespace Time +namespace TimeZone +namespace Database + +/-- +Converts a Boolean value to a corresponding `StdWall` type. +-/ +def convertWall : Bool → StdWall + | true => .standard + | false => .wall + +/-- +Converts a Boolean value to a corresponding `UTLocal` type. +-/ +def convertUt : Bool → UTLocal + | true => .ut + | false => .local + +/-- +Converts a `TZif.LeapSecond` structure to a `LeapSecond` structure. +-/ +def convertLeapSecond (tz: TZif.LeapSecond) : LeapSecond := + { transitionTime := Internal.UnitVal.mk tz.transitionTime, correction := Internal.UnitVal.mk tz.correction } + +/-- +Converts a `TZif.TZifV1` structure to a `ZoneRules` structure. +-/ +def convertTZifV1 (tz: TZif.TZifV1) : Option ZoneRules := do + let mut times : Array LocalTimeType := #[] + + for i in [0:tz.header.timecnt.toNat] do + let localType := tz.localTimeTypes.get! i + let abbreviation := tz.abbreviations.get! i + let wallflag := convertWall <| tz.stdWallIndicators.get! i + let utLocal := convertUt <| tz.utLocalIndicators.get! i + times := times.push { + gmtOffset := Internal.UnitVal.mk localType.gmtOffset + isDst := localType.isDst + abbreviation + wall := wallflag + utLocal + } + + let mut transitions := #[] + let leapSeconds := tz.leapSeconds.map convertLeapSecond + + for i in [0:tz.transitionTimes.size] do + let time := tz.transitionTimes.get! i + let time := Internal.UnitVal.mk time + let indice ← tz.transitionIndices.get? i + transitions := transitions.push { time, localTimeType := times.get! indice.toNat } + + some { leapSeconds,transitions } + +/-- +Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. +-/ +def convertTZifV2 (tz: TZif.TZifV2) : Option ZoneRules := do + convertTZifV1 tz.toTZifV1 + +/-- +Converts a `TZif.TZif` structure to a `ZoneRules` structure. +-/ +def convertTZif (tz: TZif.TZif) : Option ZoneRules := do + if let some v2 := tz.v2 then + convertTZifV2 v2 + else + convertTZifV1 tz.v1 diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean new file mode 100644 index 000000000000..82d4c53bd336 --- /dev/null +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -0,0 +1,235 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +import Lean.Data.Parsec +import Lean.Data.Parsec.ByteArray + +-- Based on: https://www.rfc-editor.org/rfc/rfc8536.html + +namespace Std +namespace Time +namespace TimeZone +namespace Database +namespace TZif +open Lean.Parsec Lean.Parsec.ByteArray + +abbrev Int32 := Int +abbrev Int64 := Int + +/-- +Represents the header of a TZif file, containing metadata about the file's structure. +-/ +structure Header where + version : UInt8 + isutcnt : UInt32 + isstdcnt : UInt32 + leapcnt : UInt32 + timecnt : UInt32 + typecnt : UInt32 + charcnt : UInt32 + deriving Repr, Inhabited + +/-- +Represents the local time type information, including offset and daylight saving details. +-/ +structure LocalTimeType where + gmtOffset : Int32 + isDst : Bool + abbreviationIndex : UInt8 + deriving Repr, Inhabited + +/-- +Represents a leap second record, including the transition time and the correction applied. +-/ +structure LeapSecond where + transitionTime : Int64 + correction : Int64 + deriving Repr, Inhabited + +/-- +Represents version 1 of the TZif format. +-/ +structure TZifV1 where + header : Header + transitionTimes : Array Int32 + transitionIndices : Array UInt8 + localTimeTypes : Array LocalTimeType + abbreviations : Array String + leapSeconds : Array LeapSecond + stdWallIndicators : Array Bool + utLocalIndicators : Array Bool + deriving Repr, Inhabited + +/-- +Represents version 2 of the TZif format, extending TZifV1 with an optional footer. +-/ +structure TZifV2 extends TZifV1 where + footer : Option String + deriving Repr, Inhabited + +/-- +Represents a TZif file, which can be either version 1 or version 2. +-/ +structure TZif where + v1: TZifV1 + v2 : Option TZifV2 + deriving Repr, Inhabited + +namespace TZif +open Lean.Parsec Lean.Parsec.ByteArray + +private def toUInt32 (bs : ByteArray) : UInt32 := + assert! bs.size == 4 + (bs.get! 0).toUInt32 <<< 0x18 ||| + (bs.get! 1).toUInt32 <<< 0x10 ||| + (bs.get! 2).toUInt32 <<< 0x8 ||| + (bs.get! 3).toUInt32 + +private def toInt32 (bs : ByteArray) : Int32 := + let n := toUInt32 bs |>.toNat + if n < (1 <<< 31) + then Int.ofNat n + else Int.negOfNat (UInt32.size - n) + +private def toInt64 (bs : ByteArray) : Int64 := + let n := ByteArray.toUInt64BE! bs |>.toNat + if n < (1 <<< 63) + then Int.ofNat n + else Int.negOfNat (UInt64.size - n) + +private def manyN (n : Nat) (p : Parser α) : Parser (Array α) := do + let mut result := #[] + for _ in [0:n] do + let x ← p + result := result.push x + return result + +private def pu64 : Parser UInt64 := ByteArray.toUInt64LE! <$> take 8 +private def pi64 : Parser Int64 := toInt64 <$> take 8 +private def pu32 : Parser UInt32 := toUInt32 <$> take 4 +private def pi32 : Parser Int32 := toInt32 <$> take 4 +private def pu8 : Parser UInt8 := any +private def pbool : Parser Bool := (· != 0) <$> pu8 + +private def parseHeader : Parser Header := + Header.mk + <$> (pstring "TZif" *> pu8) + <*> (take 15 *> pu32) + <*> pu32 + <*> pu32 + <*> pu32 + <*> pu32 + <*> pu32 + +private def parseLocalTimeType : Parser LocalTimeType := + LocalTimeType.mk + <$> pi32 + <*> pbool + <*> pu8 + +private def parseLeapSecond (p : Parser Int) : Parser LeapSecond := + LeapSecond.mk + <$> p + <*> pi32 + +private def parseTransitionTimes (size: Parser Int32) (n : UInt32) : Parser (Array Int32) := + manyN (n.toNat) size + +private def parseTransitionIndices (n : UInt32) : Parser (Array UInt8) := + manyN (n.toNat) pu8 + +private def parseLocalTimeTypes (n : UInt32) : Parser (Array LocalTimeType) := + manyN (n.toNat) parseLocalTimeType + +private def parseAbbreviations (times: Array LocalTimeType) (n : UInt32) : Parser (Array String) := do + let mut strings := #[] + let mut current := "" + let mut chars ← manyN n.toNat pu8 + + for time in times do + for indx in [time.abbreviationIndex.toNat:n.toNat] do + let char := chars.get! indx + if char = 0 then + strings := strings.push current + current := "" + break + else + current := current.push (Char.ofUInt8 char) + + return strings + +private def parseLeapSeconds (size : Parser Int) (n : UInt32) : Parser (Array LeapSecond) := + manyN (n.toNat) (parseLeapSecond size) + +private def parseIndicators (n : UInt32) : Parser (Array Bool) := + manyN (n.toNat) pbool + +private def parseTZifV1 : Parser TZifV1 := do + let header ← parseHeader + + let transitionTimes ← parseTransitionTimes pi32 header.timecnt + let transitionIndices ← parseTransitionIndices header.timecnt + let localTimeTypes ← parseLocalTimeTypes header.typecnt + let abbreviations ← parseAbbreviations localTimeTypes header.charcnt + let leapSeconds ← parseLeapSeconds pi32 header.leapcnt + let stdWallIndicators ← parseIndicators header.isstdcnt + let utLocalIndicators ← parseIndicators header.isutcnt + + return { + header + transitionTimes + transitionIndices + localTimeTypes + abbreviations + leapSeconds + stdWallIndicators + utLocalIndicators + } + +private def parseFooter : Parser (Option String) := do + let char ← pu8 + + if char = 0x0A then pure () else return none + + let tzString ← many (satisfy (λ c => c ≠ 0x0A)) + let mut str := "" + + for byte in tzString do + str := str.push (Char.ofUInt8 byte) + + return str + +private def parseTZifV2 : Parser (Option TZifV2) := optional $ do + let header ← parseHeader + + let transitionTimes ← parseTransitionTimes pi64 header.timecnt + let transitionIndices ← parseTransitionIndices header.timecnt + let localTimeTypes ← parseLocalTimeTypes header.typecnt + let abbreviations ← parseAbbreviations localTimeTypes header.charcnt + let leapSeconds ← parseLeapSeconds pi64 header.leapcnt + let stdWallIndicators ← parseIndicators header.isstdcnt + let utLocalIndicators ← parseIndicators header.isutcnt + + return { + header + transitionTimes + transitionIndices + localTimeTypes + abbreviations + leapSeconds + stdWallIndicators + utLocalIndicators + footer := ← parseFooter + } + +/-- +Parses a TZif file, which may be in either version 1 or version 2 format. +-/ +def parse : Parser TZif := do + let v1 ← parseTZifV1 + let v2 ← parseTZifV2 + return { v1, v2 } + +end TZif diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean new file mode 100644 index 000000000000..c603da209a20 --- /dev/null +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -0,0 +1,88 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.Time +import Std.Time.DateTime +import Std.Time.Internal + +namespace Std +namespace Time +namespace TimeZone +open Internal + +/-- +Represents the type of local time in relation to UTC. +-/ +inductive UTLocal + | ut + | local + deriving Repr, Inhabited + +/-- +Represents types of wall clocks or standard times. +-/ +inductive StdWall + | wall + | standard + deriving Repr, Inhabited + +/-- +Represents a type of local time, including offset and daylight saving information. +-/ +structure LocalTimeType where + gmtOffset : Second.Offset + isDst : Bool + abbreviation : String + wall : StdWall + utLocal : UTLocal + deriving Repr, Inhabited + +/-- +Represents a leap second event, including the time of the transition and the correction applied. +-/ +structure LeapSecond where + transitionTime : Second.Offset + correction : Second.Offset + deriving Repr, Inhabited + +/-- +Represents a time zone transition, mapping a time to a local time type. +-/ +structure Transition where + time : Second.Offset + localTimeType : LocalTimeType + deriving Repr, Inhabited + +/-- +Represents the rules for a time zone, abstracting away binary data and focusing on key transitions and types. +-/ +structure ZoneRules where + transitions : Array Transition + leapSeconds : Array LeapSecond + deriving Repr, Inhabited + + +namespace ZoneRules + +/-- +Finds the transition corresponding to a given timestamp in `ZoneRules`. +If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. +-/ +def findTransitionForTimestamp (zoneRules : ZoneRules) (timestamp : Timestamp) : Option Transition := + if let some idx := zoneRules.transitions.findIdx? (λ t => t.time.val > timestamp.val) + then zoneRules.transitions.get? (idx - 1) + else zoneRules.transitions.back? + +/-- +Apply leap seconds to a Timestamp +-/ +def applyLeapSeconds (timeInSeconds : Timestamp) (leapSeconds : Array LeapSecond) : Timestamp := Id.run do + let mut currentTime := timeInSeconds + for i in [:leapSeconds.size] do + let leapSec := leapSeconds.get! i + if currentTime.val >= leapSec.transitionTime.val then + currentTime := ⟨currentTime.val + leapSec.correction.val⟩ + return currentTime From 5e269e2ca3ad20679d60345f90ee5dacc6c0f41c Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 9 Aug 2024 02:55:54 -0300 Subject: [PATCH 007/118] fix: wrong import --- src/Std/Time/Duration.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 4f538f952a40..0ee4ac0aaff0 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Time.Internal -import Time.Time +import Std.Time.Internal +import Std.Time.Time namespace Std namespace Time From 8d0f19634da19db5596e13b8c032efb7db2073f1 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 9 Aug 2024 14:40:14 -0300 Subject: [PATCH 008/118] feat: add io functions for timestamp and instant --- src/Std/Time.lean | 1 + src/Std/Time/Date/Unit/Basic.lean | 2 + src/Std/Time/Date/Unit/Month.lean | 49 +++++++++++++++++++ src/Std/Time/Date/WeekDate.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 9 ++++ src/Std/Time/Duration.lean | 6 +++ src/Std/Time/Format/Basic.lean | 44 ++++++++--------- src/Std/Time/Zoned.lean | 38 +++++++++++++++ src/Std/Time/Zoned/Database/Basic.lean | 66 ++++++++++++++++---------- src/Std/Time/Zoned/Database/TzIf.lean | 6 +-- src/Std/Time/Zoned/DateTime.lean | 26 +++++----- src/Std/Time/Zoned/Offset.lean | 18 +++++-- src/Std/Time/Zoned/TimeZone.lean | 12 ++++- src/Std/Time/Zoned/ZoneRules.lean | 16 ++++++- src/Std/Time/Zoned/ZonedDateTime.lean | 15 ++++++ src/runtime/io.cpp | 21 ++++++++ 16 files changed, 259 insertions(+), 72 deletions(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 3f84a1da7b69..69cc46fce838 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -9,6 +9,7 @@ import Std.Time.Zoned import Std.Time.Format import Std.Time.DateTime import Std.Time.Notation +import Std.Time.Duration namespace Time diff --git a/src/Std/Time/Date/Unit/Basic.lean b/src/Std/Time/Date/Unit/Basic.lean index 27f5b563f887..69a15c3e4387 100644 --- a/src/Std/Time/Date/Unit/Basic.lean +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -13,6 +13,8 @@ import Std.Time.Date.Unit.WeekOfYear namespace Std namespace Time.Day.Ordinal.OfYear +set_option linter.all true + /-- Conevrts a `Year` and a `Ordinal.OfYear` to a valid day and month. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 6b8fd613ef2c..24809b0b50ad 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -13,6 +13,8 @@ namespace Std namespace Time open Internal +set_option linter.all true + namespace Month /-- `Ordinal` represents a bounded value for months, which ranges between 1 and 12. @@ -35,17 +37,64 @@ instance : OfNat Offset n := ⟨Int.ofNat n⟩ namespace Ordinal +/-- +The ordinal value representing the month of January. +-/ @[inline] def january : Ordinal := 1 + +/-- +The ordinal value representing the month of February. +-/ @[inline] def february : Ordinal := 2 + +/-- +The ordinal value representing the month of March. +-/ @[inline] def march : Ordinal := 3 + +/-- +The ordinal value representing the month of April. +-/ @[inline] def april : Ordinal := 4 + +/-- +The ordinal value representing the month of May. +-/ @[inline] def may : Ordinal := 5 + +/-- +The ordinal value representing the month of June. +-/ @[inline] def june : Ordinal := 6 + +/-- +The ordinal value representing the month of July. +-/ @[inline] def july : Ordinal := 7 + +/-- +The ordinal value representing the month of August. +-/ @[inline] def august : Ordinal := 8 + +/-- +The ordinal value representing the month of September. +-/ @[inline] def september : Ordinal := 9 + +/-- +The ordinal value representing the month of October. +-/ @[inline] def october : Ordinal := 10 + +/-- +The ordinal value representing the month of November. +-/ @[inline] def november : Ordinal := 11 + +/-- +The ordinal value representing the month of December. +-/ @[inline] def december : Ordinal := 12 /-- diff --git a/src/Std/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean index d546c5a6e320..fab98d38e713 100644 --- a/src/Std/Time/Date/WeekDate.lean +++ b/src/Std/Time/Date/WeekDate.lean @@ -32,7 +32,7 @@ def toDays (wd : WeekDate) : Day.Offset := /-- Creates a `WeekDate` from a `Day.Offset`. -/ -def fromScalar (scalar : Day.Offset) : WeekDate := +def fromDays (scalar : Day.Offset) : WeekDate := let totalDays := scalar.val let year := totalDays / 365 let week := diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index e4c5fd7612fa..03263ee6e2a6 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -20,6 +20,15 @@ def Timestamp := Second.Offset namespace Timestamp +/-- +Get the current monotonic time. +-/ +@[extern "lean_get_current_timestamp"] +protected opaque now : IO Timestamp + +/-- +A function to transform `Timestamp` to `Second.Offset` +-/ @[inline] def toSeconds (tm : Timestamp) : Second.Offset := tm diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 0ee4ac0aaff0..4ae0cb2937d6 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -31,6 +31,12 @@ structure Duration where namespace Instant +/-- +Get the current monotonic time. +-/ +@[extern "lean_get_current_time"] +protected opaque now : IO Instant + /-- Gets the difference of two `Instant` in a `Duration`. -/ diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 25ca3a5a335b..b72351c9b8de 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -177,12 +177,12 @@ private def parseFormatPart : Parser FormatPart private def specParser : Parser FormatString := (Array.toList <$> many parseFormatPart) <* eof -private def specParse (s: String) : Except String FormatString := +private def specParse (s : String) : Except String FormatString := specParser.run s -- Pretty printer -private def unabbrevMonth (month: Month.Ordinal) : String := +private def unabbrevMonth (month : Month.Ordinal) : String := match month.val, month.property with | 1, _ => "January" | 2, _ => "February" @@ -197,7 +197,7 @@ private def unabbrevMonth (month: Month.Ordinal) : String := | 11, _ => "November" | 12, _ => "December" -private def abbrevMonth (month: Month.Ordinal) : String := +private def abbrevMonth (month : Month.Ordinal) : String := match month.val, month.property with | 1, _ => "Jan" | 2, _ => "Feb" @@ -212,7 +212,7 @@ private def abbrevMonth (month: Month.Ordinal) : String := | 11, _ => "Nov" | 12, _ => "Dec" -private def abbrevDayOfWeek (day: Weekday) : String := +private def abbrevDayOfWeek (day : Weekday) : String := match day with | .sun => "Sun" | .mon => "Mon" @@ -222,7 +222,7 @@ private def abbrevDayOfWeek (day: Weekday) : String := | .fri => "Fri" | .sat => "Sat" -private def dayOfWeek (day: Weekday) : String := +private def dayOfWeek (day : Weekday) : String := match day with | .sun => "Sunday" | .mon => "Monday" @@ -237,26 +237,26 @@ private def leftPad (n : Nat) (a : Char) (s : String) : String := private def formatWithDate (date : DateTime tz) : Modifier → String | .YYYY => s!"{leftPad 4 '0' (toString date.year)}" - | .YY => s!"{leftPad 2 '0' (toString $ date.year.toNat % 100)}" + | .YY => s!"{leftPad 2 '0' (toString <| date.year.toNat % 100)}" | .MMMM => unabbrevMonth date.month | .MMM => abbrevMonth date.month - | .MM => s!"{leftPad 2 '0' (toString $ date.month.toNat)}" + | .MM => s!"{leftPad 2 '0' (toString <| date.month.toNat)}" | .M => s!"{date.month.toNat}" - | .DD => s!"{leftPad 2 '0' (toString $ date.day.toNat)}" + | .DD => s!"{leftPad 2 '0' (toString <| date.day.toNat)}" | .D => s!"{date.day.toNat}" - | .d => s!"{leftPad 2 ' ' $ toString date.day.toNat}" + | .d => s!"{leftPad 2 ' ' <| toString date.day.toNat}" | .EEEE => dayOfWeek date.weekday | .EEE => abbrevDayOfWeek date.weekday | .hh => s!"{leftPad 2 '0' (toString date.hour.toNat)}" | .h => s!"{date.hour.toNat}" - | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' $ toString hour}" + | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' <| toString hour}" | .H => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{hour}" | .AA => if date.hour.toNat < 12 then "AM" else "PM" | .aa => if date.hour.toNat < 12 then "am" else "pm" - | .mm => s!"{leftPad 2 '0' $ toString date.minute.toNat}" + | .mm => s!"{leftPad 2 '0' <| toString date.minute.toNat}" | .m => s!"{date.minute.toNat}" - | .sss => s!"{leftPad 3 '0' $ toString date.milliseconds.toNat}" - | .ss => s!"{leftPad 2 '0' $ toString date.second.toNat}" + | .sss => s!"{leftPad 3 '0' <| toString date.milliseconds.toNat}" + | .ss => s!"{leftPad 2 '0' <| toString date.second.toNat}" | .s => s!"{date.second.toNat}" | .ZZZZZ => tz.offset.toIsoString true | .ZZZZ => tz.offset.toIsoString false @@ -303,26 +303,26 @@ private def SingleFormatType : Modifier → Type private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) : String := match modifier with | .YYYY => s!"{leftPad 4 '0' (toString data.toNat)}" - | .YY => s!"{leftPad 2 '0' (toString $ data.toNat % 100)}" + | .YY => s!"{leftPad 2 '0' (toString <| data.toNat % 100)}" | .MMMM => unabbrevMonth data | .MMM => abbrevMonth data - | .MM => s!"{leftPad 2 '0' (toString $ data.toNat)}" + | .MM => s!"{leftPad 2 '0' (toString <| data.toNat)}" | .M => s!"{data.toNat}" - | .DD => s!"{leftPad 2 '0' (toString $ data.toNat)}" + | .DD => s!"{leftPad 2 '0' (toString <| data.toNat)}" | .D => s!"{data.toNat}" - | .d => s!"{leftPad 2 ' ' $ toString data.toNat}" + | .d => s!"{leftPad 2 ' ' <| toString data.toNat}" | .EEEE => dayOfWeek data | .EEE => abbrevDayOfWeek data | .hh => s!"{leftPad 2 '0' (toString data.toNat)}" | .h => s!"{data.toNat}" - | .HH => let hour := data.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' $ toString hour}" + | .HH => let hour := data.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' <| toString hour}" | .H => let hour := data.val % 12; if hour == 0 then "12" else s!"{hour}" | .AA => match data with | .am => "AM" | .pm => "PM" | .aa => match data with | .am => "am" | .pm => "pm" - | .mm => s!"{leftPad 2 '0' $ toString data.toNat}" + | .mm => s!"{leftPad 2 '0' <| toString data.toNat}" | .m => s!"{data.toNat}" - | .sss => s!"{leftPad 3 '0' $ toString data.toNat}" - | .ss => s!"{leftPad 2 '0' $ toString data.toNat}" + | .sss => s!"{leftPad 3 '0' <| toString data.toNat}" + | .ss => s!"{leftPad 2 '0' <| toString data.toNat}" | .s => s!"{data.toNat}" | .ZZZZZ => data.toIsoString true | .ZZZZ => data.toIsoString false @@ -331,7 +331,7 @@ private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) | .z => data @[simp] -def FormatType (result: Type) : FormatString → Type +def FormatType (result : Type) : FormatString → Type | .modifier entry :: xs => (SingleFormatType entry) → (FormatType result xs) | .string _ :: xs => (FormatType result xs) | [] => result diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index f327cae98b01..58cafcd02735 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -5,4 +5,42 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Zoned.DateTime +import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.ZonedDateTime +import Std.Time.Zoned.Database.Basic + +namespace Std +namespace Time +namespace TimeZone +namespace ZoneRules + +/-- +Parses a binary data into a zone rules. +-/ +def parseTZif (bin : ByteArray) : Except String ZoneRules := do + let database ← Database.TZif.TZif.parse.run bin + Database.convertTZif database + +/-- +Gets the ZoneRules from the a TZIf file. +-/ +def parseTZIfFromDisk (path : System.FilePath) : IO ZoneRules := do + let binary ← IO.FS.readBinFile path + IO.ofExcept (parseTZif binary) + +/-- +Gets the rules from local timezone. +-/ +def localRules : IO ZoneRules := do + parseTZIfFromDisk "etc/localtime" + +/-- +Gets the rules from local timezone. +-/ +def readRulesFromDisk (id : String) : IO ZoneRules := do + parseTZIfFromDisk (System.FilePath.join "/usr/share/zoneinfo/" id) + +end ZoneRules +end TimeZone +end Time +end Std diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index ef2331d1a9ea..e5a3587761ca 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -30,50 +30,66 @@ def convertUt : Bool → UTLocal /-- Converts a `TZif.LeapSecond` structure to a `LeapSecond` structure. -/ -def convertLeapSecond (tz: TZif.LeapSecond) : LeapSecond := +def convertLeapSecond (tz : TZif.LeapSecond) : LeapSecond := { transitionTime := Internal.UnitVal.mk tz.transitionTime, correction := Internal.UnitVal.mk tz.correction } +/-- +Converts a LocalTime. +-/ +def convetLocalTime (index : Nat) (tz : TZif.TZifV1) : Option LocalTimeType := do + let localType ← tz.localTimeTypes.get? index + let abbreviation ← tz.abbreviations.get? index + let wallflag ← convertWall <$> tz.stdWallIndicators.get? index + let utLocal ← convertUt <$> tz.utLocalIndicators.get? index + + return { + gmtOffset := Offset.ofSeconds <| Internal.UnitVal.mk localType.gmtOffset + isDst := localType.isDst + abbreviation + wall := wallflag + utLocal + } + +/-- +Converts a transition. +-/ +def convertTransition (times: Array LocalTimeType) (index : Nat) (tz : TZif.TZifV1) : Option Transition := do + let time := tz.transitionTimes.get! index + let time := Internal.UnitVal.mk time + let indice := tz.transitionIndices.get! index + return { time, localTimeType := times.get! indice.toNat } + /-- Converts a `TZif.TZifV1` structure to a `ZoneRules` structure. -/ -def convertTZifV1 (tz: TZif.TZifV1) : Option ZoneRules := do +def convertTZifV1 (tz : TZif.TZifV1) : Except String ZoneRules := do let mut times : Array LocalTimeType := #[] - for i in [0:tz.header.timecnt.toNat] do - let localType := tz.localTimeTypes.get! i - let abbreviation := tz.abbreviations.get! i - let wallflag := convertWall <| tz.stdWallIndicators.get! i - let utLocal := convertUt <| tz.utLocalIndicators.get! i - times := times.push { - gmtOffset := Internal.UnitVal.mk localType.gmtOffset - isDst := localType.isDst - abbreviation - wall := wallflag - utLocal - } + for i in [0:tz.header.typecnt.toNat] do + if let some result := convetLocalTime i tz + then times := times.push result + else .error s!"cannot convert local time {i} of the file" let mut transitions := #[] let leapSeconds := tz.leapSeconds.map convertLeapSecond for i in [0:tz.transitionTimes.size] do - let time := tz.transitionTimes.get! i - let time := Internal.UnitVal.mk time - let indice ← tz.transitionIndices.get? i - transitions := transitions.push { time, localTimeType := times.get! indice.toNat } + if let some result := convertTransition times i tz + then transitions := transitions.push result + else .error s!"cannot convert transtiion {i} of the file" - some { leapSeconds,transitions } + .ok { leapSeconds,transitions, localTimes := times } /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. -/ -def convertTZifV2 (tz: TZif.TZifV2) : Option ZoneRules := do +def convertTZifV2 (tz : TZif.TZifV2) : Except String ZoneRules := do convertTZifV1 tz.toTZifV1 /-- Converts a `TZif.TZif` structure to a `ZoneRules` structure. -/ -def convertTZif (tz: TZif.TZif) : Option ZoneRules := do - if let some v2 := tz.v2 then - convertTZifV2 v2 - else - convertTZifV1 tz.v1 +def convertTZif (tz : TZif.TZif) : Except String ZoneRules := do + if let some v2 := tz.v2 + then convertTZifV2 v2 + else convertTZifV1 tz.v1 diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index 82d4c53bd336..09388565c9fd 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -73,7 +73,7 @@ structure TZifV2 extends TZifV1 where Represents a TZif file, which can be either version 1 or version 2. -/ structure TZif where - v1: TZifV1 + v1 : TZifV1 v2 : Option TZifV2 deriving Repr, Inhabited @@ -134,7 +134,7 @@ private def parseLeapSecond (p : Parser Int) : Parser LeapSecond := <$> p <*> pi32 -private def parseTransitionTimes (size: Parser Int32) (n : UInt32) : Parser (Array Int32) := +private def parseTransitionTimes (size : Parser Int32) (n : UInt32) : Parser (Array Int32) := manyN (n.toNat) size private def parseTransitionIndices (n : UInt32) : Parser (Array UInt8) := @@ -143,7 +143,7 @@ private def parseTransitionIndices (n : UInt32) : Parser (Array UInt8) := private def parseLocalTimeTypes (n : UInt32) : Parser (Array LocalTimeType) := manyN (n.toNat) parseLocalTimeType -private def parseAbbreviations (times: Array LocalTimeType) (n : UInt32) : Parser (Array String) := do +private def parseAbbreviations (times : Array LocalTimeType) (n : UInt32) : Parser (Array String) := do let mut strings := #[] let mut current := "" let mut chars ← manyN n.toNat pu8 diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index dc263f7e046b..d011a292da24 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -6,6 +6,7 @@ Authors: Sofia Rodrigues prelude import Std.Time.DateTime import Std.Time.Zoned.TimeZone +import Std.Time.Zoned.ZoneRules namespace Std namespace Time @@ -16,8 +17,10 @@ It stores a `Timestamp`, a `LocalDateTime` and a `TimeZone` structure DateTime (tz : TimeZone) where private mk :: timestamp : Timestamp - date : LocalDateTime - deriving Repr, Inhabited + date : Thunk LocalDateTime + +instance : Inhabited (DateTime tz) where + default := ⟨Inhabited.default, Thunk.mk λ_ => Inhabited.default⟩ namespace DateTime @@ -26,8 +29,7 @@ Creates a new `DateTime` out of a `Timestamp` -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := - let date := (tm + tz.toSeconds).toLocalDateTime - DateTime.mk tm date + DateTime.mk tm (Thunk.mk <| λ_ => (tm + tz.toSeconds).toLocalDateTime) /-- Creates a new `Timestamp` out of a `DateTime` @@ -56,55 +58,55 @@ Getter for the `Year` inside of a `DateTime` -/ @[inline] def year (dt : DateTime tz) : Year.Offset := - dt.date.year + dt.date.get.year /-- Getter for the `Month` inside of a `DateTime` -/ @[inline] def month (dt : DateTime tz) : Month.Ordinal := - dt.date.month + dt.date.get.month /-- Getter for the `Day` inside of a `DateTime` -/ @[inline] def day (dt : DateTime tz) : Day.Ordinal := - dt.date.day + dt.date.get.day /-- Getter for the `Hour` inside of a `DateTime` -/ @[inline] def hour (dt : DateTime tz) : Hour.Ordinal := - dt.date.hour + dt.date.get.hour /-- Getter for the `Minute` inside of a `DateTime` -/ @[inline] def minute (dt : DateTime tz) : Minute.Ordinal := - dt.date.minute + dt.date.get.minute /-- Getter for the `Second` inside of a `DateTime` -/ @[inline] def second (dt : DateTime tz) : Second.Ordinal := - dt.date.second + dt.date.get.second /-- Getter for the `Milliseconds` inside of a `DateTime` -/ @[inline] def milliseconds (dt : DateTime tz) : Millisecond.Ordinal := - dt.date.time.nano.toMillisecond + dt.date.get.time.nano.toMillisecond /-- Gets the `Weekday` of a DateTime. -/ @[inline] def weekday (dt : DateTime tz) : Weekday := - dt.date.date.weekday + dt.date.get.date.weekday end DateTime end Time diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index bb0ff04a952d..b8dc1d33184e 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -12,13 +12,21 @@ namespace Time namespace TimeZone open Internal +set_option linter.all true + /-- Represents a timezone offset with an hour and second component. -/ structure Offset where private mk :: - hour: Hour.Offset - second: Second.Offset + /-- + The timezone offset in Hours. + -/ + hour : Hour.Offset + /-- + The same timezone offset in seconds. + -/ + second : Second.Offset deriving Repr, Inhabited namespace Offset @@ -46,20 +54,20 @@ def zero : Offset := /-- Creates an `Offset` from a given number of hour. -/ -def ofHours (n: Hour.Offset) : Offset := +def ofHours (n : Hour.Offset) : Offset := mk n n.toSeconds /-- Creates an `Offset` from a given number of hour and minuets. -/ -def ofHoursAndMinutes (n: Hour.Offset) (m: Minute.Offset) : Offset := +def ofHoursAndMinutes (n : Hour.Offset) (m : Minute.Offset) : Offset := let secs := n.toSeconds + m.toSeconds mk secs.toHours secs /-- Creates an `Offset` from a given number of second. -/ -def ofSeconds (n: Second.Offset) : Offset := +def ofSeconds (n : Second.Offset) : Offset := mk n.toHours n end Offset diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index cf6dcafcf190..b74140a8bff4 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -10,11 +10,19 @@ import Std.Time.Zoned.Offset namespace Std namespace Time +set_option linter.all true + /-- An enumeration representing different time zones. -/ structure TimeZone where + /-- + The `Offset` of the date time. + -/ offset : TimeZone.Offset + /-- + The name of the time zone. + -/ name : String deriving Inhabited, Repr @@ -35,13 +43,13 @@ def GMT : TimeZone := /-- Creates a `Timestamp` from a given number of hour. -/ -def ofHours (name : String) (n: Hour.Offset) : TimeZone := +def ofHours (name : String) (n : Hour.Offset) : TimeZone := TimeZone.mk (Offset.ofHours n) name /-- Creates a `Timestamp` from a given number of second. -/ -def ofSeconds (name : String) (n: Second.Offset) : TimeZone := +def ofSeconds (name : String) (n : Second.Offset) : TimeZone := TimeZone.mk (Offset.ofSeconds n) name /-- diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index c603da209a20..214d9632c65d 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -7,6 +7,8 @@ prelude import Std.Time.Time import Std.Time.DateTime import Std.Time.Internal +import Std.Time.Zoned.Offset +import Std.Time.Zoned.TimeZone namespace Std namespace Time @@ -33,13 +35,23 @@ inductive StdWall Represents a type of local time, including offset and daylight saving information. -/ structure LocalTimeType where - gmtOffset : Second.Offset + gmtOffset : TimeZone.Offset isDst : Bool abbreviation : String wall : StdWall utLocal : UTLocal deriving Repr, Inhabited +namespace LocalTimeType + +/-- +Gets the `TimeZone` offset from a `LocalTimeType`. +-/ +def getTimeZone (time : LocalTimeType) : TimeZone := + ⟨time.gmtOffset, time.abbreviation⟩ + +end LocalTimeType + /-- Represents a leap second event, including the time of the transition and the correction applied. -/ @@ -60,11 +72,11 @@ structure Transition where Represents the rules for a time zone, abstracting away binary data and focusing on key transitions and types. -/ structure ZoneRules where + localTimes : Array LocalTimeType transitions : Array Transition leapSeconds : Array LeapSecond deriving Repr, Inhabited - namespace ZoneRules /-- diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index bf392cad0b13..21f71d1879ac 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -6,10 +6,16 @@ Authors: Sofia Rodrigues prelude import Std.Time.Zoned.TimeZone import Std.Time.Zoned.DateTime +import Std.Time.Zoned.ZoneRules namespace Std namespace Time +set_option linter.all true + +/-- +The existential version of `DateTime` that instead of storing the timezone with a +-/ def ZonedDateTime := Sigma DateTime instance : CoeDep ZonedDateTime d (DateTime d.fst) where @@ -35,6 +41,15 @@ Creates a new `Timestamp` out of a `ZonedDateTime` def toTimestamp (date : ZonedDateTime) : Timestamp := date.snd.toTimestamp +/-- +Creates a new `DateTime` out of a `Timestamp` +-/ +@[inline] +def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do + let transition ← rules.findTransitionForTimestamp tm + return ofTimestamp tm transition.localTimeType.getTimeZone + + /-- Changes the `TimeZone` to a new one. -/ diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 136612e2fee5..3cdd7a778362 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -528,6 +528,27 @@ extern "C" LEAN_EXPORT obj_res lean_io_prim_handle_put_str(b_obj_arg h, b_obj_ar } } +extern "C" LEAN_EXPORT lean_obj_res lean_get_current_time() { + std::chrono::steady_clock::time_point now = std::chrono::steady_clock::now(); + std::chrono::steady_clock::duration duration = now.time_since_epoch(); + + long long secs = std::chrono::duration_cast(duration).count(); + long long nano = std::chrono::duration_cast(duration).count() % 1000000000; + + lean_object *lean_ts = lean_alloc_ctor(0, 2, 0); + lean_ctor_set(lean_ts, 0, lean_int_to_int(static_cast(secs))); + lean_ctor_set(lean_ts, 1, lean_int_to_int(static_cast(nano))); + + return lean_io_result_mk_ok(lean_ts); +} + +extern "C" LEAN_EXPORT lean_obj_res lean_get_current_timestamp() { + std::chrono::system_clock::time_point now = std::chrono::system_clock::now(); + long long timestamp = std::chrono::duration_cast(now.time_since_epoch()).count(); + lean_object *lean_timestamp = lean_int_to_int((int)timestamp); + return lean_io_result_mk_ok(lean_timestamp); +} + /* monoMsNow : BaseIO Nat */ extern "C" LEAN_EXPORT obj_res lean_io_mono_ms_now(obj_arg /* w */) { static_assert(sizeof(std::chrono::milliseconds::rep) <= sizeof(uint64), "size of std::chrono::nanoseconds::rep may not exceed 64"); From 2c8ef660bf0b5f744a5c57f81d7f5e145c19e420 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:20:47 -0300 Subject: [PATCH 009/118] feat: removes LE, adds a bunch of add and sub functions, move interval stuff to timestamp --- src/Std/Time.lean | 9 +- src/Std/Time/Date/LocalDate.lean | 138 ++++++++++- src/Std/Time/Date/Unit/Basic.lean | 15 -- src/Std/Time/Date/Unit/Day.lean | 28 ++- src/Std/Time/Date/Unit/Month.lean | 14 +- src/Std/Time/Date/Unit/WeekOfYear.lean | 3 +- src/Std/Time/Date/Unit/Year.lean | 2 +- src/Std/Time/Date/WeekDate.lean | 12 +- src/Std/Time/DateTime.lean | 4 +- src/Std/Time/DateTime/LocalDateTime.lean | 140 +++++++++-- src/Std/Time/DateTime/Timestamp.lean | 148 +++++++++-- src/Std/Time/Duration.lean | 116 +-------- src/Std/Time/Format.lean | 23 +- src/Std/Time/Format/Basic.lean | 303 +++++++++++++---------- src/Std/Time/Internal.lean | 1 - src/Std/Time/Internal/Bounded.lean | 64 +++-- src/Std/Time/Internal/LessEq.lean | 34 --- src/Std/Time/Internal/UnitVal.lean | 4 +- src/Std/Time/Notation.lean | 4 - src/Std/Time/Time.lean | 2 + src/Std/Time/Time/HourMarker.lean | 20 +- src/Std/Time/Time/LocalTime.lean | 80 +++++- src/Std/Time/Time/Unit/Basic.lean | 6 +- src/Std/Time/Time/Unit/Hour.lean | 25 +- src/Std/Time/Time/Unit/Millisecond.lean | 4 +- src/Std/Time/Time/Unit/Minute.lean | 2 +- src/Std/Time/Time/Unit/Nanosecond.lean | 2 +- src/Std/Time/Time/Unit/Second.lean | 25 +- src/Std/Time/Zoned.lean | 38 +-- src/Std/Time/Zoned/Database.lean | 44 ++++ src/Std/Time/Zoned/Database/Basic.lean | 41 ++- src/Std/Time/Zoned/Database/TZdb.lean | 76 ++++++ src/Std/Time/Zoned/Database/TzIf.lean | 97 +++++++- src/Std/Time/Zoned/DateTime.lean | 102 +++++++- src/Std/Time/Zoned/Offset.lean | 4 +- src/Std/Time/Zoned/ZoneRules.lean | 121 ++++++++- src/Std/Time/Zoned/ZonedDateTime.lean | 61 ++++- src/runtime/io.cpp | 22 +- 38 files changed, 1318 insertions(+), 516 deletions(-) delete mode 100644 src/Std/Time/Internal/LessEq.lean create mode 100644 src/Std/Time/Zoned/Database.lean create mode 100644 src/Std/Time/Zoned/Database/TZdb.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 69cc46fce838..10ce573310d8 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -11,8 +11,11 @@ import Std.Time.DateTime import Std.Time.Notation import Std.Time.Duration +namespace Std namespace Time +set_option linter.all true + /-! # Time @@ -100,7 +103,7 @@ Combines date and time into a single representation, useful for precise timestam - **`LocalDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. - **`Timestamp`**: Represents a point in time with second-level precision. It starts on the UNIX -dpoch and it should be used when you receive or need to send timestamps to another systems. +epoch and it should be used when you receive or need to send timestamps to another systems. ## Zoned date and times. Combines date, time and time zones. @@ -111,9 +114,7 @@ Combines date, time and time zones. ## Duration Represents spans of time and the difference between two points in time. -- **`Instant`**: Represents a non-negative instant in time with nanosecond precision, represents a - specific point on the time axis. -- **`Duration`**: Represents the time span or difference between two `Instant` values with nanosecond precision. +- **`Duration`**: Represents the time span or difference between two `Timestamp`s values with nanosecond precision. # Formats diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index 2a2d01861592..48de3b02d3fc 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -9,15 +9,35 @@ import Std.Time.Date.Basic namespace Std namespace Time +open Internal + +set_option linter.all true /-- Date in YMD format. -/ structure LocalDate where + + /-- + The year component of the date. It is represented as an `Offset` type from `Year`. + -/ year : Year.Offset + + /-- + The month component of the date. It is represented as an `Ordinal` type from `Month`. + -/ month : Month.Ordinal + + /-- + The day component of the date. It is represented as an `Ordinal` type from `Day`. + -/ day : Day.Ordinal + + /-- + Validates the date by ensuring that the year, month, and day form a correct and valid date. + -/ valid : year.Valid month day + deriving Repr namespace LocalDate @@ -44,7 +64,7 @@ def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordin Creates a new `LocalDate` using YO. -/ def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : LocalDate := - let ⟨⟨month, day⟩, valid⟩ := ordinal.toMonthAndDay + let ⟨⟨month, day⟩, valid⟩ := Month.Ordinal.ofOrdinal ordinal LocalDate.mk year month day valid /-- @@ -122,8 +142,122 @@ Calculate the Year.Offset from a LocalDate def yearsSince (date : LocalDate) (year : Year.Offset) : Year.Offset := date.year - year +/-- +Add `Day.Offset` to a `LocalDate`. +-/ +@[inline] +def addDays (date : LocalDate) (days : Day.Offset) : LocalDate := + let dateDays := date.toDaysSinceUNIXEpoch + ofDaysSinceUNIXEpoch (Add.add dateDays days) + +/-- +Subtracts `Day.Offset` to a `LocalDate`. +-/ +@[inline] +def subDays (date : LocalDate) (days : Day.Offset) : LocalDate := + addDays date (-days) + +/-- +Add `Month.Offset` to a `LocalDate`, it clips the day to the last valid day of that month. +-/ +def addMonthsClip (date : LocalDate) (months : Month.Offset) : LocalDate := + let yearsOffset := months.div 12 + let monthOffset := Bounded.LE.byMod months 12 (by decide) + let months := date.month.addBounds monthOffset + + let (yearsOffset, months) : Year.Offset × Month.Ordinal := by + if h₁ : months.val > 12 then + let months := months |>.truncateBottom h₁ |>.sub 12 + exact (yearsOffset.add 1, months.expandTop (by decide)) + else if h₂ : months.val < 1 then + let months := months |>.truncateTop (Int.le_sub_one_of_lt h₂) |>.add 12 + exact (yearsOffset.sub 1, months.expandBottom (by decide)) + else + exact (yearsOffset, months.truncateTop (Int.not_lt.mp h₁) |>.truncateBottom (Int.not_lt.mp h₂)) + + LocalDate.clip (date.year.add yearsOffset) months date.day + +/-- +Subtracts `Month.Offset` from a `LocalDate`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subMonthsClip (date : LocalDate) (months : Month.Offset) : LocalDate := + addMonthsClip date (-months) + +/-- +Add `Month.Offset` to a `LocalDate`, this function rolls over any excess days into the following +month. +-/ +def addMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := + let yearsOffset := months.div 12 + let monthOffset := Bounded.LE.byMod months 12 (by decide) + let months := date.month.addBounds monthOffset + + let (yearsOffset, months) : Year.Offset × Month.Ordinal := by + if h₁ : months.val > 12 then + let months := months |>.truncateBottom h₁ |>.sub 12 + exact (yearsOffset.add 1, months.expandTop (by decide)) + else if h₂ : months.val < 1 then + let months := months |>.truncateTop (Int.le_sub_one_of_lt h₂) |>.add 12 + exact (yearsOffset.sub 1, months.expandBottom (by decide)) + else + exact (yearsOffset, months.truncateTop (Int.not_lt.mp h₁) |>.truncateBottom (Int.not_lt.mp h₂)) + + let year : Year.Offset := date.year.add yearsOffset + let ⟨days, proof⟩ := Month.Ordinal.days year.isLeap months + + if h : days.val ≥ date.day.val then + let p : year.Valid months date.day := by + simp_all [Year.Offset.Valid, Month.Ordinal.Valid] + exact Int.le_trans h proof + dbg_trace s!"roll {days.val} {date.day.val}" + LocalDate.mk year months date.day p + else + let roll : Day.Offset := UnitVal.mk (date.day.val - days.toInt) + let date := LocalDate.clip (date.year.add yearsOffset) months date.day + let days := date.toDaysSinceUNIXEpoch + roll + LocalDate.ofDaysSinceUNIXEpoch days + +/-- +Subtract `Month.Offset` from a `LocalDate`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def subMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := + addMonthsRollOver date (-months) + +/-- +Add `Year.Offset` to a `LocalDate`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def addYearsRollOver (date : LocalDate) (years : Year.Offset) : LocalDate := + addMonthsRollOver date (years.mul 12) + +/-- +Add `Year.Offset` to a `LocalDate`, it clips the day to the last valid day of that month. +-/ +@[inline] +def addYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := + addMonthsClip date (years.mul 12) + +/-- +Subtract `Year.Offset` from a `LocalDate`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def subYearsRollOver (date : LocalDate) (years : Year.Offset) : LocalDate := + addMonthsRollOver date (- years.mul 12) + +/-- +Subtract `Year.Offset` from a `LocalDate`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := + addMonthsClip date (- years.mul 12) + instance : HAdd LocalDate Day.Offset LocalDate where - hAdd date day := ofDaysSinceUNIXEpoch (toDaysSinceUNIXEpoch date + day) + hAdd date day := ofDaysSinceUNIXEpoch (toDaysSinceUNIXEpoch date + day) end LocalDate end Time diff --git a/src/Std/Time/Date/Unit/Basic.lean b/src/Std/Time/Date/Unit/Basic.lean index 69a15c3e4387..d8868f8a2972 100644 --- a/src/Std/Time/Date/Unit/Basic.lean +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -9,18 +9,3 @@ import Std.Time.Date.Unit.Month import Std.Time.Date.Unit.Year import Std.Time.Date.Unit.Weekday import Std.Time.Date.Unit.WeekOfYear - -namespace Std -namespace Time.Day.Ordinal.OfYear - -set_option linter.all true - -/-- -Conevrts a `Year` and a `Ordinal.OfYear` to a valid day and month. --/ -@[inline] -def toMonthAndDay (year : Year.Offset) (ordinal : OfYear year.isLeap) : { val : Month.Ordinal × Day.Ordinal // Year.Offset.Valid year (Prod.fst val) (Prod.snd val) } := - Month.Ordinal.ofOrdinal ordinal - -end Time.Day.Ordinal.OfYear -end Std diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index ccc1609d04a7..7f9639966b77 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -22,8 +22,7 @@ namespace Day def Ordinal := Bounded.LE 1 31 deriving Repr, BEq, LE, LT -instance [Le 1 n] [Le n 31] : OfNat Ordinal n where - ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) +instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (30 : Nat))) n) instance { x y : Ordinal } : Decidable (x ≤ y) := inferInstanceAs (Decidable (x.val ≤ y.val)) @@ -36,8 +35,11 @@ is a leap year or 365. -/ def Ordinal.OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) -instance [Le 1 n] [Le n 365] : OfNat (Ordinal.OfYear leap) n where - ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr (by have := Le.p (n := n) (m := 365); split <;> omega))) +instance : OfNat (Ordinal.OfYear leap) n := by + have inst := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (364 : Nat))) n) + cases leap + · exact inst + · exact ⟨inst.ofNat.expandTop (by decide)⟩ instance : OfNat (Ordinal.OfYear true) 366 where ofNat := Bounded.LE.mk (Int.ofNat 366) (by decide) @@ -60,6 +62,7 @@ namespace Ordinal /-- Creates an `Ordinal` from a natural number, ensuring the value is within bounds. -/ +@[inline] def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 31 := by decide) : Ordinal := Bounded.LE.ofNat' data h @@ -67,12 +70,14 @@ def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 31 := by decide) : Ordinal : Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds, if its 0 then its converted to 1. -/ +@[inline] def ofFin (data : Fin 32) : Ordinal := Bounded.LE.ofFin' data (by decide) /-- Converts an `Ordinal` to an `Offset`. -/ +@[inline] def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val @@ -83,9 +88,24 @@ namespace Offset /-- Convert `Day.Offset` into `Second.Offset`. -/ +@[inline] def toSeconds (days : Offset) : Second.Offset := days.mul 86400 +/-- +Convert `Day.Offset` into `Minute.Offset`. +-/ +@[inline] +def toMinutes (days : Offset) : Minute.Offset := + days.mul 1440 + +/-- +Convert `Day.Offset` into `Hour.Offset`. +-/ +@[inline] +def toHours (days : Offset) : Hour.Offset := + days.mul 24 + end Offset end Day end Time diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 24809b0b50ad..364604a51d78 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -16,14 +16,14 @@ open Internal set_option linter.all true namespace Month + /-- `Ordinal` represents a bounded value for months, which ranges between 1 and 12. -/ def Ordinal := Bounded.LE 1 12 deriving Repr, BEq, LE -instance [Le 1 n] [Le n 12] : OfNat Ordinal n where - ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) +instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (11 : Nat))) n) instance : Inhabited Ordinal where default := 1 @@ -133,17 +133,17 @@ def toSeconds (leap : Bool) (month : Ordinal) : Second.Offset := Transforms `Month.Ordinal` into `Minute.Offset`. -/ @[inline] -def toMinute (leap : Bool) (month : Ordinal) : Minute.Offset := +def toMinutes (leap : Bool) (month : Ordinal) : Minute.Offset := toSeconds leap month - |>.div 60 + |>.ediv 60 /-- Transforms `Month.Ordinal` into `Hour.Offset`. -/ @[inline] def toHours (leap : Bool) (month : Ordinal) : Hour.Offset := - toMinute leap month - |>.div 60 + toMinutes leap month + |>.ediv 60 /-- Transforms `Month.Ordinal` into `Day.Offset`. @@ -177,7 +177,7 @@ Check if the day is valid in a month and a leap year. -/ @[inline] def Valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := - day ≤ daysWithoutProof leap month + day.val ≤ (daysWithoutProof leap month).val instance : Decidable (Valid leap month day) := dite (day ≤ daysWithoutProof leap month) isTrue isFalse diff --git a/src/Std/Time/Date/Unit/WeekOfYear.lean b/src/Std/Time/Date/Unit/WeekOfYear.lean index 65de02c2966a..e6249e528ea7 100644 --- a/src/Std/Time/Date/Unit/WeekOfYear.lean +++ b/src/Std/Time/Date/Unit/WeekOfYear.lean @@ -23,8 +23,7 @@ namespace WeekOfYear def Ordinal := Bounded.LE 1 53 deriving Repr, BEq, LE, LT -instance [Le 1 n] [Le n 53] : OfNat Ordinal n where - ofNat := Bounded.LE.mk (Int.ofNat n) (And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)) +instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (52 : Nat))) n) instance : Inhabited Ordinal where default := 1 diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index ee5332a2fe0a..8f72111aa6f6 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -52,7 +52,7 @@ def isLeap (y : Offset) : Bool := Forces the day to be on the valid range. -/ @[inline] -def Valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := +abbrev Valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := month.Valid year.isLeap day instance : Decidable (Valid year month day) := diff --git a/src/Std/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean index fab98d38e713..fe46aa6d5f87 100644 --- a/src/Std/Time/Date/WeekDate.lean +++ b/src/Std/Time/Date/WeekDate.lean @@ -12,11 +12,21 @@ namespace Std namespace Time open Internal +set_option linter.all true + /-- `WeekDate` represents a date using a combination of a week of the year and the year. -/ structure WeekDate where + + /-- + The year component of the date. It is represented as an `Offset` type from `Year`. + -/ year : Year.Offset + + /-- + The week of the year component. It is represented as an `Ordinal` type from `WeekOfYear`. + -/ week : WeekOfYear.Ordinal deriving Repr, BEq, Inhabited @@ -37,7 +47,7 @@ def fromDays (scalar : Day.Offset) : WeekDate := let year := totalDays / 365 let week := Bounded.LE.byEmod totalDays 365 (by decide) - |>.div 7 (by decide) + |>.ediv 7 (by decide) |>.add 1 { year := year, week := week } diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index b1cea7f83849..54b0af601981 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -18,11 +18,11 @@ Converts a `LocalDateTime` to a `Timestamp` -/ @[inline] def ofLocalDateTime (Local : LocalDateTime) : Timestamp := - Local.toTimestamp + Local.toUTCTimestamp /-- Converts a `Timestamp` to a `LocalDateTime` -/ @[inline] def toLocalDateTime (timestamp : Timestamp) : LocalDateTime := - LocalDateTime.ofTimestamp timestamp + LocalDateTime.ofUTCTimestamp timestamp diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean index 43de1f68fc41..975d93981ea4 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -13,39 +13,52 @@ namespace Std namespace Time open Internal +set_option linter.all true + /-- Date time format with Year, Month, Day, Hour, Minute, Seconds and Nanoseconds. -/ structure LocalDateTime where + + /-- + The `Date` component of a `LocalDate` + -/ date : LocalDate + + /-- + The `Time` component of a `LocalTime` + -/ time : LocalTime - deriving Repr, Inhabited + + deriving Inhabited namespace LocalDateTime /-- Converts a `LocalDateTime` into a `Std.Time.Timestamp` -/ -def toTimestamp (dt : LocalDateTime) : Timestamp := +def toUTCTimestamp (dt : LocalDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch - let second := dt.time.toSeconds - days.toSeconds + second + let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 + let nanos := nanos.val + dt.time.nano.val + Timestamp.ofNanoseconds (UnitVal.mk nanos) /-- Converts a UNIX `Timestamp` into a `LocalDateTime`. -/ -def ofTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do +def ofUTCTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do let leapYearEpoch := 11017 let daysPer400Y := 365 * 400 + 97 let daysPer100Y := 365 * 100 + 24 let daysPer4Y := 365 * 4 + 1 - let secs := stamp.toInt - let daysSinceEpoch := secs / 86400 + let nanos := stamp.toNanoseconds + let secs : Second.Offset := nanos.ediv 1000000000 + let daysSinceEpoch : Day.Offset := secs.ediv 86400 let boundedDaysSinceEpoch := daysSinceEpoch let mut rawDays := boundedDaysSinceEpoch - leapYearEpoch - let mut rem := Bounded.LE.byMod secs 86400 (by decide) + let mut rem := Bounded.LE.byMod secs.val 86400 (by decide) let ⟨remSecs, days⟩ := if h : rem.val ≤ -1 then @@ -58,10 +71,10 @@ def ofTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do (h, rawDays) let mut quadracentennialCycles := days / daysPer400Y; - let mut remDays := days % daysPer400Y; + let mut remDays := days.val % daysPer400Y.val; if remDays < 0 then - remDays := remDays + daysPer400Y + remDays := remDays + daysPer400Y.val quadracentennialCycles := quadracentennialCycles - 1 let mut centenialCycles := remDays / daysPer100Y; @@ -83,7 +96,7 @@ def ofTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do remDays := remDays - remYears * 365 - let mut year := 2000 + remYears + 4 * quadrennialCycles + 100 * centenialCycles + 400 * quadracentennialCycles + let mut year := 2000 + remYears + 4 * quadrennialCycles + 100 * centenialCycles + 400 * quadracentennialCycles.val let months := [31, 30, 31, 30, 31, 31, 30, 31, 30, 31, 31, 29]; let mut mon : Fin 13 := 0; @@ -104,50 +117,141 @@ def ofTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do pure (Month.Ordinal.ofNat (mon.val + 2) (by omega)) let second : Bounded.LE 0 59 := remSecs.emod 60 (by decide) - let minute : Bounded.LE 0 59 := (remSecs.div 60 (by decide)).emod 60 (by decide) - let hour : Bounded.LE 0 23 := remSecs.div 3600 (by decide) + let minute : Bounded.LE 0 59 := (remSecs.ediv 60 (by decide)).emod 60 (by decide) + let hour : Bounded.LE 0 23 := remSecs.ediv 3600 (by decide) + let nano : Bounded.LE 0 999999999 := Bounded.LE.byEmod nanos.val 1000000000 (by decide) return { date := LocalDate.clip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) - time := LocalTime.mk (hour.expandTop (by decide)) minute (second.expandTop (by decide)) 0 + time := LocalTime.ofValidHourMinuteSecondsNano (hour.expandTop (by decide)) minute (second.expandTop (by decide)) nano } +/-- +Add `Day.Offset` to a `LocalDateTime`. +-/ +@[inline] +def addDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := + { dt with date := dt.date.addDays days } + +/-- +Subtracts `Day.Offset` to a `LocalDateTime`. +-/ +@[inline] +def subDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := + { dt with date := dt.date.subDays days } + +/-- +Add `Month.Offset` to a `LocalDateTime`, it clips the day to the last valid day of that month. +-/ +def addMonthsClip (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := + { dt with date := dt.date.addMonthsClip months } + +/-- +Subtracts `Month.Offset` from a `LocalDateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subMonthsClip (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := + { dt with date := dt.date.subMonthsClip months } + +/-- +Add `Month.Offset` to a `LocalDateTime`, this function rolls over any excess days into the following +month. +-/ +def addMonthsRollOver (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := + { dt with date := dt.date.addMonthsRollOver months } + +/-- +Subtract `Month.Offset` from a `LocalDateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def subMonthsRollOver (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := + { dt with date := dt.date.subMonthsRollOver months } + +/-- +Add `Year.Offset` to a `LocalDateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def addYearsRollOver (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := + { dt with date := dt.date.addYearsRollOver years } + +/-- +Add `Year.Offset` to a `LocalDateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def addYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := + { dt with date := dt.date.addYearsClip years } + +/-- +Subtracts `Year.Offset` from a `LocalDateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def subYearsRollOver (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := + { dt with date := dt.date.subYearsRollOver years } + +/-- +Subtracts `Year.Offset` from a `LocalDateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := + { dt with date := dt.date.subYearsClip years } + +/-- +Get the current monotonic time. +-/ +def now : IO LocalDateTime := + ofUTCTimestamp <$> Timestamp.now + /-- Getter for the `Year` inside of a `LocalDateTime` -/ @[inline] def year (dt : LocalDateTime) : Year.Offset := dt.date.year + /-- Getter for the `Month` inside of a `LocalDateTime` -/ @[inline] def month (dt : LocalDateTime) : Month.Ordinal := dt.date.month + /-- Getter for the `Day` inside of a `LocalDateTime` -/ @[inline] def day (dt : LocalDateTime) : Day.Ordinal := dt.date.day + /-- Getter for the `Hour` inside of a `LocalDateTime` -/ @[inline] -def hour (dt : LocalDateTime) : Hour.Ordinal := - dt.time.hour +def hour (dt : LocalDateTime) : Hour.Ordinal dt.time.hour.fst := + dt.time.hour.snd + /-- Getter for the `Minute` inside of a `LocalDateTime` -/ @[inline] def minute (dt : LocalDateTime) : Minute.Ordinal := dt.time.minute + +/-- +Getter for the `Second` inside of a `LocalDateTime` +-/ +@[inline] +def second (dt : LocalDateTime) : Second.Ordinal dt.time.second.fst := + dt.time.second.snd + /-- Getter for the `Second` inside of a `LocalDateTime` -/ @[inline] -def second (dt : LocalDateTime) : Second.Ordinal := - dt.time.second +def nanosecond (dt : LocalDateTime) : Nanosecond.Ordinal := + dt.time.nano end LocalDateTime end Time diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 03263ee6e2a6..8f88ad231e30 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -12,36 +12,150 @@ namespace Std namespace Time open Internal +set_option linter.all true + /-- -Seconds since the UNIX Epoch. +Nanoseconds since the UNIX Epoch. -/ -def Timestamp := Second.Offset - deriving Repr, BEq, Inhabited +structure Timestamp where + + /-- + Second offset of the timestamp. + -/ + second : Second.Offset + + /-- + Nanosecond span that ranges from -999999999 and 999999999 + -/ + nano : Nanosecond.Span + + /-- + Proof that it's a valid timestamp. + -/ + proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) + deriving Repr + +instance : Inhabited Timestamp where + default := ⟨0, Bounded.LE.mk 0 (by decide), by decide⟩ namespace Timestamp /-- Get the current monotonic time. -/ -@[extern "lean_get_current_timestamp"] -protected opaque now : IO Timestamp +@[extern "lean_get_current_time"] +opaque now : IO Timestamp + +/-- +Transforms a `Timestamp` into a `Second.Offset` +-/ +@[inline] +def toSeconds (t : Timestamp) : Second.Offset := + t.second + +/-- +Negates the `Timestamp` +-/ +@[inline] +def neg (t : Timestamp) : Timestamp := by + refine ⟨-t.second, t.nano.neg, ?_⟩ + cases t.proof with + | inl n => exact Or.inr (n.imp Int.neg_le_neg Int.neg_le_neg) + | inr n => exact Or.inl (n.imp Int.neg_le_neg Int.neg_le_neg) + +/-- +Adds two timestamps. +-/ +def add (t₁ t₂ : Timestamp) : Timestamp := by + let diffSecs := t₁.second + t₂.second + let diffNano := t₁.nano.addBounds t₂.nano + + let (diffSecs, diffNano) : Second.Offset × Nanosecond.Span := + if h₀ : diffNano.val > 999999999 then + have diffNano := diffNano.truncateBottom h₀ |>.sub 999999999 + (diffSecs + 1, diffNano.expandBottom (by decide)) + else if h₁ : diffNano.val < -999999999 then + have diffNano := diffNano.truncateTop (Int.le_sub_one_of_lt h₁) |>.add 999999999 + (diffSecs - 1, diffNano.expandTop (by decide)) + else by + have h₀ := Int.le_sub_one_of_lt <| Int.not_le.mp h₀ + have h₁ := Int.le_sub_one_of_lt <| Int.not_le.mp h₁ + simp_all [Int.add_sub_cancel] + exact ⟨diffSecs, Bounded.mk diffNano.val (And.intro h₁ h₀)⟩ + + if h : diffSecs.val > 0 ∧ diffNano.val < 0 then + let truncated := diffNano.truncateTop (Int.le_sub_one_of_lt h.right) + let nano := truncated.addTop 1000000000 (by decide) + let proof₁ : 0 ≤ diffSecs - 1 := Int.le_sub_one_of_lt h.left + refine { second := UnitVal.mk (diffSecs.val - 1), nano, proof := ?_ } + simp [nano, Bounded.LE.addTop] + refine (Or.inl (And.intro proof₁ ?_)) + let h₃ := (Int.add_le_add_iff_left 1000000000).mpr diffNano.property.left + rw [Int.add_comm] + exact Int.le_trans (by decide) h₃ + else if h₁ : diffSecs.val < 0 ∧ diffNano.val > 0 then + let second := diffSecs.val + 1 + let truncated := diffNano.truncateBottom h₁.right + let nano := truncated.subBottom 1000000000 (by decide) + refine { second := UnitVal.mk second, nano, proof := ?_ } + simp [nano, truncated, Bounded.LE.subBottom, Bounded.LE.truncateBottom] + refine (Or.inr (And.intro ?_ ?_)) + · exact h₁.left + · let h₃ := Int.sub_le_sub_right diffNano.property.right 1000000000 + simp at h₃ + exact Int.le_trans h₃ (by decide) + else + let h := not_and.mp h + let h₁ := not_and.mp h₁ + refine ⟨diffSecs, diffNano, ?_⟩ + if h₂ : diffSecs.val > 0 then + exact Or.inl (And.intro (Int.le_of_lt h₂) (Int.not_lt.mp (h h₂))) + else if h₃ : diffSecs.val < 0 then + exact Or.inr (And.intro (Int.le_of_lt h₃) (Int.not_lt.mp (h₁ h₃))) + else + exact Int.le_total diffNano.val 0 |>.symm.imp (And.intro (Int.not_lt.mp h₃)) (And.intro (Int.not_lt.mp h₂)) /-- -A function to transform `Timestamp` to `Second.Offset` +Subtracts two `Timestamp`. -/ @[inline] -def toSeconds (tm : Timestamp) : Second.Offset := - tm +def sub (t₁ t₂ : Timestamp) : Timestamp := + t₁.add t₂.neg -instance : OfNat Timestamp n where - ofNat := UnitVal.ofNat n +/-- +Creates a new `Timestamp` out of `Second.Offset`. +-/ +@[inline] +def ofSeconds (s : Second.Offset) : Timestamp := by + refine ⟨s, ⟨0, by decide⟩, ?_⟩ + simp <;> exact Int.le_total s.val 0 |>.symm -instance : HAdd Timestamp Second.Offset Timestamp where - hAdd x y := UnitVal.add x y +/-- +Creates a new `Timestamp` out of `Second.Offset`. +-/ +@[inline] +def ofNanoseconds (s : Nanosecond.Offset) : Timestamp := by + refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ + cases Int.le_total s.val 0 + next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) + next n => exact Or.inl (And.intro (Int.ediv_nonneg n (by decide)) (Int.mod_nonneg 1000000000 n)) + where + mod_nonpos : ∀ {a : Int} (b : Int), (a ≤ 0) → (b ≥ 0) → 0 ≥ a.mod b + | .negSucc m, .ofNat n, _, _ => Int.neg_le_neg (Int.mod_nonneg (↑n) (Int.ofNat_le.mpr (Nat.zero_le (m + 1)))) + | 0, n, _, _ => Int.eq_iff_le_and_ge.mp (Int.zero_mod n) |>.left -instance : HSub Timestamp Second.Offset Timestamp where - hSub x y := UnitVal.sub x y +/-- +Converts a `Timestamp` from a `Nanosecond.Offset` +-/ +@[inline] +def toNanoseconds (tm : Timestamp) : Nanosecond.Offset := + let nanos := tm.toSeconds.mul 1000000000 + let nanos := nanos + (UnitVal.mk tm.nano.val) + nanos -end Timestamp -end Time -end Std +/-- +Adds a `Second.Offset` to a `Timestamp` +-/ +@[inline] +def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := + t.add (ofSeconds s) diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 4ae0cb2937d6..5226e9380062 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -5,126 +5,24 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Internal +import Std.Time.DateTime import Std.Time.Time namespace Std namespace Time open Internal -/-- -`Instant` represents a place in time with second and nanoseconds precision. --/ -structure Instant where - second : Second.Offset - nano : Nanosecond.Ordinal - valid : second.val ≥ 0 - deriving Repr +set_option linter.all true /-- -Time duration with nanosecond precision. This type allows negative duration. +Duration is just a period between two timestamps. -/ -structure Duration where - second : Second.Offset - nano : Nanosecond.Span - proof : second.val = 0 ∨ (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) - deriving Repr - -namespace Instant - -/-- -Get the current monotonic time. --/ -@[extern "lean_get_current_time"] -protected opaque now : IO Instant - -/-- -Gets the difference of two `Instant` in a `Duration`. --/ -def sub (t₁ t₂ : Instant) : Duration := - let nsec_diff := (t₁.nano).subBounds (t₂.nano) - let sec_diff := (t₁.second) - (t₂.second) - - if h₀ : sec_diff.val > 0 ∧ nsec_diff.val ≤ -1 then by - let truncated := nsec_diff.truncateTop h₀.right - let nano := truncated.addTop 1000000000 - let proof₁ : 0 ≤ sec_diff.val - 1 := Int.le_sub_one_of_lt h₀.left - refine { second := UnitVal.mk (sec_diff.val - 1), nano, proof := ?_ } - simp [nano, truncated, Bounded.LE.addTop, Bounded.LE.truncateTop] - refine Or.intro_right _ (Or.intro_left _ (And.intro proof₁ ?_)) - let h₃ := (Int.add_le_add_iff_left 1000000000).mpr nsec_diff.property.left - rw [Int.add_comm] - exact Int.le_trans (by decide) h₃ - else if h₁ : sec_diff.val < 0 ∧ nsec_diff.val ≥ 1 then by - let second := sec_diff.val + 1 - let truncated := nsec_diff.truncateBottom h₁.right - let nano := truncated.subBottom 1000000000 - refine { second := UnitVal.mk second, nano, proof := ?_ } - simp [nano, truncated, Bounded.LE.subBottom, Bounded.LE.truncateBottom] - refine Or.intro_right _ (Or.intro_right _ (And.intro ?_ ?_)) - · exact h₁.left - · let h₃ := Int.sub_le_sub_right nsec_diff.property.right 1000000000 - simp at h₃ - exact Int.le_trans h₃ (by decide) - else by - refine { second := sec_diff, nano := nsec_diff, proof := ?_ } - if h₄ : sec_diff.val > 0 then - let h₅ := Int.not_le.mp (not_and.mp h₀ h₄) - refine Or.intro_right _ (Or.intro_left _ (And.intro (Int.le_of_lt h₄) (Int.add_one_le_iff.mp h₅))) - else if h₅ : sec_diff.val < 0 then - let h₆ := Int.not_le.mp (not_and.mp h₁ h₅) - refine Or.intro_right _ (Or.intro_right _ (And.intro (Int.le_of_lt h₅) (Int.le_sub_one_of_lt h₆))) - else - let h₈ := Int.eq_iff_le_and_ge.mpr (And.intro (Int.not_lt.mp h₄) (Int.not_lt.mp h₅)) - exact Or.intro_left _ h₈ - -instance : HSub Instant Instant Duration where - hSub := Instant.sub - -/-- -Gets how much time elapsed since another `Instant` and returns a `Duration`. --/ -@[inline] -def since (instant : Instant) (since : Instant) : Duration := - Instant.sub since instant - -end Instant +def Duration := Timestamp namespace Duration /-- -Proof that for nano = 0 then second.val = 0 ∨ (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) -is valid. --/ -theorem nano_zero - {second : Second.Offset} - {nano : Nanosecond.Span} - (h : nano.val = 0) - : second.val = 0 ∨ (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) - := by - if h₄ : second.val > 0 then - exact Or.intro_right _ (Or.intro_left _ (And.intro (Int.le_of_lt h₄) (Int.eq_iff_le_and_ge.mp h).right)) - else if h₅ : second.val < 0 then - exact Or.intro_right _ (Or.intro_right _ (And.intro (Int.le_of_lt h₅) (Int.eq_iff_le_and_ge.mp h).left)) - else - let h₈ := Int.eq_iff_le_and_ge.mpr (And.intro (Int.not_lt.mp h₄) (Int.not_lt.mp h₅)) - exact Or.intro_left _ h₈ - -/-- -Returns a `Duration` representing the given number of second. --/ -def ofSeconds (second : Second.Offset) : Duration := - { second := second, nano := Bounded.LE.mk 0 (by decide), proof := nano_zero (by simp [Bounded.LE.mk]) } - -/-- -Returns a `Duration` representing the given number of minute. --/ -def ofMinutes (minute : Minute.Offset) : Duration := - { second := minute.toSeconds * 60, nano := Bounded.LE.mk 0 (by decide), proof := nano_zero (by simp [Bounded.LE.mk]) } - -/-- -Returns a `Duration` representing the given number of hour. +Calculates a `Duration` out of two `Timestamp`s. -/ -def ofHours (hour : Hour.Offset) : Duration := - { second := hour.toSeconds * 3600, nano := Bounded.LE.mk 0 (by decide), proof := nano_zero (by simp [Bounded.LE.mk]) } - -end Duration +def since (f : Timestamp) (t : Timestamp) : Duration := + t.sub f diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 2363f449d51d..1a1e7005e1f2 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -79,10 +79,7 @@ namespace LocalDate Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. -/ def fromAmericanDateString (input : String) : Except String LocalDate := do - let res ← Formats.AmericanDate.parseBuilder (λm d y => LocalDate.ofYearMonthDay y m d) input - match res with - | .some res => pure res - | none => .error "Invalid date." + Formats.AmericanDate.parseBuilder (λm d y => LocalDate.ofYearMonthDay y m d) input end LocalDate @@ -92,7 +89,7 @@ namespace LocalTime Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `LocalTime`. -/ def fromTime24Hour (input : String) : Except String LocalTime := - Formats.Time24Hour.parseBuilder (LocalTime.ofHourMinuteSeconds) input + Formats.Time24Hour.parseBuilder (λh m s => LocalTime.ofHourMinuteSeconds? h.snd m s.snd) input /-- Formats a `LocalTime` value into a 24-hour format string (`hh:mm:ss`). @@ -104,13 +101,13 @@ def toTime24Hour (input : LocalTime) : String := Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `LocalTime`. -/ def fromTime12Hour (input : String) : Except String LocalTime := - Formats.Time12Hour.parseBuilder (λh m s a => LocalTime.ofHourMinuteSeconds (HourMarker.toAbsoluteShift a h) m s) input + Formats.Time12Hour.parseBuilder (λh m s a => LocalTime.ofHourMinuteSeconds? (HourMarker.toAbsolute a h.snd) m s.snd) input /-- Formats a `LocalTime` value into a 12-hour format string (`hh:mm:ss aa`). -/ def toTime12Hour (input : LocalTime) : String := - Formats.Time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.val ≥ 12 then HourMarker.pm else HourMarker.am) + Formats.Time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.snd.val ≥ 12 then HourMarker.pm else HourMarker.am) end LocalTime @@ -183,19 +180,19 @@ def toLongDateFormatString (datetime : DateTime .GMT) : String := /-- Formats a `ZonedDateTime` value into an ISO8601 string. -/ -def toISO8601String (date : ZonedDateTime) : String := - Formats.ISO8601.format date.snd +def toISO8601String (date : DateTime tz) : String := + Formats.ISO8601.format date /-- Formats a `ZonedDateTime` value into an RFC822 format string. -/ -def toRFC822String (date : ZonedDateTime) : String := - Formats.RFC822.format date.snd +def toRFC822String (date : DateTime tz) : String := + Formats.RFC822.format date /-- Formats a `ZonedDateTime` value into an RFC850 format string. -/ -def toRFC850String (date : ZonedDateTime) : String := - Formats.RFC850.format date.snd +def toRFC850String (date : DateTime tz) : String := + Formats.RFC850.format date end DateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index b72351c9b8de..e59fbb6ae98b 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -13,69 +13,149 @@ import Lean.Data.Parsec namespace Std namespace Time open Internal - open Lean.Parsec.String open Lean.Parsec Lean LocalTime LocalDate TimeZone DateTime +set_option linter.all true + /-- The `Modifier` inductive type represents various formatting options for date and time components. These modifiers are typically used in formatting functions to generate human-readable date and time strings. - -- `YYYY`: Four-digit year (e.g., 2024). -- `YY`: Two-digit year (e.g., 24 for 2024). -- `MMMM`: Full month name (e.g., January, February). -- `MMM`: Abbreviated month name (e.g., Jan, Feb). -- `MM`: Two-digit month (e.g., 01 for January). -- `M`: One or two-digit month (e.g., 1 for January, 10 for October). -- `DD`: Two-digit day of the month (e.g., 01, 02). -- `D`: One or two-digit day of the month (e.g., 1, 2). -- `d`: One or two digit day of the month with space padding at the beggining (e.g. 1, 12). -- `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday). -- `EEE`: Abbreviated day of the week (e.g., Mon, Tue). -- `hh`: Two-digit hour in 24-hour format (e.g., 01, 02). -- `h`: One or two-digit hour in 24-hour format (e.g., 1, 2). -- `HH`: Two-digit hour in 12-hour format (e.g., 13, 14). -- `H`: One or two-digit hour in 12-hour format (e.g., 1, 2). -- `AA`: Uppercase AM/PM indicator (e.g., AM, PM). -- `aa`: Lowercase am/pm indicator (e.g., am, pm). -- `mm`: Two-digit minute (e.g., 01, 02). -- `m`: One or two-digit minute (e.g., 1, 2). -- `sss`: Three-digit milliseconds (e.g., 001, 202). -- `ss`: Two-digit second (e.g., 01, 02). -- `s`: One or two-digit second (e.g., 1, 2). -- `ZZZZZ`: Full timezone offset including hours and minutes (e.g., +03:00). -- `ZZZZ`: Timezone offset including hours and minutes without the colon (e.g., +0300). -- `ZZZ`: Like ZZZZ but with a special case "UTC" for UTC. -- `Z`: Like ZZZZZ but with a special case "Z" for UTC. -- `z`: Name of the time-zone like (Brasilia Standard Time). -/ inductive Modifier + /-- + `YYYY`: Four-digit year (e.g., 2024). + -/ | YYYY + + /-- + `YY`: Two-digit year (e.g., 24 for 2024). + -/ | YY + + /-- + `MMMM`: Full month name (e.g., January, February). + -/ | MMMM + + /-- + `MMM`: Abbreviated month name (e.g., Jan, Feb). + -/ | MMM + + /-- + `MM`: Two-digit month (e.g., 01 for January). + -/ | MM + + /-- + `M`: One or two-digit month (e.g., 1 for January, 10 for October). + -/ | M + + /-- + `DD`: Two-digit day of the month (e.g., 01, 02). + -/ | DD + + /-- + `D`: One or two-digit day of the month (e.g., 1, 2). + -/ | D + + /-- + `d`: One or two digit day of the month with space padding at the beggining (e.g. 1, 12). + -/ | d + + /-- + `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday). + -/ | EEEE + + /-- + `EEE`: Abbreviated day of the week (e.g., Mon, Tue). + -/ | EEE + + /-- + `hh`: Two-digit hour in 24-hour format (e.g., 01, 02). + -/ | hh + + /-- + `h`: One or two-digit hour in 24-hour format (e.g., 1, 2). + -/ | h + + /-- + `HH`: Two-digit hour in 12-hour format (e.g., 13, 14). + -/ | HH + + /-- + `H`: One or two-digit hour in 12-hour format (e.g., 1, 2). + -/ | H + + /-- + `AA`: Uppercase AM/PM indicator (e.g., AM, PM). + -/ | AA + + /-- + `aa`: Lowercase am/pm indicator (e.g., am, pm). + -/ | aa + + /-- + `mm`: Two-digit minute (e.g., 01, 02). + -/ | mm + + /-- + `m`: One or two-digit minute (e.g., 1, 2). + -/ | m + + /-- + `sss`: Three-digit milliseconds (e.g., 001, 202). + -/ | sss + + /-- + `ss`: Two-digit second (e.g., 01, 02). + -/ | ss + + /-- + `s`: One or two-digit second (e.g., 1, 2). + -/ | s + + /-- + `ZZZZZ`: Full timezone offset including hours and minutes (e.g., +03:00). + -/ | ZZZZZ + + /-- + `ZZZZ`: Timezone offset including hours and minutes without the colon (e.g., +0300). + -/ | ZZZZ + + /-- + `ZZZ`: Like ZZZZ but with a special case "UTC" for UTC. + -/ | ZZZ + + /-- + `Z`: Like ZZZZZ but with a special case "Z" for UTC. + -/ | Z + + /-- + `z`: Name of the time-zone like (Brasilia Standard Time). + -/ | z deriving Repr @@ -84,7 +164,9 @@ The part of a formatting string. a string is just a text and a modifier is in th 0 is the quantity of left pad and `T` the `Modifier`. -/ inductive FormatPart + /-- A string literal. -/ | string (val : String) + /-- A modifier that renders some data into text. -/ | modifier (modifier : Modifier) deriving Repr @@ -99,8 +181,13 @@ The format of date and time string. -/ abbrev FormatString := List FormatPart +/-- +If the format is aware of some timezone data it parses or if it parses any timezone. +-/ inductive Awareness + /-- The format only parses a single timezone. -/ | only : TimeZone → Awareness + /-- The format parses any timezone. -/ | any namespace Awareness @@ -109,7 +196,7 @@ instance : Coe TimeZone Awareness where coe := .only @[simp] -def type (x : Awareness) : Type := +private def type (x : Awareness) : Type := match x with | .any => ZonedDateTime | .only tz => DateTime tz @@ -119,14 +206,18 @@ instance : Inhabited (type aw) where simp [type] split <;> exact Inhabited.default -def getD (x : Awareness) (default : TimeZone) : TimeZone := +private def getD (x : Awareness) (default : TimeZone) : TimeZone := match x with | .any => default | .only tz => tz end Awareness +/-- +A specification on how to format a data or parse some string. +-/ structure Format (awareness : Awareness) where + /-- The format that is not aware of the timezone. -/ string : FormatString deriving Inhabited, Repr @@ -161,7 +252,7 @@ private def parseModifier : Parser Modifier <|> pstring "Z" *> pure .Z <|> pstring "z" *> pure .z -def isFormatStart : Char → Bool := Char.isAlpha +private def isFormatStart : Char → Bool := Char.isAlpha private def pnumber : Parser Nat := do let numbers ← manyChars digit @@ -272,32 +363,16 @@ private def formatPartWithDate (date : DateTime z) : FormatPart → String @[simp] private def SingleFormatType : Modifier → Type - | .YYYY => Year.Offset - | .YY => Year.Offset - | .MMMM => Month.Ordinal - | .MMM => Month.Ordinal - | .MM => Month.Ordinal - | .M => Month.Ordinal - | .DD => Day.Ordinal - | .D => Day.Ordinal - | .d => Day.Ordinal - | .EEEE => Weekday - | .EEE => Weekday - | .hh => Hour.Ordinal - | .h => Hour.Ordinal - | .HH => Hour.Ordinal - | .H => Hour.Ordinal - | .AA => HourMarker - | .aa => HourMarker - | .mm => Minute.Ordinal - | .m => Minute.Ordinal + | .YYYY | .YY => Year.Offset + | .MMMM | .MMM | .MM | .M => Month.Ordinal + | .DD | .D | .d => Day.Ordinal + | .EEEE | .EEE => Weekday + | .hh | .h | .HH | .H => Sigma Hour.Ordinal + | .AA | .aa => HourMarker + | .mm | .m => Minute.Ordinal | .sss => Millisecond.Ordinal - | .ss => Second.Ordinal - | .s => Second.Ordinal - | .ZZZZZ => Offset - | .ZZZZ => Offset - | .ZZZ => Offset - | .Z => Offset + | .ss | .s => Sigma Second.Ordinal + | .ZZZZZ | .ZZZZ | .ZZZ | .Z => Offset | .z => String private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) : String := @@ -313,17 +388,17 @@ private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) | .d => s!"{leftPad 2 ' ' <| toString data.toNat}" | .EEEE => dayOfWeek data | .EEE => abbrevDayOfWeek data - | .hh => s!"{leftPad 2 '0' (toString data.toNat)}" - | .h => s!"{data.toNat}" - | .HH => let hour := data.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' <| toString hour}" - | .H => let hour := data.val % 12; if hour == 0 then "12" else s!"{hour}" + | .hh => s!"{leftPad 2 '0' (toString data.snd.toNat)}" + | .h => s!"{data.snd.toNat}" + | .HH => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' <| toString hour}" + | .H => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{hour}" | .AA => match data with | .am => "AM" | .pm => "PM" | .aa => match data with | .am => "am" | .pm => "pm" | .mm => s!"{leftPad 2 '0' <| toString data.toNat}" | .m => s!"{data.toNat}" | .sss => s!"{leftPad 3 '0' <| toString data.toNat}" - | .ss => s!"{leftPad 2 '0' <| toString data.toNat}" - | .s => s!"{data.toNat}" + | .ss => s!"{leftPad 2 '0' <| toString data.snd.toNat}" + | .s => s!"{data.snd.toNat}" | .ZZZZZ => data.toIsoString true | .ZZZZ => data.toIsoString false | .ZZZ => if data.second.val = 0 then "UTC" else data.toIsoString false @@ -331,7 +406,7 @@ private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) | .z => data @[simp] -def FormatType (result : Type) : FormatString → Type +private def FormatType (result : Type) : FormatString → Type | .modifier entry :: xs => (SingleFormatType entry) → (FormatType result xs) | .string _ :: xs => (FormatType result xs) | [] => result @@ -456,21 +531,21 @@ private def parserWithFormat : (typ: Modifier) → Parser (SingleFormatType typ) | .d => transform Bounded.LE.ofInt (orElse twoDigit (λ_ => pchar ' ' *> (singleDigit))) | .EEEE => parseWeekdayUnnabrev | .EEE => parseWeekday - | .hh => transform Bounded.LE.ofInt twoDigit - | .h => transform Bounded.LE.ofInt number + | .hh => Sigma.mk true <$> transform Bounded.LE.ofInt twoDigit + | .h => Sigma.mk true <$> transform Bounded.LE.ofInt number | .HH => do - let res : Bounded.LE 1 13 ← transform Bounded.LE.ofInt twoDigit - return (res.sub 1).expandTop (by decide) + let res : Bounded.LE 0 12 ← transform Bounded.LE.ofInt twoDigit + return Sigma.mk true (res.expandTop (by decide)) | .H => do - let res : Bounded.LE 1 13 ← transform Bounded.LE.ofInt number - return (res.sub 1).expandTop (by decide) + let res : Bounded.LE 0 12 ← transform Bounded.LE.ofInt number + return Sigma.mk true (res.expandTop (by decide)) | .AA => parserUpperHourMarker | .aa => parserLowerHourMarker | .mm => transform Bounded.LE.ofInt twoDigit | .m => transform Bounded.LE.ofInt number | .sss => transform Bounded.LE.ofInt threeDigit - | .ss => transform Bounded.LE.ofInt twoDigit - | .s => transform Bounded.LE.ofInt number + | .ss => Sigma.mk true <$> transform Bounded.LE.ofInt twoDigit + | .s => Sigma.mk true <$> transform Bounded.LE.ofInt number | .ZZZZZ => timeOffset true | .ZZZZ => timeOffset false | .ZZZ => timeOrUTC "UTC" false @@ -483,23 +558,25 @@ private structure DateBuilder where year : Year.Offset := 0 month : Month.Ordinal := 1 day : Day.Ordinal := 1 - hour : Hour.Ordinal := 0 + hour : Sigma Hour.Ordinal := ⟨true, 0⟩ minute : Minute.Ordinal := 0 - second : Second.Ordinal := 0 + second : Sigma Second.Ordinal := ⟨true, 0⟩ millisecond : Millisecond.Ordinal := 0 -private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : aw.type := - let build := DateTime.ofLocalDateTime { +private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except String aw.type := + if let .isTrue p := inferInstanceAs (Decidable (ValidTime builder.hour.snd builder.minute builder.second.snd)) then + let build := DateTime.ofLocalDateTime { date := LocalDate.clip builder.year builder.month builder.day - time := LocalTime.mk builder.hour builder.minute builder.second (.ofMillisecond builder.millisecond) - } + time := LocalTime.mk builder.hour builder.minute builder.second (.ofMillisecond builder.millisecond) p + } - match aw with - | .any => - let tz₁ := TimeZone.mk builder.tz builder.tzName - ⟨tz₁, build tz₁⟩ - | .only tz => - build tz + match aw with + | .only tz => .ok (build tz) + | .any => + let tz₁ := TimeZone.mk builder.tz builder.tzName + .ok ⟨tz₁, build tz₁⟩ + else + .error "invalid leap seconds {} {} {}" private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : SingleFormatType typ) : DateBuilder := match typ with @@ -508,7 +585,7 @@ private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : Sin | .DD | .D | .d => { data with day := value } | .EEEE | .EEE => data | .hh | .h | .HH | .H => { data with hour := value } - | .AA | .aa => { data with hour := HourMarker.toAbsoluteShift value data.hour } + | .AA | .aa => { data with hour := ⟨data.hour.fst, HourMarker.toAbsolute value data.hour.snd⟩ } | .mm | .m => { data with minute := value } | .sss => { data with millisecond := value } | .ss | .s => { data with second := value } @@ -520,45 +597,11 @@ private def formatParser (date : DateBuilder) : FormatPart → Parser DateBuilde | .modifier mod => addDataInDateTime date mod <$> parserWithFormat mod | .string s => skipString s *> pure date --- API - namespace Format /-- Constructs a new `Format` specification for a date-time string. Modifiers can be combined to create custom formats, such as %YYYY, MMMM, D". - -### Supported Modifiers: -- `YYYY`: Four-digit year (e.g., 2024). -- `YY`: Two-digit year (e.g., 24 for 2024). -- `MMMM`: Full month name (e.g., January, February). -- `MMM`: Abbreviated month name (e.g., Jan, Feb). -- `MM`: Two-digit month (e.g., 01 for January). -- `M`: One or two-digit month (e.g., 1 for January, 10 for October). -- `DD`: Two-digit day of the month (e.g., 01, 02). -- `D`: One or two-digit day of the month (e.g., 1, 2). -- `d`: One or two-digit day of the month with left padding with spaces. (e.g., 1, 2). -- `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday). -- `EEE`: Abbreviated day of the week (e.g., Mon, Tue). -- `hh`: Two-digit hour in 12-hour format (e.g., 01, 02). -- `h`: One or two-digit hour in 12-hour format (e.g., 1, 2). -- `HH`: Two-digit hour in 24-hour format (e.g., 13, 14). -- `H`: One or two-digit hour in 24-hour format (e.g., 1, 2). -- `AA`: Uppercase AM/PM indicator (e.g., AM, PM). -- `aa`: Lowercase am/pm indicator (e.g., am, pm). -- `mm`: Two-digit minute (e.g., 01, 02). -- `m`: One or two-digit minute (e.g., 1, 2). -- `sss`: Three-digit millisecond (e.g., 001, 002). -- `ss`: Two-digit second (e.g., 01, 02). -- `s`: One or two-digit second (e.g., 1, 2). -- `ZZZZZ`: Full timezone offset with hours and minutes (e.g., +03:00). -- `ZZZZ`: Timezone offset with hours and minutes, without the colon (e.g., +0300). -- `ZZZ`: Like `ZZZZ`, but displays "UTC" for UTC time. -- `Z`: Like `ZZZZZ`, but displays "Z" for UTC time. -- `z`: Timezone name (e.g., Brasilia Standard Time). - -Example usage: -- `"YYYY-MM-DD HH:mm:ss ZZZZ"` → "2024-08-04 14:23:45 +0300" -/ def spec (input : String) : Except String (Format tz) := do let string ← specParser.run input @@ -601,29 +644,35 @@ def parser (format : FormatString) (aw : Awareness) : Parser (aw.type) := let rec go (date : DateBuilder) (x : FormatString) : Parser aw.type := match x with | x :: xs => formatParser date x >>= (go · xs) - | [] => pure (date.build aw) + | [] => + match date.build aw with + | .ok res => pure res + | .error err => fail err go {} format /-- Parser for a format with a builder. -/ -def builderParser (format: FormatString) (func: FormatType α format) : Parser α := - let rec go (format : FormatString) (func: FormatType α format) : Parser α := +def builderParser (format: FormatString) (func: FormatType (Option α) format) : Parser α := + let rec go (format : FormatString) (func: FormatType (Option α) format) : Parser α := match format with | .modifier x :: xs => do let res ← parserWithFormat x go xs (func res) | .string s :: xs => skipString s *> (go xs func) - | [] => pure func + | [] => + match func with + | some res => pure res + | none => fail "invalid date." go format func /-- Parses the input string into a `ZoneDateTime` -/ def parse (format : Format aw) (input : String) : Except String aw.type := - (parser format.string aw).run input + (parser format.string aw <* eof).run input -/- +/-- Parses the input string into a `ZoneDateTime`, panics if its wrong -/ def parse! (format : Format aw) (input : String) : aw.type := @@ -634,13 +683,13 @@ def parse! (format : Format aw) (input : String) : aw.type := /-- Parses and instead of using a builder to build a date, it uses a builder function instead. -/ -def parseBuilder (format : Format aw) (builder : FormatType α format.string) (input : String) : Except String α := +def parseBuilder (format : Format aw) (builder : FormatType (Option α) format.string) (input : String) : Except String α := (builderParser format.string builder).run input /-- Parses and instead of using a builder to build a date, it uses a builder function instead. -/ -def parseBuilder! [Inhabited α] (format : Format aw) (builder : FormatType α format.string) (input : String) : α := +def parseBuilder! [Inhabited α] (format : Format aw) (builder : FormatType (Option α) format.string) (input : String) : α := match parseBuilder format builder input with | .ok res => res | .error err => panic! err diff --git a/src/Std/Time/Internal.lean b/src/Std/Time/Internal.lean index 1ce45c0915d8..cf91f4cd6070 100644 --- a/src/Std/Time/Internal.lean +++ b/src/Std/Time/Internal.lean @@ -5,6 +5,5 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Internal.Bounded -import Std.Time.Internal.LessEq import Std.Time.Internal.Sign import Std.Time.Internal.UnitVal diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index 6c3c24917010..ca0f94b8463e 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -5,7 +5,6 @@ Authors: Sofia Rodrigues -/ prelude import Init.Data.Int -import Std.Time.Internal.LessEq namespace Std namespace Time @@ -47,12 +46,6 @@ integers that `lo ≤ val ≤ hi`. -/ abbrev LE := @Bounded LE.le -instance [Le lo n] [Le n hi] : OfNat (Bounded.LE lo hi) n where - ofNat := ⟨n, And.intro (Int.ofNat_le.mpr Le.p) (Int.ofNat_le.mpr Le.p)⟩ - -instance [Le lo hi] : Inhabited (Bounded.LE lo hi) where - default := ⟨lo, And.intro (Int.le_refl lo) (Int.ofNat_le.mpr Le.p)⟩ - def cast {rel : Int → Int → Prop} {lo₁ lo₂ hi₁ hi₂ : Int} (h₁ : lo₁ = lo₂) (h₂ : hi₁ = hi₂) (b : Bounded rel lo₁ hi₁) : Bounded rel lo₂ hi₂ := .mk b.val ⟨h₁ ▸ b.property.1, h₂ ▸ b.property.2⟩ @@ -72,6 +65,35 @@ def mk {rel : Int → Int → Prop} (val : Int) (proof : rel lo val ∧ rel val namespace LE +/-- +Convert a `Nat` to a `Bounded.LE` by wrapping it. +-/ +@[inline] +def ofNatWrapping { lo hi : Int } (val : Int) (h : lo ≤ hi): Bounded.LE lo hi := by + let range := hi - lo + 1 + have range_pos := Int.add_pos_of_nonneg_of_pos (b := 1) (Int.sub_nonneg_of_le h) (by decide) + have not_zero := Int.ne_iff_lt_or_gt.mpr (Or.inl range_pos) + have mod_nonneg : 0 ≤ (val - lo) % range := Int.emod_nonneg (val - lo) not_zero.symm + have add_nonneg : lo ≤ lo + (val - lo) % range := Int.le_add_of_nonneg_right mod_nonneg + have mod_range : (val - lo) % (hi - lo + 1) < range := Int.emod_lt_of_pos (a := val - lo) range_pos + refine ⟨((val - lo) % range + range) % range + lo, And.intro ?_ ?_⟩ + · simp_all [range] + rw [Int.add_comm] at add_nonneg + exact add_nonneg + · apply Int.add_le_of_le_sub_right + simp_all [range] + exact Int.le_of_lt_add_one mod_range + +instance {k : Nat} : OfNat (Bounded.LE lo (lo + k)) n where + ofNat := + let h : lo ≤ lo + k := Int.le_add_of_nonneg_right (Int.ofNat_zero_le k) + ofNatWrapping n h + +instance {k : Nat} : Inhabited (Bounded.LE lo (lo + k)) where + default := + let h : lo ≤ lo + k := Int.le_add_of_nonneg_right (Int.ofNat_zero_le k) + ofNatWrapping lo h + /-- Creates a new `Bounded` integer that the relation is less-equal. -/ @@ -123,17 +145,6 @@ def clip (val : Int) (h : lo ≤ hi) : Bounded.LE lo hi := else ⟨hi, And.intro h (Int.le_refl hi)⟩ else ⟨lo, And.intro (Int.le_refl lo) h⟩ -/-- -Convert a `Nat` to a `Bounded.LE` using the lower boundary too. --/ -@[inline] -def clip! [Le lo hi] (val : Int) : Bounded.LE lo hi := - if h₀ : lo ≤ val then - if h₁ : val ≤ hi - then ⟨val, And.intro h₀ h₁⟩ - else panic! "greater than hi" - else panic! "lower than lo" - /-- Convert a `Bounded.LE` to a Nat. -/ @@ -270,9 +281,9 @@ def add (bounded : Bounded.LE n m) (num : Int) : Bounded.LE (n + num) (m + num) Adjust the bounds of a `Bounded` by adding a constant value to the upper bounds. -/ @[inline] -def addTop (bounded : Bounded.LE n m) (num : Nat) : Bounded.LE n (m + num) := by +def addTop (bounded : Bounded.LE n m) (num : Int) (h : num ≥ 0) : Bounded.LE n (m + num) := by refine ⟨bounded.val + num, And.intro ?_ ?_⟩ - · let h := Int.add_le_add bounded.property.left (Int.ofNat_zero_le num) + · let h := Int.add_le_add bounded.property.left h simp at h exact h · exact Int.add_le_add bounded.property.right (Int.le_refl num) @@ -281,10 +292,10 @@ def addTop (bounded : Bounded.LE n m) (num : Nat) : Bounded.LE n (m + num) := by Adjust the bounds of a `Bounded` by adding a constant value to the lower bounds. -/ @[inline] -def subBottom (bounded : Bounded.LE n m) (num : Nat) : Bounded.LE (n - num) m := by +def subBottom (bounded : Bounded.LE n m) (num : Int) (h : num ≥ 0) : Bounded.LE (n - num) m := by refine ⟨bounded.val - num, And.intro ?_ ?_⟩ · exact Int.add_le_add bounded.property.left (Int.le_refl (-num)) - · let h := Int.sub_le_sub bounded.property.right (Int.ofNat_zero_le num) + · let h := Int.sub_le_sub bounded.property.right h simp at h exact h @@ -292,7 +303,7 @@ def subBottom (bounded : Bounded.LE n m) (num : Nat) : Bounded.LE (n - num) m := Adds two `Bounded` and adjust the boundaries. -/ @[inline] -def addBounds (bounded : Bounded.LE n m) (bounded₂ : Bounded.LE n m) : Bounded.LE (n + n) (m + m) := by +def addBounds (bounded : Bounded.LE n m) (bounded₂ : Bounded.LE i j) : Bounded.LE (n + i) (m + j) := by refine ⟨bounded.val + bounded₂.val, And.intro ?_ ?_⟩ · exact Int.add_le_add bounded.property.left bounded₂.property.left · exact Int.add_le_add bounded.property.right bounded₂.property.right @@ -349,9 +360,9 @@ def mul_neg (bounded : Bounded.LE n m) (num : Int) (h : num ≤ 0) : Bounded.LE Adjust the bounds of a `Bounded` by applying the div operation. -/ @[inline] -def div (bounded : Bounded.LE n m) (num : Int) (h : num > 0) : Bounded.LE (n / num) (m / num) := by +def ediv (bounded : Bounded.LE n m) (num : Int) (h : num > 0) : Bounded.LE (n / num) (m / num) := by let ⟨left, right⟩ := bounded.property - refine ⟨bounded.val / num, And.intro ?_ ?_⟩ + refine ⟨bounded.val.ediv num, And.intro ?_ ?_⟩ apply Int.ediv_le_ediv · exact h · exact left @@ -359,6 +370,9 @@ def div (bounded : Bounded.LE n m) (num : Int) (h : num > 0) : Bounded.LE (n / n · exact h · exact right +@[inline] +def eq {n : Int} : Bounded.LE n n := + ⟨n, And.intro (Int.le_refl n) (Int.le_refl n)⟩ /-- Expand the bottom of the bounded to a number `nhi` is `hi` is less or equal to the previous higher bound. -/ diff --git a/src/Std/Time/Internal/LessEq.lean b/src/Std/Time/Internal/LessEq.lean deleted file mode 100644 index c74f7d921731..000000000000 --- a/src/Std/Time/Internal/LessEq.lean +++ /dev/null @@ -1,34 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Init -import Lean.Data.Rat - -namespace Std -namespace Time -namespace Internal - -set_option linter.all true - -/-- -Type class to check if a number `n` is less or equal than `m` in compile time and give me a proof. --/ -class Le (n : Nat) (m : Nat) where - /-- The proof that n ≤ m. -/ - p : n ≤ m - -instance : Le n n where - p := Nat.le_refl n - -instance (m : Nat) : Le 0 m where - p := Nat.zero_le m - -instance {n m : Nat} [Le n m] : Le (Nat.succ n) (Nat.succ m) where - p := Nat.succ_le_succ (Le.p) - -end Internal -end Time -end Std diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 4f1ef3e71655..77ed47232939 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -60,8 +60,8 @@ def mul (unit : UnitVal a) (factor : Int) : UnitVal (a / factor) := Divides the `UnitVal` by an `Int`, resulting in a new `UnitVal` with an adjusted ratio. -/ @[inline] -def div (unit : UnitVal a) (divisor : Int) : UnitVal (a * divisor) := - ⟨unit.val / divisor⟩ +def ediv (unit : UnitVal a) (divisor : Int) : UnitVal (a * divisor) := + ⟨unit.val.ediv divisor⟩ /-- Adds two `UnitVal` values of the same ratio. diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 66ae77dfd35e..328ba2e46515 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -208,10 +208,6 @@ macro_rules let syn := Syntax.mkStrLit name do `(Std.Time.TimeZone.mk $(← parseOffset offset) $syn) -/-- -Macro for creating formats --/ - private def convertModifier : Modifier → MacroM (TSyntax `term) | .YYYY => `(Std.Time.Modifier.YYYY) | .YY => `(Std.Time.Modifier.YY) diff --git a/src/Std/Time/Time.lean b/src/Std/Time/Time.lean index 467a4fb35bb9..10d9082ee548 100644 --- a/src/Std/Time/Time.lean +++ b/src/Std/Time/Time.lean @@ -7,3 +7,5 @@ prelude import Std.Time.Time.Basic import Std.Time.Time.HourMarker import Std.Time.Time.LocalTime + +set_option linter.all true diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 9bee8ac07e3c..e22951ebc9cd 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -25,15 +25,15 @@ inductive HourMarker /-- Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker`. -/ -def HourMarker.toAbsolute (marker : HourMarker) (time : Hour.Ordinal) (h₀ : time.toInt ≤ 12) : Hour.Ordinal := +def HourMarker.toAbsolute (marker : HourMarker) (time : Hour.Ordinal l) : Hour.Ordinal l := match marker with | .am => time - | .pm => Bounded.LE.mk (time.toInt + 12) (And.intro (Int.add_le_add (c := 0) time.property.left (by decide)) (Int.add_le_add_right h₀ 12)) - -/-- -Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker` without a proof. --/ -def HourMarker.toAbsoluteShift (marker : HourMarker) (time : Hour.Ordinal) : Hour.Ordinal := - match marker with - | .am => time - | .pm => Bounded.LE.ofFin (Fin.ofNat (time.val.toNat + 12)) + | .pm => by + let mod := Int.mod_lt_of_pos (b := 24) (time.val + 12) (by decide) + let l := time.property.left + refine Bounded.LE.mk ((time.val + 12).mod 24) (And.intro ?_ ?_) + · have h : 0 + 0 ≤ time.val + 12 := Int.add_le_add l (by decide) + exact Int.mod_nonneg 24 h + · split + · exact Int.le_of_lt mod + · exact Int.le_sub_one_of_lt mod diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/LocalTime.lean index 84389d851989..bdd65b065869 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -8,40 +8,102 @@ import Std.Time.Time.Basic namespace Std namespace Time +open Internal + +set_option linter.all true /-- Represents a specific point in time, including hours, minutes, seconds, and nanoseconds. -/ structure LocalTime where - hour : Hour.Ordinal + /-- + `Hour` component of the `LocalTime` + -/ + hour : Sigma Hour.Ordinal + + /-- + `Minute` component of the `LocalTime` + -/ minute : Minute.Ordinal - second : Second.Ordinal + + /-- + `Second` component of the `LocalTime` + -/ + second : Sigma Second.Ordinal + + /-- + `Nanoseconds` component of the `LocalTime` + -/ nano : Nanosecond.Ordinal - deriving Repr, Inhabited, BEq + + /-- + The prove that if it includes a leap second than it needs to be exactly 23 : 59 : 60 + -/ + proof : (second.snd.val = 60 → hour.snd.val = 23 ∧ minute.val = 59) + ∧ (hour.snd.val = 24 → second.snd.val = 0 ∧ minute.val = 0) + +instance : Inhabited LocalTime where + default := ⟨Sigma.mk false 0, 0, Sigma.mk false 0, 0, by simp; decide⟩ namespace LocalTime +/-- +Checks if the hour is valid if it has a leap second or leap hour. +-/ +@[simp] +abbrev ValidTime (hour : Hour.Ordinal l) (minute : Minute.Ordinal) (second : Second.Ordinal l₁) : Prop := + (second.val = 60 → hour.val = 23 ∧ minute.val = 59) + ∧ (hour.val = 24 → second.val = 0 ∧ minute.val = 0) + +/-- +Creates a `LocalTime` value from hours, minutes, and seconds. +-/ +def ofHourMinuteSeconds (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (proof : ValidTime hour minute second) : LocalTime := + ⟨Sigma.mk leap₂ hour, minute, Sigma.mk leap second, 0, proof⟩ + +/-- +Creates a `LocalTime` value from hours, minutes, and seconds. Return `none` if its invalid. +-/ +def ofHourMinuteSeconds? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : Option LocalTime := + if h : ValidTime hour minute second + then some <| ofHourMinuteSeconds hour minute second h + else none + /-- Creates a `LocalTime` value from hours, minutes, and seconds, setting nanoseconds to zero. -/ -def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal) : LocalTime := - ⟨hour, minute, second, 0⟩ +def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Ordinal) (second : Second.Ordinal false) (nano : Nanosecond.Ordinal) : LocalTime := by + refine ⟨Sigma.mk false hour, minute, Sigma.mk false second, nano, ?_⟩ + constructor + exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr second.property.right)) x) + exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) + +/-- +Converts a `LocalTime` value to the total number of seconds since midnight. +-/ +def toNanoseconds (time : LocalTime) : Nanosecond.Offset := + let secs := + time.hour.snd.toOffset.toSeconds + + time.minute.toOffset.toSeconds + + time.second.snd.toOffset + let nanos := secs.mul 1000000000 + UnitVal.mk (nanos.val + time.nano.val) /-- Converts a `LocalTime` value to the total number of seconds since midnight. -/ def toSeconds (time : LocalTime) : Second.Offset := - time.hour.toOffset.toSeconds + + time.hour.snd.toOffset.toSeconds + time.minute.toOffset.toSeconds + - time.second.toOffset + time.second.snd.toOffset /-- Converts a `LocalTime` value to the total number of minutes since midnight. -/ def toMinutes (time : LocalTime) : Minute.Offset := - time.hour.toOffset.toMinutes + + time.hour.snd.toOffset.toMinutes + time.minute.toOffset + - time.second.toOffset.toMinutes + time.second.snd.toOffset.toMinutes end LocalTime end Time diff --git a/src/Std/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean index e1213290bb23..4d526bf0673e 100644 --- a/src/Std/Time/Time/Unit/Basic.lean +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -20,12 +20,12 @@ set_option linter.all true /-- Convert `Second.Offset` to `Minute.Offset` -/ @[inline] def toMinutes (offset : Second.Offset) : Minute.Offset := - offset.div 60 + offset.ediv 60 /-- Convert `Second.Offset` to `Hour.Offset` -/ @[inline] def toHours (offset : Second.Offset) : Hour.Offset := - offset.div 3600 + offset.ediv 3600 end Second.Offset @@ -34,7 +34,7 @@ namespace Minute.Offset /-- Convert `Minute.Offset` to `Hour.Offset` -/ @[inline] def toHours (offset : Minute.Offset) : Hour.Offset := - offset.div 60 + offset.ediv 60 end Minute.Offset diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 46834ed32358..5d73e8c10a6a 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -20,12 +20,19 @@ set_option linter.all true `Ordinal` represents a bounded value for hour, which ranges between 0 and 24. It is bounded to 24 because 24:00:00 is a valid date with leap seconds. -/ -def Ordinal := Bounded.LE 0 24 - deriving Repr, BEq, LE, LT +def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 24 else 23)) -instance [Le n 24] : OfNat Ordinal n where ofNat := Bounded.LE.ofNat n Le.p +instance : Repr (Ordinal l) where + reprPrec r l := reprPrec r.val l -instance : Inhabited Ordinal where default := 0 +instance : OfNat (Ordinal leap) n := by + have inst := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (23 : Nat))) n) + cases leap + · exact inst + · exact ⟨inst.ofNat.expandTop (by decide)⟩ + +instance : OfNat (Ordinal true) 24 where + ofNat := Bounded.LE.mk (Int.ofNat 24) (by decide) /-- `Offset` represents an offset in hour. It is defined as an `Int`. @@ -41,21 +48,23 @@ namespace Ordinal Creates an `Ordinal` from a natural number, ensuring the value is within bounds. -/ @[inline] -def ofNat (data : Nat) (h : data ≤ 24 := by decide) : Ordinal := +def ofNat (data : Nat) (h : data ≤ (if leap then 24 else 23)) : Ordinal leap := Bounded.LE.ofNat data h /-- Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. -/ @[inline] -def ofFin (data : Fin 25) : Ordinal := - Bounded.LE.ofFin data +def ofFin (data : Fin (if leap then 25 else 24)) : Ordinal leap := + match leap with + | true => Bounded.LE.ofFin data + | false => Bounded.LE.ofFin data /-- Converts an `Ordinal` to an `Offset`. -/ @[inline] -def toOffset (ordinal : Ordinal) : Offset := +def toOffset (ordinal : Ordinal leap) : Offset := UnitVal.ofInt ordinal.val end Ordinal diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 37ec556c6457..f71362f464c0 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -12,13 +12,15 @@ namespace Time namespace Millisecond open Internal +set_option linter.all true + /-- `Ordinal` represents a bounded value for milliseconds, which ranges between 0 and 999. -/ def Ordinal := Bounded.LE 0 999 deriving Repr, BEq, LE, LT -instance : OfNat Ordinal n where ofNat := Bounded.LE.ofFin (Fin.ofNat n) +instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (999 : Nat))) n) instance : Inhabited Ordinal where default := 0 diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index d889483c93dd..c9cc2cfd1df9 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -21,7 +21,7 @@ set_option linter.all true def Ordinal := Bounded.LE 0 59 deriving Repr, BEq, LE -instance [Le n 59] : OfNat Ordinal n where ofNat := Bounded.LE.ofNat n Le.p +instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) n) instance : Inhabited Ordinal where default := 0 diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index a970aebb4f91..f0d3195e68b6 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -37,7 +37,7 @@ def OfDay := Bounded.LE 0 86400000000000 Convert to `Millisecond.Ordinal` -/ def toMillisecond (nano : Ordinal) : Millisecond.Ordinal := - nano.div 1000000 (by decide) + nano.ediv 1000000 (by decide) /-- Convert from `Millisecond.Ordinal` diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 89a03a354825..c02270553e6d 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -19,12 +19,19 @@ set_option linter.all true `Ordinal` represents a bounded value for second, which ranges between 0 and 60. This accounts for potential leap second. -/ -def Ordinal := Bounded.LE 0 60 - deriving Repr, BEq, LE +def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 60 else 59)) -instance [Le n 60] : OfNat Ordinal n where ofNat := Bounded.LE.ofNat n Le.p +instance : Repr (Ordinal l) where + reprPrec r l := reprPrec r.val l -instance : Inhabited Ordinal where default := 0 +instance : OfNat (Ordinal leap) n := by + have inst := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) n) + cases leap + · exact inst + · exact ⟨inst.ofNat.expandTop (by decide)⟩ + +instance : OfNat (Ordinal true) 60 where + ofNat := Bounded.LE.mk (Int.ofNat 60) (by decide) /-- `Offset` represents an offset in second. It is defined as an `Int`. @@ -40,21 +47,23 @@ namespace Ordinal Creates an `Ordinal` from a natural number, ensuring the value is within bounds. -/ @[inline] -def ofNat (data : Nat) (h : data ≤ 60 := by decide) : Ordinal := +def ofNat (data : Nat) (h : data ≤ (if leap then 60 else 59)) : Ordinal leap := Bounded.LE.ofNat data h /-- Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. -/ @[inline] -def ofFin (data : Fin 61) : Ordinal := - Bounded.LE.ofFin data +def ofFin (data : Fin (if leap then 61 else 60)) : Ordinal leap := + match leap with + | true => Bounded.LE.ofFin data + | false => Bounded.LE.ofFin data /-- Converts an `Ordinal` to an `Offset`. -/ @[inline] -def toOffset (ordinal : Ordinal) : Offset := +def toOffset (ordinal : Ordinal leap) : Offset := UnitVal.ofInt ordinal.val end Ordinal diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 58cafcd02735..6c72e09ed09f 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -5,42 +5,10 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Zoned.DateTime +import Std.Time.Zoned.Database import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.ZonedDateTime +import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.Database.Basic -namespace Std -namespace Time -namespace TimeZone -namespace ZoneRules - -/-- -Parses a binary data into a zone rules. --/ -def parseTZif (bin : ByteArray) : Except String ZoneRules := do - let database ← Database.TZif.TZif.parse.run bin - Database.convertTZif database - -/-- -Gets the ZoneRules from the a TZIf file. --/ -def parseTZIfFromDisk (path : System.FilePath) : IO ZoneRules := do - let binary ← IO.FS.readBinFile path - IO.ofExcept (parseTZif binary) - -/-- -Gets the rules from local timezone. --/ -def localRules : IO ZoneRules := do - parseTZIfFromDisk "etc/localtime" - -/-- -Gets the rules from local timezone. --/ -def readRulesFromDisk (id : String) : IO ZoneRules := do - parseTZIfFromDisk (System.FilePath.join "/usr/share/zoneinfo/" id) - -end ZoneRules -end TimeZone -end Time -end Std +set_option linter.all true diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean new file mode 100644 index 000000000000..08c2856712e1 --- /dev/null +++ b/src/Std/Time/Zoned/Database.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +import Std.Time.Zoned.Database.Basic +import Std.Time.Zoned.Database.TZdb +import Std.Time.Zoned.ZonedDateTime +import Std.Time.Zoned.ZoneRules +import Std.Time.Zoned.TimeZone +import Std.Time.Zoned.DateTime +import Std.Time.DateTime + +namespace Std +namespace Time +namespace Database +open TimeZone.ZoneRules + +/-- +Apply leap seconds rules and transition rules on a UTC Timestamp to make it aware of the timezone. +-/ +def applyLeapSecondsOnUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO Timestamp := do + (applyLeapSeconds tm ·.leapSeconds) <$> Database.localRules db + +/-- +Gets the TimeZone at the local timezone. +-/ +def getLocalTimeZoneAt [Database α] (db : α) (tm : Timestamp) : IO TimeZone := do + (IO.ofExcept <| timezoneAt · tm) =<< Database.localRules db + +/-- +Gets the TimeZone at a timezone. +-/ +def getTimeZoneAt [Database α] (db : α) (id : String) (tm : Timestamp) : IO TimeZone := do + (IO.ofExcept <| timezoneAt · tm) =<< Database.load db id + +/-- +Get the local ZonedDataTime given a UTC `Timestamp`. +-/ +def ofTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime := do + let rules ← Database.localRules db + let tz ← IO.ofExcept <| timezoneAt rules tm + let tm := applyLeapSeconds tm rules.leapSeconds + return ZonedDateTime.ofTimestamp tm tz diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index e5a3587761ca..82ecc7c1850e 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -10,8 +10,24 @@ import Std.Time.DateTime namespace Std namespace Time + +set_option linter.all true + +/-- +A timezone database that we can read the `ZoneRules` of some area by it's id. +-/ +class Database (α : Type) where + /-- + Loads a `ZoneRules` by it's id. + -/ + load : α → String → IO TimeZone.ZoneRules + + /-- + Loads a `ZoneRules` that is setted by the local computer. + -/ + localRules : α → IO TimeZone.ZoneRules + namespace TimeZone -namespace Database /-- Converts a Boolean value to a corresponding `StdWall` type. @@ -36,11 +52,11 @@ def convertLeapSecond (tz : TZif.LeapSecond) : LeapSecond := /-- Converts a LocalTime. -/ -def convetLocalTime (index : Nat) (tz : TZif.TZifV1) : Option LocalTimeType := do +def convetLocalTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do let localType ← tz.localTimeTypes.get? index - let abbreviation ← tz.abbreviations.get? index - let wallflag ← convertWall <$> tz.stdWallIndicators.get? index - let utLocal ← convertUt <$> tz.utLocalIndicators.get? index + let abbreviation ← tz.abbreviations.getD index "Unknown" + let wallflag := convertWall (tz.stdWallIndicators.getD index true) + let utLocal := convertUt (tz.utLocalIndicators.getD index true) return { gmtOffset := Offset.ofSeconds <| Internal.UnitVal.mk localType.gmtOffset @@ -48,6 +64,7 @@ def convetLocalTime (index : Nat) (tz : TZif.TZifV1) : Option LocalTimeType := d abbreviation wall := wallflag utLocal + identifier } /-- @@ -62,11 +79,11 @@ def convertTransition (times: Array LocalTimeType) (index : Nat) (tz : TZif.TZif /-- Converts a `TZif.TZifV1` structure to a `ZoneRules` structure. -/ -def convertTZifV1 (tz : TZif.TZifV1) : Except String ZoneRules := do +def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := do let mut times : Array LocalTimeType := #[] for i in [0:tz.header.typecnt.toNat] do - if let some result := convetLocalTime i tz + if let some result := convetLocalTime i tz id then times := times.push result else .error s!"cannot convert local time {i} of the file" @@ -83,13 +100,13 @@ def convertTZifV1 (tz : TZif.TZifV1) : Except String ZoneRules := do /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. -/ -def convertTZifV2 (tz : TZif.TZifV2) : Except String ZoneRules := do - convertTZifV1 tz.toTZifV1 +def convertTZifV2 (tz : TZif.TZifV2) (id : String) : Except String ZoneRules := do + convertTZifV1 tz.toTZifV1 id /-- Converts a `TZif.TZif` structure to a `ZoneRules` structure. -/ -def convertTZif (tz : TZif.TZif) : Except String ZoneRules := do +def convertTZif (tz : TZif.TZif) (id : String) : Except String ZoneRules := do if let some v2 := tz.v2 - then convertTZifV2 v2 - else convertTZifV1 tz.v1 + then convertTZifV2 v2 id + else convertTZifV1 tz.v1 id diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean new file mode 100644 index 000000000000..d42eb346b19e --- /dev/null +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -0,0 +1,76 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +import Std.Time.Zoned.Database.Basic +import Std.Time.Zoned.ZoneRules +import Std.Time.Zoned.TimeZone +import Std.Time.DateTime + +namespace Std +namespace Time +namespace Database + +/-- +Reads the TZDb from the system from the TZDb database usually available in ""/usr/share/zoneinfo/".", +-/ +structure TZdb where + localPath : System.FilePath := "/etc/localtime" + zonesPath : System.FilePath := "/usr/share/zoneinfo/" + +namespace TZdb +open TimeZone + +/-- +Returns a default TZdb with all the fields completed with the default location for timezones in most +linux distributions and MacOS. +-/ +@[inline] +def default : TZdb := {} + +/-- +Parses a binary data into a zone rules. +-/ +def parseTZif (bin : ByteArray) (id : String) : Except String ZoneRules := do + let database ← TZif.parse.run bin + convertTZif database id + +/-- +Gets the ZoneRules from the a TZIf file. +-/ +def parseTZIfFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := do + let binary ← IO.FS.readBinFile path + IO.ofExcept (parseTZif binary id) + +/-- +Creates an id out of a `FilePath`, e.g, "/var/db/timezone/zoneinfo/America/Sao_Paulo" returns +"America/Sao_Paulo". +-/ +def idFromPath (path : System.FilePath) : Option String := do + let res := path.components.toArray + let last ← res.get? (res.size - 1) + let last₁ ← res.get? (res.size - 2) + + if last₁ = some "zoneinfo" + then last + else last₁ ++ "/" ++ last +/-- +Gets the rules from local +-/ +def localRules (path : System.FilePath) : IO ZoneRules := do + let localtimePath ← IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } + + if let some id := idFromPath localtimePath + then parseTZIfFromDisk path id + else throw (IO.userError "cannot read the id of the path.") + +/-- +Gets the rules from local +-/ +def readRulesFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := do + parseTZIfFromDisk (System.FilePath.join path id) id + +instance : Database TZdb where + load db id := readRulesFromDisk db.zonesPath id + localRules db := localRules db.localPath diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index 09388565c9fd..e6ff6881b684 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -11,23 +11,52 @@ import Lean.Data.Parsec.ByteArray namespace Std namespace Time namespace TimeZone -namespace Database namespace TZif open Lean.Parsec Lean.Parsec.ByteArray -abbrev Int32 := Int -abbrev Int64 := Int +set_option linter.all true + +private abbrev Int32 := Int + +private abbrev Int64 := Int /-- Represents the header of a TZif file, containing metadata about the file's structure. -/ structure Header where + /-- + The version of the TZif file format. + -/ version : UInt8 + + /-- + The count of UT local indicators in the file. + -/ isutcnt : UInt32 + + /-- + The count of standard/wall indicators in the file. + -/ isstdcnt : UInt32 + + /-- + The number of leap second records. + -/ leapcnt : UInt32 + + /-- + The number of transition times in the file. + -/ timecnt : UInt32 + + /-- + The number of local time types in the file. + -/ typecnt : UInt32 + + /-- + The total number of characters used in abbreviations. + -/ charcnt : UInt32 deriving Repr, Inhabited @@ -35,8 +64,19 @@ structure Header where Represents the local time type information, including offset and daylight saving details. -/ structure LocalTimeType where + /-- + The GMT offset in seconds for this local time type. + -/ gmtOffset : Int32 + + /-- + Indicates if this local time type observes daylight saving time. + -/ isDst : Bool + + /-- + The index into the abbreviation string table for this time type. + -/ abbreviationIndex : UInt8 deriving Repr, Inhabited @@ -44,7 +84,14 @@ structure LocalTimeType where Represents a leap second record, including the transition time and the correction applied. -/ structure LeapSecond where + /-- + The transition time of the leap second event. + -/ transitionTime : Int64 + + /-- + The correction applied during the leap second event in seconds. + -/ correction : Int64 deriving Repr, Inhabited @@ -52,13 +99,44 @@ structure LeapSecond where Represents version 1 of the TZif format. -/ structure TZifV1 where + /-- + The header information of the TZif file. + -/ header : Header + + /-- + The array of transition times in seconds since the epoch. + -/ transitionTimes : Array Int32 + + /-- + The array of local time type indices corresponding to each transition time. + -/ transitionIndices : Array UInt8 + + /-- + The array of local time type structures. + -/ localTimeTypes : Array LocalTimeType + + /-- + The array of abbreviation strings used by local time types. + -/ abbreviations : Array String + + /-- + The array of leap second records. + -/ leapSeconds : Array LeapSecond + + /-- + The array indicating whether each transition time uses wall clock time or standard time. + -/ stdWallIndicators : Array Bool + + /-- + The array indicating whether each transition time uses universal time or local time. + -/ utLocalIndicators : Array Bool deriving Repr, Inhabited @@ -66,6 +144,9 @@ structure TZifV1 where Represents version 2 of the TZif format, extending TZifV1 with an optional footer. -/ structure TZifV2 extends TZifV1 where + /-- + An optional footer for additional metadata in version 2. + -/ footer : Option String deriving Repr, Inhabited @@ -73,13 +154,17 @@ structure TZifV2 extends TZifV1 where Represents a TZif file, which can be either version 1 or version 2. -/ structure TZif where + /-- + The data for version 1 of the TZif file. + -/ v1 : TZifV1 + + /-- + Optionally, the data for version 2 of the TZif file. + -/ v2 : Option TZifV2 deriving Repr, Inhabited -namespace TZif -open Lean.Parsec Lean.Parsec.ByteArray - private def toUInt32 (bs : ByteArray) : UInt32 := assert! bs.size == 4 (bs.get! 0).toUInt32 <<< 0x18 ||| diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index d011a292da24..cfbbe8721a96 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -11,12 +11,23 @@ import Std.Time.Zoned.ZoneRules namespace Std namespace Time +set_option linter.all true + /-- It stores a `Timestamp`, a `LocalDateTime` and a `TimeZone` -/ structure DateTime (tz : TimeZone) where private mk :: + + /-- + `Timestamp` represents the exact moment in time. + -/ timestamp : Timestamp + + /-- + `Date` is a `Thunk` containing the `LocalDateTime` that represents the local date and time, it's + used for accessing data like `day` and `month` without having to recompute the data everytime. + -/ date : Thunk LocalDateTime instance : Inhabited (DateTime tz) where @@ -29,7 +40,7 @@ Creates a new `DateTime` out of a `Timestamp` -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := - DateTime.mk tm (Thunk.mk <| λ_ => (tm + tz.toSeconds).toLocalDateTime) + DateTime.mk tm (Thunk.mk <| λ_ => (tm.addSeconds tz.toSeconds).toLocalDateTime) /-- Creates a new `Timestamp` out of a `DateTime` @@ -50,8 +61,88 @@ Creates a new DateTime out of a `LocalDateTime` -/ @[inline] def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : DateTime tz := - let tm := date.toTimestamp - tz.toSeconds - DateTime.mk tm date + DateTime.ofTimestamp date.toUTCTimestamp tz + +/-- +Gets the current `DateTime`. +-/ +@[inline] +def now : IO (DateTime tz) := do + let loca ← LocalDateTime.now + return ofLocalDateTime loca tz + + +/-- +Add `Day.Offset` to a `DateTime`. +-/ +@[inline] +def addDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addDays days) tz + +/-- +Subtracts `Day.Offset` to a `DateTime`. +-/ +@[inline] +def subDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subDays days) tz + +/-- +Add `Month.Offset` to a `DateTime`, it clips the day to the last valid day of that month. +-/ +def addMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addMonthsClip months) tz + +/-- +Subtracts `Month.Offset` from a `DateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subMonthsClip months) tz + +/-- +Add `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following +month. +-/ +def addMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addMonthsRollOver months) tz + +/-- +Subtract `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def subMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subMonthsRollOver months) tz + +/-- +Add `Year.Offset` to a `DateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def addYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addYearsRollOver years) tz + +/-- +Add `Year.Offset` to a `DateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def addYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addYearsClip years) tz + +/-- +Subtract `Year.Offset` from a `DateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def subYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subYearsRollOver years) tz + +/-- +Subtract `Year.Offset` from to a `DateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subYearsClip years) tz /-- Getter for the `Year` inside of a `DateTime` @@ -78,7 +169,7 @@ def day (dt : DateTime tz) : Day.Ordinal := Getter for the `Hour` inside of a `DateTime` -/ @[inline] -def hour (dt : DateTime tz) : Hour.Ordinal := +def hour (dt : DateTime tz) : Hour.Ordinal dt.date.get.time.hour.fst := dt.date.get.hour /-- @@ -92,7 +183,7 @@ def minute (dt : DateTime tz) : Minute.Ordinal := Getter for the `Second` inside of a `DateTime` -/ @[inline] -def second (dt : DateTime tz) : Second.Ordinal := +def second (dt : DateTime tz) : Second.Ordinal dt.date.get.time.second.fst := dt.date.get.second /-- @@ -101,6 +192,7 @@ Getter for the `Milliseconds` inside of a `DateTime` @[inline] def milliseconds (dt : DateTime tz) : Millisecond.Ordinal := dt.date.get.time.nano.toMillisecond + /-- Gets the `Weekday` of a DateTime. -/ diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index b8dc1d33184e..0e8df17fb8d4 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -38,8 +38,8 @@ and minute components are separated by a colon (e.g., "+01:00" or "+0100"). def toIsoString (offset : Offset) (colon : Bool) : String := let sign := Sign.ofInt offset.second.val let time := offset.second.apply sign - let hour : Hour.Offset := time.div 3600 - let minute := Int.div (Int.mod time.val 3600) 60 + let hour : Hour.Offset := time.ediv 3600 + let minute := Int.ediv (Int.mod time.val 3600) 60 let hourStr := if hour.val < 10 then s!"0{hour.val}" else toString hour.val let minuteStr := if minute < 10 then s!"0{minute}" else toString minute if colon then s!"{sign}{hourStr}:{minuteStr}" diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 214d9632c65d..c484773de50d 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -15,11 +15,20 @@ namespace Time namespace TimeZone open Internal +set_option linter.all true + /-- Represents the type of local time in relation to UTC. -/ inductive UTLocal + /-- + Universal Time (UT), often referred to as UTC. + -/ | ut + + /-- + Local time that is not necessarily UTC. + -/ | local deriving Repr, Inhabited @@ -27,7 +36,14 @@ inductive UTLocal Represents types of wall clocks or standard times. -/ inductive StdWall + /-- + Time based on a wall clock, which can include daylight saving adjustments. + -/ | wall + + /-- + Standard time without adjustments for daylight saving. + -/ | standard deriving Repr, Inhabited @@ -35,11 +51,35 @@ inductive StdWall Represents a type of local time, including offset and daylight saving information. -/ structure LocalTimeType where + /-- + The offset from GMT for this local time. + -/ gmtOffset : TimeZone.Offset + + /-- + Indicates if daylight saving time is observed. + -/ isDst : Bool + + /-- + The abbreviation for this local time type (e.g., "EST", "PDT"). + -/ abbreviation : String + + /-- + Indicates if the time is wall clock or standard time. + -/ wall : StdWall + + /-- + Distinguishes between universal time and local time. + -/ utLocal : UTLocal + + /-- + ID of the timezone + -/ + identifier : String deriving Repr, Inhabited namespace LocalTimeType @@ -56,7 +96,14 @@ end LocalTimeType Represents a leap second event, including the time of the transition and the correction applied. -/ structure LeapSecond where + /-- + The time when the leap second transition occurs. + -/ transitionTime : Second.Offset + + /-- + The correction applied during the leap second event. + -/ correction : Second.Offset deriving Repr, Inhabited @@ -64,7 +111,14 @@ structure LeapSecond where Represents a time zone transition, mapping a time to a local time type. -/ structure Transition where + /-- + The specific time of the transition event. + -/ time : Second.Offset + + /-- + The local time type associated with this transition. + -/ localTimeType : LocalTimeType deriving Repr, Inhabited @@ -72,11 +126,42 @@ structure Transition where Represents the rules for a time zone, abstracting away binary data and focusing on key transitions and types. -/ structure ZoneRules where + /-- + The array of local time types for the time zone. + -/ localTimes : Array LocalTimeType + + /-- + The array of transitions for the time zone. + -/ transitions : Array Transition + + /-- + The array of leap seconds affecting the time zone. + -/ leapSeconds : Array LeapSecond deriving Repr, Inhabited +namespace Transition + +/-- +Create a TimeZone from a Transition. +-/ +def createTimeZoneFromTransition (transition : Transition) : TimeZone := + let offset := transition.localTimeType.gmtOffset + let name := transition.localTimeType.abbreviation + TimeZone.mk offset name + +/-- +Applies the transition to a Timestamp. +-/ +def apply (timestamp : Timestamp) (transition : Transition) : Timestamp := + let offsetInSeconds := transition.localTimeType.gmtOffset.hour.mul 3600 |>.add transition.localTimeType.gmtOffset.second + let localTimestamp := timestamp.addSeconds offsetInSeconds + localTimestamp + +end Transition + namespace ZoneRules /-- @@ -84,17 +169,35 @@ Finds the transition corresponding to a given timestamp in `ZoneRules`. If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. -/ def findTransitionForTimestamp (zoneRules : ZoneRules) (timestamp : Timestamp) : Option Transition := - if let some idx := zoneRules.transitions.findIdx? (λ t => t.time.val > timestamp.val) + let value := timestamp.toSeconds + if let some idx := zoneRules.transitions.findIdx? (λ t => t.time.val > value.val) then zoneRules.transitions.get? (idx - 1) else zoneRules.transitions.back? /-- -Apply leap seconds to a Timestamp +Find the current `TimeZone` out of a `Transition` in a `ZoneRules` +-/ +def timezoneAt (rules : ZoneRules) (tm : Timestamp) : Except String TimeZone := + if let some transition := rules.findTransitionForTimestamp tm + then .ok transition.createTimeZoneFromTransition + else .error "cannot find local timezone." + +/-- +Applies a `LeapSecond` sequence in a `Timestamp`. +-/ +def applyLeapSeconds (tm : Timestamp) (leapSeconds : Array LeapSecond) : Timestamp := Id.run do + let mut currentTime := tm + for i in [:leapSeconds.size] do + let leapSec := leapSeconds.get! i + if currentTime.second.val >= leapSec.transitionTime.val then + currentTime := tm.addSeconds (UnitVal.mk leapSec.correction.val) + return currentTime + +/-- +Adjust a UTC timestamp according to `ZoneRules`. -/ -def applyLeapSeconds (timeInSeconds : Timestamp) (leapSeconds : Array LeapSecond) : Timestamp := Id.run do - let mut currentTime := timeInSeconds - for i in [:leapSeconds.size] do - let leapSec := leapSeconds.get! i - if currentTime.val >= leapSec.transitionTime.val then - currentTime := ⟨currentTime.val + leapSec.correction.val⟩ - return currentTime +def apply (rules : ZoneRules) (tm : Timestamp) : Timestamp := + let tm := applyLeapSeconds tm rules.leapSeconds + if let some transition := findTransitionForTimestamp rules tm + then transition.apply tm + else tm diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 21f71d1879ac..45114c192ded 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -49,7 +49,6 @@ def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDate let transition ← rules.findTransitionForTimestamp tm return ofTimestamp tm transition.localTimeType.getTimeZone - /-- Changes the `TimeZone` to a new one. -/ @@ -64,6 +63,14 @@ Creates a new `ZonedDateTime` out of a `LocalDateTime` def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : ZonedDateTime := ⟨tz, DateTime.ofLocalDateTime date tz⟩ +/-- +Gets the current `ZonedDataTime`. +-/ +@[inline] +def now (tz : TimeZone) : IO ZonedDateTime := do + let loca ← LocalDateTime.now + return ofLocalDateTime loca tz + /-- Getter for the `Year` inside of a `ZonedDateTime` -/ @@ -89,8 +96,8 @@ def day (zdt : ZonedDateTime) : Day.Ordinal := Getter for the `Hour` inside of a `ZonedDateTime` -/ @[inline] -def hour (zdt : ZonedDateTime) : Hour.Ordinal := - zdt.snd.hour +def hour (zdt : ZonedDateTime) : Hour.Ordinal zdt.snd.date.get.time.hour.fst := + zdt.snd.date.get.time.hour.snd /-- Getter for the `Minute` inside of a `ZonedDateTime` @@ -103,8 +110,8 @@ def minute (zdt : ZonedDateTime) : Minute.Ordinal := Getter for the `Second` inside of a `ZonedDateTime` -/ @[inline] -def second (zdt : ZonedDateTime) : Second.Ordinal := - zdt.snd.second +def second (zdt : ZonedDateTime) : Second.Ordinal zdt.snd.date.get.time.second.fst := + zdt.snd.date.get.time.second.snd /-- Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` @@ -112,6 +119,7 @@ Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` @[inline] def offset (zdt : ZonedDateTime) : TimeZone.Offset := zdt.fst.offset + /-- Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` -/ @@ -126,6 +134,49 @@ Returns the weekday. def weekday (zdt : ZonedDateTime) : Weekday := zdt.snd.weekday +/-- +Add `Month.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. +-/ +def addMonthsClip (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addMonthsClip months) + +/-- +Subtracts `Month.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subMonthsClip (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subMonthsClip months) + +/-- +Add `Month.Offset` to a `ZonedDateTime`, this function rolls over any excess days into the following +month. +-/ +def addMonthsRollOver (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addMonthsRollOver months) + +/-- +Add `Month.Offset` to a `ZonedDateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def subMonthsRollOver (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subMonthsRollOver months) + +/-- +Add `Year.Offset` to a `ZonedDateTime`, this function rolls over any excess days into the following +month. +-/ +@[inline] +def addYearsRollOver (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addYearsRollOver years) + +/-- +Add `Year.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def addYearsClip (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addYearsClip years) + end ZonedDateTime end Time end Std diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 3cdd7a778362..035c5a600593 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -528,27 +528,23 @@ extern "C" LEAN_EXPORT obj_res lean_io_prim_handle_put_str(b_obj_arg h, b_obj_ar } } -extern "C" LEAN_EXPORT lean_obj_res lean_get_current_time() { - std::chrono::steady_clock::time_point now = std::chrono::steady_clock::now(); - std::chrono::steady_clock::duration duration = now.time_since_epoch(); +/* Std.Time.Timestamp.now : IO Timestamp */ +extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { + using namespace std::chrono; - long long secs = std::chrono::duration_cast(duration).count(); - long long nano = std::chrono::duration_cast(duration).count() % 1000000000; + std::chrono::system_clock::time_point now = std::chrono::system_clock::now(); + long long timestamp = std::chrono::duration_cast(now.time_since_epoch()).count(); + + long long secs = timestamp / 1000000000; + long long nano = timestamp % 1000000000; - lean_object *lean_ts = lean_alloc_ctor(0, 2, 0); + lean_object *lean_ts = lean_alloc_ctor(1, 2, 0); lean_ctor_set(lean_ts, 0, lean_int_to_int(static_cast(secs))); lean_ctor_set(lean_ts, 1, lean_int_to_int(static_cast(nano))); return lean_io_result_mk_ok(lean_ts); } -extern "C" LEAN_EXPORT lean_obj_res lean_get_current_timestamp() { - std::chrono::system_clock::time_point now = std::chrono::system_clock::now(); - long long timestamp = std::chrono::duration_cast(now.time_since_epoch()).count(); - lean_object *lean_timestamp = lean_int_to_int((int)timestamp); - return lean_io_result_mk_ok(lean_timestamp); -} - /* monoMsNow : BaseIO Nat */ extern "C" LEAN_EXPORT obj_res lean_io_mono_ms_now(obj_arg /* w */) { static_assert(sizeof(std::chrono::milliseconds::rep) <= sizeof(uint64), "size of std::chrono::nanoseconds::rep may not exceed 64"); From 858415c16c4762748704ca16f186d4e94a130f10 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:25:51 -0300 Subject: [PATCH 010/118] chore: fix comment --- src/Std/Time/Zoned/ZonedDateTime.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 45114c192ded..d617de3fc2d6 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -14,7 +14,8 @@ namespace Time set_option linter.all true /-- -The existential version of `DateTime` that instead of storing the timezone with a +An existential version of `DateTime` that encapsulates a `DateTime` value without explicitly storing +the time zone. -/ def ZonedDateTime := Sigma DateTime From 882fb0e827b5407dd25fb0f1a0059c995d8159e3 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:28:49 -0300 Subject: [PATCH 011/118] style: fix comment whitespaces --- src/Std/Time/Format/Basic.lean | 5 ++++- src/Std/Time/Internal/UnitVal.lean | 1 + src/Std/Time/Time/LocalTime.lean | 1 + src/Std/Time/Zoned/Database/Basic.lean | 1 + src/Std/Time/Zoned/Database/TzIf.lean | 6 ++++++ src/Std/Time/Zoned/DateTime.lean | 1 - src/Std/Time/Zoned/Offset.lean | 2 ++ src/Std/Time/Zoned/TimeZone.lean | 2 ++ src/Std/Time/Zoned/ZoneRules.lean | 4 ++++ 9 files changed, 21 insertions(+), 2 deletions(-) diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index e59fbb6ae98b..2bf45b1e0985 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -217,7 +217,10 @@ end Awareness A specification on how to format a data or parse some string. -/ structure Format (awareness : Awareness) where - /-- The format that is not aware of the timezone. -/ + + /-- + The format that is not aware of the timezone. + -/ string : FormatString deriving Inhabited, Repr diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 77ed47232939..8df4ba8d2440 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -20,6 +20,7 @@ set_option linter.all true A structure representing a unit of a given ratio type `α`. -/ structure UnitVal (α : Rat) where + /-- Value inside the UnitVal Value -/ diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/LocalTime.lean index bdd65b065869..0e6c3324fe60 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -16,6 +16,7 @@ set_option linter.all true Represents a specific point in time, including hours, minutes, seconds, and nanoseconds. -/ structure LocalTime where + /-- `Hour` component of the `LocalTime` -/ diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 82ecc7c1850e..d3d4df060add 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -17,6 +17,7 @@ set_option linter.all true A timezone database that we can read the `ZoneRules` of some area by it's id. -/ class Database (α : Type) where + /-- Loads a `ZoneRules` by it's id. -/ diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index e6ff6881b684..5929eba9ab64 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -24,6 +24,7 @@ private abbrev Int64 := Int Represents the header of a TZif file, containing metadata about the file's structure. -/ structure Header where + /-- The version of the TZif file format. -/ @@ -64,6 +65,7 @@ structure Header where Represents the local time type information, including offset and daylight saving details. -/ structure LocalTimeType where + /-- The GMT offset in seconds for this local time type. -/ @@ -84,6 +86,7 @@ structure LocalTimeType where Represents a leap second record, including the transition time and the correction applied. -/ structure LeapSecond where + /-- The transition time of the leap second event. -/ @@ -99,6 +102,7 @@ structure LeapSecond where Represents version 1 of the TZif format. -/ structure TZifV1 where + /-- The header information of the TZif file. -/ @@ -144,6 +148,7 @@ structure TZifV1 where Represents version 2 of the TZif format, extending TZifV1 with an optional footer. -/ structure TZifV2 extends TZifV1 where + /-- An optional footer for additional metadata in version 2. -/ @@ -154,6 +159,7 @@ structure TZifV2 extends TZifV1 where Represents a TZif file, which can be either version 1 or version 2. -/ structure TZif where + /-- The data for version 1 of the TZif file. -/ diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index cfbbe8721a96..c5702e1980db 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -71,7 +71,6 @@ def now : IO (DateTime tz) := do let loca ← LocalDateTime.now return ofLocalDateTime loca tz - /-- Add `Day.Offset` to a `DateTime`. -/ diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 0e8df17fb8d4..ef894316de11 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -19,10 +19,12 @@ Represents a timezone offset with an hour and second component. -/ structure Offset where private mk :: + /-- The timezone offset in Hours. -/ hour : Hour.Offset + /-- The same timezone offset in seconds. -/ diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index b74140a8bff4..349dacb0a1cc 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -16,10 +16,12 @@ set_option linter.all true An enumeration representing different time zones. -/ structure TimeZone where + /-- The `Offset` of the date time. -/ offset : TimeZone.Offset + /-- The name of the time zone. -/ diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index c484773de50d..4189fbe1023d 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -51,6 +51,7 @@ inductive StdWall Represents a type of local time, including offset and daylight saving information. -/ structure LocalTimeType where + /-- The offset from GMT for this local time. -/ @@ -96,6 +97,7 @@ end LocalTimeType Represents a leap second event, including the time of the transition and the correction applied. -/ structure LeapSecond where + /-- The time when the leap second transition occurs. -/ @@ -111,6 +113,7 @@ structure LeapSecond where Represents a time zone transition, mapping a time to a local time type. -/ structure Transition where + /-- The specific time of the transition event. -/ @@ -126,6 +129,7 @@ structure Transition where Represents the rules for a time zone, abstracting away binary data and focusing on key transitions and types. -/ structure ZoneRules where + /-- The array of local time types for the time zone. -/ From 5c01abee3ea74c4b9836db53422f62432890be38 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:29:59 -0300 Subject: [PATCH 012/118] style: fix whitespaces between declarations --- src/Std/Time/Internal/UnitVal.lean | 2 -- src/Std/Time/Notation.lean | 1 - 2 files changed, 3 deletions(-) diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 8df4ba8d2440..952acb9bc29d 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -85,8 +85,6 @@ Apply the `Sign` to a value. def apply (u1 : UnitVal α) (sign : Sign) : UnitVal α := ⟨u1.val * sign.val⟩ - - /-- Converts an `Offset` to another unit type. -/ diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 328ba2e46515..ca8b2f1c75c4 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -48,7 +48,6 @@ Category of dates like `DD-MM-YYYY`. -/ declare_syntax_cat date - /-- Date in `DD-MM-YYYY` format. -/ From cf9b5671679fd816366b2460da5f0220c3b5671f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:36:35 -0300 Subject: [PATCH 013/118] style: fix some comments and add prelude on the top --- src/Std/Time/Time/HourMarker.lean | 2 +- src/Std/Time/Zoned/Database.lean | 1 + src/Std/Time/Zoned/Database/Basic.lean | 1 + src/Std/Time/Zoned/Database/TZdb.lean | 17 ++++++++--------- src/Std/Time/Zoned/Database/TzIf.lean | 2 ++ 5 files changed, 13 insertions(+), 10 deletions(-) diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index e22951ebc9cd..6249c27eed27 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -14,7 +14,7 @@ set_option linter.all true /-- `HourMarker` represents the two 12-hour periods of the day: `am` for hour between 12:00 AM and -11:59 AM, and `pm` for hour between 12:00 PM and 11:59 PM. +11:59 AM, and `pm` for hours between 12:00 PM and 11:59 PM. -/ inductive HourMarker /-- Ante meridiem. -/ diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index 08c2856712e1..cd9a3405d80e 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ +prelude import Std.Time.Zoned.Database.Basic import Std.Time.Zoned.Database.TZdb import Std.Time.Zoned.ZonedDateTime diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index d3d4df060add..1f2708dddec6 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ +prelude import Std.Time.Zoned.Database.TzIf import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.TimeZone diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index d42eb346b19e..0b01ddff4e82 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -13,7 +13,7 @@ namespace Time namespace Database /-- -Reads the TZDb from the system from the TZDb database usually available in ""/usr/share/zoneinfo/".", +Represents a Time Zone Database (TZdb) configuration with paths to local and general timezone data. -/ structure TZdb where localPath : System.FilePath := "/etc/localtime" @@ -23,29 +23,27 @@ namespace TZdb open TimeZone /-- -Returns a default TZdb with all the fields completed with the default location for timezones in most -linux distributions and MacOS. +Returns a default `TZdb` instance with common timezone data paths for most Linux distributions and macOS. -/ @[inline] def default : TZdb := {} /-- -Parses a binary data into a zone rules. +Parses binary timezone data into zone rules based on a given timezone ID. -/ def parseTZif (bin : ByteArray) (id : String) : Except String ZoneRules := do let database ← TZif.parse.run bin convertTZif database id /-- -Gets the ZoneRules from the a TZIf file. +Reads a TZif file from disk and retrieves the zone rules for the specified timezone ID. -/ def parseTZIfFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := do let binary ← IO.FS.readBinFile path IO.ofExcept (parseTZif binary id) /-- -Creates an id out of a `FilePath`, e.g, "/var/db/timezone/zoneinfo/America/Sao_Paulo" returns -"America/Sao_Paulo". +Extracts a timezone ID from a file path. -/ def idFromPath (path : System.FilePath) : Option String := do let res := path.components.toArray @@ -55,8 +53,9 @@ def idFromPath (path : System.FilePath) : Option String := do if last₁ = some "zoneinfo" then last else last₁ ++ "/" ++ last + /-- -Gets the rules from local +Retrieves the timezone rules from the local timezone data file. -/ def localRules (path : System.FilePath) : IO ZoneRules := do let localtimePath ← IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } @@ -66,7 +65,7 @@ def localRules (path : System.FilePath) : IO ZoneRules := do else throw (IO.userError "cannot read the id of the path.") /-- -Gets the rules from local +Reads timezone rules from disk based on the provided file path and timezone ID. -/ def readRulesFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := do parseTZIfFromDisk (System.FilePath.join path id) id diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index 5929eba9ab64..db3838609298 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -3,8 +3,10 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ +prelude import Lean.Data.Parsec import Lean.Data.Parsec.ByteArray +import Init.Data.Range -- Based on: https://www.rfc-editor.org/rfc/rfc8536.html From 5d058ad92f5162ba0d7b956468fd6801bdcf15e4 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:46:26 -0300 Subject: [PATCH 014/118] feat: add parsing functions --- src/Std/Time/Format.lean | 55 +++++++++++++++++++++++++++++++++++++--- 1 file changed, 51 insertions(+), 4 deletions(-) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 1a1e7005e1f2..425e6bf3e4c1 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -81,6 +81,31 @@ Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalD def fromAmericanDateString (input : String) : Except String LocalDate := do Formats.AmericanDate.parseBuilder (λm d y => LocalDate.ofYearMonthDay y m d) input +/-- +Converts a Date in the American format (`MM/DD/YYYY`) into a `String`. +-/ +def toAmericanDateString (input : LocalDate) : String := + Formats.AmericanDate.formatBuilder input.month input.day input.year + +/-- +Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. +-/ +def fromSQLDateString (input : String) : Except String LocalDate := do + Formats.SQLDate.parseBuilder (λy m d => LocalDate.ofYearMonthDay y m d) input + +/-- +Converts a Date in the SQL format (`MM/DD/YYYY`) into a `String`. +-/ +def toSQLDateString (input : LocalDate) : String := + Formats.SQLDate.formatBuilder input.year input.month input.day + +/-- +Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `LocalDate`. +-/ +def parse (input : String) : Except String LocalDate := + fromAmericanDateString input + <|> fromSQLDateString input + end LocalDate namespace LocalTime @@ -109,12 +134,19 @@ Formats a `LocalTime` value into a 12-hour format string (`hh:mm:ss aa`). def toTime12Hour (input : LocalTime) : String := Formats.Time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.snd.val ≥ 12 then HourMarker.pm else HourMarker.am) +/-- +Parses a `String` in the `Time12Hour` or `Time24Hour` format and returns a `LocalTime`. +-/ +def parse (input : String) : Except String LocalTime := + fromTime12Hour input + <|> fromTime24Hour input + end LocalTime namespace ZonedDateTime /-- -Parses a `String` in the `ISO8601` format and returns a `DateTime`. +Parses a `String` in the `ISO8601` format and returns a `ZonedDateTime`. -/ def fromISO8601String (input : String) : Except String ZonedDateTime := Formats.ISO8601.parse input @@ -126,7 +158,7 @@ def toISO8601String (date : ZonedDateTime) : String := Formats.ISO8601.format date.snd /-- -Parses a `String` in the `RFC822` format and returns a `DateTime`. +Parses a `String` in the `RFC822` format and returns a `ZonedDateTime`. -/ def fromRFC822String (input : String) : Except String ZonedDateTime := Formats.RFC822.parse input @@ -138,7 +170,7 @@ def toRFC822String (date : ZonedDateTime) : String := Formats.RFC822.format date.snd /-- -Parses a `String` in the `RFC850` format and returns a `DateTime`. +Parses a `String` in the `RFC850` format and returns a `ZonedDateTime`. -/ def fromRFC850String (input : String) : Except String ZonedDateTime := Formats.RFC850.parse input @@ -149,6 +181,14 @@ Formats a `ZonedDateTime` value into an RFC850 format string. def toRFC850String (date : ZonedDateTime) : String := Formats.RFC850.format date.snd +/-- +Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. +-/ +def parse (input : String) : Except String ZonedDateTime := + fromISO8601String input + <|> fromRFC822String input + <|> fromRFC850String input + end ZonedDateTime namespace DateTime @@ -178,7 +218,7 @@ def toLongDateFormatString (datetime : DateTime .GMT) : String := Formats.LongDateFormat.format datetime /-- -Formats a `ZonedDateTime` value into an ISO8601 string. +Formats a `DateTime` value into an ISO8601 string. -/ def toISO8601String (date : DateTime tz) : String := Formats.ISO8601.format date @@ -195,4 +235,11 @@ Formats a `ZonedDateTime` value into an RFC850 format string. def toRFC850String (date : DateTime tz) : String := Formats.RFC850.format date +/-- +Parses a `String` in the `AscTime` or `LongDate` format and returns a `DateTime`. +-/ +def parse (date : String) : Except String (DateTime .GMT) := + fromAscTimeString date + <|> fromLongDateFormatString date + end DateTime From 4e3b9b1c5eb8c75f104243873d32e04d684e2f2b Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:50:37 -0300 Subject: [PATCH 015/118] style: rename function and whitespace --- src/Std/Time/Internal/Bounded.lean | 2 +- src/Std/Time/Zoned/Database/Basic.lean | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index ca0f94b8463e..3f40b29a8395 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -16,7 +16,7 @@ set_option linter.all true in A `Bounded` is represented by an `Int` that is constrained by a lower and higher bounded using some relation `rel`. It includes all the integers that `rel lo val ∧ rel val hi`. -/ -def Bounded (rel : Int → Int → Prop) (lo : Int) (hi : Int) := { val : Int // rel lo val ∧ rel val hi} +def Bounded (rel : Int → Int → Prop) (lo : Int) (hi : Int) := { val : Int // rel lo val ∧ rel val hi } namespace Bounded diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 1f2708dddec6..eb25ee3c1e0c 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -54,7 +54,7 @@ def convertLeapSecond (tz : TZif.LeapSecond) : LeapSecond := /-- Converts a LocalTime. -/ -def convetLocalTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do +def convertLocalTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do let localType ← tz.localTimeTypes.get? index let abbreviation ← tz.abbreviations.getD index "Unknown" let wallflag := convertWall (tz.stdWallIndicators.getD index true) @@ -85,7 +85,7 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := let mut times : Array LocalTimeType := #[] for i in [0:tz.header.typecnt.toNat] do - if let some result := convetLocalTime i tz id + if let some result := convertLocalTime i tz id then times := times.push result else .error s!"cannot convert local time {i} of the file" From abfd3be55f42e373a819c170d414486be3f253aa Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 18:55:34 -0300 Subject: [PATCH 016/118] style: add prelude import --- src/Std/Time/Zoned/Database/TZdb.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 0b01ddff4e82..b8dc57a8074d 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ +prelude import Std.Time.Zoned.Database.Basic import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.TimeZone From bd236c1c28bd7eff585536b51226b4f1f23e1539 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 14 Aug 2024 21:40:50 -0300 Subject: [PATCH 017/118] style: reorder imports and comments --- src/Std/Time.lean | 2 - src/Std/Time/Date/Unit/Day.lean | 6 +-- src/Std/Time/Date/Unit/Month.lean | 5 +-- src/Std/Time/Date/Unit/WeekOfYear.lean | 5 +-- src/Std/Time/Date/Unit/Weekday.lean | 10 ++++- src/Std/Time/Date/Unit/Year.lean | 3 +- src/Std/Time/DateTime.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 56 +++++++++++++++++++++++++ src/Std/Time/Duration.lean | 2 - src/Std/Time/Format.lean | 2 +- src/Std/Time/Format/Basic.lean | 4 +- src/Std/Time/Internal.lean | 1 - src/Std/Time/Internal/Sign.lean | 46 -------------------- src/Std/Time/Internal/UnitVal.lean | 11 +---- src/Std/Time/Notation.lean | 4 +- src/Std/Time/Time.lean | 2 - src/Std/Time/Time/LocalTime.lean | 13 ++++-- src/Std/Time/Time/Unit/Hour.lean | 2 +- src/Std/Time/Time/Unit/Millisecond.lean | 2 +- src/Std/Time/Time/Unit/Minute.lean | 2 +- src/Std/Time/Time/Unit/Nanosecond.lean | 1 - src/Std/Time/Time/Unit/Second.lean | 1 - src/Std/Time/Zoned.lean | 2 - src/Std/Time/Zoned/Database.lean | 9 ++-- src/Std/Time/Zoned/Database/Basic.lean | 4 +- src/Std/Time/Zoned/Database/TZdb.lean | 17 ++++++-- src/Std/Time/Zoned/Database/TzIf.lean | 2 +- src/Std/Time/Zoned/DateTime.lean | 1 - src/Std/Time/Zoned/Offset.lean | 4 +- src/Std/Time/Zoned/TimeZone.lean | 1 - src/Std/Time/Zoned/ZoneRules.lean | 3 -- src/Std/Time/Zoned/ZonedDateTime.lean | 1 - 32 files changed, 109 insertions(+), 117 deletions(-) delete mode 100644 src/Std/Time/Internal/Sign.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 10ce573310d8..38f592318ca6 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -14,8 +14,6 @@ import Std.Time.Duration namespace Std namespace Time -set_option linter.all true - /-! # Time diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 7f9639966b77..c7c7eb73235e 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -4,18 +4,16 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time -import Std.Time.Internal import Lean.Data.Rat +import Std.Time.Time namespace Std namespace Time +namespace Day open Lean Internal set_option linter.all true -namespace Day - /-- `Ordinal` represents a bounded value for days, which ranges between 1 and 31. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 364604a51d78..54be3a7de2f8 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -4,19 +4,16 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat -import Std.Time.Time.Basic import Std.Time.Date.Unit.Day namespace Std namespace Time +namespace Month open Internal set_option linter.all true -namespace Month - /-- `Ordinal` represents a bounded value for months, which ranges between 1 and 12. -/ diff --git a/src/Std/Time/Date/Unit/WeekOfYear.lean b/src/Std/Time/Date/Unit/WeekOfYear.lean index e6249e528ea7..948f1fa73902 100644 --- a/src/Std/Time/Date/Unit/WeekOfYear.lean +++ b/src/Std/Time/Date/Unit/WeekOfYear.lean @@ -4,19 +4,16 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat import Std.Time.Date.Unit.Day -import Std.Time.Date.Unit.Month namespace Std namespace Time +namespace WeekOfYear open Internal set_option linter.all true -namespace WeekOfYear - /-- `Ordinal` represents a bounded value for weeks, which ranges between 1 and 53. -/ diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index 74390715112b..72fefa8e3e22 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat import Std.Time.Date.Unit.Day -import Std.Time.Date.Unit.Month namespace Std namespace Time @@ -107,6 +105,14 @@ def next : Weekday → Weekday | .sat => .fri | .sun => .sat +/-- +Check if it's a Weekend. +-/ +def weekend : Weekday → Bool + | .sat => true + | .sun => true + | _ => false + end Weekday end Time end Std diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index 8f72111aa6f6..e3ffa44db328 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -11,12 +11,11 @@ import Std.Time.Date.Unit.Month namespace Std namespace Time +namespace Year open Internal set_option linter.all true -namespace Year - /-- `Offset` represents an offset in years. It is defined as an `Int`. -/ diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 54b0af601981..6752ef284af7 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.DateTime.LocalDateTime import Std.Time.DateTime.Timestamp +import Std.Time.DateTime.LocalDateTime namespace Std namespace Time diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 8f88ad231e30..aed5065c7090 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -7,6 +7,7 @@ prelude import Std.Time.Internal import Init.Data.Int import Std.Time.Time +import Std.Time.Date namespace Std namespace Time @@ -159,3 +160,58 @@ Adds a `Second.Offset` to a `Timestamp` @[inline] def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := t.add (ofSeconds s) + +/-- +Subtracts a `Second.Offset` from a `Timestamp` +-/ +@[inline] +def subSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := + t.sub (ofSeconds s) + +/-- +Adds a `Minute.Offset` to a `Timestamp` +-/ +@[inline] +def addMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := + let seconds := m.mul 60 + t.addSeconds seconds + +/-- +Subtracts a `Minute.Offset` from a `Timestamp` +-/ +@[inline] +def subMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := + let seconds := m.mul 60 + t.subSeconds seconds + +/-- +Adds an `Hour.Offset` to a `Timestamp` +-/ +@[inline] +def addHours (t : Timestamp) (h : Hour.Offset) : Timestamp := + let seconds := h.mul 3600 + t.addSeconds seconds + +/-- +Subtracts an `Hour.Offset` from a `Timestamp` +-/ +@[inline] +def subHours (t : Timestamp) (h : Hour.Offset) : Timestamp := + let seconds := h.mul 3600 + t.subSeconds seconds + +/-- +Adds a `Day.Offset` to a `Timestamp` +-/ +@[inline] +def addDays (t : Timestamp) (d : Day.Offset) : Timestamp := + let seconds := d.mul 86400 + t.addSeconds seconds + +/-- +Subtracts a `Day.Offset` from a `Timestamp` +-/ +@[inline] +def subDays (t : Timestamp) (d : Day.Offset) : Timestamp := + let seconds := d.mul 86400 + t.subSeconds seconds diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 5226e9380062..6d2625dc592b 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Std.Time.DateTime -import Std.Time.Time namespace Std namespace Time diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 425e6bf3e4c1..fa3d78839b83 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Format.Basic import Std.Time.Notation +import Std.Time.Format.Basic namespace Std namespace Time diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 2bf45b1e0985..75c757e6886a 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude +import Lean.Data.Parsec import Std.Time.Date import Std.Time.Time -import Std.Time.DateTime import Std.Time.Zoned -import Lean.Data.Parsec +import Std.Time.DateTime namespace Std namespace Time diff --git a/src/Std/Time/Internal.lean b/src/Std/Time/Internal.lean index cf91f4cd6070..596a2a0489c7 100644 --- a/src/Std/Time/Internal.lean +++ b/src/Std/Time/Internal.lean @@ -5,5 +5,4 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Internal.Bounded -import Std.Time.Internal.Sign import Std.Time.Internal.UnitVal diff --git a/src/Std/Time/Internal/Sign.lean b/src/Std/Time/Internal/Sign.lean deleted file mode 100644 index d3127637409b..000000000000 --- a/src/Std/Time/Internal/Sign.lean +++ /dev/null @@ -1,46 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Std.Time.Internal.Bounded - -namespace Std -namespace Time -namespace Internal - -set_option linter.all true - -/-- -A `Sign` is a type that can have three values -1, 0 and 1 that represents the sign of a value. --/ -def Sign := Bounded.LE (-1) 1 - -instance : ToString Sign where - toString - | ⟨-1, _⟩ => "-" - | _ => "+" - -namespace Sign - -/-- -Gets the `Sign` out of a integer. --/ -@[inline] -def ofInt (val : Int) : Sign := by - refine ⟨val.sign, ?_⟩ - simp [Int.sign] - split <;> simp - -/-- -Applies the sign to a value. --/ -@[inline] -def apply (sign : Sign) (val : Int) : Int := - val * sign.val - -end Sign -end Internal -end Time -end Std diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 952acb9bc29d..fd56590a63b7 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Init +import Init.Data import Lean.Data.Rat -import Std.Time.Internal.Sign namespace Std namespace Time @@ -78,16 +77,10 @@ Subtracts one `UnitVal` value from another of the same ratio. def sub (u1 : UnitVal α) (u2 : UnitVal α) : UnitVal α := ⟨u1.val - u2.val⟩ -/-- -Apply the `Sign` to a value. --/ -@[inline] -def apply (u1 : UnitVal α) (sign : Sign) : UnitVal α := - ⟨u1.val * sign.val⟩ - /-- Converts an `Offset` to another unit type. -/ +@[inline] def convert (val : UnitVal a) : UnitVal b := let ratio := a.div b ofInt <| val.toInt * ratio.num / ratio.den diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index ca8b2f1c75c4..f0c7cdb7a244 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -6,8 +6,8 @@ Authors: Sofia Rodrigues prelude import Std.Time.Date import Std.Time.Time -import Std.Time.DateTime import Std.Time.Zoned +import Std.Time.DateTime import Std.Time.Format.Basic import Lean.Elab.Command @@ -15,8 +15,6 @@ namespace Std namespace Time open Lean.Elab Term Lean Parser Command Std -set_option linter.all true - /-- Category of units that are valid inside a date. -/ diff --git a/src/Std/Time/Time.lean b/src/Std/Time/Time.lean index 10d9082ee548..467a4fb35bb9 100644 --- a/src/Std/Time/Time.lean +++ b/src/Std/Time/Time.lean @@ -7,5 +7,3 @@ prelude import Std.Time.Time.Basic import Std.Time.Time.HourMarker import Std.Time.Time.LocalTime - -set_option linter.all true diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/LocalTime.lean index 0e6c3324fe60..3a5295a53a42 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -80,7 +80,7 @@ def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Or exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) /-- -Converts a `LocalTime` value to the total number of seconds since midnight. +Converts a `LocalTime` value to the total number of seconds. -/ def toNanoseconds (time : LocalTime) : Nanosecond.Offset := let secs := @@ -91,7 +91,7 @@ def toNanoseconds (time : LocalTime) : Nanosecond.Offset := UnitVal.mk (nanos.val + time.nano.val) /-- -Converts a `LocalTime` value to the total number of seconds since midnight. +Converts a `LocalTime` value to the total number of seconds. -/ def toSeconds (time : LocalTime) : Second.Offset := time.hour.snd.toOffset.toSeconds + @@ -99,13 +99,20 @@ def toSeconds (time : LocalTime) : Second.Offset := time.second.snd.toOffset /-- -Converts a `LocalTime` value to the total number of minutes since midnight. +Converts a `LocalTime` value to the total number of minutes. -/ def toMinutes (time : LocalTime) : Minute.Offset := time.hour.snd.toOffset.toMinutes + time.minute.toOffset + time.second.snd.toOffset.toMinutes +/-- +Converts a `LocalTime` value to the total number of hours. +-/ +def toHours (time : LocalTime) : Hour.Offset := + let hour : Hour.Offset := time.minute.toOffset.ediv 60 + time.hour.snd.toOffset + hour + time.second.snd.toOffset.toHours + end LocalTime end Time end Std diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 5d73e8c10a6a..6e018f5be6b6 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat +import Std.Time.Internal import Std.Time.Time.Unit.Minute import Std.Time.Time.Unit.Second diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index f71362f464c0..f1978d404380 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat +import Std.Time.Internal namespace Std namespace Time diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index c9cc2cfd1df9..ac215e4c87e8 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat +import Std.Time.Internal import Std.Time.Time.Unit.Second namespace Std diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index f0d3195e68b6..e8f08c129e25 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat import Std.Time.Time.Unit.Millisecond diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index c02270553e6d..6d9f44359a3f 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Internal import Lean.Data.Rat import Std.Time.Time.Unit.Nanosecond diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 6c72e09ed09f..86a87bc12412 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -10,5 +10,3 @@ import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.Database.Basic - -set_option linter.all true diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index cd9a3405d80e..a4a57af6a1c1 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -4,19 +4,16 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Zoned.Database.Basic -import Std.Time.Zoned.Database.TZdb import Std.Time.Zoned.ZonedDateTime -import Std.Time.Zoned.ZoneRules -import Std.Time.Zoned.TimeZone -import Std.Time.Zoned.DateTime -import Std.Time.DateTime +import Std.Time.Zoned.Database.Basic namespace Std namespace Time namespace Database open TimeZone.ZoneRules +set_option linter.all true + /-- Apply leap seconds rules and transition rules on a UTC Timestamp to make it aware of the timezone. -/ diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index eb25ee3c1e0c..57052c80b2d1 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Zoned.Database.TzIf import Std.Time.Zoned.ZoneRules -import Std.Time.Zoned.TimeZone -import Std.Time.DateTime +import Std.Time.Zoned.Database.TzIf namespace Std namespace Time diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index b8dc57a8074d..7b36c9ed781a 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -4,20 +4,31 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Zoned.Database.Basic -import Std.Time.Zoned.ZoneRules -import Std.Time.Zoned.TimeZone import Std.Time.DateTime +import Std.Time.Zoned.TimeZone +import Std.Time.Zoned.ZoneRules +import Std.Time.Zoned.Database.Basic namespace Std namespace Time namespace Database +set_option linter.all true + /-- Represents a Time Zone Database (TZdb) configuration with paths to local and general timezone data. -/ structure TZdb where + /-- + The path to the local timezone file. This is typically a symlink to a file within the timezone + database that corresponds to the current local time zone. + -/ localPath : System.FilePath := "/etc/localtime" + + /-- + The path to the directory containing all available time zone files. These files define various + time zones and their rules. + -/ zonesPath : System.FilePath := "/usr/share/zoneinfo/" namespace TZdb diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index db3838609298..d60facf09f48 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude +import Init.Data.Range import Lean.Data.Parsec import Lean.Data.Parsec.ByteArray -import Init.Data.Range -- Based on: https://www.rfc-editor.org/rfc/rfc8536.html diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index c5702e1980db..a253e0287d42 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -6,7 +6,6 @@ Authors: Sofia Rodrigues prelude import Std.Time.DateTime import Std.Time.Zoned.TimeZone -import Std.Time.Zoned.ZoneRules namespace Std namespace Time diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index ef894316de11..7ac1a4ace1a3 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -5,7 +5,6 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Time -import Std.Time.Internal namespace Std namespace Time @@ -38,8 +37,7 @@ Converts an `Offset` to a string in ISO 8601 format. The `colon` parameter deter and minute components are separated by a colon (e.g., "+01:00" or "+0100"). -/ def toIsoString (offset : Offset) (colon : Bool) : String := - let sign := Sign.ofInt offset.second.val - let time := offset.second.apply sign + let (sign, time) := if offset.second.val > 0 then ("+", offset.second) else ("-", -offset.second) let hour : Hour.Offset := time.ediv 3600 let minute := Int.ediv (Int.mod time.val 3600) 60 let hourStr := if hour.val < 10 then s!"0{hour.val}" else toString hour.val diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index 349dacb0a1cc..e508053a5810 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time import Std.Time.Zoned.Offset namespace Std diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 4189fbe1023d..435ea28b6565 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -4,10 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Time import Std.Time.DateTime -import Std.Time.Internal -import Std.Time.Zoned.Offset import Std.Time.Zoned.TimeZone namespace Std diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index d617de3fc2d6..4f4ecaeb7510 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Zoned.TimeZone import Std.Time.Zoned.DateTime import Std.Time.Zoned.ZoneRules From 4032ad2c170ed94928110457677264590de98e2f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 15 Aug 2024 08:25:18 -0300 Subject: [PATCH 018/118] feat: isDST flag for timezones --- src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Notation.lean | 4 ++-- src/Std/Time/Zoned/TimeZone.lean | 17 +++++++++++------ src/Std/Time/Zoned/ZoneRules.lean | 4 ++-- 4 files changed, 16 insertions(+), 11 deletions(-) diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 75c757e6886a..4c025051e4f5 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -576,7 +576,7 @@ private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except match aw with | .only tz => .ok (build tz) | .any => - let tz₁ := TimeZone.mk builder.tz builder.tzName + let tz₁ := TimeZone.mk builder.tz builder.tzName false .ok ⟨tz₁, build tz₁⟩ else .error "invalid leap seconds {} {} {}" diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index f0c7cdb7a244..ec84558a90d9 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -199,11 +199,11 @@ macro_rules | `(date% $zoned:zoned) => parseZoned zoned | `(offset% $offset:offset) => parseOffset offset | `(timezone% $str $offset:offset) => do - do `(Std.Time.TimeZone.mk $(← parseOffset offset) $str) + do `(Std.Time.TimeZone.mk $(← parseOffset offset) $str false) | `(timezone% $name/$sub $offset:offset) => do let name := s!"{name.getId.toString}/{sub.getId.toString}" let syn := Syntax.mkStrLit name - do `(Std.Time.TimeZone.mk $(← parseOffset offset) $syn) + do `(Std.Time.TimeZone.mk $(← parseOffset offset) $syn false) private def convertModifier : Modifier → MacroM (TSyntax `term) | .YYYY => `(Std.Time.Modifier.YYYY) diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index e508053a5810..e24a90d9d270 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -25,6 +25,11 @@ structure TimeZone where The name of the time zone. -/ name : String + + /-- + Day light saving flag. + -/ + isDST : Bool deriving Inhabited, Repr namespace TimeZone @@ -33,25 +38,25 @@ namespace TimeZone A zeroed `Timezone` representing UTC (no offset). -/ def UTC : TimeZone := - TimeZone.mk (Offset.zero) "Coordinated Universal Time" + TimeZone.mk (Offset.zero) "Coordinated Universal Time" false /-- A zeroed `Timezone` representing GMT (no offset). -/ def GMT : TimeZone := - TimeZone.mk (Offset.zero) "Greenwich Mean Time" + TimeZone.mk (Offset.zero) "Greenwich Mean Time" false /-- Creates a `Timestamp` from a given number of hour. -/ -def ofHours (name : String) (n : Hour.Offset) : TimeZone := - TimeZone.mk (Offset.ofHours n) name +def ofHours (name : String) (n : Hour.Offset) (isDST : Bool := false) : TimeZone := + TimeZone.mk (Offset.ofHours n) name isDST /-- Creates a `Timestamp` from a given number of second. -/ -def ofSeconds (name : String) (n : Second.Offset) : TimeZone := - TimeZone.mk (Offset.ofSeconds n) name +def ofSeconds (name : String) (n : Second.Offset) (isDST : Bool := false) : TimeZone := + TimeZone.mk (Offset.ofSeconds n) name isDST /-- Gets the number of seconds in a timezone offset. diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 435ea28b6565..13d657ed526a 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -86,7 +86,7 @@ namespace LocalTimeType Gets the `TimeZone` offset from a `LocalTimeType`. -/ def getTimeZone (time : LocalTimeType) : TimeZone := - ⟨time.gmtOffset, time.abbreviation⟩ + ⟨time.gmtOffset, time.abbreviation, time.isDst⟩ end LocalTimeType @@ -151,7 +151,7 @@ Create a TimeZone from a Transition. def createTimeZoneFromTransition (transition : Transition) : TimeZone := let offset := transition.localTimeType.gmtOffset let name := transition.localTimeType.abbreviation - TimeZone.mk offset name + TimeZone.mk offset name transition.localTimeType.isDst /-- Applies the transition to a Timestamp. From 2376e628538335a41062597bd9126cf3c1f50bde Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 15 Aug 2024 16:22:06 -0300 Subject: [PATCH 019/118] feat: a lot of conversion functions --- src/Std/Time/Date/LocalDate.lean | 3 + src/Std/Time/Date/Unit/Day.lean | 21 +++ src/Std/Time/DateTime.lean | 72 +++++++- src/Std/Time/DateTime/LocalDateTime.lean | 4 +- src/Std/Time/DateTime/Timestamp.lean | 3 + src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Notation.lean | 16 +- src/Std/Time/Time/LocalTime.lean | 24 +++ src/Std/Time/Time/Unit/Basic.lean | 1 + src/Std/Time/Zoned.lean | 66 ++++++++ src/Std/Time/Zoned/Database.lean | 5 +- src/Std/Time/Zoned/DateTime.lean | 26 ++- src/Std/Time/Zoned/Offset.lean | 2 +- src/Std/Time/Zoned/TimeZone.lean | 2 +- src/Std/Time/Zoned/ZonedDateTime.lean | 11 +- tests/lean/timeFormats.lean | 203 +++++++++++++++++++++++ tests/lean/timeParse.lean | 177 ++++++++++++++++++++ tests/lean/timeTzifParse.lean | 53 ++++++ 18 files changed, 663 insertions(+), 28 deletions(-) create mode 100644 tests/lean/timeFormats.lean create mode 100644 tests/lean/timeParse.lean create mode 100644 tests/lean/timeTzifParse.lean diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index 48de3b02d3fc..087883e6c2d0 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -40,6 +40,9 @@ structure LocalDate where deriving Repr +instance : BEq LocalDate where + beq x y := x.day == y.day && x.month == y.month && x.year == y.year + namespace LocalDate /-- diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index c7c7eb73235e..401b78768319 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -104,6 +104,27 @@ Convert `Day.Offset` into `Hour.Offset`. def toHours (days : Offset) : Hour.Offset := days.mul 24 +/-- +Convert `Second.Offset` into `Day.Offset`. +-/ +@[inline] +def ofSeconds (secs : Second.Offset) : Offset := + secs.ediv 86400 + +/-- +Convert `Minute.Offset` into `Day.Offset`. +-/ +@[inline] +def ofMinutes (minutes : Minute.Offset) : Offset := + minutes.ediv 1440 + +/-- +Convert `Hour.Offset` into `Day.Offset`. +-/ +@[inline] +def ofHours (hours : Hour.Offset) : Offset := + hours.ediv 24 + end Offset end Day end Time diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 6752ef284af7..296b26e81204 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -17,8 +17,8 @@ set_option linter.all true Converts a `LocalDateTime` to a `Timestamp` -/ @[inline] -def ofLocalDateTime (Local : LocalDateTime) : Timestamp := - Local.toUTCTimestamp +def ofLocalDateTime (ldt : LocalDateTime) : Timestamp := + ldt.toLocalTimestamp /-- Converts a `Timestamp` to a `LocalDateTime` @@ -26,3 +26,71 @@ Converts a `Timestamp` to a `LocalDateTime` @[inline] def toLocalDateTime (timestamp : Timestamp) : LocalDateTime := LocalDateTime.ofUTCTimestamp timestamp + +/-- +Converts a `LocalDate` to a `Timestamp` +-/ +@[inline] +def ofLocalDate (ld : LocalDate) : Timestamp := + let days := ld.toDaysSinceUNIXEpoch + let secs := days.toSeconds + Timestamp.ofSeconds secs + +/-- +Converts a `Timestamp` to a `LocalDate` +-/ +@[inline] +def toLocalDate (timestamp : Timestamp) : LocalDate := + let secs := timestamp.toSeconds + let days := Day.Offset.ofSeconds secs + LocalDate.ofDaysSinceUNIXEpoch days + +/-- +Converts a `LocalTime` to a `Timestamp` +-/ +@[inline] +def ofLocalTime (lt : LocalTime) : Timestamp := + let nanos := lt.toNanoseconds + Timestamp.ofNanoseconds nanos + +/-- +Converts a `Timestamp` to a `LocalTime` +-/ +@[inline] +def toLocalTime (timestamp : Timestamp) : LocalTime := + let nanos := timestamp.toNanoseconds + LocalTime.ofNanoseconds nanos + +end Timestamp + +namespace LocalDateTime + +/-- +Converts a `LocalDate` to a `Timestamp` +-/ +@[inline] +def ofLocalDate (ld : LocalDate) : LocalDateTime := + LocalDateTime.ofUTCTimestamp (Timestamp.ofLocalDate ld) + +/-- +Converts a `LocalDateTime` to a `LocalDate` +-/ +@[inline] +def toLocalDate (ldt : LocalDateTime) : LocalDate := + Timestamp.toLocalDate ldt.toLocalTimestamp + +/-- +Converts a `LocalTime` to a `LocalDateTime` +-/ +@[inline] +def ofLocalTime (lt : LocalTime) : LocalDateTime := + LocalDateTime.ofUTCTimestamp (Timestamp.ofLocalTime lt) + +/-- +Converts a `LocalDateTime` to a `LocalTime` +-/ +@[inline] +def toLocalTime (ldt : LocalDateTime) : LocalTime := + Timestamp.toLocalTime ldt.toLocalTimestamp + +end LocalDateTime diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean index 975d93981ea4..fe4847feb4dc 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -30,14 +30,14 @@ structure LocalDateTime where -/ time : LocalTime - deriving Inhabited + deriving Inhabited, BEq namespace LocalDateTime /-- Converts a `LocalDateTime` into a `Std.Time.Timestamp` -/ -def toUTCTimestamp (dt : LocalDateTime) : Timestamp := +def toLocalTimestamp (dt : LocalDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 let nanos := nanos.val + dt.time.nano.val diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index aed5065c7090..1264fd8fbc34 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -36,6 +36,9 @@ structure Timestamp where proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) deriving Repr +instance : BEq Timestamp where + beq x y := x.second == y.second && y.nano == x.nano + instance : Inhabited Timestamp where default := ⟨0, Bounded.LE.mk 0 (by decide), by decide⟩ diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 4c025051e4f5..f5645f232466 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -322,7 +322,7 @@ private def dayOfWeek (day : Weekday) : String := | .mon => "Monday" | .tue => "Tuesday" | .wed => "Wednesday" - | .thu => "Thusday" + | .thu => "Thursday" | .fri => "Friday" | .sat => "Saturday" diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index ec84558a90d9..019f4cce8d02 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -73,9 +73,9 @@ syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component private def parseTime : TSyntax `time -> MacroM (TSyntax `term) | `(time| $hour:date_component:$minute:date_component:$second:date_component) => do - `(Std.Time.LocalTime.mk $(← parseComponent hour) $(← parseComponent minute) $(← parseComponent second) 0) + `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent hour)⟩ $(← parseComponent minute) ⟨true, $(← parseComponent second)⟩ 0 (by decide)) | `(time| $hour:date_component:$minute:date_component:$second:date_component.$nanos:date_component) => do - `(Std.Time.LocalTime.mk $(← parseComponent hour) $(← parseComponent minute) $(← parseComponent second) $(← parseComponent nanos)) + `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent hour)⟩ $(← parseComponent minute) ⟨true, $(← parseComponent second)⟩ $(← parseComponent nanos) (by decide)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -95,9 +95,9 @@ syntax date_component : datetime private def parseDateTime : TSyntax `datetime -> MacroM (TSyntax `term) | `(datetime| $date:date:$time:time) => do - `(Std.Time.LocalDateTime.mk $(← parseDate date) $(← parseTime time)) + `(Std.Time.LocalDateTime.mk $(← parseDate date) $(← parseTime time)) | `(datetime|$tm:date_component) => do - `(Std.Time.LocalDateTime.ofTimestamp $(← parseComponent tm)) + `(Std.Time.LocalDateTime.ofUTCTimestamp $(← parseComponent tm)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -132,7 +132,7 @@ private def parseZone : TSyntax `zone -> MacroM (TSyntax `term) | `(zone| UTC) => do `(Std.Time.TimeZone.UTC) | `(zone| GMT) => do `(Std.Time.TimeZone.GMT) | `(zone| Z) => do `(Std.Time.TimeZone.UTC) - | `(zone| $offset:offset) => do `(Std.Time.TimeZone.mk $(← parseOffset offset) "Offset") + | `(zone| $offset:offset) => do `(Std.Time.TimeZone.mk $(← parseOffset offset) "Unknown" false) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -147,7 +147,7 @@ syntax datetime zone : zoned private def parseZoned : TSyntax `zoned -> MacroM (TSyntax `term) | `(zoned| $timestamp:num $zone) => do - `(Std.Time.DateTime.ofTimestamp $timestamp $(← parseZone zone)) + `(Std.Time.DateTime.ofUTCTimestamp $timestamp $(← parseZone zone)) | `(zoned| $datetime:datetime $zone) => do `(Std.Time.DateTime.ofLocalDateTime $(← parseDateTime datetime) $(← parseZone zone)) | syn => Macro.throwErrorAt syn "unsupported syntax" @@ -175,7 +175,7 @@ syntax (name := dateDate) "date%" date : term /-- Macro for defining times. -/ -syntax "date%" time : term +syntax "time%" time : term /-- Macro for creating offsets @@ -194,9 +194,9 @@ syntax "timezone%" str offset : term macro_rules | `(date% $date:date) => parseDate date - | `(date% $time:time) => parseTime time | `(date% $datetime:datetime) => parseDateTime datetime | `(date% $zoned:zoned) => parseZoned zoned + | `(time% $time:time) => parseTime time | `(offset% $offset:offset) => parseOffset offset | `(timezone% $str $offset:offset) => do do `(Std.Time.TimeZone.mk $(← parseOffset offset) $str false) diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/LocalTime.lean index 3a5295a53a42..0c37f76826d7 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -46,6 +46,11 @@ structure LocalTime where instance : Inhabited LocalTime where default := ⟨Sigma.mk false 0, 0, Sigma.mk false 0, 0, by simp; decide⟩ + +instance : BEq LocalTime where + beq x y := x.hour.snd.val == y.hour.snd.val && x.minute == y.minute + && x.second.snd.val == y.second.snd.val && x.nano.val == y.nano.val + namespace LocalTime /-- @@ -113,6 +118,25 @@ def toHours (time : LocalTime) : Hour.Offset := let hour : Hour.Offset := time.minute.toOffset.ediv 60 time.hour.snd.toOffset + hour + time.second.snd.toOffset.toHours +/-- +Creates a `LocalTime` value from a total number of nanoseconds. +-/ +def ofNanoseconds (nanos : Nanosecond.Offset) : LocalTime := + have totalSeconds := nanos.ediv 1000000000 + have remainingNanos := Bounded.LE.byEmod nanos.val 1000000000 (by decide) + have hours := Bounded.LE.byEmod (totalSeconds.val / 3600) 24 (by decide) + have minutes := (Bounded.LE.byEmod totalSeconds.val 3600 (by decide)).ediv 60 (by decide) + have seconds := Bounded.LE.byEmod totalSeconds.val 60 (by decide) + let nanos := Bounded.LE.byEmod nanos.val 1000000000 (by decide) + ofValidHourMinuteSecondsNano hours minutes seconds nanos + +/-- +Creates a `LocalTime` value from a total number of seconds. +-/ +@[inline] +def ofSeconds (secs : Second.Offset) : LocalTime := + ofNanoseconds (secs.mul 1000000000) + end LocalTime end Time end Std diff --git a/src/Std/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean index 4d526bf0673e..ae449b9c271c 100644 --- a/src/Std/Time/Time/Unit/Basic.lean +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -9,6 +9,7 @@ import Std.Time.Time.Unit.Hour import Std.Time.Time.Unit.Minute import Std.Time.Time.Unit.Second import Std.Time.Time.Unit.Nanosecond +import Std.Time.Time.Unit.Millisecond namespace Std namespace Time diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 86a87bc12412..514656ffc201 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -10,3 +10,69 @@ import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.Database.Basic + +namespace Std +namespace Time +namespace DateTime + +/-- +Converts a `LocalDate` to a `DateTime` +-/ +@[inline] +def ofLocalDate (ld : LocalDate) (tz : TimeZone) : DateTime tz := + DateTime.ofUTCTimestamp (Timestamp.ofLocalDate ld) tz + +/-- +Converts a `DateTime` to a `LocalDate` +-/ +@[inline] +def toLocalDate (dt : DateTime tz) : LocalDate := + Timestamp.toLocalDate dt.toTimestamp + +/-- +Converts a `LocalTime` to a `DateTime` +-/ +@[inline] +def ofLocalTime (lt : LocalTime) (tz : TimeZone) : DateTime tz := + DateTime.ofUTCTimestamp (Timestamp.ofLocalTime lt) tz + +/-- +Converts a `DateTime` to a `LocalTime` +-/ +@[inline] +def toLocalTime (dt : DateTime tz) : LocalTime := + dt.date.get.time + +end DateTime + +namespace ZonedDateTime + +/-- +Converts a `LocalDate` to a `ZonedDateTime` +-/ +@[inline] +def ofLocalDate (ld : LocalDate) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofUTCTimestamp (Timestamp.ofLocalDate ld) tz⟩ + +/-- +Converts a `ZonedDateTime` to a `LocalDate` +-/ +@[inline] +def toLocalDate (dt : ZonedDateTime) : LocalDate := + DateTime.toLocalDate dt.snd + +/-- +Converts a `LocalTime` to a `ZonedDateTime` +-/ +@[inline] +def ofLocalTime (lt : LocalTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofUTCTimestamp (Timestamp.ofLocalTime lt) tz⟩ + +/-- +Converts a `ZonedDateTime` to a `LocalTime` +-/ +@[inline] +def toLocalTime (dt : ZonedDateTime) : LocalTime := + DateTime.toLocalTime dt.snd + +end ZonedDateTime diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index a4a57af6a1c1..500af50e67eb 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -6,6 +6,7 @@ Authors: Sofia Rodrigues prelude import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.Database.Basic +import Std.Time.Zoned.Database.TZdb namespace Std namespace Time @@ -35,8 +36,8 @@ def getTimeZoneAt [Database α] (db : α) (id : String) (tm : Timestamp) : IO Ti /-- Get the local ZonedDataTime given a UTC `Timestamp`. -/ -def ofTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime := do +def ofUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime := do let rules ← Database.localRules db let tz ← IO.ofExcept <| timezoneAt rules tm let tm := applyLeapSeconds tm rules.leapSeconds - return ZonedDateTime.ofTimestamp tm tz + return ZonedDateTime.ofUTCTimestamp tm tz diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index a253e0287d42..173d7502889d 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -19,7 +19,7 @@ structure DateTime (tz : TimeZone) where private mk :: /-- - `Timestamp` represents the exact moment in time. + `Timestamp` represents the exact moment in time. It's a UTC related `Timestamp`. -/ timestamp : Timestamp @@ -29,20 +29,23 @@ structure DateTime (tz : TimeZone) where -/ date : Thunk LocalDateTime +instance : BEq (DateTime tz) where + beq x y := x.timestamp == y.timestamp + instance : Inhabited (DateTime tz) where default := ⟨Inhabited.default, Thunk.mk λ_ => Inhabited.default⟩ namespace DateTime /-- -Creates a new `DateTime` out of a `Timestamp` +Creates a new `DateTime` out of a `Timestamp` that is in a `TimeZone`. -/ @[inline] -def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := +def ofUTCTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := DateTime.mk tm (Thunk.mk <| λ_ => (tm.addSeconds tz.toSeconds).toLocalDateTime) /-- -Creates a new `Timestamp` out of a `DateTime` +Creates a new zone aware `Timestamp` out of a `DateTime`. -/ @[inline] def toTimestamp (date : DateTime tz) : Timestamp := @@ -53,14 +56,16 @@ Changes the `TimeZone` to a new one. -/ @[inline] def convertTimeZone (date : DateTime tz) (tz₁ : TimeZone) : DateTime tz₁ := - ofTimestamp (date.toTimestamp) tz₁ + ofUTCTimestamp date.timestamp tz₁ + /-- -Creates a new DateTime out of a `LocalDateTime` +Creates a new `DateTime` out of a `LocalDateTime` -/ @[inline] def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : DateTime tz := - DateTime.ofTimestamp date.toUTCTimestamp tz + let tm := Timestamp.ofLocalDateTime date + DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk <| λ_ => date) /-- Gets the current `DateTime`. @@ -142,6 +147,13 @@ Subtract `Year.Offset` from to a `DateTime`, it clips the day to the last valid def subYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := ofLocalDateTime (dt.timestamp.toLocalDateTime.subYearsClip years) tz +/-- +Converts a `Timestamp` to a `LocalDateTime` +-/ +@[inline] +def toLocalDateTime (timestamp : Timestamp) : LocalDateTime := + LocalDateTime.ofUTCTimestamp timestamp + /-- Getter for the `Year` inside of a `DateTime` -/ diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 7ac1a4ace1a3..5ca3dcbbf50c 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -28,7 +28,7 @@ structure Offset where The same timezone offset in seconds. -/ second : Second.Offset - deriving Repr, Inhabited + deriving Repr, Inhabited, BEq namespace Offset diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index e24a90d9d270..0dafb4742575 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -30,7 +30,7 @@ structure TimeZone where Day light saving flag. -/ isDST : Bool - deriving Inhabited, Repr + deriving Inhabited, Repr, BEq namespace TimeZone diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 4f4ecaeb7510..d7eef4efcddb 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -18,6 +18,9 @@ the time zone. -/ def ZonedDateTime := Sigma DateTime +instance : BEq ZonedDateTime where + beq x y := x.fst == y.fst && x.snd.timestamp == y.snd.timestamp + instance : CoeDep ZonedDateTime d (DateTime d.fst) where coe := d.snd @@ -31,8 +34,8 @@ open DateTime Creates a new `ZonedDateTime` out of a `Timestamp` -/ @[inline] -def ofTimestamp (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp tm tz⟩ +def ofUTCTimestamp (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofUTCTimestamp tm tz⟩ /-- Creates a new `Timestamp` out of a `ZonedDateTime` @@ -47,14 +50,14 @@ Creates a new `DateTime` out of a `Timestamp` @[inline] def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do let transition ← rules.findTransitionForTimestamp tm - return ofTimestamp tm transition.localTimeType.getTimeZone + return ofUTCTimestamp tm transition.localTimeType.getTimeZone /-- Changes the `TimeZone` to a new one. -/ @[inline] def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := - ofTimestamp (date.toTimestamp) tz₁ + ofUTCTimestamp (date.toTimestamp) tz₁ /-- Creates a new `ZonedDateTime` out of a `LocalDateTime` diff --git a/tests/lean/timeFormats.lean b/tests/lean/timeFormats.lean new file mode 100644 index 000000000000..cd42a368516c --- /dev/null +++ b/tests/lean/timeFormats.lean @@ -0,0 +1,203 @@ +import Std.Time +open Std.Time + +def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ssZZZ" +def RFC1123 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" +def ShortDate : Format .any := date-spec% "MM/DD/YYYY" +def LongDate : Format .any := date-spec% "MMMM D, YYYY" +def ShortDateTime : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss" +def LongDateTime : Format .any := date-spec% "MMMM D, YYYY h:mm aa" +def Time24Hour : Format .any := date-spec% "hh:mm:ss" +def Time12Hour : Format .any := date-spec% "HH:mm:ss aa" +def FullDayTimeZone : Format .any := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss ZZZZ" +def CustomDayTime : Format .any := date-spec% "EEE D MMM YYYY hh:mm" + +-- Dates + +def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 +def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 + +def date₁ := date% 16-06-2014:03:03:03(brTZ) +def time₁ := time% 14:11:01 +def time₂ := time% 03:11:01 + +/-- +info: "Monday, June 16, 2014 03:03:03 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format date₁ + +def tm := date₁.timestamp +def date₂ := DateTime.ofUTCTimestamp tm brTZ + +/-- +info: "Monday, June 16, 2014 03:03:03 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format date₂ + +def tm₃ := date₁.toTimestamp +def date₃ := DateTime.ofUTCTimestamp tm₃ brTZ + +/-- +info: "Monday, June 16, 2014 03:03:03 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format date₃ + +-- Section for testing timezone conversion. + +-- the timestamp is always related to UTC. + +/-- +Timestamp: 1723739292 +GMT: Thursday, 15 August 2024 16:28:12 +BR: 15 August 2024 13:28:12 GMT-03:00 +-/ +def tm₄ : Second.Offset := 1723739292 + +def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) brTZ +def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) jpTZ +def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) .UTC + +/-- +info: "Thursday, August 15, 2024 13:28:12 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format dateBR + +/-- +info: "Friday, August 16, 2024 01:28:12 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format dateJP + +/-- +info: "Thursday, August 15, 2024 13:28:12 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateUTC.convertTimeZone brTZ) + +/-- +info: "Thursday, August 15, 2024 13:28:12 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateJP.convertTimeZone brTZ) + +/-- +info: "Thursday, August 15, 2024 16:28:12 -0000" +-/ +#guard_msgs in +#eval FullDayTimeZone.format dateUTC + +/-- +info: "Thursday, August 15, 2024 16:28:12 -0000" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateBR.convertTimeZone .UTC) + +/-- +info: "Thursday, August 15, 2024 16:28:12 -0000" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateJP.convertTimeZone .UTC) + +/-- +info: "Friday, August 16, 2024 01:28:12 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format dateJP + +/-- +info: "Friday, August 16, 2024 01:28:12 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateBR.convertTimeZone jpTZ) + +/-- +info: "Friday, August 16, 2024 01:28:12 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateUTC.convertTimeZone jpTZ) + +/-- +TM: 1723730627 +GMT: Thursday, 15 August 2024 14:03:47 +Your time zone: 15 Aguust 2024 11:03:47 GMT-03:00 +-/ +def localTm : Second.Offset := 1723730627 + +/-- +This localDate is relative to the local time. +-/ +def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSeconds localTm) + +/-- +info: "08/15/2024 14:03:47" +-/ +#guard_msgs in +#eval ShortDateTime.formatBuilder localDate.month localDate.day localDate.year localDate.time.hour localDate.minute localDate.time.second + +def dateBR₁ := DateTime.ofLocalDateTime localDate brTZ +def dateJP₁ := DateTime.ofLocalDateTime localDate jpTZ +def dateUTC₁ := DateTime.ofLocalDateTime localDate .UTC + +/-- +info: "Thursday, August 15, 2024 14:03:47 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format dateBR₁ + +/-- +info: "Thursday, August 15, 2024 14:03:47 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format dateJP₁ + +/-- +info: "Thursday, August 15, 2024 23:03:47 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateUTC₁.convertTimeZone jpTZ) + +/-- +info: "Friday, August 16, 2024 02:03:47 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateBR₁.convertTimeZone jpTZ) + +/-- +info: "Thursday, August 15, 2024 14:03:47 +0900" +-/ +#guard_msgs in +#eval FullDayTimeZone.format (dateJP₁.convertTimeZone jpTZ) + +/-- +info: "Monday, June 16, 2014 03:03:03 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format date₂ + +/-- +info: "14:11:01" +-/ +#guard_msgs in +#eval Time24Hour.formatBuilder time₁.hour time₁.minute time₁.second + +/-- +info: "02:11:01 pm" +-/ +#guard_msgs in +#eval Time12Hour.formatBuilder time₁.hour time₁.minute time₁.second (if time₁.hour.snd.val > 12 then HourMarker.pm else HourMarker.am) + +/-- +info: "03:11:01 am" +-/ +#guard_msgs in +#eval Time12Hour.formatBuilder time₂.hour time₂.minute time₂.second (if time₂.hour.snd.val > 12 then HourMarker.pm else HourMarker.am) + +/-- +info: "06/16/2014" +-/ +#guard_msgs in +#eval ShortDate.formatBuilder date₁.month date₁.day date₁.year diff --git a/tests/lean/timeParse.lean b/tests/lean/timeParse.lean new file mode 100644 index 000000000000..2dc6785d5bd9 --- /dev/null +++ b/tests/lean/timeParse.lean @@ -0,0 +1,177 @@ +import Std.Time +open Std.Time + +def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ssZZZ" +def RFC1123 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" +def ShortDate : Format .any := date-spec% "MM/DD/YYYY" +def LongDate : Format .any := date-spec% "MMMM D, YYYY" +def ShortDateTime : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss" +def LongDateTime : Format .any := date-spec% "MMMM D, YYYY h:mm aa" +def Time24Hour : Format .any := date-spec% "hh:mm:ss" +def Time12Hour : Format .any := date-spec% "HH:mm:ss aa" +def FullDayTimeZone : Format .any := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss ZZZZ" +def CustomDayTime : Format .any := date-spec% "EEE D MMM YYYY hh:mm" + +-- Dates + +def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 +def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 + +def date₁ := date% 16-06-2014:03:03:03(brTZ) +def time₁ := time% 14:11:01 +def time₂ := time% 03:11:01 + +/-- +info: "2014-06-16T03:03:03-0300" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" + ISO8601UTC.format t.snd + +def tm := date₁.timestamp +def date₂ := DateTime.ofUTCTimestamp tm brTZ + +/-- +info: "2014-06-16T03:03:03-0300" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" + ISO8601UTC.format t.snd + +def tm₃ := date₁.toTimestamp +def date₃ := DateTime.ofUTCTimestamp tm₃ brTZ + +/-- +info: "2014-06-16T03:03:03-0300" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" + ISO8601UTC.format t.snd + +-- Section for testing timezone conversion. + +-- the timestamp is always related to UTC. + +/-- +Timestamp: 1723739292 +GMT: Thursday, 15 August 2024 16:28:12 +BR: 15 August 2024 13:28:12 GMT-03:00 +-/ +def tm₄ : Second.Offset := 1723739292 + +def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) brTZ +def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) jpTZ +def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) .UTC + +/-- +info: "2024-08-15T13:28:12-0300" +-/ +#guard_msgs in +#eval + let t := FullDayTimeZone.parse! "Thursday, August 15, 2024 13:28:12 -0300" + ISO8601UTC.format t.snd + +/-- +info: "2024-08-16T01:28:12+0900" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" + ISO8601UTC.format t.snd + +/-- +info: "2024-08-15T13:28:12-0300" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 13:28:12 -0300" + ISO8601UTC.format t.snd + +/-- +info: "2024-08-15T13:28:12-0300" +-/ +#guard_msgs in +#eval + let t1 : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 13:28:12 -0300" + ISO8601UTC.format (t1.snd.convertTimeZone brTZ) + +/-- +info: "Thu 15 Aug 2024 16:28" +-/ +#guard_msgs in +#eval + let t2 : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 16:28:12 -0000" + CustomDayTime.format t2.snd + +/-- +info: "2024-08-16T01:28:12+0900" +-/ +#guard_msgs in +#eval + let t5 : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" + ISO8601UTC.format t5.snd + +/-- +info: "2024-08-16T01:28:12+0900" +-/ +#guard_msgs in +#eval + let t6 : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" + ISO8601UTC.format (t6.snd.convertTimeZone jpTZ) + +/-- +info: "2024-08-16T01:28:12+0900" +-/ +#guard_msgs in +#eval + let t7 : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" + ISO8601UTC.format (t7.snd.convertTimeZone jpTZ) + +/-- +TM: 1723730627 +GMT: Thursday, 15 August 2024 14:03:47 +Your time zone: 15 Aguust 2024 11:03:47 GMT-03:00 +-/ +def localTm : Second.Offset := 1723730627 + +/-- +This localDate is relative to the local time. +-/ +def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSeconds localTm) + +/-- +info: "08/15/2024 14:03:47" +-/ +#guard_msgs in +#eval ShortDateTime.formatBuilder localDate.month localDate.day localDate.year localDate.time.hour localDate.minute localDate.time.second + +def dateBR₁ := DateTime.ofLocalDateTime localDate brTZ +def dateJP₁ := DateTime.ofLocalDateTime localDate jpTZ +def dateUTC₁ := DateTime.ofLocalDateTime localDate .UTC + +/-- +info: "2024-08-15T14:03:47-0300" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 14:03:47 -0300" + ISO8601UTC.format t.snd + +/-- +info: "2024-08-15T14:03:47+0900" +-/ +#guard_msgs in +#eval + let t1 : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 14:03:47 +0900" + ISO8601UTC.format t1.snd + +/-- +info: "2014-06-16T03:03:03-0300" +-/ +#guard_msgs in +#eval + let t2 : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" + ISO8601UTC.format t2.snd diff --git a/tests/lean/timeTzifParse.lean b/tests/lean/timeTzifParse.lean new file mode 100644 index 000000000000..37cacd9b1e33 --- /dev/null +++ b/tests/lean/timeTzifParse.lean @@ -0,0 +1,53 @@ +import Std.Time + +def file := ByteArray.mk #[84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0, 0, 0, 3, 0, 0, 0, 12, 150, 170, 114, 180, 184, 15, 73, 224, 184, 253, 64, 160, 185, 241, 52, 48, 186, 222, 116, 32, 218, 56, 174, 48, 218, 235, 250, 48, 220, 25, 225, 176, 220, 185, 89, 32, 221, 251, 21, 48, 222, 155, 222, 32, 223, 221, 154, 48, 224, 84, 51, 32, 244, 90, 9, 48, 245, 5, 94, 32, 246, 192, 100, 48, 247, 14, 30, 160, 248, 81, 44, 48, 248, 199, 197, 32, 250, 10, 210, 176, 250, 168, 248, 160, 251, 236, 6, 48, 252, 139, 125, 160, 29, 201, 142, 48, 30, 120, 215, 160, 31, 160, 53, 176, 32, 51, 207, 160, 33, 129, 105, 48, 34, 11, 200, 160, 35, 88, 16, 176, 35, 226, 112, 32, 37, 55, 242, 176, 37, 212, 199, 32, 39, 33, 15, 48, 39, 189, 227, 160, 41, 0, 241, 48, 41, 148, 139, 32, 42, 234, 13, 176, 43, 107, 50, 160, 44, 192, 181, 48, 45, 102, 196, 32, 46, 160, 151, 48, 47, 70, 166, 32, 48, 128, 121, 48, 49, 29, 77, 160, 50, 87, 32, 176, 51, 6, 106, 32, 52, 56, 84, 48, 52, 248, 193, 32, 54, 32, 31, 48, 54, 207, 104, 160, 55, 246, 198, 176, 56, 184, 133, 32, 57, 223, 227, 48, 58, 143, 44, 160, 59, 200, 255, 176, 60, 111, 14, 160, 61, 196, 145, 48, 62, 78, 240, 160, 63, 145, 254, 48, 64, 46, 210, 160, 65, 134, 248, 48, 66, 23, 239, 32, 67, 81, 194, 48, 67, 247, 209, 32, 69, 77, 83, 176, 69, 224, 237, 160, 71, 17, 134, 48, 71, 183, 149, 32, 72, 250, 162, 176, 73, 151, 119, 32, 74, 218, 132, 176, 75, 128, 147, 160, 76, 186, 102, 176, 77, 96, 117, 160, 78, 154, 72, 176, 79, 73, 146, 32, 80, 131, 101, 48, 81, 32, 57, 160, 82, 99, 71, 48, 83, 0, 27, 160, 84, 67, 41, 48, 84, 233, 56, 32, 86, 35, 11, 48, 86, 201, 26, 32, 88, 2, 237, 48, 88, 168, 252, 32, 89, 226, 207, 48, 90, 136, 222, 32, 91, 222, 96, 176, 92, 104, 192, 32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255, 213, 208, 0, 8, 76, 77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0, 0, 0, 3, 0, 0, 0, 12, 255, 255, 255, 255, 150, 170, 114, 180, 255, 255, 255, 255, 184, 15, 73, 224, 255, 255, 255, 255, 184, 253, 64, 160, 255, 255, 255, 255, 185, 241, 52, 48, 255, 255, 255, 255, 186, 222, 116, 32, 255, 255, 255, 255, 218, 56, 174, 48, 255, 255, 255, 255, 218, 235, 250, 48, 255, 255, 255, 255, 220, 25, 225, 176, 255, 255, 255, 255, 220, 185, 89, 32, 255, 255, 255, 255, 221, 251, 21, 48, 255, 255, 255, 255, 222, 155, 222, 32, 255, 255, 255, 255, 223, 221, 154, 48, 255, 255, 255, 255, 224, 84, 51, 32, 255, 255, 255, 255, 244, 90, 9, 48, 255, 255, 255, 255, 245, 5, 94, 32, 255, 255, 255, 255, 246, 192, 100, 48, 255, 255, 255, 255, 247, 14, 30, 160, 255, 255, 255, 255, 248, 81, 44, 48, 255, 255, 255, 255, 248, 199, 197, 32, 255, 255, 255, 255, 250, 10, 210, 176, 255, 255, 255, 255, 250, 168, 248, 160, 255, 255, 255, 255, 251, 236, 6, 48, 255, 255, 255, 255, 252, 139, 125, 160, 0, 0, 0, 0, 29, 201, 142, 48, 0, 0, 0, 0, 30, 120, 215, 160, 0, 0, 0, 0, 31, 160, 53, 176, 0, 0, 0, 0, 32, 51, 207, 160, 0, 0, 0, 0, 33, 129, 105, 48, 0, 0, 0, 0, 34, 11, 200, 160, 0, 0, 0, 0, 35, 88, 16, 176, 0, 0, 0, 0, 35, 226, 112, 32, 0, 0, 0, 0, 37, 55, 242, 176, 0, 0, 0, 0, 37, 212, 199, 32, 0, 0, 0, 0, 39, 33, 15, 48, 0, 0, 0, 0, 39, 189, 227, 160, 0, 0, 0, 0, 41, 0, 241, 48, 0, 0, 0, 0, 41, 148, 139, 32, 0, 0, 0, 0, 42, 234, 13, 176, 0, 0, 0, 0, 43, 107, 50, 160, 0, 0, 0, 0, 44, 192, 181, 48, 0, 0, 0, 0, 45, 102, 196, 32, 0, 0, 0, 0, 46, 160, 151, 48, 0, 0, 0, 0, 47, 70, 166, 32, 0, 0, 0, 0, 48, 128, 121, 48, 0, 0, 0, 0, 49, 29, 77, 160, 0, 0, 0, 0, 50, 87, 32, 176, 0, 0, 0, 0, 51, 6, 106, 32, 0, 0, 0, 0, 52, 56, 84, 48, 0, 0, 0, 0, 52, 248, 193, 32, 0, 0, 0, 0, 54, 32, 31, 48, 0, 0, 0, 0, 54, 207, 104, 160, 0, 0, 0, 0, 55, 246, 198, 176, 0, 0, 0, 0, 56, 184, 133, 32, 0, 0, 0, 0, 57, 223, 227, 48, 0, 0, 0, 0, 58, 143, 44, 160, 0, 0, 0, 0, 59, 200, 255, 176, 0, 0, 0, 0, 60, 111, 14, 160, 0, 0, 0, 0, 61, 196, 145, 48, 0, 0, 0, 0, 62, 78, 240, 160, 0, 0, 0, 0, 63, 145, 254, 48, 0, 0, 0, 0, 64, 46, 210, 160, 0, 0, 0, 0, 65, 134, 248, 48, 0, 0, 0, 0, 66, 23, 239, 32, 0, 0, 0, 0, 67, 81, 194, 48, 0, 0, 0, 0, 67, 247, 209, 32, 0, 0, 0, 0, 69, 77, 83, 176, 0, 0, 0, 0, 69, 224, 237, 160, 0, 0, 0, 0, 71, 17, 134, 48, 0, 0, 0, 0, 71, 183, 149, 32, 0, 0, 0, 0, 72, 250, 162, 176, 0, 0, 0, 0, 73, 151, 119, 32, 0, 0, 0, 0, 74, 218, 132, 176, 0, 0, 0, 0, 75, 128, 147, 160, 0, 0, 0, 0, 76, 186, 102, 176, 0, 0, 0, 0, 77, 96, 117, 160, 0, 0, 0, 0, 78, 154, 72, 176, 0, 0, 0, 0, 79, 73, 146, 32, 0, 0, 0, 0, 80, 131, 101, 48, 0, 0, 0, 0, 81, 32, 57, 160, 0, 0, 0, 0, 82, 99, 71, 48, 0, 0, 0, 0, 83, 0, 27, 160, 0, 0, 0, 0, 84, 67, 41, 48, 0, 0, 0, 0, 84, 233, 56, 32, 0, 0, 0, 0, 86, 35, 11, 48, 0, 0, 0, 0, 86, 201, 26, 32, 0, 0, 0, 0, 88, 2, 237, 48, 0, 0, 0, 0, 88, 168, 252, 32, 0, 0, 0, 0, 89, 226, 207, 48, 0, 0, 0, 0, 90, 136, 222, 32, 0, 0, 0, 0, 91, 222, 96, 176, 0, 0, 0, 0, 92, 104, 192, 32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255, 213, 208, 0, 8, 76, 77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 10, 60, 45, 48, 51, 62, 51, 10] + +def code := Std.Time.TimeZone.TZif.parse.run file |>.toOption |>.get! + +/-- +info: { version := 50, isutcnt := 0, isstdcnt := 0, leapcnt := 0, timecnt := 91, typecnt := 3, charcnt := 12 } +-/ +#guard_msgs in +#eval code.v1.header + +/-- +info: 0 +-/ +#guard_msgs in +#eval code.v1.leapSeconds.size + +/-- +info: 3 +-/ +#guard_msgs in +#eval code.v1.abbreviations.size + +/-- +info: 91 +-/ +#guard_msgs in +#eval code.v1.transitionIndices.size + +/-- +info: 91 +-/ +#guard_msgs in +#eval code.v1.transitionTimes.size + +/-- +info: 3 +-/ +#guard_msgs in +#eval code.v1.localTimeTypes.size + +/-- +info: 0 +-/ +#guard_msgs in +#eval code.v1.stdWallIndicators.size + +/-- +info: 0 +-/ +#guard_msgs in +#eval code.v1.utLocalIndicators.size From d53c100f0b12cd3469a39d53e58bf7c2f507de57 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 16 Aug 2024 12:34:46 -0300 Subject: [PATCH 020/118] style: fix some comments --- src/Std/Time/Date/LocalDate.lean | 50 ++++++++++---------- src/Std/Time/Date/Unit/Basic.lean | 8 ++++ src/Std/Time/Date/Unit/Day.lean | 8 ++-- src/Std/Time/Date/Unit/Month.lean | 19 ++++---- src/Std/Time/Date/Unit/WeekOfYear.lean | 6 ++- src/Std/Time/Date/Unit/Weekday.lean | 11 +++-- src/Std/Time/Date/Unit/Year.lean | 10 ++-- src/Std/Time/Date/WeekDate.lean | 5 +- src/Std/Time/DateTime/LocalDateTime.lean | 60 ++++++++++++------------ src/Std/Time/DateTime/Timestamp.lean | 20 ++++---- src/Std/Time/Internal/Bounded.lean | 2 + src/Std/Time/Internal/UnitVal.lean | 2 +- src/Std/Time/Time/HourMarker.lean | 18 ++++--- src/Std/Time/Time/Unit/Basic.lean | 20 ++++++-- src/Std/Time/Time/Unit/Hour.lean | 22 +++++---- src/Std/Time/Time/Unit/Millisecond.lean | 14 ++++-- src/Std/Time/Time/Unit/Minute.lean | 22 +++++---- src/Std/Time/Time/Unit/Nanosecond.lean | 14 +++--- src/Std/Time/Time/Unit/Second.lean | 6 +-- src/Std/Time/Zoned/Database/TZdb.lean | 1 + src/Std/Time/Zoned/Database/TzIf.lean | 2 +- 21 files changed, 178 insertions(+), 142 deletions(-) diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index 087883e6c2d0..2107e43a8493 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -14,7 +14,8 @@ open Internal set_option linter.all true /-- -Date in YMD format. +`LocalDate` represents a date in the Year-Month-Day (YMD) format. It encapsulates the year, month, +and day components, with validation to ensure the date is valid. -/ structure LocalDate where @@ -46,7 +47,8 @@ instance : BEq LocalDate where namespace LocalDate /-- -Force the date to be valid. +Creates a `LocalDate` by clipping the day to ensure validity. This function forces the date to be +valid by adjusting the day to fit within the valid range to fit the given month and year. -/ def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : LocalDate := let ⟨day, valid⟩ := month.clipDay year.isLeap day @@ -56,7 +58,7 @@ instance : Inhabited LocalDate where default := clip 0 1 1 /-- -Creates a new `LocalDate` using YMD. +Creates a new `LocalDate` from year, month, and day components. -/ def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option LocalDate := if valid : year.Valid month day @@ -64,14 +66,14 @@ def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordin else none /-- -Creates a new `LocalDate` using YO. +Creates a `LocalDate` from a year and a day ordinal within that year. -/ def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : LocalDate := let ⟨⟨month, day⟩, valid⟩ := Month.Ordinal.ofOrdinal ordinal LocalDate.mk year month day valid /-- -Creates a new `LocalDate` using the `Day.Offset` which `0` corresponds the UNIX Epoch. +Creates a `LocalDate` from the number of days since the UNIX epoch (January 1st, 1970). -/ def ofDaysSinceUNIXEpoch (day : Day.Offset) : LocalDate := let z := day.toInt + 719468 @@ -88,7 +90,7 @@ def ofDaysSinceUNIXEpoch (day : Day.Offset) : LocalDate := .clip y (.clip m (by decide)) (.clip (d + 1) (by decide)) /-- -Returns the `Weekday` of a `LocalDate`. +Calculates the `Weekday` of a given `LocalDate` using Zeller's Congruence for the Gregorian calendar. -/ def weekday (date : LocalDate) : Weekday := let q := date.day.toInt @@ -107,7 +109,7 @@ def weekday (date : LocalDate) : Weekday := .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ /-- -Returns the `Weekday` of a `LocalDate` using Zeller's Congruence for the Julian calendar. +Calculates the `Weekday` of a given `LocalDate` using Zeller's Congruence for the Julian calendar. -/ def weekdayJulian (date : LocalDate) : Weekday := let month := date.month.toInt @@ -126,7 +128,7 @@ def weekdayJulian (date : LocalDate) : Weekday := .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ /-- -Convert `LocalDate` to `Day.Offset` since the UNIX Epoch. +Converts a `LocalDate` to the number of days since the UNIX epoch. -/ def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := let y : Int := if date.month.toInt > 2 then date.year else date.year.toInt - 1 @@ -140,13 +142,13 @@ def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := .ofInt (era * 146097 + doe - 719468) /-- -Calculate the Year.Offset from a LocalDate +Calculates the difference in years between a `LocalDate` and a given year. -/ def yearsSince (date : LocalDate) (year : Year.Offset) : Year.Offset := date.year - year /-- -Add `Day.Offset` to a `LocalDate`. +Adds a given number of days to a `LocalDate`. -/ @[inline] def addDays (date : LocalDate) (days : Day.Offset) : LocalDate := @@ -154,14 +156,14 @@ def addDays (date : LocalDate) (days : Day.Offset) : LocalDate := ofDaysSinceUNIXEpoch (Add.add dateDays days) /-- -Subtracts `Day.Offset` to a `LocalDate`. +Subtracts a given number of days from a `LocalDate`. -/ @[inline] def subDays (date : LocalDate) (days : Day.Offset) : LocalDate := addDays date (-days) /-- -Add `Month.Offset` to a `LocalDate`, it clips the day to the last valid day of that month. +Adds a given number of months to a `LocalDate`, clipping the day to the last valid day of the month. -/ def addMonthsClip (date : LocalDate) (months : Month.Offset) : LocalDate := let yearsOffset := months.div 12 @@ -188,8 +190,7 @@ def subMonthsClip (date : LocalDate) (months : Month.Offset) : LocalDate := addMonthsClip date (-months) /-- -Add `Month.Offset` to a `LocalDate`, this function rolls over any excess days into the following -month. +Adds a given number of months to a `LocalDate`, rolling over any excess days into the following month. -/ def addMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := let yearsOffset := months.div 12 @@ -222,38 +223,35 @@ def addMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := LocalDate.ofDaysSinceUNIXEpoch days /-- -Subtract `Month.Offset` from a `LocalDate`, this function rolls over any excess days into the following -month. +Subtracts `Month.Offset` from a `LocalDate`, rolling over excess days as needed. -/ @[inline] def subMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := addMonthsRollOver date (-months) /-- -Add `Year.Offset` to a `LocalDate`, this function rolls over any excess days into the following -month. +Adds `Year.Offset` to a `LocalDate`, rolling over excess days to the next month, or next year. -/ @[inline] def addYearsRollOver (date : LocalDate) (years : Year.Offset) : LocalDate := addMonthsRollOver date (years.mul 12) /-- -Add `Year.Offset` to a `LocalDate`, it clips the day to the last valid day of that month. +Subtracts `Year.Offset` from a `LocalDate`, rolling over excess days to the next month. -/ @[inline] -def addYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := - addMonthsClip date (years.mul 12) +def subYearsRollOver (date : LocalDate) (years : Year.Offset) : LocalDate := + addMonthsRollOver date (- years.mul 12) /-- -Subtract `Year.Offset` from a `LocalDate`, this function rolls over any excess days into the following -month. +Adds `Year.Offset` to a `LocalDate`, clipping the day to the last valid day of the month. -/ @[inline] -def subYearsRollOver (date : LocalDate) (years : Year.Offset) : LocalDate := - addMonthsRollOver date (- years.mul 12) +def addYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := + addMonthsClip date (years.mul 12) /-- -Subtract `Year.Offset` from a `LocalDate`, it clips the day to the last valid day of that month. +Subtracts `Year.Offset` from a `LocalDate`, clipping the day to the last valid day of the month. -/ @[inline] def subYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := diff --git a/src/Std/Time/Date/Unit/Basic.lean b/src/Std/Time/Date/Unit/Basic.lean index d8868f8a2972..d64675eddd1f 100644 --- a/src/Std/Time/Date/Unit/Basic.lean +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -9,3 +9,11 @@ import Std.Time.Date.Unit.Month import Std.Time.Date.Unit.Year import Std.Time.Date.Unit.Weekday import Std.Time.Date.Unit.WeekOfYear + +/-! +This module defines various units used for measuring, counting, and converting between days, months, +years, weekdays, and weeks of the year. + +The units are organized into types representing these time-related concepts, with operations provided +to facilitate conversions and manipulations between them. +-/ diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 401b78768319..cdabaacdef9e 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat import Std.Time.Time namespace Std @@ -20,7 +19,8 @@ set_option linter.all true def Ordinal := Bounded.LE 1 31 deriving Repr, BEq, LE, LT -instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (30 : Nat))) n) +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 1 (1 + (30 : Nat))) n) instance { x y : Ordinal } : Decidable (x ≤ y) := inferInstanceAs (Decidable (x.val ≤ y.val)) @@ -28,8 +28,8 @@ instance { x y : Ordinal } : Decidable (x ≤ y) := instance : Inhabited Ordinal where default := 1 /-- -`Ordinal.OfYear` represents a bounded value for days, which ranges between 0 and 366 if the year -is a leap year or 365. +`Ordinal.OfYear` represents the number of days in a year, accounting for leap years. It ensures that +the day is within the correct bounds—either 1 to 365 for regular years or 1 to 366 for leap years. -/ def Ordinal.OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 54be3a7de2f8..ae9eba8b2e26 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -20,9 +20,11 @@ set_option linter.all true def Ordinal := Bounded.LE 1 12 deriving Repr, BEq, LE -instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (11 : Nat))) n) +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 1 (1 + (11 : Nat))) n) -instance : Inhabited Ordinal where default := 1 +instance : Inhabited Ordinal where + default := 1 /-- `Offset` represents an offset in months. It is defined as an `Int`. @@ -30,7 +32,8 @@ instance : Inhabited Ordinal where default := 1 def Offset : Type := Int deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString -instance : OfNat Offset n := ⟨Int.ofNat n⟩ +instance : OfNat Offset n := + ⟨Int.ofNat n⟩ namespace Ordinal @@ -151,14 +154,15 @@ def toDays (leap : Bool) (month : Ordinal) : Day.Offset := |>.convert /-- -Size in days of each month if the year is not leap. +Size in days of each month if the year is not a leap year. -/ @[inline] def monthSizesNonLeap : { val : Array Day.Ordinal // val.size = 12 } := ⟨#[31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31], by simp⟩ /-- -Gets the number of days in a month without a proof of validity of the ordinal in a month and year. +Gets the number of days in a month without requiring a proof of the validity of the ordinal in a +month and year. -/ def daysWithoutProof (leap : Bool) (month : Ordinal) : Day.Ordinal := if month.val = 2 then @@ -180,14 +184,14 @@ instance : Decidable (Valid leap month day) := dite (day ≤ daysWithoutProof leap month) isTrue isFalse /-- -Gets the number of days in a month along side a proof of it's validity. +Gets the number of days in a month along with a proof of its validity. -/ @[inline] def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // Valid leap month day } := ⟨daysWithoutProof leap month, Int.le_refl ((daysWithoutProof leap month).val)⟩ /-- -Clips the day to be on the valid range. +Clips the day to be within the valid range. -/ @[inline] def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : Day.Ordinal // Valid leap month day } := @@ -199,7 +203,6 @@ def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : /-- Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. -/ -@[inline] def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } := Id.run do let rec go (idx : Fin 12) (cumulative : Fin 366) := let month := Month.Ordinal.ofFin idx.succ diff --git a/src/Std/Time/Date/Unit/WeekOfYear.lean b/src/Std/Time/Date/Unit/WeekOfYear.lean index 948f1fa73902..93d907b0d325 100644 --- a/src/Std/Time/Date/Unit/WeekOfYear.lean +++ b/src/Std/Time/Date/Unit/WeekOfYear.lean @@ -20,9 +20,11 @@ set_option linter.all true def Ordinal := Bounded.LE 1 53 deriving Repr, BEq, LE, LT -instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (52 : Nat))) n) +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 1 (1 + (52 : Nat))) n) -instance : Inhabited Ordinal where default := 1 +instance : Inhabited Ordinal where + default := 1 /-- `Offset` represents an offset in weeks. It is defined as an `Int`. diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index 72fefa8e3e22..9ea4664c2358 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -13,8 +13,9 @@ open Internal set_option linter.all true -/-- Defines the enumeration for days of the week. Each variant corresponds to a day of the week, -from Monday to Sunday. -/ +/-- +Defines the enumeration for days of the week. Each variant corresponds to a day of the week, from Monday to Sunday. +-/ inductive Weekday /-- Monday. -/ | mon @@ -35,8 +36,8 @@ inductive Weekday namespace Weekday /-- -Converts a `Fin 7` representing a day index into a corresponding `Weekday`. This function is -useful for mapping numerical representations to days of the week. +Converts a `Fin 7` representing a day index into a corresponding `Weekday`. This function is useful +for mapping numerical representations to days of the week. -/ def ofFin : Fin 7 → Weekday | 0 => .mon @@ -72,7 +73,7 @@ def toFin : Weekday → Nat | .sun => 6 /-- -Converts a `Nat` to a `Option Weekday`. +Converts a `Nat` to an `Option Weekday`. -/ def ofNat? : Nat → Option Weekday | 0 => some .mon diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index e3ffa44db328..4534d230d747 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -17,7 +17,7 @@ open Internal set_option linter.all true /-- -`Offset` represents an offset in years. It is defined as an `Int`. +`Offset` represents a year offset, defined as an `Int`. -/ def Offset : Type := Int deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString @@ -27,28 +27,28 @@ instance : OfNat Offset n := ⟨Int.ofNat n⟩ namespace Offset /-- -Convert the `Year` offset to an `Int`. +Converts the `Year` offset to an `Int`. -/ @[inline] def toInt (offset : Offset) : Int := offset /-- -Convert the `Year` offset to a `Month` Offset. +Converts the `Year` offset to a `Month` offset (i.e., multiplies by 12). -/ @[inline] def toMonths (val : Offset) : Month.Offset := val.mul 12 /-- -Checks if the `Year` is a Gregorian Leap Year. +Determines if a year is a Gregorian Leap Year. -/ @[inline] def isLeap (y : Offset) : Bool := y.toInt.mod 4 = 0 ∧ (y.toInt.mod 100 ≠ 0 ∨ y.toInt.mod 400 = 0) /-- -Forces the day to be on the valid range. +Checks if the given date is valid for the specified year, month, and day. -/ @[inline] abbrev Valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := diff --git a/src/Std/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean index fe46aa6d5f87..18ce67ea831e 100644 --- a/src/Std/Time/Date/WeekDate.lean +++ b/src/Std/Time/Date/WeekDate.lean @@ -45,10 +45,7 @@ Creates a `WeekDate` from a `Day.Offset`. def fromDays (scalar : Day.Offset) : WeekDate := let totalDays := scalar.val let year := totalDays / 365 - let week := - Bounded.LE.byEmod totalDays 365 (by decide) - |>.ediv 7 (by decide) - |>.add 1 + let week := Bounded.LE.byEmod totalDays 365 (by decide) |>.ediv 7 (by decide) |>.add 1 { year := year, week := week } end WeekDate diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean index fe4847feb4dc..160b6db28488 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -16,7 +16,7 @@ open Internal set_option linter.all true /-- -Date time format with Year, Month, Day, Hour, Minute, Seconds and Nanoseconds. +Represents a date and time with components for Year, Month, Day, Hour, Minute, Second, and Nanosecond. -/ structure LocalDateTime where @@ -35,7 +35,7 @@ structure LocalDateTime where namespace LocalDateTime /-- -Converts a `LocalDateTime` into a `Std.Time.Timestamp` +Converts a `LocalDateTime` to a `Timestamp` -/ def toLocalTimestamp (dt : LocalDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch @@ -44,7 +44,7 @@ def toLocalTimestamp (dt : LocalDateTime) : Timestamp := Timestamp.ofNanoseconds (UnitVal.mk nanos) /-- -Converts a UNIX `Timestamp` into a `LocalDateTime`. +Converts a UNIX `Timestamp` to a `LocalDateTime`. -/ def ofUTCTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do let leapYearEpoch := 11017 @@ -127,21 +127,22 @@ def ofUTCTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do } /-- -Add `Day.Offset` to a `LocalDateTime`. +Adds a `Day.Offset` to a `LocalDateTime`. -/ @[inline] def addDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := { dt with date := dt.date.addDays days } /-- -Subtracts `Day.Offset` to a `LocalDateTime`. +Subtracts a `Day.Offset` from a `LocalDateTime`. -/ @[inline] def subDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := { dt with date := dt.date.subDays days } /-- -Add `Month.Offset` to a `LocalDateTime`, it clips the day to the last valid day of that month. +Adds a `Month.Offset` to a `LocalDateTime`, adjusting the day to the last valid day of the resulting +month. -/ def addMonthsClip (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := { dt with date := dt.date.addMonthsClip months } @@ -154,105 +155,106 @@ def subMonthsClip (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime : { dt with date := dt.date.subMonthsClip months } /-- -Add `Month.Offset` to a `LocalDateTime`, this function rolls over any excess days into the following -month. +Adds a `Month.Offset` to a `LocalDateTime`, rolling over excess days to the following month if needed. -/ def addMonthsRollOver (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := { dt with date := dt.date.addMonthsRollOver months } /-- -Subtract `Month.Offset` from a `LocalDateTime`, this function rolls over any excess days into the following -month. +Subtracts a `Month.Offset` from a `LocalDateTime`, adjusting the day to the last valid day of the +resulting month. -/ @[inline] def subMonthsRollOver (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := { dt with date := dt.date.subMonthsRollOver months } /-- -Add `Year.Offset` to a `LocalDateTime`, this function rolls over any excess days into the following -month. +Adds a `Month.Offset` to a `LocalDateTime`, rolling over excess days to the following month if needed. -/ @[inline] def addYearsRollOver (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := { dt with date := dt.date.addYearsRollOver years } /-- -Add `Year.Offset` to a `LocalDateTime`, it clips the day to the last valid day of that month. +Subtracts a `Month.Offset` from a `LocalDateTime`, rolling over excess days to the following month if +needed. -/ @[inline] def addYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := { dt with date := dt.date.addYearsClip years } /-- -Subtracts `Year.Offset` from a `LocalDateTime`, this function rolls over any excess days into the following -month. +Subtracts a `Year.Offset` from a `LocalDateTime`, this function rolls over any excess days into the +following month. -/ @[inline] def subYearsRollOver (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := { dt with date := dt.date.subYearsRollOver years } /-- -Subtracts `Year.Offset` from a `LocalDateTime`, it clips the day to the last valid day of that month. +Subtracts a `Year.Offset` from a `LocalDateTime`, adjusting the day to the last valid day of the +resulting month. -/ @[inline] def subYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := { dt with date := dt.date.subYearsClip years } /-- -Get the current monotonic time. --/ -def now : IO LocalDateTime := - ofUTCTimestamp <$> Timestamp.now - -/-- -Getter for the `Year` inside of a `LocalDateTime` +Getter for the `Year` inside of a `LocalDateTime`. -/ @[inline] def year (dt : LocalDateTime) : Year.Offset := dt.date.year /-- -Getter for the `Month` inside of a `LocalDateTime` +Getter for the `Month` inside of a `LocalDateTime`. -/ @[inline] def month (dt : LocalDateTime) : Month.Ordinal := dt.date.month /-- -Getter for the `Day` inside of a `LocalDateTime` +Getter for the `Day` inside of a `LocalDateTime`. -/ @[inline] def day (dt : LocalDateTime) : Day.Ordinal := dt.date.day /-- -Getter for the `Hour` inside of a `LocalDateTime` +Getter for the `Hour` inside of a `LocalDateTime`. -/ @[inline] def hour (dt : LocalDateTime) : Hour.Ordinal dt.time.hour.fst := dt.time.hour.snd /-- -Getter for the `Minute` inside of a `LocalDateTime` +Getter for the `Minute` inside of a `LocalDateTime`. -/ @[inline] def minute (dt : LocalDateTime) : Minute.Ordinal := dt.time.minute /-- -Getter for the `Second` inside of a `LocalDateTime` +Getter for the `Second` inside of a `LocalDateTime`. -/ @[inline] def second (dt : LocalDateTime) : Second.Ordinal dt.time.second.fst := dt.time.second.snd /-- -Getter for the `Second` inside of a `LocalDateTime` +Getter for the `Second` inside of a `LocalDateTime`. -/ @[inline] def nanosecond (dt : LocalDateTime) : Nanosecond.Ordinal := dt.time.nano +/-- +Get the current monotonic time. +-/ +@[inline] +def now : IO LocalDateTime := + ofUTCTimestamp <$> Timestamp.now + end LocalDateTime end Time end Std diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 1264fd8fbc34..0284f65cf3c3 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -16,7 +16,7 @@ open Internal set_option linter.all true /-- -Nanoseconds since the UNIX Epoch. +Represents a point in time relative to the UNIX Epoch, with nanoseconds precision. -/ structure Timestamp where @@ -31,7 +31,7 @@ structure Timestamp where nano : Nanosecond.Span /-- - Proof that it's a valid timestamp. + Proof that the timestamp is valid, ensuring that the `second` and `nano` values are correctly related. -/ proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) deriving Repr @@ -45,20 +45,20 @@ instance : Inhabited Timestamp where namespace Timestamp /-- -Get the current monotonic time. +Fetches the current timestamp from the system. -/ @[extern "lean_get_current_time"] opaque now : IO Timestamp /-- -Transforms a `Timestamp` into a `Second.Offset` +Converts a `Timestamp` into its equivalent `Second.Offset`. -/ @[inline] def toSeconds (t : Timestamp) : Second.Offset := t.second /-- -Negates the `Timestamp` +Negates a `Timestamp`, flipping its second and nanosecond values. -/ @[inline] def neg (t : Timestamp) : Timestamp := by @@ -68,7 +68,7 @@ def neg (t : Timestamp) : Timestamp := by | inr n => exact Or.inl (n.imp Int.neg_le_neg Int.neg_le_neg) /-- -Adds two timestamps. +Adds two timestamps together, handling any carry-over in nanoseconds. -/ def add (t₁ t₂ : Timestamp) : Timestamp := by let diffSecs := t₁.second + t₂.second @@ -109,18 +109,16 @@ def add (t₁ t₂ : Timestamp) : Timestamp := by simp at h₃ exact Int.le_trans h₃ (by decide) else - let h := not_and.mp h - let h₁ := not_and.mp h₁ refine ⟨diffSecs, diffNano, ?_⟩ if h₂ : diffSecs.val > 0 then - exact Or.inl (And.intro (Int.le_of_lt h₂) (Int.not_lt.mp (h h₂))) + exact Or.inl (And.intro (Int.le_of_lt h₂) (Int.not_lt.mp (not_and.mp h h₂))) else if h₃ : diffSecs.val < 0 then - exact Or.inr (And.intro (Int.le_of_lt h₃) (Int.not_lt.mp (h₁ h₃))) + exact Or.inr (And.intro (Int.le_of_lt h₃) (Int.not_lt.mp (not_and.mp h₁ h₃))) else exact Int.le_total diffNano.val 0 |>.symm.imp (And.intro (Int.not_lt.mp h₃)) (And.intro (Int.not_lt.mp h₂)) /-- -Subtracts two `Timestamp`. +Subtracts one `Timestamp` from another. -/ @[inline] def sub (t₁ t₂ : Timestamp) : Timestamp := diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index 3f40b29a8395..1cb3d5d042cf 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -356,6 +356,7 @@ def mul_neg (bounded : Bounded.LE n m) (num : Int) (h : num ≤ 0) : Bounded.LE refine ⟨bounded.val * num, And.intro ?_ ?_⟩ · exact Int.mul_le_mul_of_nonpos_right bounded.property.right h · exact Int.mul_le_mul_of_nonpos_right bounded.property.left h + /-- Adjust the bounds of a `Bounded` by applying the div operation. -/ @@ -373,6 +374,7 @@ def ediv (bounded : Bounded.LE n m) (num : Int) (h : num > 0) : Bounded.LE (n / @[inline] def eq {n : Int} : Bounded.LE n n := ⟨n, And.intro (Int.le_refl n) (Int.le_refl n)⟩ + /-- Expand the bottom of the bounded to a number `nhi` is `hi` is less or equal to the previous higher bound. -/ diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index fd56590a63b7..5039147b6041 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -21,7 +21,7 @@ A structure representing a unit of a given ratio type `α`. structure UnitVal (α : Rat) where /-- - Value inside the UnitVal Value + Value inside the UnitVal Value. -/ val : Int deriving Inhabited, BEq diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 6249c27eed27..02e2242ec68b 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -25,15 +25,13 @@ inductive HourMarker /-- Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker`. -/ -def HourMarker.toAbsolute (marker : HourMarker) (time : Hour.Ordinal l) : Hour.Ordinal l := +def HourMarker.toAbsolute (marker : HourMarker) (time : Bounded.LE 0 12) : Hour.Ordinal l := match marker with - | .am => time + | .am => by + refine time.expandTop ?_ + split <;> decide | .pm => by - let mod := Int.mod_lt_of_pos (b := 24) (time.val + 12) (by decide) - let l := time.property.left - refine Bounded.LE.mk ((time.val + 12).mod 24) (And.intro ?_ ?_) - · have h : 0 + 0 ≤ time.val + 12 := Int.add_le_add l (by decide) - exact Int.mod_nonneg 24 h - · split - · exact Int.le_of_lt mod - · exact Int.le_sub_one_of_lt mod + have time := time.add 12 |>.emod 24 (by decide) + cases l + · exact time + · exact time.expandTop (by decide) diff --git a/src/Std/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean index ae449b9c271c..47e8cf276c27 100644 --- a/src/Std/Time/Time/Unit/Basic.lean +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -11,6 +11,14 @@ import Std.Time.Time.Unit.Second import Std.Time.Time.Unit.Nanosecond import Std.Time.Time.Unit.Millisecond +/-! +This module defines various units used for measuring, counting, and converting between hours, minutes, +second, nanosecond, millisecond and nanoseconds. + +The units are organized into types representing these time-related concepts, with operations provided +to facilitate conversions and manipulations between them. +-/ + namespace Std namespace Time namespace Second.Offset @@ -18,12 +26,16 @@ open Internal set_option linter.all true -/-- Convert `Second.Offset` to `Minute.Offset` -/ +/-- +Converts a `Second.Offset` to a `Minute.Offset`. +-/ @[inline] def toMinutes (offset : Second.Offset) : Minute.Offset := offset.ediv 60 -/-- Convert `Second.Offset` to `Hour.Offset` -/ +/-- +Converts a `Second.Offset` to an `Hour.Offset`. +-/ @[inline] def toHours (offset : Second.Offset) : Hour.Offset := offset.ediv 3600 @@ -32,7 +44,9 @@ end Second.Offset namespace Minute.Offset -/-- Convert `Minute.Offset` to `Hour.Offset` -/ +/-- +Converts a `Minute.Offset` to an `Hour.Offset`. +-/ @[inline] def toHours (offset : Minute.Offset) : Hour.Offset := offset.ediv 60 diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 6e018f5be6b6..f5b6d185bd64 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -17,8 +17,8 @@ open Internal set_option linter.all true /-- -`Ordinal` represents a bounded value for hour, which ranges between 0 and 24. It is bounded to 24 -because 24:00:00 is a valid date with leap seconds. +`Ordinal` represents a bounded value for hours, ranging from 0 to 24 or 23. The upper bound is 24 to +account for valid timestamps like 24:00:00 with leap seconds. -/ def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 24 else 23)) @@ -35,24 +35,27 @@ instance : OfNat (Ordinal true) 24 where ofNat := Bounded.LE.mk (Int.ofNat 24) (by decide) /-- -`Offset` represents an offset in hour. It is defined as an `Int`. +`Offset` represents an offset in hours, defined as an `Int`. This can be used to express durations +or differences in hours. -/ def Offset : Type := UnitVal 3600 deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString -instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +instance : OfNat Offset n := + ⟨UnitVal.ofNat n⟩ namespace Ordinal /-- -Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +Creates an `Ordinal` from a natural number, ensuring the value is within the valid bounds for hours. -/ @[inline] def ofNat (data : Nat) (h : data ≤ (if leap then 24 else 23)) : Ordinal leap := Bounded.LE.ofNat data h /-- -Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +Creates an `Ordinal` from a `Fin` value, ensuring the value is within bounds depending on whether +leap seconds are considered. -/ @[inline] def ofFin (data : Fin (if leap then 25 else 24)) : Ordinal leap := @@ -61,25 +64,24 @@ def ofFin (data : Fin (if leap then 25 else 24)) : Ordinal leap := | false => Bounded.LE.ofFin data /-- -Converts an `Ordinal` to an `Offset`. +Converts an `Ordinal` to an `Offset`, which represents the duration in hours as an integer value. -/ @[inline] def toOffset (ordinal : Ordinal leap) : Offset := UnitVal.ofInt ordinal.val end Ordinal - namespace Offset /-- -Convert the `Hour` offset to a `Second` Offset. +Converts an `Hour.Offset` to a `Second.Offset`. -/ @[inline] def toSeconds (val : Offset) : Second.Offset := val.mul 3600 /-- -Convert the `Hour` offset to a `Minute` Offset. +Converts an `Hour.Offset` to a `Minute.Offset`. -/ @[inline] def toMinutes (val : Offset) : Minute.Offset := diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index f1978d404380..1b995a486763 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -15,22 +15,26 @@ open Internal set_option linter.all true /-- -`Ordinal` represents a bounded value for milliseconds, which ranges between 0 and 999. +`Ordinal` represents a bounded value for milliseconds, ranging from 0 to 999 milliseconds. -/ def Ordinal := Bounded.LE 0 999 deriving Repr, BEq, LE, LT -instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (999 : Nat))) n) +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 0 (0 + (999 : Nat))) n) -instance : Inhabited Ordinal where default := 0 +instance : Inhabited Ordinal where + default := 0 /-- -`Offset` represents an offset in milliseconds. It is defined as an `Int`. +`Offset` represents a duration offset in milliseconds. It is defined as an `Int` value, +where each unit corresponds to one millisecond. -/ def Offset : Type := UnitVal (1 / 1000) deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString -instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +instance : OfNat Offset n := + ⟨UnitVal.ofNat n⟩ end Millisecond end Time diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index ac215e4c87e8..ce9fad7176f9 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -16,41 +16,45 @@ open Internal set_option linter.all true /-- -`Ordinal` represents a bounded value for minute, which ranges between 0 and 59. +`Ordinal` represents a bounded value for minutes, ranging from 0 to 59. This is useful for representing the minute component of a time. -/ def Ordinal := Bounded.LE 0 59 deriving Repr, BEq, LE -instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) n) +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) n) -instance : Inhabited Ordinal where default := 0 +instance : Inhabited Ordinal where + default := 0 /-- -`Offset` represents an offset in minute. It is defined as an `Int`. +`Offset` represents a duration offset in minutes. It is defined as an `Int` value. -/ def Offset : Type := UnitVal 60 deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString -instance : OfNat Offset n := ⟨UnitVal.ofInt <| Int.ofNat n⟩ +instance : OfNat Offset n := + ⟨UnitVal.ofInt <| Int.ofNat n⟩ namespace Ordinal /-- -Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +Creates an `Ordinal` from a natural number, ensuring the value is within the valid range (0 to 59). -/ @[inline] def ofNat (data : Nat) (h : data ≤ 59 := by decide) : Ordinal := Bounded.LE.ofNat data h /-- -Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +Creates an `Ordinal` from a `Fin` value, ensuring it is within the valid range (0 to 59). -/ @[inline] def ofFin (data : Fin 60) : Ordinal := Bounded.LE.ofFin data /-- -Converts an `Ordinal` to an `Offset`. +Converts an `Ordinal` to an `Offset`. The `Ordinal` value is converted to an `Offset` by interpreting +it as the number of minutes. -/ @[inline] def toOffset (ordinal : Ordinal) : Offset := @@ -61,7 +65,7 @@ end Ordinal namespace Offset /-- -Converts an `Offset` to `Second.Offset`. +Converts a `Minute.Offset` to `Second.Offset`. -/ @[inline] def toSeconds (val : Offset) : Second.Offset := diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index e8f08c129e25..de186ff053e4 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -15,14 +15,16 @@ open Internal set_option linter.all true /-- -`Ordinal` represents a bounded value for nanoseconds, which ranges between 0 and 999999999. +`Ordinal` represents a nanosecond value that is bounded between 0 and 999,999,999 nanoseconds. -/ def Ordinal := Bounded.LE 0 999999999 deriving Repr, BEq, LE, LT -instance : OfNat Ordinal n where ofNat := Bounded.LE.ofFin (Fin.ofNat n) +instance : OfNat Ordinal n where + ofNat := Bounded.LE.ofFin (Fin.ofNat n) -instance : Inhabited Ordinal where default := 0 +instance : Inhabited Ordinal where + default := 0 namespace Ordinal @@ -33,13 +35,13 @@ def OfDay := Bounded.LE 0 86400000000000 deriving Repr, BEq, LE, LT /-- -Convert to `Millisecond.Ordinal` +Converts a `Nanosecond.Ordinal` value to `Millisecond.Ordinal`. -/ def toMillisecond (nano : Ordinal) : Millisecond.Ordinal := nano.ediv 1000000 (by decide) /-- -Convert from `Millisecond.Ordinal` +Converts a `Millisecond.Ordinal` value to `Nanosecond.Ordinal`. -/ def ofMillisecond (nano : Millisecond.Ordinal) : Nanosecond.Ordinal := nano.mul_pos 1000000 (by decide) @@ -48,7 +50,7 @@ def ofMillisecond (nano : Millisecond.Ordinal) : Nanosecond.Ordinal := end Ordinal /-- -`Offset` represents an offset in nanoseconds. It is defined as an `Int`. +`Offset` represents a time offset in nanoseconds and is defined as an `Int`. -/ def Offset : Type := UnitVal (1 / 1000000000) deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 6d9f44359a3f..e6e2c5a47000 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -15,8 +15,8 @@ open Internal set_option linter.all true /-- -`Ordinal` represents a bounded value for second, which ranges between 0 and 60. -This accounts for potential leap second. +`Ordinal` represents a bounded value for second, which ranges between 0 and 59 or 60. This accounts +for potential leap second. -/ def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 60 else 59)) @@ -33,7 +33,7 @@ instance : OfNat (Ordinal true) 60 where ofNat := Bounded.LE.mk (Int.ofNat 60) (by decide) /-- -`Offset` represents an offset in second. It is defined as an `Int`. +`Offset` represents an offset in seconds. It is defined as an `Int`. -/ def Offset : Type := UnitVal 1 deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 7b36c9ed781a..70ca1f2d8b47 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -19,6 +19,7 @@ set_option linter.all true Represents a Time Zone Database (TZdb) configuration with paths to local and general timezone data. -/ structure TZdb where + /-- The path to the local timezone file. This is typically a symlink to a file within the timezone database that corresponds to the current local time zone. diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index d60facf09f48..081c7179a144 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -294,7 +294,7 @@ private def parseFooter : Parser (Option String) := do return str -private def parseTZifV2 : Parser (Option TZifV2) := optional $ do +private def parseTZifV2 : Parser (Option TZifV2) := optional do let header ← parseHeader let transitionTimes ← parseTransitionTimes pi64 header.timecnt From f72455142acbbdec841c40c2d6bedff7860d07b0 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 16 Aug 2024 14:23:50 -0300 Subject: [PATCH 021/118] feat: add some functions and tests --- src/Std/Time/Date/LocalDate.lean | 19 -------- src/Std/Time/Date/Unit/Month.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 6 +++ src/Std/Time/Format.lean | 9 +++- src/Std/Time/Format/Basic.lean | 46 ++++++++++++------- src/Std/Time/Internal/Bounded.lean | 10 ++++ tests/lean/timeLocalDateTime.lean | 69 ++++++++++++++++++++++++++++ tests/lean/timeParse.lean | 64 +++++++++++++++++++------- 8 files changed, 171 insertions(+), 54 deletions(-) create mode 100644 tests/lean/timeLocalDateTime.lean diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index 2107e43a8493..d40eafc26a4f 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -108,25 +108,6 @@ def weekday (date : LocalDate) : Weekday := .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ -/-- -Calculates the `Weekday` of a given `LocalDate` using Zeller's Congruence for the Julian calendar. --/ -def weekdayJulian (date : LocalDate) : Weekday := - let month := date.month.toInt - let year := date.year.toInt - - let q := date.day.toInt - let m := if month < 3 then month + 12 else month - let y := if month < 3 then year - 1 else year - - let k := y % 100 - let j := y / 100 - - let h := (q + (13 * (m + 1)) / 5 + k + (k / 4) + 5 - j) % 7 - let d := (h + 5 - 1) % 7 - - .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ - /-- Converts a `LocalDate` to the number of days since the UNIX epoch. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index ae9eba8b2e26..695b78893ae3 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -98,7 +98,7 @@ The ordinal value representing the month of December. @[inline] def december : Ordinal := 12 /-- -Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +Creates an `Ordinal` from a `Nat`, ensuring the value is within bounds. -/ @[inline] def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 12 := by decide) : Ordinal := diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 0284f65cf3c3..0ebfa82c032c 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -42,6 +42,12 @@ instance : BEq Timestamp where instance : Inhabited Timestamp where default := ⟨0, Bounded.LE.mk 0 (by decide), by decide⟩ +instance : OfNat Timestamp n where + ofNat := by + refine ⟨UnitVal.mk n, ⟨0, by decide⟩, ?_⟩ + simp <;> exact Int.le_total s.val 0 |>.symm + exact Int.le_total 0 n + namespace Timestamp /-- diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index fa3d78839b83..3987667d759b 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -6,6 +6,7 @@ Authors: Sofia Rodrigues prelude import Std.Time.Notation import Std.Time.Format.Basic +import Std.Time.Internal.Bounded namespace Std namespace Time @@ -125,8 +126,12 @@ def toTime24Hour (input : LocalTime) : String := /-- Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `LocalTime`. -/ -def fromTime12Hour (input : String) : Except String LocalTime := - Formats.Time12Hour.parseBuilder (λh m s a => LocalTime.ofHourMinuteSeconds? (HourMarker.toAbsolute a h.snd) m s.snd) input +def fromTime12Hour (input : String) : Except String LocalTime := do + let builder h m s a : Option LocalTime := do + let value ← Internal.Bounded.ofInt? h.snd.val + LocalTime.ofHourMinuteSeconds? (leap₂ := false) (HourMarker.toAbsolute a value) m s.snd + + Formats.Time12Hour.parseBuilder builder input /-- Formats a `LocalTime` value into a 12-hour format string (`hh:mm:ss aa`). diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index f5645f232466..3e7eba99fecb 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -565,12 +565,22 @@ private structure DateBuilder where minute : Minute.Ordinal := 0 second : Sigma Second.Ordinal := ⟨true, 0⟩ millisecond : Millisecond.Ordinal := 0 + marker : Option HourMarker := none -private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except String aw.type := - if let .isTrue p := inferInstanceAs (Decidable (ValidTime builder.hour.snd builder.minute builder.second.snd)) then +private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except String aw.type := do + + let hour ← + if let some marker := builder.marker then + match Bounded.ofInt? builder.hour.snd.val with + | some res => pure ⟨true, marker.toAbsolute res⟩ + | none => .error "The 24-hour is out of the range and cannot be transformed into a 12-hour with a marker." + else + pure builder.hour + + if let .isTrue p := inferInstanceAs (Decidable (ValidTime hour.snd builder.minute builder.second.snd)) then let build := DateTime.ofLocalDateTime { date := LocalDate.clip builder.year builder.month builder.day - time := LocalTime.mk builder.hour builder.minute builder.second (.ofMillisecond builder.millisecond) p + time := LocalTime.mk hour builder.minute builder.second (.ofMillisecond builder.millisecond) p } match aw with @@ -581,23 +591,27 @@ private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except else .error "invalid leap seconds {} {} {}" -private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : SingleFormatType typ) : DateBuilder := +private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : SingleFormatType typ) : Except String DateBuilder := match typ with - | .YYYY | .YY => { data with year := value } - | .MMMM | .MMM | .MM | .M => { data with month := value } - | .DD | .D | .d => { data with day := value } - | .EEEE | .EEE => data - | .hh | .h | .HH | .H => { data with hour := value } - | .AA | .aa => { data with hour := ⟨data.hour.fst, HourMarker.toAbsolute value data.hour.snd⟩ } - | .mm | .m => { data with minute := value } - | .sss => { data with millisecond := value } - | .ss | .s => { data with second := value } + | .AA | .aa => do pure { data with marker := some value } + | .YYYY | .YY => .ok { data with year := value } + | .MMMM | .MMM | .MM | .M => .ok { data with month := value } + | .DD | .D | .d => .ok { data with day := value } + | .EEEE | .EEE => .ok data + | .hh | .h | .HH | .H => .ok { data with hour := value } + | .mm | .m => .ok { data with minute := value } + | .sss => .ok { data with millisecond := value } + | .ss | .s => .ok { data with second := value } | .ZZZZZ | .ZZZZ | .ZZZ - | .Z => { data with tz := value } - | .z => { data with tzName := value } + | .Z => .ok { data with tz := value } + | .z => .ok { data with tzName := value } private def formatParser (date : DateBuilder) : FormatPart → Parser DateBuilder - | .modifier mod => addDataInDateTime date mod <$> parserWithFormat mod + | .modifier mod => do + let res ← addDataInDateTime date mod <$> parserWithFormat mod + match res with + | .error err => fail err + | .ok res => pure res | .string s => skipString s *> pure date namespace Format diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index 1cb3d5d042cf..61f80855988b 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -63,6 +63,16 @@ Creates a new `Bounded` Integer. def mk {rel : Int → Int → Prop} (val : Int) (proof : rel lo val ∧ rel val hi) : @Bounded rel lo hi := ⟨val, proof⟩ +/-- +Convert a `Int` to a `Bounded` if it checks. +-/ +@[inline] +def ofInt? [DecidableRel rel] (val : Int) : Option (Bounded rel lo hi) := + if h : rel lo ↑val ∧ rel ↑val hi then + some ⟨val, h⟩ + else + none + namespace LE /-- diff --git a/tests/lean/timeLocalDateTime.lean b/tests/lean/timeLocalDateTime.lean new file mode 100644 index 000000000000..48d532582442 --- /dev/null +++ b/tests/lean/timeLocalDateTime.lean @@ -0,0 +1,69 @@ +import Std.Time +open Std.Time + +def ShortDateTime : Format .any := date-spec% "DD/MM/YYYY hh:mm:ss" +def ShortDate : Format .any := date-spec% "DD/MM/YYYY" +def format (localDate : LocalDateTime) : String := ShortDateTime.formatBuilder localDate.day localDate.month localDate.year localDate.time.hour localDate.minute localDate.time.second +def format₂ (localDate : LocalDate) : String := ShortDate.formatBuilder localDate.day localDate.month localDate.year + +def date₁ := date% 19-11-1993:09:08:07 +def date₂ := date% 09-05-1993:12:59:59 +def date₃ := date% 16-08-2024 +def date₄ := date% 16-08-1500 + +def tm₁ := 753700087 +def tm₂ := 736952399 + +/-- +info: "19/11/1993 09:08:07" +-/ +#guard_msgs in +#eval format date₁ + +/-- +info: "09/05/1993 12:59:59" +-/ +#guard_msgs in +#eval format date₂ + +/-- +info: 753700087 +-/ +#guard_msgs in +#eval date₁.toLocalTimestamp.toSeconds + +/-- +info: 736952399 +-/ +#guard_msgs in +#eval date₂.toLocalTimestamp.toSeconds + +/-- +info: "09/05/1993 12:59:59" +-/ +#guard_msgs in +#eval LocalDateTime.ofUTCTimestamp 736952399 |> format + +/-- +info: 736952399 +-/ +#guard_msgs in +#eval LocalDateTime.toLocalTimestamp date₂ |>.toSeconds + +/-- +info: "16/08/2024" +-/ +#guard_msgs in +#eval LocalDate.ofDaysSinceUNIXEpoch 19951 |> format₂ + +/-- +info: 19951 +-/ +#guard_msgs in +#eval LocalDate.toDaysSinceUNIXEpoch date₃ + +/-- +info: Std.Time.Weekday.fri +-/ +#guard_msgs in +#eval LocalDate.weekday date₃ diff --git a/tests/lean/timeParse.lean b/tests/lean/timeParse.lean index 2dc6785d5bd9..e375d93668c2 100644 --- a/tests/lean/timeParse.lean +++ b/tests/lean/timeParse.lean @@ -6,12 +6,14 @@ def RFC1123 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" def ShortDate : Format .any := date-spec% "MM/DD/YYYY" def LongDate : Format .any := date-spec% "MMMM D, YYYY" def ShortDateTime : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss" -def LongDateTime : Format .any := date-spec% "MMMM D, YYYY h:mm aa" +def LongDateTime : Format .any := date-spec% "MMMM D, YYYY h:mm AA" def Time24Hour : Format .any := date-spec% "hh:mm:ss" def Time12Hour : Format .any := date-spec% "HH:mm:ss aa" def FullDayTimeZone : Format .any := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss ZZZZ" def CustomDayTime : Format .any := date-spec% "EEE D MMM YYYY hh:mm" +def Full12HourWrong : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss aa Z" + -- Dates def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 @@ -26,7 +28,7 @@ info: "2014-06-16T03:03:03-0300" -/ #guard_msgs in #eval - let t : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" + let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03-0300" ISO8601UTC.format t.snd def tm := date₁.timestamp @@ -37,22 +39,20 @@ info: "2014-06-16T03:03:03-0300" -/ #guard_msgs in #eval - let t : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" + let t : ZonedDateTime := RFC1123.parse! "Mon, 16 Jun 2014 03:03:03 -0300" ISO8601UTC.format t.snd def tm₃ := date₁.toTimestamp def date₃ := DateTime.ofUTCTimestamp tm₃ brTZ /-- -info: "2014-06-16T03:03:03-0300" +info: "2014-06-16T00:00:00UTC" -/ #guard_msgs in #eval - let t : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" + let t : ZonedDateTime := ShortDate.parse! "06/16/2014" ISO8601UTC.format t.snd --- Section for testing timezone conversion. - -- the timestamp is always related to UTC. /-- @@ -75,27 +75,27 @@ info: "2024-08-15T13:28:12-0300" ISO8601UTC.format t.snd /-- -info: "2024-08-16T01:28:12+0900" +info: "2024-08-16T01:28:00UTC" -/ #guard_msgs in #eval - let t : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" + let t : ZonedDateTime := LongDateTime.parse! "August 16, 2024 1:28 AM" ISO8601UTC.format t.snd /-- -info: "2024-08-15T13:28:12-0300" +info: "00-1-12-31T22:28:12+0900" -/ #guard_msgs in #eval - let t : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 13:28:12 -0300" - ISO8601UTC.format t.snd + let t : ZonedDateTime := Time24Hour.parse! "13:28:12" + ISO8601UTC.format (t.snd.convertTimeZone jpTZ) /-- -info: "2024-08-15T13:28:12-0300" +info: "00-1-12-31T09:28:12-0300" -/ #guard_msgs in #eval - let t1 : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 13:28:12 -0300" + let t1 : ZonedDateTime := Time12Hour.parse! "12:28:12 am" ISO8601UTC.format (t1.snd.convertTimeZone brTZ) /-- @@ -107,11 +107,11 @@ info: "Thu 15 Aug 2024 16:28" CustomDayTime.format t2.snd /-- -info: "2024-08-16T01:28:12+0900" +info: "2024-08-16T13:28:00UTC" -/ #guard_msgs in #eval - let t5 : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" + let t5 : ZonedDateTime := CustomDayTime.parse! "Thu 16 Aug 2024 13:28" ISO8601UTC.format t5.snd /-- @@ -175,3 +175,35 @@ info: "2014-06-16T03:03:03-0300" #eval let t2 : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" ISO8601UTC.format t2.snd + +/-- +info: Except.ok "1993-05-10T10:30:23+0300" +-/ +#guard_msgs in +#eval + let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 am +03:00" + (ISO8601UTC.format ·.snd) <$> t2 + +/-- +info: Except.ok "1993-05-10T22:30:23+0300" +-/ +#guard_msgs in +#eval + let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 pm +03:00" + (ISO8601UTC.format ·.snd) <$> t2 + +/-- +info: Except.error "offset 29: The 24-hour is out of the range and cannot be transformed into a 12-hour with a marker." +-/ +#guard_msgs in +#eval + let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 am +03:00" + (ISO8601UTC.format ·.snd) <$> t2 + +/-- +info: Except.error "offset 29: The 24-hour is out of the range and cannot be transformed into a 12-hour with a marker." +-/ +#guard_msgs in +#eval + let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 pm +03:00" + (ISO8601UTC.format ·.snd) <$> t2 From 248ed5ed426a6d53add241cff58ccb1b8554bfc2 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 16 Aug 2024 14:42:49 -0300 Subject: [PATCH 022/118] feat: add missing functions --- src/Std/Time/Format.lean | 37 +++++++++++++++++++++++++++ src/Std/Time/Time/HourMarker.lean | 2 +- src/Std/Time/Zoned/DateTime.lean | 4 +-- src/Std/Time/Zoned/ZonedDateTime.lean | 7 +++++ 4 files changed, 47 insertions(+), 3 deletions(-) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 3987667d759b..0a1eaa57ec82 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -196,6 +196,43 @@ def parse (input : String) : Except String ZonedDateTime := end ZonedDateTime +namespace LocalDateTime + +/-- +Parses a `String` in the `AscTime` format and returns a `LocalDateTime` object in the GMT time zone. +-/ +def fromAscTimeString (input : String) : Except String LocalDateTime := + Formats.AscTime.parse input + |>.map DateTime.toLocalDateTime + +/-- +Formats a `LocalDateTime` value into an AscTime format string. +-/ +def toAscTimeString (ldt : LocalDateTime) : String := + Formats.AscTime.format (DateTime.ofLocalDateTime ldt .UTC) + +/-- +Parses a `String` in the `LongDateFormat` and returns a `LocalDateTime` object in the GMT time zone. +-/ +def fromLongDateFormatString (input : String) : Except String LocalDateTime := + Formats.LongDateFormat.parse input + |>.map DateTime.toLocalDateTime + +/-- +Formats a `LocalDateTime` value into a LongDateFormat string. +-/ +def toLongDateFormatString (ldt : LocalDateTime) : String := + Formats.LongDateFormat.format (DateTime.ofLocalDateTime ldt .UTC) + +/-- +Parses a `String` in the `AscTime` or `LongDate` format and returns a `LocalDateTime`. +-/ +def parse (date : String) : Except String LocalDateTime := + fromAscTimeString date + <|> fromLongDateFormatString date + +end LocalDateTime + namespace DateTime /-- diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 02e2242ec68b..3f87610f4bf5 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -13,7 +13,7 @@ open Internal set_option linter.all true /-- -`HourMarker` represents the two 12-hour periods of the day: `am` for hour between 12:00 AM and +`HourMarker` represents the two 12-hour periods of the day: `am` for hours between 12:00 AM and 11:59 AM, and `pm` for hours between 12:00 PM and 11:59 PM. -/ inductive HourMarker diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 173d7502889d..712e6700008b 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -151,8 +151,8 @@ def subYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := Converts a `Timestamp` to a `LocalDateTime` -/ @[inline] -def toLocalDateTime (timestamp : Timestamp) : LocalDateTime := - LocalDateTime.ofUTCTimestamp timestamp +def toLocalDateTime (dt : DateTime tz) : LocalDateTime := + dt.date.get /-- Getter for the `Year` inside of a `DateTime` diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index d7eef4efcddb..9ecd035e7804 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -66,6 +66,13 @@ Creates a new `ZonedDateTime` out of a `LocalDateTime` def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : ZonedDateTime := ⟨tz, DateTime.ofLocalDateTime date tz⟩ +/-- +Converts a `ZonedDateTime` to a `LocalDateTime` +-/ +@[inline] +def toLocalDateTime (dt : ZonedDateTime) : LocalDateTime := + DateTime.toLocalDateTime dt.snd + /-- Gets the current `ZonedDataTime`. -/ From 58a446b045d50b14f368d3b23ea3d8641473fa02 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 16 Aug 2024 14:43:40 -0300 Subject: [PATCH 023/118] fix: wrong comments --- src/Std/Time/Format.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 0a1eaa57ec82..7b6602a83f89 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -266,13 +266,13 @@ def toISO8601String (date : DateTime tz) : String := Formats.ISO8601.format date /-- -Formats a `ZonedDateTime` value into an RFC822 format string. +Formats a `DateTime` value into an RFC822 format string. -/ def toRFC822String (date : DateTime tz) : String := Formats.RFC822.format date /-- -Formats a `ZonedDateTime` value into an RFC850 format string. +Formats a `DateTime` value into an RFC850 format string. -/ def toRFC850String (date : DateTime tz) : String := Formats.RFC850.format date From c480249a1194da8d528813a7814d2ed2059c72c9 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 19 Aug 2024 08:57:08 -0300 Subject: [PATCH 024/118] fix: some small fixes and add things to duration --- src/Std/Time/Date/Unit/Day.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 3 +- src/Std/Time/Duration.lean | 26 ++++++- src/Std/Time/Format.lean | 78 ++++++++++----------- src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Zoned/Database/Basic.lean | 2 +- src/runtime/io.cpp | 2 +- tests/lean/{ => run}/timeFormats.lean | 0 tests/lean/{ => run}/timeLocalDateTime.lean | 0 tests/lean/{ => run}/timeParse.lean | 0 tests/lean/{ => run}/timeTzifParse.lean | 0 11 files changed, 69 insertions(+), 46 deletions(-) rename tests/lean/{ => run}/timeFormats.lean (100%) rename tests/lean/{ => run}/timeLocalDateTime.lean (100%) rename tests/lean/{ => run}/timeParse.lean (100%) rename tests/lean/{ => run}/timeTzifParse.lean (100%) diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index cdabaacdef9e..dfcf055bf68a 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -22,7 +22,7 @@ def Ordinal := Bounded.LE 1 31 instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (30 : Nat))) n) -instance { x y : Ordinal } : Decidable (x ≤ y) := +instance {x y : Ordinal} : Decidable (x ≤ y) := inferInstanceAs (Decidable (x.val ≤ y.val)) instance : Inhabited Ordinal where default := 1 diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 0ebfa82c032c..2da8da394c96 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -74,7 +74,8 @@ def neg (t : Timestamp) : Timestamp := by | inr n => exact Or.inl (n.imp Int.neg_le_neg Int.neg_le_neg) /-- -Adds two timestamps together, handling any carry-over in nanoseconds. +Adds two timestamps together, handling any carry-over in nanoseconds. It should not be used for `Timestamp`. +The subtraction of two `Timestamp` returns a duration but the addition does not make sense at all. -/ def add (t₁ t₂ : Timestamp) : Timestamp := by let diffSecs := t₁.second + t₂.second diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 6d2625dc592b..194163c38b33 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -16,11 +16,33 @@ set_option linter.all true Duration is just a period between two timestamps. -/ def Duration := Timestamp + deriving Repr, BEq + +instance : ToString Duration where + toString s := + let (sign, secs, nanos) := + if s.second.val > 0 then ("" ,s.second.val, s.nano.val) + else if s.second.val < 0 then ("-", -s.second.val, -s.nano.val) + else if s.nano.val < 0 then ("-", -s.second.val, -s.nano.val) else ("", s.second.val, s.nano.val) + sign ++ toString secs ++ "." ++ toString nanos ++ "s" namespace Duration /-- Calculates a `Duration` out of two `Timestamp`s. -/ -def since (f : Timestamp) (t : Timestamp) : Duration := - t.sub f +def since (f : Timestamp) : IO Duration := do + let cur ← Timestamp.now + return cur.sub f + +/-- +Adds a `Duration` to a `Timestamp`. +-/ +def add (f : Timestamp) (d : Duration) : Timestamp := + f.add d + +/-- +Checks if the duration is zero seconds ands zero nanoseconds. +-/ +def isZero (d : Duration) : Bool := + d.second.val = 0 ∧ d.nano.val = 0 diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 7b6602a83f89..d54cf6f57721 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -18,59 +18,59 @@ set_option linter.all true The ISO8601 format, which is always 24 or 27 characters long, used for representing date and time in a standardized format. The format follows the pattern `YYYY-MM-DD'T'hh:mm:ss'Z'`. -/ -def ISO8601 : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" +def iso8601 : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" /-- -The AmericanDate format, which follows the pattern `MM/DD/YYYY`. +The americanDate format, which follows the pattern `MM/DD/YYYY`. -/ -def AmericanDate : Format .any := date-spec% "MM-DD-YYYY" +def americanDate : Format .any := date-spec% "MM-DD-YYYY" /-- -The EuropeanDate format, which follows the pattern `DD/MM/YYYY`. +The europeanDate format, which follows the pattern `DD-MM-YYYY`. -/ -def EuropeanDate : Format .any := date-spec% "DD-MM-YYYY" +def europeanDate : Format .any := date-spec% "DD-MM-YYYY" /-- -The Time12Hour format, which follows the pattern `HH:mm:ss aa` for representing time -in a 12-hour clock format with an AM/PM marker. +The time12Hour format, which follows the pattern `HH:mm:ss AA` for representing time +in a 12-hour clock format with an upper case AM/PM marker. -/ -def Time12Hour : Format .any := date-spec% "HH:mm:ss aa" +def time12Hour : Format .any := date-spec% "HH:mm:ss AA" /-- The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time in a 24-hour clock format. -/ -def Time24Hour : Format .any := date-spec% "hh:mm:ss" +def time24Hour : Format .any := date-spec% "hh:mm:ss" /-- The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used in SQL databases to represent dates. -/ -def SQLDate : Format .any := date-spec% "YYYY-MM-DD" +def sqlDate : Format .any := date-spec% "YYYY-MM-DD" /-- The LongDateFormat, which follows the pattern `EEEE, MMMM D, YYYY hh:mm:ss` for representing a full date and time with the day of the week and month name. -/ -def LongDateFormat : Format (.only .GMT) := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss" +def longDateFormat : Format (.only .GMT) := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss" /-- The AscTime format, which follows the pattern `EEE MMM d hh:mm:ss YYYY`. This format is often used in older systems for logging and time-stamping events. -/ -def AscTime : Format (.only .GMT) := date-spec% "EEE MMM d hh:mm:ss YYYY" +def ascTime : Format (.only .GMT) := date-spec% "EEE MMM d hh:mm:ss YYYY" /-- The RFC822 format, which follows the pattern `EEE, DD MMM YYYY hh:mm:ss ZZZ`. This format is used in email headers and HTTP headers. -/ -def RFC822 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" +def rfc822 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" /-- The RFC850 format, which follows the pattern `EEEE, DD-MMM-YY hh:mm:ss ZZZ`. This format is an older standard for representing date and time in headers. -/ -def RFC850 : Format .any := date-spec% "EEEE, DD-MMM-YY hh:mm:ss ZZZ" +def rfc850 : Format .any := date-spec% "EEEE, DD-MMM-YY hh:mm:ss ZZZ" end Formats @@ -80,25 +80,25 @@ namespace LocalDate Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. -/ def fromAmericanDateString (input : String) : Except String LocalDate := do - Formats.AmericanDate.parseBuilder (λm d y => LocalDate.ofYearMonthDay y m d) input + Formats.americanDate.parseBuilder (λm d y => LocalDate.ofYearMonthDay y m d) input /-- Converts a Date in the American format (`MM/DD/YYYY`) into a `String`. -/ def toAmericanDateString (input : LocalDate) : String := - Formats.AmericanDate.formatBuilder input.month input.day input.year + Formats.americanDate.formatBuilder input.month input.day input.year /-- Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. -/ def fromSQLDateString (input : String) : Except String LocalDate := do - Formats.SQLDate.parseBuilder (λy m d => LocalDate.ofYearMonthDay y m d) input + Formats.sqlDate.parseBuilder (λy m d => LocalDate.ofYearMonthDay y m d) input /-- Converts a Date in the SQL format (`MM/DD/YYYY`) into a `String`. -/ def toSQLDateString (input : LocalDate) : String := - Formats.SQLDate.formatBuilder input.year input.month input.day + Formats.sqlDate.formatBuilder input.year input.month input.day /-- Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `LocalDate`. @@ -115,13 +115,13 @@ namespace LocalTime Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `LocalTime`. -/ def fromTime24Hour (input : String) : Except String LocalTime := - Formats.Time24Hour.parseBuilder (λh m s => LocalTime.ofHourMinuteSeconds? h.snd m s.snd) input + Formats.time24Hour.parseBuilder (λh m s => LocalTime.ofHourMinuteSeconds? h.snd m s.snd) input /-- Formats a `LocalTime` value into a 24-hour format string (`hh:mm:ss`). -/ def toTime24Hour (input : LocalTime) : String := - Formats.Time24Hour.formatBuilder input.hour input.minute input.second + Formats.time24Hour.formatBuilder input.hour input.minute input.second /-- Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `LocalTime`. @@ -131,13 +131,13 @@ def fromTime12Hour (input : String) : Except String LocalTime := do let value ← Internal.Bounded.ofInt? h.snd.val LocalTime.ofHourMinuteSeconds? (leap₂ := false) (HourMarker.toAbsolute a value) m s.snd - Formats.Time12Hour.parseBuilder builder input + Formats.time12Hour.parseBuilder builder input /-- Formats a `LocalTime` value into a 12-hour format string (`hh:mm:ss aa`). -/ def toTime12Hour (input : LocalTime) : String := - Formats.Time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.snd.val ≥ 12 then HourMarker.pm else HourMarker.am) + Formats.time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.snd.val ≥ 12 then HourMarker.pm else HourMarker.am) /-- Parses a `String` in the `Time12Hour` or `Time24Hour` format and returns a `LocalTime`. @@ -154,37 +154,37 @@ namespace ZonedDateTime Parses a `String` in the `ISO8601` format and returns a `ZonedDateTime`. -/ def fromISO8601String (input : String) : Except String ZonedDateTime := - Formats.ISO8601.parse input + Formats.iso8601.parse input /-- Formats a `ZonedDateTime` value into an ISO8601 string. -/ def toISO8601String (date : ZonedDateTime) : String := - Formats.ISO8601.format date.snd + Formats.iso8601.format date.snd /-- Parses a `String` in the `RFC822` format and returns a `ZonedDateTime`. -/ def fromRFC822String (input : String) : Except String ZonedDateTime := - Formats.RFC822.parse input + Formats.rfc822.parse input /-- Formats a `ZonedDateTime` value into an RFC822 format string. -/ def toRFC822String (date : ZonedDateTime) : String := - Formats.RFC822.format date.snd + Formats.rfc822.format date.snd /-- Parses a `String` in the `RFC850` format and returns a `ZonedDateTime`. -/ def fromRFC850String (input : String) : Except String ZonedDateTime := - Formats.RFC850.parse input + Formats.rfc850.parse input /-- Formats a `ZonedDateTime` value into an RFC850 format string. -/ def toRFC850String (date : ZonedDateTime) : String := - Formats.RFC850.format date.snd + Formats.rfc850.format date.snd /-- Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. @@ -202,27 +202,27 @@ namespace LocalDateTime Parses a `String` in the `AscTime` format and returns a `LocalDateTime` object in the GMT time zone. -/ def fromAscTimeString (input : String) : Except String LocalDateTime := - Formats.AscTime.parse input + Formats.ascTime.parse input |>.map DateTime.toLocalDateTime /-- Formats a `LocalDateTime` value into an AscTime format string. -/ def toAscTimeString (ldt : LocalDateTime) : String := - Formats.AscTime.format (DateTime.ofLocalDateTime ldt .UTC) + Formats.ascTime.format (DateTime.ofLocalDateTime ldt .UTC) /-- Parses a `String` in the `LongDateFormat` and returns a `LocalDateTime` object in the GMT time zone. -/ def fromLongDateFormatString (input : String) : Except String LocalDateTime := - Formats.LongDateFormat.parse input + Formats.longDateFormat.parse input |>.map DateTime.toLocalDateTime /-- Formats a `LocalDateTime` value into a LongDateFormat string. -/ def toLongDateFormatString (ldt : LocalDateTime) : String := - Formats.LongDateFormat.format (DateTime.ofLocalDateTime ldt .UTC) + Formats.longDateFormat.format (DateTime.ofLocalDateTime ldt .UTC) /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `LocalDateTime`. @@ -239,43 +239,43 @@ namespace DateTime Parses a `String` in the `AscTime` format and returns a `DateTime` object in the GMT time zone. -/ def fromAscTimeString (input : String) : Except String (DateTime .GMT) := - Formats.AscTime.parse input + Formats.ascTime.parse input /-- Formats a `DateTime` value into an AscTime format string. -/ def toAscTimeString (datetime : DateTime .GMT) : String := - Formats.AscTime.format datetime + Formats.ascTime.format datetime /-- Parses a `String` in the `LongDateFormat` and returns a `DateTime` object in the GMT time zone. -/ def fromLongDateFormatString (input : String) : Except String (DateTime .GMT) := - Formats.LongDateFormat.parse input + Formats.longDateFormat.parse input /-- Formats a `DateTime` value into a LongDateFormat string. -/ def toLongDateFormatString (datetime : DateTime .GMT) : String := - Formats.LongDateFormat.format datetime + Formats.longDateFormat.format datetime /-- Formats a `DateTime` value into an ISO8601 string. -/ def toISO8601String (date : DateTime tz) : String := - Formats.ISO8601.format date + Formats.iso8601.format date /-- Formats a `DateTime` value into an RFC822 format string. -/ def toRFC822String (date : DateTime tz) : String := - Formats.RFC822.format date + Formats.rfc822.format date /-- Formats a `DateTime` value into an RFC850 format string. -/ def toRFC850String (date : DateTime tz) : String := - Formats.RFC850.format date + Formats.rfc850.format date /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `DateTime`. diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 3e7eba99fecb..0ce16cc990da 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -618,7 +618,7 @@ namespace Format /-- Constructs a new `Format` specification for a date-time string. Modifiers can be combined to create -custom formats, such as %YYYY, MMMM, D". +custom formats, such as "YYYY, MMMM, D". -/ def spec (input : String) : Except String (Format tz) := do let string ← specParser.run input diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 57052c80b2d1..a4b7911d30e6 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -95,7 +95,7 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := then transitions := transitions.push result else .error s!"cannot convert transtiion {i} of the file" - .ok { leapSeconds,transitions, localTimes := times } + .ok { leapSeconds, transitions, localTimes := times } /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 035c5a600593..d0163e403be8 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -538,7 +538,7 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { long long secs = timestamp / 1000000000; long long nano = timestamp % 1000000000; - lean_object *lean_ts = lean_alloc_ctor(1, 2, 0); + lean_object *lean_ts = lean_alloc_ctor(0, 2, 0); lean_ctor_set(lean_ts, 0, lean_int_to_int(static_cast(secs))); lean_ctor_set(lean_ts, 1, lean_int_to_int(static_cast(nano))); diff --git a/tests/lean/timeFormats.lean b/tests/lean/run/timeFormats.lean similarity index 100% rename from tests/lean/timeFormats.lean rename to tests/lean/run/timeFormats.lean diff --git a/tests/lean/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean similarity index 100% rename from tests/lean/timeLocalDateTime.lean rename to tests/lean/run/timeLocalDateTime.lean diff --git a/tests/lean/timeParse.lean b/tests/lean/run/timeParse.lean similarity index 100% rename from tests/lean/timeParse.lean rename to tests/lean/run/timeParse.lean diff --git a/tests/lean/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean similarity index 100% rename from tests/lean/timeTzifParse.lean rename to tests/lean/run/timeTzifParse.lean From 9abb7b6877a7d07d721440fccd98f856ba6a0dcb Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 20 Aug 2024 10:13:50 -0300 Subject: [PATCH 025/118] fix: small bugs and tests --- src/Std/Time.lean | 2 +- src/Std/Time/Date/LocalDate.lean | 51 +++++++++----- src/Std/Time/Date/Unit/Day.lean | 17 +++++ src/Std/Time/Date/Unit/Month.lean | 13 +++- src/Std/Time/Date/Unit/Weekday.lean | 89 ++++++++++++------------ src/Std/Time/Format.lean | 16 ++++- src/Std/Time/Format/Basic.lean | 97 ++++++++++++++------------- src/Std/Time/Notation.lean | 28 +++++--- tests/lean/run/timeFormats.lean | 26 ++++++- tests/lean/run/timeLocalDateTime.lean | 24 +++++-- tests/lean/run/timeParse.lean | 4 +- 11 files changed, 236 insertions(+), 131 deletions(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 38f592318ca6..2cda878c1577 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -160,7 +160,7 @@ The supported formats include: In order to help the user build dates easily, there are a lot of macros available for creating dates. -- **`date% DD-MM-YYYY`**: Defines a date in the `DD-MM-YYYY` format. +- **`date% YYYY-MM-DD`**: Defines a date in the `YYYY-MM-DD` format. - **`date% HH:mm:ss`**: Defines a time in the `HH:mm:ss` format. - **`date% HH:mm:ss.sssssss`**: Defines a time in the `HH:mm:ss.sssssss` format, including fractional seconds. - **`date% YYYY-MM-DD:HH:mm:ss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss` format. diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index d40eafc26a4f..45771592dc4d 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -6,10 +6,12 @@ Authors: Sofia Rodrigues prelude import Std.Time.Internal import Std.Time.Date.Basic +import Lean.Data.Rat namespace Std namespace Time open Internal +open Lean set_option linter.all true @@ -54,6 +56,21 @@ def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Loca let ⟨day, valid⟩ := month.clipDay year.isLeap day LocalDate.mk year month day valid +/-- +Creates a `LocalDate` by rolling over the extra days to the next month. +-/ +def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : LocalDate := by + let ⟨max, valid, gt⟩ := month.days year.isLeap + let max : Bounded.LE 28 31 := max.truncateBottom gt + if h : day.val > max.val then + let day := day.truncateBottom h + let off := day.sub max.val + simp at off + sorry + else + sorry + + instance : Inhabited LocalDate where default := clip 0 1 1 @@ -77,17 +94,16 @@ Creates a `LocalDate` from the number of days since the UNIX epoch (January 1st, -/ def ofDaysSinceUNIXEpoch (day : Day.Offset) : LocalDate := let z := day.toInt + 719468 - let era := (if z ≥ 0 then z else z - 146096) / 146097 + let era := (if z ≥ 0 then z else z - 146096).div 146097 let doe := z - era * 146097 - let yoe := (doe - doe / 1460 + doe / 36524 - doe / 146096) / 365 + let yoe := (doe - doe.div 1460 + doe.div 36524 - doe.div 146096).div 365 let y := yoe + era * 400 - let doy := doe - (365 * yoe + yoe / 4 - yoe / 100) - let mp : Int := (5 * doy + 2) / 153 - let d := doy - (153 * mp + 2) / 5 + let doy := doe - (365 * yoe + yoe.div 4 - yoe.div 100) + let mp := (5 * doy + 2).div 153 + let d := doy - (153 * mp + 2).div 5 + 1 let m := mp + (if mp < 10 then 3 else -9) let y := y + (if m <= 2 then 1 else 0) - - .clip y (.clip m (by decide)) (.clip (d + 1) (by decide)) + .clip y (.clip m (by decide)) (.clip d (by decide)) /-- Calculates the `Weekday` of a given `LocalDate` using Zeller's Congruence for the Gregorian calendar. @@ -101,9 +117,9 @@ def weekday (date : LocalDate) : Weekday := let m := if m < 2 then m + 12 else m let k := y % 100 - let j := y / 100 - let part := q + (13 * (m + 1)) / 5 + k + (k / 4) - let h := if y ≥ 1582 then part + (j/4) - 2*j else part + 5 - j + let j := y.div 100 + let part := q + (13 * (m + 1)).div 5 + k + (k.div 4) + let h := part + (j.div 4) - 2*j let d := (h + 5) % 7 .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ @@ -113,15 +129,17 @@ Converts a `LocalDate` to the number of days since the UNIX epoch. -/ def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := let y : Int := if date.month.toInt > 2 then date.year else date.year.toInt - 1 - let era : Int := (if y ≥ 0 then y else y - 399) / 400 + let era : Int := (if y ≥ 0 then y else y - 399).div 400 let yoe : Int := y - era * 400 let m : Int := date.month.toInt let d : Int := date.day.toInt - let doy := (153 * (m + (if m > 2 then -3 else 9)) + 2) / 5 + d - 1 - let doe := yoe * 365 + yoe / 4 - yoe / 100 + doy + let doy := (153 * (m + (if m > 2 then -3 else 9)) + 2).div 5 + d - 1 + let doe := yoe * 365 + yoe.div 4 - yoe.div 100 + doy .ofInt (era * 146097 + doe - 719468) + + /-- Calculates the difference in years between a `LocalDate` and a given year. -/ @@ -134,7 +152,7 @@ Adds a given number of days to a `LocalDate`. @[inline] def addDays (date : LocalDate) (days : Day.Offset) : LocalDate := let dateDays := date.toDaysSinceUNIXEpoch - ofDaysSinceUNIXEpoch (Add.add dateDays days) + ofDaysSinceUNIXEpoch (dateDays + days) /-- Subtracts a given number of days from a `LocalDate`. @@ -189,13 +207,12 @@ def addMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := exact (yearsOffset, months.truncateTop (Int.not_lt.mp h₁) |>.truncateBottom (Int.not_lt.mp h₂)) let year : Year.Offset := date.year.add yearsOffset - let ⟨days, proof⟩ := Month.Ordinal.days year.isLeap months + let ⟨days, proof, _⟩ := Month.Ordinal.days year.isLeap months if h : days.val ≥ date.day.val then - let p : year.Valid months date.day := by + have p : year.Valid months date.day := by simp_all [Year.Offset.Valid, Month.Ordinal.Valid] exact Int.le_trans h proof - dbg_trace s!"roll {days.val} {date.day.val}" LocalDate.mk year months date.day p else let roll : Day.Offset := UnitVal.mk (date.day.val - days.toInt) diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index dfcf055bf68a..51ada9fc337f 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -25,6 +25,9 @@ instance : OfNat Ordinal n := instance {x y : Ordinal} : Decidable (x ≤ y) := inferInstanceAs (Decidable (x.val ≤ y.val)) +instance {x y : Ordinal} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + instance : Inhabited Ordinal where default := 1 /-- @@ -83,6 +86,20 @@ end Ordinal namespace Offset +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + UnitVal.mk data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + UnitVal.mk data + /-- Convert `Day.Offset` into `Second.Offset`. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 695b78893ae3..9a09c5c05007 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -173,6 +173,13 @@ def daysWithoutProof (leap : Bool) (month : Ordinal) : Day.Ordinal := rw [← p] at r exact months.get r +theorem all_greater_than_27 (leap : Bool) (i: Month.Ordinal) : daysWithoutProof leap i > 27 := by + simp [daysWithoutProof, monthSizesNonLeap, Bounded.LE.sub, Bounded.LE.add, Bounded.LE.toFin] + match i with + | ⟨2, _⟩ => split <;> (simp; try split); all_goals decide + | ⟨1, _⟩ | ⟨3, _⟩ | ⟨4, _⟩ | ⟨5, _⟩ | ⟨6, _⟩ | ⟨7, _⟩ + | ⟨8, _⟩ | ⟨9, _⟩ | ⟨10, _⟩ | ⟨11, _⟩ | ⟨12, _⟩ => simp; decide + /-- Check if the day is valid in a month and a leap year. -/ @@ -187,8 +194,8 @@ instance : Decidable (Valid leap month day) := Gets the number of days in a month along with a proof of its validity. -/ @[inline] -def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // Valid leap month day } := - ⟨daysWithoutProof leap month, Int.le_refl ((daysWithoutProof leap month).val)⟩ +def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // Valid leap month day ∧ day.val > 27 } := by + refine ⟨daysWithoutProof leap month, ⟨Int.le_refl ((daysWithoutProof leap month).val), all_greater_than_27 leap month⟩⟩ /-- Clips the day to be within the valid range. @@ -206,7 +213,7 @@ Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } := Id.run do let rec go (idx : Fin 12) (cumulative : Fin 366) := let month := Month.Ordinal.ofFin idx.succ - let ⟨days, valid⟩ := days leap month + let ⟨days, valid, _⟩ := days leap month if h : cumulative.val < ordinal.val ∧ ordinal.val ≤ cumulative.val + days.val then let bounded := Bounded.LE.mk ordinal.val h |>.sub cumulative diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index 9ea4664c2358..f1110aeb0d08 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -18,19 +18,19 @@ Defines the enumeration for days of the week. Each variant corresponds to a day -/ inductive Weekday /-- Monday. -/ - | mon + | monday /-- Tuesday. -/ - | tue + | tuesday /-- Wednesday. -/ - | wed + | wednesday /-- Thursday. -/ - | thu + | thursday /-- Friday. -/ - | fri + | friday /-- Saturday. -/ - | sat + | saturday /-- Sunday. -/ - | sun + | sunday deriving Repr, Inhabited, BEq namespace Weekday @@ -40,49 +40,49 @@ Converts a `Fin 7` representing a day index into a corresponding `Weekday`. This for mapping numerical representations to days of the week. -/ def ofFin : Fin 7 → Weekday - | 0 => .mon - | 1 => .tue - | 2 => .wed - | 3 => .thu - | 4 => .fri - | 5 => .sat - | 6 => .sun + | 0 => .monday + | 1 => .tuesday + | 2 => .wednesday + | 3 => .thursday + | 4 => .friday + | 5 => .saturday + | 6 => .sunday /-- Converts a `Weekday` to a `Nat`. -/ def toNat : Weekday → Nat - | .mon => 0 - | .tue => 1 - | .wed => 2 - | .thu => 3 - | .fri => 4 - | .sat => 5 - | .sun => 6 + | .monday => 0 + | .tuesday => 1 + | .wednesday => 2 + | .thursday => 3 + | .friday => 4 + | .saturday => 5 + | .sunday => 6 /-- Converts a `Weekday` to a `Fin`. -/ def toFin : Weekday → Nat - | .mon => 0 - | .tue => 1 - | .wed => 2 - | .thu => 3 - | .fri => 4 - | .sat => 5 - | .sun => 6 + | .monday => 0 + | .tuesday => 1 + | .wednesday => 2 + | .thursday => 3 + | .friday => 4 + | .saturday => 5 + | .sunday => 6 /-- Converts a `Nat` to an `Option Weekday`. -/ def ofNat? : Nat → Option Weekday - | 0 => some .mon - | 1 => some .tue - | 2 => some .wed - | 3 => some .thu - | 4 => some .fri - | 5 => some .sat - | 6 => some .sun + | 0 => some .monday + | 1 => some .tuesday + | 2 => some .wednesday + | 3 => some .thursday + | 4 => some .friday + | 5 => some .saturday + | 6 => some .sunday | _ => none /-- @@ -98,20 +98,19 @@ def ofNat! (n : Nat) : Weekday := Gets the next `Weekday`. -/ def next : Weekday → Weekday - | .mon => .sun - | .tue => .mon - | .wed => .tue - | .thu => .wed - | .fri => .thu - | .sat => .fri - | .sun => .sat - + | .monday => .tuesday + | .tuesday => .wednesday + | .wednesday => .thursday + | .thursday => .friday + | .friday => .saturday + | .saturday => .sunday + | .sunday => .monday /-- Check if it's a Weekend. -/ def weekend : Weekday → Bool - | .sat => true - | .sun => true + | .saturday => true + | .sunday => true | _ => false end Weekday diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index d54cf6f57721..852d0e796375 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -89,13 +89,13 @@ def toAmericanDateString (input : LocalDate) : String := Formats.americanDate.formatBuilder input.month input.day input.year /-- -Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. +Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. -/ def fromSQLDateString (input : String) : Except String LocalDate := do Formats.sqlDate.parseBuilder (λy m d => LocalDate.ofYearMonthDay y m d) input /-- -Converts a Date in the SQL format (`MM/DD/YYYY`) into a `String`. +Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. -/ def toSQLDateString (input : LocalDate) : String := Formats.sqlDate.formatBuilder input.year input.month input.day @@ -107,6 +107,9 @@ def parse (input : String) : Except String LocalDate := fromAmericanDateString input <|> fromSQLDateString input +instance : ToString LocalDate where + toString := toSQLDateString + end LocalDate namespace LocalTime @@ -146,6 +149,9 @@ def parse (input : String) : Except String LocalTime := fromTime12Hour input <|> fromTime24Hour input +instance : ToString LocalTime where + toString := toTime24Hour + end LocalTime namespace ZonedDateTime @@ -194,6 +200,9 @@ def parse (input : String) : Except String ZonedDateTime := <|> fromRFC822String input <|> fromRFC850String input +instance : ToString ZonedDateTime where + toString := toRFC822String + end ZonedDateTime namespace LocalDateTime @@ -284,4 +293,7 @@ def parse (date : String) : Except String (DateTime .GMT) := fromAscTimeString date <|> fromLongDateFormatString date +instance : ToString (DateTime tz) where + toString := toRFC822String + end DateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 0ce16cc990da..eddee8071f05 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -308,49 +308,54 @@ private def abbrevMonth (month : Month.Ordinal) : String := private def abbrevDayOfWeek (day : Weekday) : String := match day with - | .sun => "Sun" - | .mon => "Mon" - | .tue => "Tue" - | .wed => "Wed" - | .thu => "Thu" - | .fri => "Fri" - | .sat => "Sat" + | .sunday => "Sun" + | .monday => "Mon" + | .tuesday => "Tue" + | .wednesday => "Wed" + | .thursday => "Thu" + | .friday => "Fri" + | .saturday => "Sat" private def dayOfWeek (day : Weekday) : String := match day with - | .sun => "Sunday" - | .mon => "Monday" - | .tue => "Tuesday" - | .wed => "Wednesday" - | .thu => "Thursday" - | .fri => "Friday" - | .sat => "Saturday" + | .sunday => "Sunday" + | .monday => "Monday" + | .tuesday => "Tuesday" + | .wednesday => "Wednesday" + | .thursday => "Thursday" + | .friday => "Friday" + | .saturday => "Saturday" private def leftPad (n : Nat) (a : Char) (s : String) : String := "".pushn a (n - s.length) ++ s +private def leftPadNum (n : Nat) (s : Int) : String := + let str := if s < 0 then toString (-s) else toString s + let start := if s < 0 then "-" else "" + start ++ "".pushn '0' (n - str.length) ++ str + private def formatWithDate (date : DateTime tz) : Modifier → String - | .YYYY => s!"{leftPad 4 '0' (toString date.year)}" - | .YY => s!"{leftPad 2 '0' (toString <| date.year.toNat % 100)}" + | .YYYY => s!"{leftPadNum 4 (date.year)}" + | .YY => s!"{leftPadNum 2 (date.year.toNat % 100)}" | .MMMM => unabbrevMonth date.month | .MMM => abbrevMonth date.month - | .MM => s!"{leftPad 2 '0' (toString <| date.month.toNat)}" + | .MM => s!"{leftPadNum 2 (date.month.toNat)}" | .M => s!"{date.month.toNat}" - | .DD => s!"{leftPad 2 '0' (toString <| date.day.toNat)}" + | .DD => s!"{leftPadNum 2 (date.day.toNat)}" | .D => s!"{date.day.toNat}" | .d => s!"{leftPad 2 ' ' <| toString date.day.toNat}" | .EEEE => dayOfWeek date.weekday | .EEE => abbrevDayOfWeek date.weekday - | .hh => s!"{leftPad 2 '0' (toString date.hour.toNat)}" + | .hh => s!"{leftPadNum 2 (date.hour.toNat)}" | .h => s!"{date.hour.toNat}" - | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' <| toString hour}" + | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPadNum 2 <| hour}" | .H => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{hour}" | .AA => if date.hour.toNat < 12 then "AM" else "PM" | .aa => if date.hour.toNat < 12 then "am" else "pm" - | .mm => s!"{leftPad 2 '0' <| toString date.minute.toNat}" + | .mm => s!"{leftPadNum 2 <| date.minute.toNat}" | .m => s!"{date.minute.toNat}" - | .sss => s!"{leftPad 3 '0' <| toString date.milliseconds.toNat}" - | .ss => s!"{leftPad 2 '0' <| toString date.second.toNat}" + | .sss => s!"{leftPadNum 3 <| date.milliseconds.toNat}" + | .ss => s!"{leftPadNum 2 <| date.second.toNat}" | .s => s!"{date.second.toNat}" | .ZZZZZ => tz.offset.toIsoString true | .ZZZZ => tz.offset.toIsoString false @@ -380,27 +385,27 @@ private def SingleFormatType : Modifier → Type private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) : String := match modifier with - | .YYYY => s!"{leftPad 4 '0' (toString data.toNat)}" - | .YY => s!"{leftPad 2 '0' (toString <| data.toNat % 100)}" + | .YYYY => s!"{leftPadNum 4 (data.toNat)}" + | .YY => s!"{leftPadNum 2 (data.toNat % 100)}" | .MMMM => unabbrevMonth data | .MMM => abbrevMonth data - | .MM => s!"{leftPad 2 '0' (toString <| data.toNat)}" + | .MM => s!"{leftPadNum 2 (data.toNat)}" | .M => s!"{data.toNat}" - | .DD => s!"{leftPad 2 '0' (toString <| data.toNat)}" + | .DD => s!"{leftPadNum 2 (data.toNat)}" | .D => s!"{data.toNat}" | .d => s!"{leftPad 2 ' ' <| toString data.toNat}" | .EEEE => dayOfWeek data | .EEE => abbrevDayOfWeek data - | .hh => s!"{leftPad 2 '0' (toString data.snd.toNat)}" + | .hh => s!"{leftPadNum 2 (data.snd.toNat)}" | .h => s!"{data.snd.toNat}" - | .HH => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{leftPad 2 '0' <| toString hour}" + | .HH => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{leftPadNum 2 hour}" | .H => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{hour}" | .AA => match data with | .am => "AM" | .pm => "PM" | .aa => match data with | .am => "am" | .pm => "pm" - | .mm => s!"{leftPad 2 '0' <| toString data.toNat}" + | .mm => s!"{leftPadNum 2 data.toNat}" | .m => s!"{data.toNat}" - | .sss => s!"{leftPad 3 '0' <| toString data.toNat}" - | .ss => s!"{leftPad 2 '0' <| toString data.snd.toNat}" + | .sss => s!"{leftPadNum 3 data.toNat}" + | .ss => s!"{leftPadNum 2 data.snd.toNat}" | .s => s!"{data.snd.toNat}" | .ZZZZZ => data.toIsoString true | .ZZZZ => data.toIsoString false @@ -457,22 +462,22 @@ private def parseMonthUnabbrev : Parser Month.Ordinal <|> (pstring "December" *> pure 12) private def parseWeekday : Parser Weekday - := (pstring "Mon" *> pure Weekday.mon) - <|> (pstring "Tue" *> pure Weekday.tue) - <|> (pstring "Wed" *> pure Weekday.wed) - <|> (pstring "Thu" *> pure Weekday.thu) - <|> (pstring "Fri" *> pure Weekday.fri) - <|> (pstring "Sat" *> pure Weekday.sat) - <|> (pstring "Sun" *> pure Weekday.sun) + := (pstring "Mon" *> pure Weekday.monday) + <|> (pstring "Tue" *> pure Weekday.tuesday) + <|> (pstring "Wed" *> pure Weekday.wednesday) + <|> (pstring "Thu" *> pure Weekday.thursday) + <|> (pstring "Fri" *> pure Weekday.friday) + <|> (pstring "Sat" *> pure Weekday.saturday) + <|> (pstring "Sun" *> pure Weekday.sunday) private def parseWeekdayUnnabrev : Parser Weekday - := (pstring "Monday" *> pure Weekday.mon) - <|> (pstring "Tuesday" *> pure Weekday.tue) - <|> (pstring "Wednesday" *> pure Weekday.wed) - <|> (pstring "Thursday" *> pure Weekday.thu) - <|> (pstring "Friday" *> pure Weekday.fri) - <|> (pstring "Saturday" *> pure Weekday.sat) - <|> (pstring "Sunday" *> pure Weekday.sun) + := (pstring "Monday" *> pure Weekday.monday) + <|> (pstring "Tuesday" *> pure Weekday.tuesday) + <|> (pstring "Wednesday" *> pure Weekday.wednesday) + <|> (pstring "Thursday" *> pure Weekday.thursday) + <|> (pstring "Friday" *> pure Weekday.friday) + <|> (pstring "Saturday" *> pure Weekday.saturday) + <|> (pstring "Sunday" *> pure Weekday.sunday) private def parserUpperHourMarker : Parser HourMarker := (pstring "AM" *> pure HourMarker.am) diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 019f4cce8d02..3a1aac6135b6 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -35,8 +35,18 @@ Variables inside of the date. -/ syntax ident : date_component -private def parseComponent : TSyntax `date_component -> MacroM (TSyntax `term) - | `(date_component| $num:num) => `($num) +private def parseComponent (lower : Option Int := none) (upper : Option Int := none) : TSyntax `date_component -> MacroM (TSyntax `term) + | `(date_component| $num:num) => do + let res := Int.ofNat num.getNat + if let some lower := lower then + if res < lower then + Macro.throwErrorAt num s!"the number should be bigger than {lower}" + + if let some upper := upper then + if res > upper then + Macro.throwErrorAt num s!"the number should be lower than {upper}" + + `($num) | `(date_component| $str:str) => `($(Syntax.mkNumLit str.getString)) | `(date_component| $ident:ident) => `($ident:ident) | syn => Macro.throwErrorAt syn "unsupported syntax" @@ -52,8 +62,8 @@ Date in `DD-MM-YYYY` format. syntax date_component noWs "-" noWs date_component noWs "-" noWs date_component : date private def parseDate : TSyntax `date -> MacroM (TSyntax `term) - | `(date|$day:date_component-$month:date_component-$year:date_component) => do - `(Std.Time.LocalDate.mk $(← parseComponent year) $(← parseComponent month) $(← parseComponent day) (by decide)) + | `(date|$year:date_component-$month:date_component-$day:date_component) => do + `(Std.Time.LocalDate.mk $(← parseComponent none none year) $(← parseComponent (some 1) (some 12) month) $(← parseComponent (some 1) (some 31) day) (by decide)) | syn => Macro.throwErrorAt syn "unsupported type" /-- @@ -73,9 +83,9 @@ syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component private def parseTime : TSyntax `time -> MacroM (TSyntax `term) | `(time| $hour:date_component:$minute:date_component:$second:date_component) => do - `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent hour)⟩ $(← parseComponent minute) ⟨true, $(← parseComponent second)⟩ 0 (by decide)) + `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ 0 (by decide)) | `(time| $hour:date_component:$minute:date_component:$second:date_component.$nanos:date_component) => do - `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent hour)⟩ $(← parseComponent minute) ⟨true, $(← parseComponent second)⟩ $(← parseComponent nanos) (by decide)) + `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ $(← parseComponent (some 0) (some 999) nanos) (by decide)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -97,7 +107,7 @@ private def parseDateTime : TSyntax `datetime -> MacroM (TSyntax `term) | `(datetime| $date:date:$time:time) => do `(Std.Time.LocalDateTime.mk $(← parseDate date) $(← parseTime time)) | `(datetime|$tm:date_component) => do - `(Std.Time.LocalDateTime.ofUTCTimestamp $(← parseComponent tm)) + `(Std.Time.LocalDateTime.ofUTCTimestamp $(← parseComponent none none tm)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -112,9 +122,9 @@ syntax ("+" <|> "-") date_component ":" date_component : offset private def parseOffset : TSyntax `offset -> MacroM (TSyntax `term) | `(offset| +$hour:date_component:$minutes:date_component) => do - `(Std.Time.TimeZone.Offset.ofHoursAndMinutes $(← parseComponent hour) $(← parseComponent minutes)) + `(Std.Time.TimeZone.Offset.ofHoursAndMinutes $(← parseComponent (some 0) (some 24) hour) $(← parseComponent (some 0) (some 59) minutes)) | `(offset| -$hour:date_component:$minutes:date_component) => do - `(Std.Time.TimeZone.Offset.ofHoursAndMinutes (- $(← parseComponent hour)) (-$(← parseComponent minutes))) + `(Std.Time.TimeZone.Offset.ofHoursAndMinutes (- $(← parseComponent (some 0) (some 24) hour)) (-$(← parseComponent (some 0) (some 59) minutes))) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index cd42a368516c..5831b90ea7be 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -17,7 +17,7 @@ def CustomDayTime : Format .any := date-spec% "EEE D MMM YYYY hh:mm" def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 -def date₁ := date% 16-06-2014:03:03:03(brTZ) +def date₁ := date% 2014-06-16:03:03:03(brTZ) def time₁ := time% 14:11:01 def time₂ := time% 03:11:01 @@ -201,3 +201,27 @@ info: "06/16/2014" -/ #guard_msgs in #eval ShortDate.formatBuilder date₁.month date₁.day date₁.year + +/-- +info: "0053-06-19" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-700000⟩) .UTC) + +/-- +info: "-0002-09-16" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-720000⟩) .UTC) + +/-- +info: "-0084-07-28" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-750000⟩) .UTC) + +/-- +info: "-0221-09-04" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 48d532582442..3d7fa3de5a30 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -1,4 +1,5 @@ import Std.Time +import Init open Std.Time def ShortDateTime : Format .any := date-spec% "DD/MM/YYYY hh:mm:ss" @@ -6,10 +7,10 @@ def ShortDate : Format .any := date-spec% "DD/MM/YYYY" def format (localDate : LocalDateTime) : String := ShortDateTime.formatBuilder localDate.day localDate.month localDate.year localDate.time.hour localDate.minute localDate.time.second def format₂ (localDate : LocalDate) : String := ShortDate.formatBuilder localDate.day localDate.month localDate.year -def date₁ := date% 19-11-1993:09:08:07 -def date₂ := date% 09-05-1993:12:59:59 -def date₃ := date% 16-08-2024 -def date₄ := date% 16-08-1500 +def date₁ := date% 1993-11-19:09:08:07 +def date₂ := date% 1993-05-09:12:59:59 +def date₃ := date% 2024-08-16 +def date₄ := date% 1500-08-16 def tm₁ := 753700087 def tm₂ := 736952399 @@ -63,7 +64,20 @@ info: 19951 #eval LocalDate.toDaysSinceUNIXEpoch date₃ /-- -info: Std.Time.Weekday.fri +info: Std.Time.Weekday.friday -/ #guard_msgs in #eval LocalDate.weekday date₃ + + +#eval Id.run do + let mut res := #[] + + for i in [0:10000] do + let i := Int.ofNat i - 999975 + let date := LocalDate.ofDaysSinceUNIXEpoch (Day.Offset.ofInt i) + let num := date.toDaysSinceUNIXEpoch + if i ≠ num.val then + res := res.push i + + return res diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index e375d93668c2..1d77674253f7 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -19,7 +19,7 @@ def Full12HourWrong : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss aa Z" def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 -def date₁ := date% 16-06-2014:03:03:03(brTZ) +def date₁ := date% 2014-06-16:03:03:03(brTZ) def time₁ := time% 14:11:01 def time₂ := time% 03:11:01 @@ -88,7 +88,7 @@ info: "00-1-12-31T22:28:12+0900" #guard_msgs in #eval let t : ZonedDateTime := Time24Hour.parse! "13:28:12" - ISO8601UTC.format (t.snd.convertTimeZone jpTZ) + ISO8601UTC.format (t.snd.convertTimeZone dateUTC) /-- info: "00-1-12-31T09:28:12-0300" From e9be73540d0f8ab8a4deff20345c6c987f26d401 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 20 Aug 2024 12:13:54 -0300 Subject: [PATCH 026/118] fix: a bunch of problems and added tests --- src/Std/Time/Date/LocalDate.lean | 31 ++++++++------ src/Std/Time/Date/Unit/Month.lean | 6 +-- src/Std/Time/Format.lean | 18 +++++++++ src/Std/Time/Format/Basic.lean | 58 +++++++++++++-------------- tests/lean/run/timeFormats.lean | 36 +++++++++++++++++ tests/lean/run/timeLocalDateTime.lean | 5 ++- tests/lean/run/timeParse.lean | 6 +-- 7 files changed, 112 insertions(+), 48 deletions(-) diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index 45771592dc4d..34a329f413ac 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -41,8 +41,6 @@ structure LocalDate where -/ valid : year.Valid month day - deriving Repr - instance : BEq LocalDate where beq x y := x.day == y.day && x.month == y.month && x.year == y.year @@ -60,16 +58,27 @@ def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Loca Creates a `LocalDate` by rolling over the extra days to the next month. -/ def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : LocalDate := by - let ⟨max, valid, gt⟩ := month.days year.isLeap - let max : Bounded.LE 28 31 := max.truncateBottom gt + let max : Day.Ordinal := month.days year.isLeap + have p := month.all_greater_than_27 year.isLeap if h : day.val > max.val then - let day := day.truncateBottom h - let off := day.sub max.val - simp at off - sorry + if h₁ : month.val > 11 then + let eq : month.val = 12 := Int.eq_iff_le_and_ge.mpr (And.intro month.property.right h₁) + let h : max.val = 31 := by simp [max, Month.Ordinal.days, Month.Ordinal.daysWithoutProof, Bounded.LE.sub , Bounded.LE.add, Bounded.LE.toFin, eq]; rfl + let h₂ := Int.le_trans day.property.right (Int.eq_iff_le_and_ge.mp h |>.right) |> Int.not_lt.mpr + contradiction + else + let max₂ : Bounded.LE 28 31 := max.truncateBottom p + let sub := Int.sub_nonneg_of_le h + simp [←Int.sub_sub] at sub + let roll := day.addBounds (max₂.neg) |>.truncateBottom (Int.add_le_of_le_sub_left sub) + let day : Day.Ordinal := roll.expandTop (by decide) + let h₂ : roll.val ≤ 27 := Int.le_trans roll.property.right (by decide) + let month := month.truncateTop (Int.not_lt.mp h₁) |>.addTop 1 (by decide) + refine ⟨year, month, day, ?_⟩ + exact Int.le_of_lt (Int.le_trans (Int.add_le_add_right h₂ 1) (Month.Ordinal.all_greater_than_27 year.isLeap month)) else - sorry - + let h := Int.not_lt.mp h + exact ⟨year, month, day, h⟩ instance : Inhabited LocalDate where default := clip 0 1 1 @@ -138,8 +147,6 @@ def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := .ofInt (era * 146097 + doe - 719468) - - /-- Calculates the difference in years between a `LocalDate` and a given year. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 9a09c5c05007..03c785d75850 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -183,8 +183,7 @@ theorem all_greater_than_27 (leap : Bool) (i: Month.Ordinal) : daysWithoutProof /-- Check if the day is valid in a month and a leap year. -/ -@[inline] -def Valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := +abbrev Valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := day.val ≤ (daysWithoutProof leap month).val instance : Decidable (Valid leap month day) := @@ -205,7 +204,8 @@ def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : let max : Day.Ordinal := month.days leap if h : day.val > max.val then ⟨max, Int.le_refl max.val⟩ - else ⟨⟨day.val, day.property⟩, Int.not_lt.mp h⟩ + else ⟨day, Int.not_lt.mp h⟩ + /-- Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 852d0e796375..736cad6b3190 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -110,6 +110,9 @@ def parse (input : String) : Except String LocalDate := instance : ToString LocalDate where toString := toSQLDateString +instance : Repr LocalDate where + reprPrec data _ := toString data + end LocalDate namespace LocalTime @@ -152,6 +155,9 @@ def parse (input : String) : Except String LocalTime := instance : ToString LocalTime where toString := toTime24Hour +instance : Repr LocalTime where + reprPrec data _ := toString data + end LocalTime namespace ZonedDateTime @@ -203,6 +209,9 @@ def parse (input : String) : Except String ZonedDateTime := instance : ToString ZonedDateTime where toString := toRFC822String +instance : Repr ZonedDateTime where + reprPrec data _ := toString data + end ZonedDateTime namespace LocalDateTime @@ -240,6 +249,12 @@ def parse (date : String) : Except String LocalDateTime := fromAscTimeString date <|> fromLongDateFormatString date +instance : ToString LocalDateTime where + toString := toLongDateFormatString + +instance : Repr LocalDateTime where + reprPrec s _ := toLongDateFormatString s + end LocalDateTime namespace DateTime @@ -296,4 +311,7 @@ def parse (date : String) : Except String (DateTime .GMT) := instance : ToString (DateTime tz) where toString := toRFC822String +instance : Repr (DateTime tz) where + reprPrec data _ := toString data + end DateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index eddee8071f05..e46b2452463d 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -336,27 +336,27 @@ private def leftPadNum (n : Nat) (s : Int) : String := private def formatWithDate (date : DateTime tz) : Modifier → String | .YYYY => s!"{leftPadNum 4 (date.year)}" - | .YY => s!"{leftPadNum 2 (date.year.toNat % 100)}" + | .YY => s!"{leftPadNum 2 (date.year.toInt % 100)}" | .MMMM => unabbrevMonth date.month | .MMM => abbrevMonth date.month - | .MM => s!"{leftPadNum 2 (date.month.toNat)}" - | .M => s!"{date.month.toNat}" - | .DD => s!"{leftPadNum 2 (date.day.toNat)}" - | .D => s!"{date.day.toNat}" - | .d => s!"{leftPad 2 ' ' <| toString date.day.toNat}" + | .MM => s!"{leftPadNum 2 (date.month.toInt)}" + | .M => s!"{date.month.toInt}" + | .DD => s!"{leftPadNum 2 (date.day.toInt)}" + | .D => s!"{date.day.toInt}" + | .d => s!"{leftPad 2 ' ' <| toString date.day.toInt}" | .EEEE => dayOfWeek date.weekday | .EEE => abbrevDayOfWeek date.weekday - | .hh => s!"{leftPadNum 2 (date.hour.toNat)}" - | .h => s!"{date.hour.toNat}" + | .hh => s!"{leftPadNum 2 (date.hour.toInt)}" + | .h => s!"{date.hour.toInt}" | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPadNum 2 <| hour}" | .H => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{hour}" - | .AA => if date.hour.toNat < 12 then "AM" else "PM" - | .aa => if date.hour.toNat < 12 then "am" else "pm" - | .mm => s!"{leftPadNum 2 <| date.minute.toNat}" - | .m => s!"{date.minute.toNat}" - | .sss => s!"{leftPadNum 3 <| date.milliseconds.toNat}" - | .ss => s!"{leftPadNum 2 <| date.second.toNat}" - | .s => s!"{date.second.toNat}" + | .AA => if date.hour.toInt < 12 then "AM" else "PM" + | .aa => if date.hour.toInt < 12 then "am" else "pm" + | .mm => s!"{leftPadNum 2 <| date.minute.toInt}" + | .m => s!"{date.minute.toInt}" + | .sss => s!"{leftPadNum 3 <| date.milliseconds.toInt}" + | .ss => s!"{leftPadNum 2 <| date.second.toInt}" + | .s => s!"{date.second.toInt}" | .ZZZZZ => tz.offset.toIsoString true | .ZZZZ => tz.offset.toIsoString false | .ZZZ => if tz.offset.second.val = 0 then "UTC" else tz.offset.toIsoString false @@ -385,28 +385,28 @@ private def SingleFormatType : Modifier → Type private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) : String := match modifier with - | .YYYY => s!"{leftPadNum 4 (data.toNat)}" - | .YY => s!"{leftPadNum 2 (data.toNat % 100)}" + | .YYYY => s!"{leftPadNum 4 (data.toInt)}" + | .YY => s!"{leftPadNum 2 (data.toInt % 100)}" | .MMMM => unabbrevMonth data | .MMM => abbrevMonth data - | .MM => s!"{leftPadNum 2 (data.toNat)}" - | .M => s!"{data.toNat}" - | .DD => s!"{leftPadNum 2 (data.toNat)}" - | .D => s!"{data.toNat}" - | .d => s!"{leftPad 2 ' ' <| toString data.toNat}" + | .MM => s!"{leftPadNum 2 (data.toInt)}" + | .M => s!"{data.toInt}" + | .DD => s!"{leftPadNum 2 (data.toInt)}" + | .D => s!"{data.toInt}" + | .d => s!"{leftPad 2 ' ' <| toString data.toInt}" | .EEEE => dayOfWeek data | .EEE => abbrevDayOfWeek data - | .hh => s!"{leftPadNum 2 (data.snd.toNat)}" - | .h => s!"{data.snd.toNat}" + | .hh => s!"{leftPadNum 2 (data.snd.toInt)}" + | .h => s!"{data.snd.toInt}" | .HH => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{leftPadNum 2 hour}" | .H => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{hour}" | .AA => match data with | .am => "AM" | .pm => "PM" | .aa => match data with | .am => "am" | .pm => "pm" - | .mm => s!"{leftPadNum 2 data.toNat}" - | .m => s!"{data.toNat}" - | .sss => s!"{leftPadNum 3 data.toNat}" - | .ss => s!"{leftPadNum 2 data.snd.toNat}" - | .s => s!"{data.snd.toNat}" + | .mm => s!"{leftPadNum 2 data.toInt}" + | .m => s!"{data.toInt}" + | .sss => s!"{leftPadNum 3 data.toInt}" + | .ss => s!"{leftPadNum 2 data.snd.toInt}" + | .s => s!"{data.snd.toInt}" | .ZZZZZ => data.toIsoString true | .ZZZZ => data.toIsoString false | .ZZZ => if data.second.val = 0 then "UTC" else data.toIsoString false diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 5831b90ea7be..cbf091c3440c 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -225,3 +225,39 @@ info: "-0221-09-04" -/ #guard_msgs in #eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) + +/-- +info: -0221-09-04 +-/ +#guard_msgs in +#eval (LocalDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) + +/-- +info: "-0221-09-04" +-/ +#guard_msgs in +#eval toString (LocalDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) + +/-- +info: 2002-07-14 +-/ +#guard_msgs in +#eval date% 2002-07-14 + +/-- +info: 14:13:12 +-/ +#guard_msgs in +#eval time% 14:13:12 + +/-- +info: Sunday, July 14, 2002 14:13:12 +-/ +#guard_msgs in +#eval date% 2002-07-14:14:13:12 + +/-- +info: Sun, 14 Jul 2002 14:13:12 +0900 +-/ +#guard_msgs in +#eval date% 2002-07-14:14:13:12+09:00 diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 3d7fa3de5a30..63b475f477d0 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -69,7 +69,10 @@ info: Std.Time.Weekday.friday #guard_msgs in #eval LocalDate.weekday date₃ - +/-- +info: #[] +-/ +#guard_msgs in #eval Id.run do let mut res := #[] diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 1d77674253f7..3f15137a0369 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -83,15 +83,15 @@ info: "2024-08-16T01:28:00UTC" ISO8601UTC.format t.snd /-- -info: "00-1-12-31T22:28:12+0900" +info: "-0001-12-31T22:28:12+0900" -/ #guard_msgs in #eval let t : ZonedDateTime := Time24Hour.parse! "13:28:12" - ISO8601UTC.format (t.snd.convertTimeZone dateUTC) + ISO8601UTC.format (t.snd.convertTimeZone jpTZ) /-- -info: "00-1-12-31T09:28:12-0300" +info: "-0001-12-31T09:28:12-0300" -/ #guard_msgs in #eval From 89cfd2c56047532cd6460dad787b336f553dfb18 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 20 Aug 2024 23:29:19 -0300 Subject: [PATCH 027/118] refactor: some structures and add test --- src/Std/Time/DateTime.lean | 4 ++-- src/Std/Time/DateTime/LocalDateTime.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 14 +++++++------- src/Std/Time/Duration.lean | 4 ++-- src/Std/Time/Format.lean | 10 +++++----- src/Std/Time/Zoned/Database.lean | 4 ++-- src/Std/Time/Zoned/ZoneRules.lean | 7 ++++--- src/Std/Time/Zoned/ZonedDateTime.lean | 1 + tests/lean/run/timeFormats.lean | 8 ++++---- tests/lean/run/timeParse.lean | 8 ++++---- tests/lean/run/timeTzifParse.lean | 18 ++++++++++++++++++ 11 files changed, 50 insertions(+), 30 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 296b26e81204..210147afc167 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -34,7 +34,7 @@ Converts a `LocalDate` to a `Timestamp` def ofLocalDate (ld : LocalDate) : Timestamp := let days := ld.toDaysSinceUNIXEpoch let secs := days.toSeconds - Timestamp.ofSeconds secs + Timestamp.ofSecondsSinceUnixEpoch secs /-- Converts a `Timestamp` to a `LocalDate` @@ -51,7 +51,7 @@ Converts a `LocalTime` to a `Timestamp` @[inline] def ofLocalTime (lt : LocalTime) : Timestamp := let nanos := lt.toNanoseconds - Timestamp.ofNanoseconds nanos + Timestamp.ofNanosecondsSinceUnixEpoch nanos /-- Converts a `Timestamp` to a `LocalTime` diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean index 160b6db28488..387fd4eb552d 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -41,7 +41,7 @@ def toLocalTimestamp (dt : LocalDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 let nanos := nanos.val + dt.time.nano.val - Timestamp.ofNanoseconds (UnitVal.mk nanos) + Timestamp.ofNanosecondsSinceUnixEpoch (UnitVal.mk nanos) /-- Converts a UNIX `Timestamp` to a `LocalDateTime`. diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 2da8da394c96..b8794048645e 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -67,7 +67,7 @@ def toSeconds (t : Timestamp) : Second.Offset := Negates a `Timestamp`, flipping its second and nanosecond values. -/ @[inline] -def neg (t : Timestamp) : Timestamp := by +protected def neg (t : Timestamp) : Timestamp := by refine ⟨-t.second, t.nano.neg, ?_⟩ cases t.proof with | inl n => exact Or.inr (n.imp Int.neg_le_neg Int.neg_le_neg) @@ -77,7 +77,7 @@ def neg (t : Timestamp) : Timestamp := by Adds two timestamps together, handling any carry-over in nanoseconds. It should not be used for `Timestamp`. The subtraction of two `Timestamp` returns a duration but the addition does not make sense at all. -/ -def add (t₁ t₂ : Timestamp) : Timestamp := by +protected def add (t₁ t₂ : Timestamp) : Timestamp := by let diffSecs := t₁.second + t₂.second let diffNano := t₁.nano.addBounds t₂.nano @@ -128,14 +128,14 @@ def add (t₁ t₂ : Timestamp) : Timestamp := by Subtracts one `Timestamp` from another. -/ @[inline] -def sub (t₁ t₂ : Timestamp) : Timestamp := +protected def sub (t₁ t₂ : Timestamp) : Timestamp := t₁.add t₂.neg /-- Creates a new `Timestamp` out of `Second.Offset`. -/ @[inline] -def ofSeconds (s : Second.Offset) : Timestamp := by +def ofSecondsSinceUnixEpoch (s : Second.Offset) : Timestamp := by refine ⟨s, ⟨0, by decide⟩, ?_⟩ simp <;> exact Int.le_total s.val 0 |>.symm @@ -143,7 +143,7 @@ def ofSeconds (s : Second.Offset) : Timestamp := by Creates a new `Timestamp` out of `Second.Offset`. -/ @[inline] -def ofNanoseconds (s : Nanosecond.Offset) : Timestamp := by +def ofNanosecondsSinceUnixEpoch (s : Nanosecond.Offset) : Timestamp := by refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ cases Int.le_total s.val 0 next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) @@ -167,14 +167,14 @@ Adds a `Second.Offset` to a `Timestamp` -/ @[inline] def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.add (ofSeconds s) + t.add (ofSecondsSinceUnixEpoch s) /-- Subtracts a `Second.Offset` from a `Timestamp` -/ @[inline] def subSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.sub (ofSeconds s) + t.sub (ofSecondsSinceUnixEpoch s) /-- Adds a `Minute.Offset` to a `Timestamp` diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 194163c38b33..8de8ed2aa370 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -33,13 +33,13 @@ Calculates a `Duration` out of two `Timestamp`s. -/ def since (f : Timestamp) : IO Duration := do let cur ← Timestamp.now - return cur.sub f + return Std.Time.Timestamp.sub cur f /-- Adds a `Duration` to a `Timestamp`. -/ def add (f : Timestamp) (d : Duration) : Timestamp := - f.add d + Std.Time.Timestamp.add f d /-- Checks if the duration is zero seconds ands zero nanoseconds. diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 736cad6b3190..33ed01567cc6 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -111,7 +111,7 @@ instance : ToString LocalDate where toString := toSQLDateString instance : Repr LocalDate where - reprPrec data _ := toString data + reprPrec data := Repr.addAppParen (toString data) end LocalDate @@ -156,7 +156,7 @@ instance : ToString LocalTime where toString := toTime24Hour instance : Repr LocalTime where - reprPrec data _ := toString data + reprPrec data := Repr.addAppParen (toString data) end LocalTime @@ -210,7 +210,7 @@ instance : ToString ZonedDateTime where toString := toRFC822String instance : Repr ZonedDateTime where - reprPrec data _ := toString data + reprPrec data := Repr.addAppParen (toString data) end ZonedDateTime @@ -253,7 +253,7 @@ instance : ToString LocalDateTime where toString := toLongDateFormatString instance : Repr LocalDateTime where - reprPrec s _ := toLongDateFormatString s + reprPrec data := Repr.addAppParen (toString data) end LocalDateTime @@ -312,6 +312,6 @@ instance : ToString (DateTime tz) where toString := toRFC822String instance : Repr (DateTime tz) where - reprPrec data _ := toString data + reprPrec data := Repr.addAppParen (toString data) end DateTime diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index 500af50e67eb..b6f2d8cd8e23 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -19,7 +19,7 @@ set_option linter.all true Apply leap seconds rules and transition rules on a UTC Timestamp to make it aware of the timezone. -/ def applyLeapSecondsOnUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO Timestamp := do - (applyLeapSeconds tm ·.leapSeconds) <$> Database.localRules db + (applyLeapSeconds tm ·) <$> Database.localRules db /-- Gets the TimeZone at the local timezone. @@ -39,5 +39,5 @@ Get the local ZonedDataTime given a UTC `Timestamp`. def ofUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime := do let rules ← Database.localRules db let tz ← IO.ofExcept <| timezoneAt rules tm - let tm := applyLeapSeconds tm rules.leapSeconds + let tm := applyLeapSeconds tm rules return ZonedDateTime.ofUTCTimestamp tm tz diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 13d657ed526a..337dd16a4125 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -186,8 +186,9 @@ def timezoneAt (rules : ZoneRules) (tm : Timestamp) : Except String TimeZone := /-- Applies a `LeapSecond` sequence in a `Timestamp`. -/ -def applyLeapSeconds (tm : Timestamp) (leapSeconds : Array LeapSecond) : Timestamp := Id.run do +def applyLeapSeconds (tm : Timestamp) (leapSeconds : ZoneRules) : Timestamp := Id.run do let mut currentTime := tm + let leapSeconds := leapSeconds.leapSeconds for i in [:leapSeconds.size] do let leapSec := leapSeconds.get! i if currentTime.second.val >= leapSec.transitionTime.val then @@ -197,8 +198,8 @@ def applyLeapSeconds (tm : Timestamp) (leapSeconds : Array LeapSecond) : Timesta /-- Adjust a UTC timestamp according to `ZoneRules`. -/ -def apply (rules : ZoneRules) (tm : Timestamp) : Timestamp := - let tm := applyLeapSeconds tm rules.leapSeconds +protected def applyToTimestamp (rules : ZoneRules) (tm : Timestamp) : Timestamp := + let tm := applyLeapSeconds tm rules if let some transition := findTransitionForTimestamp rules tm then transition.apply tm else tm diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 9ecd035e7804..b6d6fa79b120 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -50,6 +50,7 @@ Creates a new `DateTime` out of a `Timestamp` @[inline] def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do let transition ← rules.findTransitionForTimestamp tm + let tm := rules.applyLeapSeconds tm return ofUTCTimestamp tm transition.localTimeType.getTimeZone /-- diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index cbf091c3440c..e58f284e436b 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -56,9 +56,9 @@ BR: 15 August 2024 13:28:12 GMT-03:00 -/ def tm₄ : Second.Offset := 1723739292 -def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) brTZ -def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) jpTZ -def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) .UTC +def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) brTZ +def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpTZ +def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC /-- info: "Thursday, August 15, 2024 13:28:12 -0300" @@ -130,7 +130,7 @@ def localTm : Second.Offset := 1723730627 /-- This localDate is relative to the local time. -/ -def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSeconds localTm) +def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) /-- info: "08/15/2024 14:03:47" diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 3f15137a0369..152bed720106 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -62,9 +62,9 @@ BR: 15 August 2024 13:28:12 GMT-03:00 -/ def tm₄ : Second.Offset := 1723739292 -def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) brTZ -def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) jpTZ -def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSeconds tm₄) .UTC +def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) brTZ +def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpTZ +def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC /-- info: "2024-08-15T13:28:12-0300" @@ -140,7 +140,7 @@ def localTm : Second.Offset := 1723730627 /-- This localDate is relative to the local time. -/ -def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSeconds localTm) +def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) /-- info: "08/15/2024 14:03:47" diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index 37cacd9b1e33..48816f44fc35 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -1,9 +1,15 @@ import Std.Time +open Std.Time def file := ByteArray.mk #[84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0, 0, 0, 3, 0, 0, 0, 12, 150, 170, 114, 180, 184, 15, 73, 224, 184, 253, 64, 160, 185, 241, 52, 48, 186, 222, 116, 32, 218, 56, 174, 48, 218, 235, 250, 48, 220, 25, 225, 176, 220, 185, 89, 32, 221, 251, 21, 48, 222, 155, 222, 32, 223, 221, 154, 48, 224, 84, 51, 32, 244, 90, 9, 48, 245, 5, 94, 32, 246, 192, 100, 48, 247, 14, 30, 160, 248, 81, 44, 48, 248, 199, 197, 32, 250, 10, 210, 176, 250, 168, 248, 160, 251, 236, 6, 48, 252, 139, 125, 160, 29, 201, 142, 48, 30, 120, 215, 160, 31, 160, 53, 176, 32, 51, 207, 160, 33, 129, 105, 48, 34, 11, 200, 160, 35, 88, 16, 176, 35, 226, 112, 32, 37, 55, 242, 176, 37, 212, 199, 32, 39, 33, 15, 48, 39, 189, 227, 160, 41, 0, 241, 48, 41, 148, 139, 32, 42, 234, 13, 176, 43, 107, 50, 160, 44, 192, 181, 48, 45, 102, 196, 32, 46, 160, 151, 48, 47, 70, 166, 32, 48, 128, 121, 48, 49, 29, 77, 160, 50, 87, 32, 176, 51, 6, 106, 32, 52, 56, 84, 48, 52, 248, 193, 32, 54, 32, 31, 48, 54, 207, 104, 160, 55, 246, 198, 176, 56, 184, 133, 32, 57, 223, 227, 48, 58, 143, 44, 160, 59, 200, 255, 176, 60, 111, 14, 160, 61, 196, 145, 48, 62, 78, 240, 160, 63, 145, 254, 48, 64, 46, 210, 160, 65, 134, 248, 48, 66, 23, 239, 32, 67, 81, 194, 48, 67, 247, 209, 32, 69, 77, 83, 176, 69, 224, 237, 160, 71, 17, 134, 48, 71, 183, 149, 32, 72, 250, 162, 176, 73, 151, 119, 32, 74, 218, 132, 176, 75, 128, 147, 160, 76, 186, 102, 176, 77, 96, 117, 160, 78, 154, 72, 176, 79, 73, 146, 32, 80, 131, 101, 48, 81, 32, 57, 160, 82, 99, 71, 48, 83, 0, 27, 160, 84, 67, 41, 48, 84, 233, 56, 32, 86, 35, 11, 48, 86, 201, 26, 32, 88, 2, 237, 48, 88, 168, 252, 32, 89, 226, 207, 48, 90, 136, 222, 32, 91, 222, 96, 176, 92, 104, 192, 32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255, 213, 208, 0, 8, 76, 77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0, 0, 0, 3, 0, 0, 0, 12, 255, 255, 255, 255, 150, 170, 114, 180, 255, 255, 255, 255, 184, 15, 73, 224, 255, 255, 255, 255, 184, 253, 64, 160, 255, 255, 255, 255, 185, 241, 52, 48, 255, 255, 255, 255, 186, 222, 116, 32, 255, 255, 255, 255, 218, 56, 174, 48, 255, 255, 255, 255, 218, 235, 250, 48, 255, 255, 255, 255, 220, 25, 225, 176, 255, 255, 255, 255, 220, 185, 89, 32, 255, 255, 255, 255, 221, 251, 21, 48, 255, 255, 255, 255, 222, 155, 222, 32, 255, 255, 255, 255, 223, 221, 154, 48, 255, 255, 255, 255, 224, 84, 51, 32, 255, 255, 255, 255, 244, 90, 9, 48, 255, 255, 255, 255, 245, 5, 94, 32, 255, 255, 255, 255, 246, 192, 100, 48, 255, 255, 255, 255, 247, 14, 30, 160, 255, 255, 255, 255, 248, 81, 44, 48, 255, 255, 255, 255, 248, 199, 197, 32, 255, 255, 255, 255, 250, 10, 210, 176, 255, 255, 255, 255, 250, 168, 248, 160, 255, 255, 255, 255, 251, 236, 6, 48, 255, 255, 255, 255, 252, 139, 125, 160, 0, 0, 0, 0, 29, 201, 142, 48, 0, 0, 0, 0, 30, 120, 215, 160, 0, 0, 0, 0, 31, 160, 53, 176, 0, 0, 0, 0, 32, 51, 207, 160, 0, 0, 0, 0, 33, 129, 105, 48, 0, 0, 0, 0, 34, 11, 200, 160, 0, 0, 0, 0, 35, 88, 16, 176, 0, 0, 0, 0, 35, 226, 112, 32, 0, 0, 0, 0, 37, 55, 242, 176, 0, 0, 0, 0, 37, 212, 199, 32, 0, 0, 0, 0, 39, 33, 15, 48, 0, 0, 0, 0, 39, 189, 227, 160, 0, 0, 0, 0, 41, 0, 241, 48, 0, 0, 0, 0, 41, 148, 139, 32, 0, 0, 0, 0, 42, 234, 13, 176, 0, 0, 0, 0, 43, 107, 50, 160, 0, 0, 0, 0, 44, 192, 181, 48, 0, 0, 0, 0, 45, 102, 196, 32, 0, 0, 0, 0, 46, 160, 151, 48, 0, 0, 0, 0, 47, 70, 166, 32, 0, 0, 0, 0, 48, 128, 121, 48, 0, 0, 0, 0, 49, 29, 77, 160, 0, 0, 0, 0, 50, 87, 32, 176, 0, 0, 0, 0, 51, 6, 106, 32, 0, 0, 0, 0, 52, 56, 84, 48, 0, 0, 0, 0, 52, 248, 193, 32, 0, 0, 0, 0, 54, 32, 31, 48, 0, 0, 0, 0, 54, 207, 104, 160, 0, 0, 0, 0, 55, 246, 198, 176, 0, 0, 0, 0, 56, 184, 133, 32, 0, 0, 0, 0, 57, 223, 227, 48, 0, 0, 0, 0, 58, 143, 44, 160, 0, 0, 0, 0, 59, 200, 255, 176, 0, 0, 0, 0, 60, 111, 14, 160, 0, 0, 0, 0, 61, 196, 145, 48, 0, 0, 0, 0, 62, 78, 240, 160, 0, 0, 0, 0, 63, 145, 254, 48, 0, 0, 0, 0, 64, 46, 210, 160, 0, 0, 0, 0, 65, 134, 248, 48, 0, 0, 0, 0, 66, 23, 239, 32, 0, 0, 0, 0, 67, 81, 194, 48, 0, 0, 0, 0, 67, 247, 209, 32, 0, 0, 0, 0, 69, 77, 83, 176, 0, 0, 0, 0, 69, 224, 237, 160, 0, 0, 0, 0, 71, 17, 134, 48, 0, 0, 0, 0, 71, 183, 149, 32, 0, 0, 0, 0, 72, 250, 162, 176, 0, 0, 0, 0, 73, 151, 119, 32, 0, 0, 0, 0, 74, 218, 132, 176, 0, 0, 0, 0, 75, 128, 147, 160, 0, 0, 0, 0, 76, 186, 102, 176, 0, 0, 0, 0, 77, 96, 117, 160, 0, 0, 0, 0, 78, 154, 72, 176, 0, 0, 0, 0, 79, 73, 146, 32, 0, 0, 0, 0, 80, 131, 101, 48, 0, 0, 0, 0, 81, 32, 57, 160, 0, 0, 0, 0, 82, 99, 71, 48, 0, 0, 0, 0, 83, 0, 27, 160, 0, 0, 0, 0, 84, 67, 41, 48, 0, 0, 0, 0, 84, 233, 56, 32, 0, 0, 0, 0, 86, 35, 11, 48, 0, 0, 0, 0, 86, 201, 26, 32, 0, 0, 0, 0, 88, 2, 237, 48, 0, 0, 0, 0, 88, 168, 252, 32, 0, 0, 0, 0, 89, 226, 207, 48, 0, 0, 0, 0, 90, 136, 222, 32, 0, 0, 0, 0, 91, 222, 96, 176, 0, 0, 0, 0, 92, 104, 192, 32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255, 213, 208, 0, 8, 76, 77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 10, 60, 45, 48, 51, 62, 51, 10] def code := Std.Time.TimeZone.TZif.parse.run file |>.toOption |>.get! +def rules := + match TimeZone.convertTZif code "America/Sao_Paulo" with + | .ok res => res + | .error err => panic! err + /-- info: { version := 50, isutcnt := 0, isstdcnt := 0, leapcnt := 0, timecnt := 91, typecnt := 3, charcnt := 12 } -/ @@ -51,3 +57,15 @@ info: 0 -/ #guard_msgs in #eval code.v1.utLocalIndicators.size + +/-- +info: some (Tue, 30 Dec 1969 21:00:00 -0300) +-/ +#guard_msgs in +#eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules + +/-- +info: some (Mon, 10 Dec 2012 00:35:47 -0200) +-/ +#guard_msgs in +#eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From b193aacdf0d7f9e2eb357c10599c4c6512069d91 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 01:04:06 -0300 Subject: [PATCH 028/118] feat: change some native functions --- src/Std/Time/Date/Unit/Month.lean | 1 - src/Std/Time/DateTime/LocalDateTime.lean | 7 ----- src/Std/Time/Zoned.lean | 38 +++++++++++++++++++++--- src/Std/Time/Zoned/Database.lean | 2 +- src/Std/Time/Zoned/DateTime.lean | 13 ++------ src/Std/Time/Zoned/TimeZone.lean | 6 ++++ src/Std/Time/Zoned/ZonedDateTime.lean | 16 +++------- src/runtime/io.cpp | 29 ++++++++++++++++++ tests/lean/run/timeFormats.lean | 10 +++---- tests/lean/run/timeParse.lean | 10 +++---- 10 files changed, 86 insertions(+), 46 deletions(-) diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 03c785d75850..0c786f20b835 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -206,7 +206,6 @@ def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : then ⟨max, Int.le_refl max.val⟩ else ⟨day, Int.not_lt.mp h⟩ - /-- Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. -/ diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean index 387fd4eb552d..43cce12a2d6a 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -248,13 +248,6 @@ Getter for the `Second` inside of a `LocalDateTime`. def nanosecond (dt : LocalDateTime) : Nanosecond.Ordinal := dt.time.nano -/-- -Get the current monotonic time. --/ -@[inline] -def now : IO LocalDateTime := - ofUTCTimestamp <$> Timestamp.now - end LocalDateTime end Time end Std diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 514656ffc201..2d424ddf8bc0 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -13,6 +13,27 @@ import Std.Time.Zoned.Database.Basic namespace Std namespace Time + +namespace LocalDateTime + +/-- +Creaates a new `LocalDateTime` out of a `Timestamp` and a `TimeZone`. +-/ +def ofTimestamp (stamp : Timestamp) (tz : TimeZone) : LocalDateTime := + let stamp := stamp.subSeconds tz.toSeconds + LocalDateTime.ofUTCTimestamp stamp + +/-- +Get the current monotonic time. +-/ +@[inline] +def now : IO LocalDateTime := do + let tm ← Timestamp.now + let tz ← TimeZone.getCurrentTimezone + return LocalDateTime.ofTimestamp tm tz + +end LocalDateTime + namespace DateTime /-- @@ -20,7 +41,7 @@ Converts a `LocalDate` to a `DateTime` -/ @[inline] def ofLocalDate (ld : LocalDate) (tz : TimeZone) : DateTime tz := - DateTime.ofUTCTimestamp (Timestamp.ofLocalDate ld) tz + DateTime.ofTimestamp (Timestamp.ofLocalDate ld) tz /-- Converts a `DateTime` to a `LocalDate` @@ -34,7 +55,7 @@ Converts a `LocalTime` to a `DateTime` -/ @[inline] def ofLocalTime (lt : LocalTime) (tz : TimeZone) : DateTime tz := - DateTime.ofUTCTimestamp (Timestamp.ofLocalTime lt) tz + DateTime.ofTimestamp (Timestamp.ofLocalTime lt) tz /-- Converts a `DateTime` to a `LocalTime` @@ -47,12 +68,21 @@ end DateTime namespace ZonedDateTime +/-- +Gets the current `DateTime`. +-/ +@[inline] +def now : IO ZonedDateTime := do + let date ← Timestamp.now + let tz ← TimeZone.getCurrentTimezone + return ofTimestamp date tz + /-- Converts a `LocalDate` to a `ZonedDateTime` -/ @[inline] def ofLocalDate (ld : LocalDate) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofUTCTimestamp (Timestamp.ofLocalDate ld) tz⟩ + ⟨tz, DateTime.ofTimestamp (Timestamp.ofLocalDate ld) tz⟩ /-- Converts a `ZonedDateTime` to a `LocalDate` @@ -66,7 +96,7 @@ Converts a `LocalTime` to a `ZonedDateTime` -/ @[inline] def ofLocalTime (lt : LocalTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofUTCTimestamp (Timestamp.ofLocalTime lt) tz⟩ + ⟨tz, DateTime.ofTimestamp (Timestamp.ofLocalTime lt) tz⟩ /-- Converts a `ZonedDateTime` to a `LocalTime` diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index b6f2d8cd8e23..289247565bad 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -40,4 +40,4 @@ def ofUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime : let rules ← Database.localRules db let tz ← IO.ofExcept <| timezoneAt rules tm let tm := applyLeapSeconds tm rules - return ZonedDateTime.ofUTCTimestamp tm tz + return ZonedDateTime.ofTimestamp tm tz diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 712e6700008b..0b93507d733f 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -41,7 +41,7 @@ namespace DateTime Creates a new `DateTime` out of a `Timestamp` that is in a `TimeZone`. -/ @[inline] -def ofUTCTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := +def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := DateTime.mk tm (Thunk.mk <| λ_ => (tm.addSeconds tz.toSeconds).toLocalDateTime) /-- @@ -56,8 +56,7 @@ Changes the `TimeZone` to a new one. -/ @[inline] def convertTimeZone (date : DateTime tz) (tz₁ : TimeZone) : DateTime tz₁ := - ofUTCTimestamp date.timestamp tz₁ - + ofTimestamp date.timestamp tz₁ /-- Creates a new `DateTime` out of a `LocalDateTime` @@ -67,14 +66,6 @@ def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : DateTime tz := let tm := Timestamp.ofLocalDateTime date DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk <| λ_ => date) -/-- -Gets the current `DateTime`. --/ -@[inline] -def now : IO (DateTime tz) := do - let loca ← LocalDateTime.now - return ofLocalDateTime loca tz - /-- Add `Day.Offset` to a `DateTime`. -/ diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index 0dafb4742575..ffa6bfd3df23 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -34,6 +34,12 @@ structure TimeZone where namespace TimeZone +/-- +Fetches the current timestamp from the system. +-/ +@[extern "lean_get_timezone_offset"] +opaque getCurrentTimezone : IO TimeZone + /-- A zeroed `Timezone` representing UTC (no offset). -/ diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index b6d6fa79b120..a5e183fc010f 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -34,8 +34,8 @@ open DateTime Creates a new `ZonedDateTime` out of a `Timestamp` -/ @[inline] -def ofUTCTimestamp (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofUTCTimestamp tm tz⟩ +def ofTimestamp (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofTimestamp tm tz⟩ /-- Creates a new `Timestamp` out of a `ZonedDateTime` @@ -51,14 +51,14 @@ Creates a new `DateTime` out of a `Timestamp` def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do let transition ← rules.findTransitionForTimestamp tm let tm := rules.applyLeapSeconds tm - return ofUTCTimestamp tm transition.localTimeType.getTimeZone + return ofTimestamp tm transition.localTimeType.getTimeZone /-- Changes the `TimeZone` to a new one. -/ @[inline] def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := - ofUTCTimestamp (date.toTimestamp) tz₁ + ofTimestamp (date.toTimestamp) tz₁ /-- Creates a new `ZonedDateTime` out of a `LocalDateTime` @@ -74,14 +74,6 @@ Converts a `ZonedDateTime` to a `LocalDateTime` def toLocalDateTime (dt : ZonedDateTime) : LocalDateTime := DateTime.toLocalDateTime dt.snd -/-- -Gets the current `ZonedDataTime`. --/ -@[inline] -def now (tz : TimeZone) : IO ZonedDateTime := do - let loca ← LocalDateTime.now - return ofLocalDateTime loca tz - /-- Getter for the `Year` inside of a `ZonedDateTime` -/ diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index d0163e403be8..f6dc99ca6859 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -545,6 +545,35 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } +/* Std.Time.TimeZone.getCurrentTimezone : IO Timezone */ +extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { + using namespace std::chrono; + + auto now = system_clock::now(); + auto now_time_t = system_clock::to_time_t(now); + + std::tm tm_info; +#if defined(_MSC_VER) // For Microsoft Visual C++ + localtime_s(&tm_info, &now_time_t); // localtime_s for thread safety +#else + localtime_r(&now_time_t, &tm_info); // localtime_r for POSIX systems +#endif + + int offset_hour = tm_info.tm_gmtoff / 3600; + int offset_seconds = tm_info.tm_gmtoff; + + lean_object *lean_offset = lean_alloc_ctor(0, 2, 0); + lean_ctor_set(lean_offset, 0, lean_int_to_int(static_cast(offset_hour))); + lean_ctor_set(lean_offset, 1, lean_int_to_int(static_cast(offset_seconds))); + + lean_object *lean_tz = lean_alloc_ctor(0, 2, 1); + lean_ctor_set(lean_tz, 0, lean_offset); + lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked("Unknown")); + lean_ctor_set_uint8(lean_tz, sizeof(void*)*2, tm_info.tm_isdst); + + return lean_io_result_mk_ok(lean_tz); +} + /* monoMsNow : BaseIO Nat */ extern "C" LEAN_EXPORT obj_res lean_io_mono_ms_now(obj_arg /* w */) { static_assert(sizeof(std::chrono::milliseconds::rep) <= sizeof(uint64), "size of std::chrono::nanoseconds::rep may not exceed 64"); diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index e58f284e436b..9e9810adf09a 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -28,7 +28,7 @@ info: "Monday, June 16, 2014 03:03:03 -0300" #eval FullDayTimeZone.format date₁ def tm := date₁.timestamp -def date₂ := DateTime.ofUTCTimestamp tm brTZ +def date₂ := DateTime.ofTimestamp tm brTZ /-- info: "Monday, June 16, 2014 03:03:03 -0300" @@ -37,7 +37,7 @@ info: "Monday, June 16, 2014 03:03:03 -0300" #eval FullDayTimeZone.format date₂ def tm₃ := date₁.toTimestamp -def date₃ := DateTime.ofUTCTimestamp tm₃ brTZ +def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- info: "Monday, June 16, 2014 03:03:03 -0300" @@ -56,9 +56,9 @@ BR: 15 August 2024 13:28:12 GMT-03:00 -/ def tm₄ : Second.Offset := 1723739292 -def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) brTZ -def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpTZ -def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC +def dateBR := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) brTZ +def dateJP := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpTZ +def dateUTC := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC /-- info: "Thursday, August 15, 2024 13:28:12 -0300" diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 152bed720106..e584f126d3d8 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -32,7 +32,7 @@ info: "2014-06-16T03:03:03-0300" ISO8601UTC.format t.snd def tm := date₁.timestamp -def date₂ := DateTime.ofUTCTimestamp tm brTZ +def date₂ := DateTime.ofTimestamp tm brTZ /-- info: "2014-06-16T03:03:03-0300" @@ -43,7 +43,7 @@ info: "2014-06-16T03:03:03-0300" ISO8601UTC.format t.snd def tm₃ := date₁.toTimestamp -def date₃ := DateTime.ofUTCTimestamp tm₃ brTZ +def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- info: "2014-06-16T00:00:00UTC" @@ -62,9 +62,9 @@ BR: 15 August 2024 13:28:12 GMT-03:00 -/ def tm₄ : Second.Offset := 1723739292 -def dateBR := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) brTZ -def dateJP := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpTZ -def dateUTC := DateTime.ofUTCTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC +def dateBR := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) brTZ +def dateJP := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpTZ +def dateUTC := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC /-- info: "2024-08-15T13:28:12-0300" From a48f4ff64cb2b43db5d9e6417d4505e3ef6f4f12 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 01:05:15 -0300 Subject: [PATCH 029/118] fix: bug with timezone change in localdatetime --- src/Std/Time/Zoned.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 2d424ddf8bc0..04742721bd08 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -20,7 +20,7 @@ namespace LocalDateTime Creaates a new `LocalDateTime` out of a `Timestamp` and a `TimeZone`. -/ def ofTimestamp (stamp : Timestamp) (tz : TimeZone) : LocalDateTime := - let stamp := stamp.subSeconds tz.toSeconds + let stamp := stamp.addSeconds tz.toSeconds LocalDateTime.ofUTCTimestamp stamp /-- From 7b5754e23cb37f35e01759d2db8e24e150919e18 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 01:35:01 -0300 Subject: [PATCH 030/118] feat: add format functions --- src/Std/Time/Format.lean | 77 ++++++++++++++++++++++++++++++++ src/Std/Time/Format/Basic.lean | 10 +++++ src/Std/Time/Time/LocalTime.lean | 13 +++++- tests/lean/run/timeFormats.lean | 12 +++++ 4 files changed, 111 insertions(+), 1 deletion(-) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 33ed01567cc6..131e36164078 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -76,6 +76,24 @@ end Formats namespace LocalDate +/-- +Formats a `LocalDate` using a specific format. +-/ +def format (date : LocalDate) (format : String) : String := + let format : Except String (Format .any) := Format.spec format + match format with + | .error err => s!"error: {err}" + | .ok res => + let res := res.formatGeneric fun + | .YYYY | .YY => some date.year + | .MMMM | .MMM | .MM | .M => some date.month + | .DD | .D | .d => some date.day + | .EEEE | .EEE => some date.weekday + | _ => none + match res with + | some res => res + | none => "invalid time" + /-- Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. -/ @@ -117,6 +135,24 @@ end LocalDate namespace LocalTime +/-- +Formats a `LocalTime` using a specific format. +-/ +def format (time : LocalTime) (format : String) : String := + let format : Except String (Format .any) := Format.spec format + match format with + | .error err => s!"error: {err}" + | .ok res => + let res := res.formatGeneric fun + | .HH | .H => some time.hour + | .mm | .m => some time.minute + | .sss => some (Internal.Bounded.LE.ofNat 0 (by decide)) + | .ss | .s => some time.second + | _ => none + match res with + | some res => res + | none => "invalid time" + /-- Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `LocalTime`. -/ @@ -162,6 +198,16 @@ end LocalTime namespace ZonedDateTime + +/-- +Formats a `ZonedDateTime` using a specific format. +-/ +def format (data: ZonedDateTime) (format : String) : String := + let format : Except String (Format .any) := Format.spec format + match format with + | .error err => s!"error: {err}" + | .ok res => res.format data.snd + /-- Parses a `String` in the `ISO8601` format and returns a `ZonedDateTime`. -/ @@ -216,6 +262,28 @@ end ZonedDateTime namespace LocalDateTime +/-- +Formats a `LocalDateTime` using a specific format. +-/ +def format (date : LocalDateTime) (format : String) : String := + let format : Except String (Format .any) := Format.spec format + match format with + | .error err => s!"error: {err}" + | .ok res => + let res := res.formatGeneric fun + | .YYYY | .YY => some date.year + | .MMMM | .MMM | .MM | .M => some date.month + | .DD | .D | .d => some date.day + | .EEEE | .EEE => some date.date.weekday + | .HH | .H => some date.time.hour + | .mm | .m => some date.time.minute + | .sss => some (Internal.Bounded.LE.ofNat 0 (by decide)) + | .ss | .s => some date.time.second + | _ => none + match res with + | some res => res + | none => "invalid time" + /-- Parses a `String` in the `AscTime` format and returns a `LocalDateTime` object in the GMT time zone. -/ @@ -259,6 +327,15 @@ end LocalDateTime namespace DateTime +/-- +Formats a `DateTime` using a specific format. +-/ +def format (data: DateTime tz) (format : String) : String := + let format : Except String (Format .any) := Format.spec format + match format with + | .error err => s!"error: {err}" + | .ok res => res.format data + /-- Parses a `String` in the `AscTime` format and returns a `DateTime` object in the GMT time zone. -/ diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index e46b2452463d..3e73a425d812 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -659,6 +659,16 @@ def formatBuilder (format : Format aw) : FormatType String format.string := | [] => data go "" format.string +/-- +Formats the date using the format into a String. +-/ +def formatGeneric (format : Format aw) (getInfo : (typ : Modifier) → Option (SingleFormatType typ)) : Option String := + let rec go (data : String) : (format : FormatString) → Option String + | .modifier x :: xs => do go (data ++ formatPart x (← getInfo x)) xs + | .string x :: xs => go (data ++ x) xs + | [] => data + go "" format.string + /-- Parser for a ZonedDateTime. -/ diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/LocalTime.lean index 0c37f76826d7..46843167fa5d 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -85,7 +85,18 @@ def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Or exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) /-- -Converts a `LocalTime` value to the total number of seconds. +Converts a `LocalTime` value to the total number of milliseconds. +-/ +def toMilliseconds (time : LocalTime) : Millisecond.Offset := + let secs := + time.hour.snd.toOffset.toSeconds + + time.minute.toOffset.toSeconds + + time.second.snd.toOffset + let millis := secs.mul 1000 + UnitVal.mk (millis.val + time.nano.val) + +/-- +Converts a `LocalTime` value to the total number of nanoseconds. -/ def toNanoseconds (time : LocalTime) : Nanosecond.Offset := let secs := diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 9e9810adf09a..f40ee9e7366f 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -261,3 +261,15 @@ info: Sun, 14 Jul 2002 14:13:12 +0900 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12+09:00 + +/-- +info: "2002-07-14" +-/ +#guard_msgs in +#eval (date% 2002-07-14:14:13:12+09:00).format "YYYY-MM-DD" + +/-- +info: "14-13-12" +-/ +#guard_msgs in +#eval (date% 2002-07-14:14:13:12+09:00).format "hh-mm-ss" From 842fb3b9fc1f41d3dd1b9f1d4a7fe3e4571c1116 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 01:54:07 -0300 Subject: [PATCH 031/118] feat: add a bunch of arithmetic functions --- src/Std/Time/Date/LocalDate.lean | 5 +- src/Std/Time/DateTime/LocalDateTime.lean | 101 +++++++++++++++++++++++ src/Std/Time/DateTime/Timestamp.lean | 28 +++++++ src/Std/Time/Duration.lean | 6 ++ src/Std/Time/Time/LocalTime.lean | 81 ++++++++++++++++++ 5 files changed, 220 insertions(+), 1 deletion(-) diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index 34a329f413ac..f34eec60d00b 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -263,7 +263,10 @@ def subYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := addMonthsClip date (- years.mul 12) instance : HAdd LocalDate Day.Offset LocalDate where - hAdd date day := ofDaysSinceUNIXEpoch (toDaysSinceUNIXEpoch date + day) + hAdd := addDays + +instance : HSub LocalDate Day.Offset LocalDate where + hSub := subDays end LocalDate end Time diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean index 43cce12a2d6a..61057549540c 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -133,6 +133,9 @@ Adds a `Day.Offset` to a `LocalDateTime`. def addDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := { dt with date := dt.date.addDays days } +instance : HAdd LocalDateTime Day.Offset LocalDateTime where + hAdd := addDays + /-- Subtracts a `Day.Offset` from a `LocalDateTime`. -/ @@ -140,6 +143,9 @@ Subtracts a `Day.Offset` from a `LocalDateTime`. def subDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := { dt with date := dt.date.subDays days } +instance : HSub LocalDateTime Day.Offset LocalDateTime where + hSub := subDays + /-- Adds a `Month.Offset` to a `LocalDateTime`, adjusting the day to the last valid day of the resulting month. @@ -199,6 +205,77 @@ resulting month. def subYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := { dt with date := dt.date.subYearsClip years } + +/-- +Adds an `Hour.Offset` to a `LocalDateTime`, adjusting the date if the hour overflows. +-/ +@[inline] +def addHour (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := + let totalSeconds := dt.time.toSeconds + hours.toSeconds + let days := totalSeconds.ediv 86400 + let newTime := dt.time.addSeconds (hours.toSeconds) + { dt with date := dt.date.addDays days, time := newTime } + +/-- +Subtracts an `Hour.Offset` from a `LocalDateTime`, adjusting the date if the hour underflows. +-/ +@[inline] +def subHour (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := + addHour dt (-hours) + +/-- +Adds a `Minute.Offset` to a `LocalDateTime`, adjusting the hour and date if the minutes overflow. +-/ +@[inline] +def addMinute (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := + let totalSeconds := dt.time.toSeconds + minutes.toSeconds + let days := totalSeconds.ediv 86400 + let newTime := dt.time.addSeconds (minutes.toSeconds) + { dt with date := dt.date.addDays days, time := newTime } + +/-- +Subtracts a `Minute.Offset` from a `LocalDateTime`, adjusting the hour and date if the minutes underflow. +-/ +@[inline] +def subMinute (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := + addMinute dt (-minutes) + +/-- +Adds a `Second.Offset` to a `LocalDateTime`, adjusting the minute, hour, and date if the seconds overflow. +-/ +@[inline] +def addSecond (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := + let totalSeconds := dt.time.toSeconds + seconds + let days := totalSeconds.ediv 86400 + let newTime := dt.time.addSeconds seconds + { dt with date := dt.date.addDays days, time := newTime } + +/-- +Subtracts a `Second.Offset` from a `LocalDateTime`, adjusting the minute, hour, and date if the seconds underflow. +-/ +@[inline] +def subSecond (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := + addSecond dt (-seconds) + +/-- +Adds a `Nanosecond.Offset` to a `LocalDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds overflow. +-/ +@[inline] +def addNanosecond (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := + let nano : Nanosecond.Offset := UnitVal.mk dt.time.nano.val + let totalNanos := nano + nanos + let extraSeconds : Second.Offset := totalNanos.ediv 1000000000 + let nano := Bounded.LE.byEmod totalNanos.val 1000000000 (by decide) + let newTime := dt.time.addSeconds extraSeconds + { dt with time := { newTime with nano } } + +/-- +Subtracts a `Nanosecond.Offset` from a `LocalDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds underflow. +-/ +@[inline] +def subNanosecond (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := + addNanosecond dt (-nanos) + /-- Getter for the `Year` inside of a `LocalDateTime`. -/ @@ -248,6 +325,30 @@ Getter for the `Second` inside of a `LocalDateTime`. def nanosecond (dt : LocalDateTime) : Nanosecond.Ordinal := dt.time.nano +instance : HAdd LocalDateTime Hour.Offset LocalDateTime where + hAdd := addHour + +instance : HSub LocalDateTime Hour.Offset LocalDateTime where + hSub := subHour + +instance : HAdd LocalDateTime Minute.Offset LocalDateTime where + hAdd := addMinute + +instance : HSub LocalDateTime Minute.Offset LocalDateTime where + hSub := subMinute + +instance : HAdd LocalDateTime Second.Offset LocalDateTime where + hAdd := addSecond + +instance : HSub LocalDateTime Second.Offset LocalDateTime where + hSub := subSecond + +instance : HAdd LocalDateTime Nanosecond.Offset LocalDateTime where + hAdd := addNanosecond + +instance : HSub LocalDateTime Nanosecond.Offset LocalDateTime where + hSub := subNanosecond + end LocalDateTime end Time end Std diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index b8794048645e..b8bc2aad74f1 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -223,3 +223,31 @@ Subtracts a `Day.Offset` from a `Timestamp` def subDays (t : Timestamp) (d : Day.Offset) : Timestamp := let seconds := d.mul 86400 t.subSeconds seconds + +instance : HAdd Timestamp Day.Offset Timestamp where + hAdd := addDays + +instance : HSub Timestamp Day.Offset Timestamp where + hSub := subDays + +instance : HAdd Timestamp Hour.Offset Timestamp where + hAdd := addHours + +instance : HSub Timestamp Hour.Offset Timestamp where + hSub := subHours + +instance : HAdd Timestamp Minute.Offset Timestamp where + hAdd := addMinutes + +instance : HSub Timestamp Minute.Offset Timestamp where + hSub := subMinutes + +instance : HAdd Timestamp Second.Offset Timestamp where + hAdd := addSeconds + +instance : HSub Timestamp Second.Offset Timestamp where + hSub := subSeconds + +end Timestamp +end Time +end Std diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 8de8ed2aa370..06b5105b9a61 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -46,3 +46,9 @@ Checks if the duration is zero seconds ands zero nanoseconds. -/ def isZero (d : Duration) : Bool := d.second.val = 0 ∧ d.nano.val = 0 + +instance : HAdd Timestamp Duration Timestamp where + hAdd := add + +instance : HAdd Timestamp Timestamp Duration where + hAdd := Std.Time.Timestamp.sub diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/LocalTime.lean index 46843167fa5d..fbf84c2e124b 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -148,6 +148,87 @@ Creates a `LocalTime` value from a total number of seconds. def ofSeconds (secs : Second.Offset) : LocalTime := ofNanoseconds (secs.mul 1000000000) +/-- +Adds seconds to a `LocalTime`. +-/ +@[inline] +def addSeconds (time : LocalTime) (secondsToAdd : Second.Offset) : LocalTime := + let totalSeconds := time.toSeconds + secondsToAdd + ofSeconds totalSeconds + +/-- +Subtracts seconds from a `LocalTime`. +-/ +@[inline] +def subSeconds (time : LocalTime) (secondsToSub : Second.Offset) : LocalTime := + addSeconds time (-secondsToSub) + +/-- +Adds minutes to a `LocalTime`. +-/ +@[inline] +def addMinutes (time : LocalTime) (minutesToAdd : Minute.Offset) : LocalTime := + let totalMinutes := time.toMinutes + minutesToAdd + ofSeconds (totalMinutes.toSeconds) + +/-- +Subtracts minutes from a `LocalTime`. +-/ +@[inline] +def subMinutes (time : LocalTime) (minutesToSub : Minute.Offset) : LocalTime := + addMinutes time (-minutesToSub) + +/-- +Adds hours to a `LocalTime`. +-/ +def addHours (time : LocalTime) (hoursToAdd : Hour.Offset) : LocalTime := + let totalHours := time.toHours + hoursToAdd + ofSeconds (totalHours.toSeconds) + +/-- +Subtracts hours from a `LocalTime`. +-/ +@[inline] +def subHours (time : LocalTime) (hoursToSub : Hour.Offset) : LocalTime := + addHours time (-hoursToSub) + +/-- +Adds nanoseconds to a `LocalTime`. +-/ +def addNanoseconds (time : LocalTime) (nanosToAdd : Nanosecond.Offset) : LocalTime := + let totalNanos := time.toNanoseconds + nanosToAdd + ofNanoseconds totalNanos + +/-- +Subtracts nanoseconds from a `LocalTime`. +-/ +def subNanoseconds (time : LocalTime) (nanosToSub : Nanosecond.Offset) : LocalTime := + addNanoseconds time (-nanosToSub) + +instance : HAdd LocalTime Nanosecond.Offset LocalTime where + hAdd := addNanoseconds + +instance : HSub LocalTime Nanosecond.Offset LocalTime where + hSub := subNanoseconds + +instance : HAdd LocalTime Second.Offset LocalTime where + hAdd := addSeconds + +instance : HSub LocalTime Second.Offset LocalTime where + hSub := subSeconds + +instance : HAdd LocalTime Minute.Offset LocalTime where + hAdd := addMinutes + +instance : HSub LocalTime Minute.Offset LocalTime where + hSub := subMinutes + +instance : HAdd LocalTime Hour.Offset LocalTime where + hAdd := addHours + +instance : HSub LocalTime Hour.Offset LocalTime where + hSub := subHours + end LocalTime end Time end Std From 3b83c6c6e99e85bbbe297e6cf56ef6f6870ff0c1 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 02:15:40 -0300 Subject: [PATCH 032/118] feat: add more sub and add functions --- src/Std/Time/DateTime.lean | 2 +- src/Std/Time/DateTime/LocalDateTime.lean | 54 +++++++-------- src/Std/Time/DateTime/Timestamp.lean | 22 +++++- src/Std/Time/Duration.lean | 4 +- src/Std/Time/Zoned/DateTime.lean | 86 ++++++++++++++++++++++++ src/Std/Time/Zoned/ZonedDateTime.lean | 76 +++++++++++++++++++++ 6 files changed, 213 insertions(+), 31 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 210147afc167..9a48d6b5ed13 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -58,7 +58,7 @@ Converts a `Timestamp` to a `LocalTime` -/ @[inline] def toLocalTime (timestamp : Timestamp) : LocalTime := - let nanos := timestamp.toNanoseconds + let nanos := timestamp.toNanosecondsSinceUnixEpoch LocalTime.ofNanoseconds nanos end Timestamp diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/LocalDateTime.lean index 61057549540c..3f8239e52b65 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/LocalDateTime.lean @@ -52,7 +52,7 @@ def ofUTCTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do let daysPer100Y := 365 * 100 + 24 let daysPer4Y := 365 * 4 + 1 - let nanos := stamp.toNanoseconds + let nanos := stamp.toNanosecondsSinceUnixEpoch let secs : Second.Offset := nanos.ediv 1000000000 let daysSinceEpoch : Day.Offset := secs.ediv 86400 let boundedDaysSinceEpoch := daysSinceEpoch @@ -133,9 +133,6 @@ Adds a `Day.Offset` to a `LocalDateTime`. def addDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := { dt with date := dt.date.addDays days } -instance : HAdd LocalDateTime Day.Offset LocalDateTime where - hAdd := addDays - /-- Subtracts a `Day.Offset` from a `LocalDateTime`. -/ @@ -143,9 +140,6 @@ Subtracts a `Day.Offset` from a `LocalDateTime`. def subDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := { dt with date := dt.date.subDays days } -instance : HSub LocalDateTime Day.Offset LocalDateTime where - hSub := subDays - /-- Adds a `Month.Offset` to a `LocalDateTime`, adjusting the day to the last valid day of the resulting month. @@ -210,7 +204,7 @@ def subYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := Adds an `Hour.Offset` to a `LocalDateTime`, adjusting the date if the hour overflows. -/ @[inline] -def addHour (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := +def addHours (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := let totalSeconds := dt.time.toSeconds + hours.toSeconds let days := totalSeconds.ediv 86400 let newTime := dt.time.addSeconds (hours.toSeconds) @@ -220,14 +214,14 @@ def addHour (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := Subtracts an `Hour.Offset` from a `LocalDateTime`, adjusting the date if the hour underflows. -/ @[inline] -def subHour (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := - addHour dt (-hours) +def subHours (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := + addHours dt (-hours) /-- Adds a `Minute.Offset` to a `LocalDateTime`, adjusting the hour and date if the minutes overflow. -/ @[inline] -def addMinute (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := +def addMinutes (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := let totalSeconds := dt.time.toSeconds + minutes.toSeconds let days := totalSeconds.ediv 86400 let newTime := dt.time.addSeconds (minutes.toSeconds) @@ -237,14 +231,14 @@ def addMinute (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := Subtracts a `Minute.Offset` from a `LocalDateTime`, adjusting the hour and date if the minutes underflow. -/ @[inline] -def subMinute (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := - addMinute dt (-minutes) +def subMinutes (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := + addMinutes dt (-minutes) /-- Adds a `Second.Offset` to a `LocalDateTime`, adjusting the minute, hour, and date if the seconds overflow. -/ @[inline] -def addSecond (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := +def addSeconds (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := let totalSeconds := dt.time.toSeconds + seconds let days := totalSeconds.ediv 86400 let newTime := dt.time.addSeconds seconds @@ -254,14 +248,14 @@ def addSecond (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := Subtracts a `Second.Offset` from a `LocalDateTime`, adjusting the minute, hour, and date if the seconds underflow. -/ @[inline] -def subSecond (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := - addSecond dt (-seconds) +def subSeconds (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := + addSeconds dt (-seconds) /-- Adds a `Nanosecond.Offset` to a `LocalDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds overflow. -/ @[inline] -def addNanosecond (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := +def addNanoseconds (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := let nano : Nanosecond.Offset := UnitVal.mk dt.time.nano.val let totalNanos := nano + nanos let extraSeconds : Second.Offset := totalNanos.ediv 1000000000 @@ -273,8 +267,8 @@ def addNanosecond (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTi Subtracts a `Nanosecond.Offset` from a `LocalDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds underflow. -/ @[inline] -def subNanosecond (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := - addNanosecond dt (-nanos) +def subNanoseconds (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := + addNanoseconds dt (-nanos) /-- Getter for the `Year` inside of a `LocalDateTime`. @@ -325,29 +319,35 @@ Getter for the `Second` inside of a `LocalDateTime`. def nanosecond (dt : LocalDateTime) : Nanosecond.Ordinal := dt.time.nano +instance : HAdd LocalDateTime Day.Offset LocalDateTime where + hAdd := addDays + +instance : HSub LocalDateTime Day.Offset LocalDateTime where + hSub := subDays + instance : HAdd LocalDateTime Hour.Offset LocalDateTime where - hAdd := addHour + hAdd := addHours instance : HSub LocalDateTime Hour.Offset LocalDateTime where - hSub := subHour + hSub := subHours instance : HAdd LocalDateTime Minute.Offset LocalDateTime where - hAdd := addMinute + hAdd := addMinutes instance : HSub LocalDateTime Minute.Offset LocalDateTime where - hSub := subMinute + hSub := subMinutes instance : HAdd LocalDateTime Second.Offset LocalDateTime where - hAdd := addSecond + hAdd := addSeconds instance : HSub LocalDateTime Second.Offset LocalDateTime where - hSub := subSecond + hSub := subSeconds instance : HAdd LocalDateTime Nanosecond.Offset LocalDateTime where - hAdd := addNanosecond + hAdd := addNanoseconds instance : HSub LocalDateTime Nanosecond.Offset LocalDateTime where - hSub := subNanosecond + hSub := subNanoseconds end LocalDateTime end Time diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index b8bc2aad74f1..bf334e8f7d4f 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -157,11 +157,25 @@ def ofNanosecondsSinceUnixEpoch (s : Nanosecond.Offset) : Timestamp := by Converts a `Timestamp` from a `Nanosecond.Offset` -/ @[inline] -def toNanoseconds (tm : Timestamp) : Nanosecond.Offset := +def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := let nanos := tm.toSeconds.mul 1000000000 let nanos := nanos + (UnitVal.mk tm.nano.val) nanos +/-- +Adds a `Nanosecond.Offset` to a `Timestamp` +-/ +@[inline] +def addNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := + t.add (ofNanosecondsSinceUnixEpoch s) + +/-- +Adds a `Nanosecond.Offset` to a `Timestamp` +-/ +@[inline] +def subNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := + t.sub (ofNanosecondsSinceUnixEpoch s) + /-- Adds a `Second.Offset` to a `Timestamp` -/ @@ -248,6 +262,12 @@ instance : HAdd Timestamp Second.Offset Timestamp where instance : HSub Timestamp Second.Offset Timestamp where hSub := subSeconds +instance : HAdd Timestamp Nanosecond.Offset Timestamp where + hAdd := addNanoseconds + +instance : HSub Timestamp Nanosecond.Offset Timestamp where + hSub := subNanoseconds + end Timestamp end Time end Std diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 06b5105b9a61..cee824fd135a 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -50,5 +50,5 @@ def isZero (d : Duration) : Bool := instance : HAdd Timestamp Duration Timestamp where hAdd := add -instance : HAdd Timestamp Timestamp Duration where - hAdd := Std.Time.Timestamp.sub +instance : HSub Timestamp Timestamp Duration where + hSub := Std.Time.Timestamp.sub diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 0b93507d733f..37a8ebd218b4 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -66,6 +66,62 @@ def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : DateTime tz := let tm := Timestamp.ofLocalDateTime date DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk <| λ_ => date) +/-- +Add `Hour.Offset` to a `DateTime`. +-/ +@[inline] +def addHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addHours hours) tz + +/-- +Subtract `Hour.Offset` from a `DateTime`. +-/ +@[inline] +def subHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subHours hours) tz + +/-- +Add `Minute.Offset` to a `DateTime`. +-/ +@[inline] +def addMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addMinutes minutes) tz + +/-- +Subtract `Minute.Offset` from a `DateTime`. +-/ +@[inline] +def subMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subMinutes minutes) tz + +/-- +Add `Second.Offset` to a `DateTime`. +-/ +@[inline] +def addSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addSeconds seconds) tz + +/-- +Subtract `Second.Offset` from a `DateTime`. +-/ +@[inline] +def subSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subSeconds seconds) tz + +/-- +Add `Nanosecond.Offset` to a `DateTime`. +-/ +@[inline] +def addNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.addNanoseconds nanoseconds) tz + +/-- +Subtract `Nanosecond.Offset` from a `DateTime`. +-/ +@[inline] +def subNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := + ofLocalDateTime (dt.timestamp.toLocalDateTime.subNanoseconds nanoseconds) tz + /-- Add `Day.Offset` to a `DateTime`. -/ @@ -201,6 +257,36 @@ Gets the `Weekday` of a DateTime. def weekday (dt : DateTime tz) : Weekday := dt.date.get.date.weekday +instance : HAdd (DateTime tz) (Day.Offset) (DateTime tz) where + hAdd := addDays + +instance : HSub (DateTime tz) (Day.Offset) (DateTime tz) where + hSub := subDays + +instance : HAdd (DateTime tz) (Hour.Offset) (DateTime tz) where + hAdd := addHours + +instance : HSub (DateTime tz) (Hour.Offset) (DateTime tz) where + hSub := subHours + +instance : HAdd (DateTime tz) (Minute.Offset) (DateTime tz) where + hAdd := addMinutes + +instance : HSub (DateTime tz) (Minute.Offset) (DateTime tz) where + hSub := subMinutes + +instance : HAdd (DateTime tz) (Second.Offset) (DateTime tz) where + hAdd := addSeconds + +instance : HSub (DateTime tz) (Second.Offset) (DateTime tz) where + hSub := subSeconds + +instance : HAdd (DateTime tz) (Nanosecond.Offset) (DateTime tz) where + hAdd := addNanoseconds + +instance : HSub (DateTime tz) (Nanosecond.Offset) (DateTime tz) where + hSub := subNanoseconds + end DateTime end Time end Std diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index a5e183fc010f..181f83b39c7b 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -137,6 +137,19 @@ Returns the weekday. def weekday (zdt : ZonedDateTime) : Weekday := zdt.snd.weekday +/-- +Add `Day.Offset` to a `ZonedDateTime`. +-/ +def addDays (dt : ZonedDateTime) (days : Day.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addDays days) + +/-- +Subtracts `Day.Offset` to a `ZonedDateTime`. +-/ +@[inline] +def subDays (dt : ZonedDateTime) (days : Day.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subDays days) + /-- Add `Month.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. -/ @@ -180,6 +193,69 @@ Add `Year.Offset` to a `ZonedDateTime`, it clips the day to the last valid day o def addYearsClip (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := Sigma.mk dt.fst (dt.snd.addYearsClip years) +/-- +Add `Hour.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +-/ +def addHours (dt : ZonedDateTime) (hours : Hour.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addHours hours) + +/-- +Subtract `Hour.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +-/ +@[inline] +def subHours (dt : ZonedDateTime) (hours : Hour.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subHours hours) + +/-- +Add `Minute.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +-/ +def addMinutes (dt : ZonedDateTime) (minutes : Minute.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addMinutes minutes) + +/-- +Subtract `Minute.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +-/ +@[inline] +def subMinutes (dt : ZonedDateTime) (minutes : Minute.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subMinutes minutes) + +/-- +Add `Second.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +-/ +def addSeconds (dt : ZonedDateTime) (seconds : Second.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addSeconds seconds) + +/-- +Subtract `Second.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +-/ +@[inline] +def subSeconds (dt : ZonedDateTime) (seconds : Second.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subSeconds seconds) + +instance : HAdd ZonedDateTime (Day.Offset) ZonedDateTime where + hAdd := addDays + +instance : HSub ZonedDateTime (Day.Offset) ZonedDateTime where + hSub := subDays + +instance : HAdd ZonedDateTime (Hour.Offset) ZonedDateTime where + hAdd := addHours + +instance : HSub ZonedDateTime (Hour.Offset) ZonedDateTime where + hSub := subHours + +instance : HAdd ZonedDateTime (Minute.Offset) ZonedDateTime where + hAdd := addMinutes + +instance : HSub ZonedDateTime (Minute.Offset) ZonedDateTime where + hSub := subMinutes + +instance : HAdd ZonedDateTime (Second.Offset) ZonedDateTime where + hAdd := addSeconds + +instance : HSub ZonedDateTime (Second.Offset) ZonedDateTime where + hSub := subSeconds + end ZonedDateTime end Time end Std From d1bef1816ba7df28973f0426511d3a7146bf29ff Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 02:37:25 -0300 Subject: [PATCH 033/118] test: add test for time operations --- src/Std/Time/Time/LocalTime.lean | 17 +- tests/lean/run/timeClassOperations.lean | 154 ++++++++++++++ tests/lean/run/timeOperations.lean | 268 ++++++++++++++++++++++++ 3 files changed, 432 insertions(+), 7 deletions(-) create mode 100644 tests/lean/run/timeClassOperations.lean create mode 100644 tests/lean/run/timeOperations.lean diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/LocalTime.lean index fbf84c2e124b..073aa2ae437a 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/LocalTime.lean @@ -146,7 +146,10 @@ Creates a `LocalTime` value from a total number of seconds. -/ @[inline] def ofSeconds (secs : Second.Offset) : LocalTime := - ofNanoseconds (secs.mul 1000000000) + have hours := Bounded.LE.byEmod (secs.val / 3600) 24 (by decide) + have minutes := (Bounded.LE.byEmod secs.val 3600 (by decide)).ediv 60 (by decide) + have seconds := Bounded.LE.byEmod secs.val 60 (by decide) + ofValidHourMinuteSecondsNano hours minutes seconds 0 /-- Adds seconds to a `LocalTime`. @@ -168,8 +171,8 @@ Adds minutes to a `LocalTime`. -/ @[inline] def addMinutes (time : LocalTime) (minutesToAdd : Minute.Offset) : LocalTime := - let totalMinutes := time.toMinutes + minutesToAdd - ofSeconds (totalMinutes.toSeconds) + let total := time.toSeconds + minutesToAdd.toSeconds + ofSeconds total /-- Subtracts minutes from a `LocalTime`. @@ -182,8 +185,8 @@ def subMinutes (time : LocalTime) (minutesToSub : Minute.Offset) : LocalTime := Adds hours to a `LocalTime`. -/ def addHours (time : LocalTime) (hoursToAdd : Hour.Offset) : LocalTime := - let totalHours := time.toHours + hoursToAdd - ofSeconds (totalHours.toSeconds) + let total := time.toSeconds + hoursToAdd.toSeconds + ofSeconds total /-- Subtracts hours from a `LocalTime`. @@ -196,8 +199,8 @@ def subHours (time : LocalTime) (hoursToSub : Hour.Offset) : LocalTime := Adds nanoseconds to a `LocalTime`. -/ def addNanoseconds (time : LocalTime) (nanosToAdd : Nanosecond.Offset) : LocalTime := - let totalNanos := time.toNanoseconds + nanosToAdd - ofNanoseconds totalNanos + let total := time.toNanoseconds + nanosToAdd + ofNanoseconds total /-- Subtracts nanoseconds from a `LocalTime`. diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean new file mode 100644 index 000000000000..ca1fcbbd00f5 --- /dev/null +++ b/tests/lean/run/timeClassOperations.lean @@ -0,0 +1,154 @@ +import Std.Time +open Std.Time + +def date := date% 1970-01-20 + +/-- +info: 1970-02-01 +-/ +#guard_msgs in +#eval date + (12 : Day.Offset) + +/-- +info: 1970-01-08 +-/ +#guard_msgs in +#eval date - (12 : Day.Offset) + +def datetime := date% 2000-01-20:03:02:01 + +/-- +info: Thursday, January 20, 2000 04:02:01 +-/ +#guard_msgs in +#eval datetime + (1 : Hour.Offset) + +/-- +info: Thursday, January 20, 2000 02:02:01 +-/ +#guard_msgs in +#eval datetime - (1 : Hour.Offset) + +/-- +info: Thursday, January 20, 2000 03:12:01 +-/ +#guard_msgs in +#eval datetime + (10 : Minute.Offset) + +/-- +info: Thursday, January 20, 2000 02:52:01 +-/ +#guard_msgs in +#eval datetime - (10 : Minute.Offset) + +/-- +info: Thursday, January 20, 2000 03:03:01 +-/ +#guard_msgs in +#eval datetime + (60 : Second.Offset) + +/-- +info: Thursday, January 20, 2000 03:01:01 +-/ +#guard_msgs in +#eval datetime - (60 : Second.Offset) + +/-- +info: Friday, January 21, 2000 03:02:01 +-/ +#guard_msgs in +#eval datetime + (1 : Day.Offset) + +/-- +info: Wednesday, January 19, 2000 03:02:01 +-/ +#guard_msgs in +#eval datetime - (1 : Day.Offset) + +def time := time% 13:02:01 + +/-- +info: 14:02:01 +-/ +#guard_msgs in +#eval time + (1 : Hour.Offset) + +/-- +info: 12:02:01 +-/ +#guard_msgs in +#eval time - (1 : Hour.Offset) + +/-- +info: 13:12:01 +-/ +#guard_msgs in +#eval time + (10 : Minute.Offset) + +/-- +info: 12:52:01 +-/ +#guard_msgs in +#eval time - (10 : Minute.Offset) + +/-- +info: 13:03:01 +-/ +#guard_msgs in +#eval time + (60 : Second.Offset) + +/-- +info: 13:01:01 +-/ +#guard_msgs in +#eval time - (60 : Second.Offset) + +def datetimetz := date% 2000-01-20:03:02:01-03:00 + +/-- +info: Sat, 22 Jan 2000 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz + (2 : Day.Offset) + +/-- +info: Wed, 19 Jan 2000 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz - (1 : Day.Offset) + +/-- +info: Thu, 20 Jan 2000 07:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz + (1 : Hour.Offset) + +/-- +info: Thu, 20 Jan 2000 05:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz - (1 : Hour.Offset) + +/-- +info: Thu, 20 Jan 2000 06:12:01 -0300 +-/ +#guard_msgs in +#eval datetimetz + (10 : Minute.Offset) + +/-- +info: Thu, 20 Jan 2000 05:52:01 -0300 +-/ +#guard_msgs in +#eval datetimetz - (10 : Minute.Offset) + +/-- +info: Thu, 20 Jan 2000 06:03:01 -0300 +-/ +#guard_msgs in +#eval datetimetz + (60 : Second.Offset) + +/-- +info: Thu, 20 Jan 2000 06:01:01 -0300 +-/ +#guard_msgs in +#eval datetimetz - (60 : Second.Offset) diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean new file mode 100644 index 000000000000..f7536f030c04 --- /dev/null +++ b/tests/lean/run/timeOperations.lean @@ -0,0 +1,268 @@ +import Std.Time +open Std.Time + +def date := date% 1970-01-20 + +/-- +info: 1970-02-01 +-/ +#guard_msgs in +#eval date.addDays 12 + +/-- +info: 1970-02-20 +-/ +#guard_msgs in +#eval date.addMonthsClip 1 + +/-- +info: 1971-01-20 +-/ +#guard_msgs in +#eval date.addYearsRollOver 1 + +/-- +info: 1969-01-20 +-/ +#guard_msgs in +#eval date.subYearsClip 1 + +/-- +info: 1969-12-20 +-/ +#guard_msgs in +#eval date.subMonthsClip 1 + +def datetime := date% 2000-01-20:03:02:01 + +/-- +info: Thursday, January 20, 2000 04:02:01 +-/ +#guard_msgs in +#eval datetime.addHours 1 + +/-- +info: Thursday, January 20, 2000 02:02:01 +-/ +#guard_msgs in +#eval datetime.subHours 1 + +/-- +info: Thursday, January 20, 2000 03:12:01 +-/ +#guard_msgs in +#eval datetime.addMinutes 10 + +/-- +info: Thursday, January 20, 2000 02:52:01 +-/ +#guard_msgs in +#eval datetime.subMinutes 10 + +/-- +info: Thursday, January 20, 2000 03:03:01 +-/ +#guard_msgs in +#eval datetime.addSeconds 60 + +/-- +info: Thursday, January 20, 2000 03:01:01 +-/ +#guard_msgs in +#eval datetime.subSeconds 60 + +/-- +info: Friday, January 21, 2000 03:02:01 +-/ +#guard_msgs in +#eval datetime.addDays 1 + +/-- +info: Wednesday, January 19, 2000 03:02:01 +-/ +#guard_msgs in +#eval datetime.subDays 1 + +/-- +info: Friday, February 20, 2000 03:02:01 +-/ +#guard_msgs in +#eval datetime.addMonthsClip 1 + +/-- +info: Monday, December 20, 1999 03:02:01 +-/ +#guard_msgs in +#eval datetime.subMonthsClip 1 + +/-- +info: Friday, February 20, 2000 03:02:01 +-/ +#guard_msgs in +#eval datetime.addMonthsRollOver 1 + +/-- +info: Monday, December 20, 1999 03:02:01 +-/ +#guard_msgs in +#eval datetime.subMonthsRollOver 1 + +/-- +info: Saturday, January 20, 2001 03:02:01 +-/ +#guard_msgs in +#eval datetime.addYearsClip 1 + +/-- +info: Wednesday, January 20, 1999 03:02:01 +-/ +#guard_msgs in +#eval datetime.subYearsClip 1 + +/-- +info: Saturday, January 20, 2001 03:02:01 +-/ +#guard_msgs in +#eval datetime.addYearsRollOver 1 + +/-- +info: Wednesday, January 20, 1999 03:02:01 +-/ +#guard_msgs in +#eval datetime.subYearsRollOver 1 + +def time := time% 13:02:01 + +/-- +info: 14:02:01 +-/ +#guard_msgs in +#eval time.addHours 1 + +/-- +info: 12:02:01 +-/ +#guard_msgs in +#eval time.subHours 1 + +/-- +info: 13:12:01 +-/ +#guard_msgs in +#eval time.addMinutes 10 + +/-- +info: 12:52:01 +-/ +#guard_msgs in +#eval time.subMinutes 10 + +/-- +info: 13:03:01 +-/ +#guard_msgs in +#eval time.addSeconds 60 + +/-- +info: 13:01:01 +-/ +#guard_msgs in +#eval time.subSeconds 60 + +def datetimetz := date% 2000-01-20:03:02:01-03:00 + +/-- +info: Sat, 22 Jan 2000 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addDays 2 + +/-- +info: Wed, 19 Jan 2000 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.subDays 1 + +/-- +info: Fri, 20 Feb 2000 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addMonthsClip 1 + +/-- +info: Mon, 20 Dec 1999 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.subMonthsClip 1 + +/-- +info: Fri, 20 Feb 2000 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addMonthsRollOver 1 + +/-- +info: Mon, 20 Dec 1999 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.subMonthsRollOver 1 + +/-- +info: Sat, 20 Jan 2001 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addYearsClip 1 + +/-- +info: Sat, 20 Jan 2001 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addYearsClip 1 + +/-- +info: Sat, 20 Jan 2001 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addYearsRollOver 1 + +/-- +info: Wed, 20 Jan 1999 06:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.subYearsRollOver 1 + +/-- +info: Thu, 20 Jan 2000 07:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addHours 1 + +/-- +info: Thu, 20 Jan 2000 05:02:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.subHours 1 + +/-- +info: Thu, 20 Jan 2000 06:12:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addMinutes 10 + +/-- +info: Thu, 20 Jan 2000 05:52:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.subMinutes 10 + +/-- +info: Thu, 20 Jan 2000 06:03:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.addSeconds 60 + +/-- +info: Thu, 20 Jan 2000 06:01:01 -0300 +-/ +#guard_msgs in +#eval datetimetz.subSeconds 60 From ee0578f184917a2d50c46ac5659258e37a3b3320 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 02:39:44 -0300 Subject: [PATCH 034/118] chore: remove useless comments --- src/runtime/io.cpp | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index f6dc99ca6859..18a81ac5a7d7 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -553,10 +553,10 @@ extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { auto now_time_t = system_clock::to_time_t(now); std::tm tm_info; -#if defined(_MSC_VER) // For Microsoft Visual C++ - localtime_s(&tm_info, &now_time_t); // localtime_s for thread safety +#if defined(_MSC_VER) + localtime_s(&tm_info, &now_time_t); #else - localtime_r(&now_time_t, &tm_info); // localtime_r for POSIX systems + localtime_r(&now_time_t, &tm_info); #endif int offset_hour = tm_info.tm_gmtoff / 3600; From ab1d8cd316c5833423175da8d78b21cc8b10911b Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 23 Aug 2024 03:24:16 -0300 Subject: [PATCH 035/118] feat: changed function names and added ofSeconds and ofNanoseconds to Duration --- src/Std/Time/DateTime.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 4 ++-- src/Std/Time/Duration.lean | 12 ++++++++++++ src/Std/Time/Zoned/ZoneRules.lean | 2 +- 4 files changed, 16 insertions(+), 4 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 9a48d6b5ed13..ae6583f3989e 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -41,7 +41,7 @@ Converts a `Timestamp` to a `LocalDate` -/ @[inline] def toLocalDate (timestamp : Timestamp) : LocalDate := - let secs := timestamp.toSeconds + let secs := timestamp.toSecondsSinceUnixEpoch let days := Day.Offset.ofSeconds secs LocalDate.ofDaysSinceUNIXEpoch days diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index bf334e8f7d4f..7853374ead90 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -60,7 +60,7 @@ opaque now : IO Timestamp Converts a `Timestamp` into its equivalent `Second.Offset`. -/ @[inline] -def toSeconds (t : Timestamp) : Second.Offset := +def toSecondsSinceUnixEpoch (t : Timestamp) : Second.Offset := t.second /-- @@ -158,7 +158,7 @@ Converts a `Timestamp` from a `Nanosecond.Offset` -/ @[inline] def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := - let nanos := tm.toSeconds.mul 1000000000 + let nanos := tm.toSecondsSinceUnixEpoch.mul 1000000000 let nanos := nanos + (UnitVal.mk tm.nano.val) nanos diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index cee824fd135a..26ce8e0e91eb 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -28,6 +28,18 @@ instance : ToString Duration where namespace Duration +/-- +Creates a new `Duration` out of `Second.Offset`. +-/ +def ofSeconds (secs : Second.Offset) : Duration := + Timestamp.ofSecondsSinceUnixEpoch secs + +/-- +Creates a new `Duration` out of `Nanosecond.Offset`. +-/ +def ofNanosecond (secs : Nanosecond.Offset) : Duration := + Timestamp.ofNanosecondsSinceUnixEpoch secs + /-- Calculates a `Duration` out of two `Timestamp`s. -/ diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 337dd16a4125..ccfdfb714cc0 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -170,7 +170,7 @@ Finds the transition corresponding to a given timestamp in `ZoneRules`. If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. -/ def findTransitionForTimestamp (zoneRules : ZoneRules) (timestamp : Timestamp) : Option Transition := - let value := timestamp.toSeconds + let value := timestamp.toSecondsSinceUnixEpoch if let some idx := zoneRules.transitions.findIdx? (λ t => t.time.val > value.val) then zoneRules.transitions.get? (idx - 1) else zoneRules.transitions.back? From cafb98fc2de86383373efaf30a271379e216c441 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 26 Aug 2024 18:44:17 -0300 Subject: [PATCH 036/118] chore: reverting the branch :/ From c82e856d2fd9967361c9885a0ce946e96fe6cca1 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 26 Aug 2024 18:44:40 -0300 Subject: [PATCH 037/118] chore: reverting the branch :/ From 091a96902af4330750f198acc212e1f6231dc66f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 26 Aug 2024 19:19:21 -0300 Subject: [PATCH 038/118] refactor: changes Rat to Std.Internal and Parsec imports that was broken --- src/Lean/Data.lean | 1 - src/Lean/Meta/Tactic/LinearArith/Solver.lean | 3 ++- src/{Lean/Data => Std/Internal}/Rat.lean | 6 ++++-- src/Std/Time/Date/LocalDate.lean | 3 ++- src/Std/Time/Date/Unit/Month.lean | 3 ++- src/Std/Time/Date/Unit/WeekOfYear.lean | 3 ++- src/Std/Time/Date/Unit/Weekday.lean | 3 ++- src/Std/Time/Date/Unit/Year.lean | 3 ++- src/Std/Time/Format/Basic.lean | 6 +++--- src/Std/Time/Internal/UnitVal.lean | 4 ++-- src/Std/Time/Time/Unit/Hour.lean | 3 ++- src/Std/Time/Time/Unit/Millisecond.lean | 3 ++- src/Std/Time/Time/Unit/Minute.lean | 3 ++- src/Std/Time/Time/Unit/Nanosecond.lean | 3 ++- src/Std/Time/Time/Unit/Second.lean | 3 ++- src/Std/Time/Zoned/Database/TzIf.lean | 6 +++--- tests/lean/hpow.lean | 3 ++- tests/lean/ppNumericTypes.lean | 3 ++- tests/lean/rat1.lean | 3 ++- 19 files changed, 40 insertions(+), 25 deletions(-) rename src/{Lean/Data => Std/Internal}/Rat.lean (98%) diff --git a/src/Lean/Data.lean b/src/Lean/Data.lean index 7a1854d20760..d29a2c1e0bc4 100644 --- a/src/Lean/Data.lean +++ b/src/Lean/Data.lean @@ -29,4 +29,3 @@ import Lean.Data.Xml import Lean.Data.NameTrie import Lean.Data.RBTree import Lean.Data.RBMap -import Lean.Data.Rat diff --git a/src/Lean/Meta/Tactic/LinearArith/Solver.lean b/src/Lean/Meta/Tactic/LinearArith/Solver.lean index 8a560658d798..066462885916 100644 --- a/src/Lean/Meta/Tactic/LinearArith/Solver.lean +++ b/src/Lean/Meta/Tactic/LinearArith/Solver.lean @@ -6,9 +6,10 @@ Authors: Leonardo de Moura prelude import Init.Data.Ord import Init.Data.Array.DecidableEq -import Lean.Data.Rat +import Std.Internal.Rat namespace Lean.Meta.Linear +open Std.Internal structure Var where id : Nat diff --git a/src/Lean/Data/Rat.lean b/src/Std/Internal/Rat.lean similarity index 98% rename from src/Lean/Data/Rat.lean rename to src/Std/Internal/Rat.lean index 03fe31ddee87..b8ee378c2aef 100644 --- a/src/Lean/Data/Rat.lean +++ b/src/Std/Internal/Rat.lean @@ -8,7 +8,8 @@ import Init.NotationExtra import Init.Data.ToString.Macro import Init.Data.Int.DivMod import Init.Data.Nat.Gcd -namespace Lean +namespace Std +namespace Internal /-! Rational numbers for implementing decision procedures. @@ -144,4 +145,5 @@ instance : Coe Int Rat where coe num := { num } end Rat -end Lean +end Internal +end Std diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/LocalDate.lean index f34eec60d00b..537f3f141708 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/LocalDate.lean @@ -6,10 +6,11 @@ Authors: Sofia Rodrigues prelude import Std.Time.Internal import Std.Time.Date.Basic -import Lean.Data.Rat +import Std.Internal.Rat namespace Std namespace Time +open Std.Internal open Internal open Lean diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 0c786f20b835..45ba1e703ae2 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Date.Unit.Day namespace Std namespace Time namespace Month +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Date/Unit/WeekOfYear.lean b/src/Std/Time/Date/Unit/WeekOfYear.lean index 93d907b0d325..0df5a2db38a0 100644 --- a/src/Std/Time/Date/Unit/WeekOfYear.lean +++ b/src/Std/Time/Date/Unit/WeekOfYear.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Date.Unit.Day namespace Std namespace Time namespace WeekOfYear +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index f1110aeb0d08..c86f3495aca9 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -4,11 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Date.Unit.Day namespace Std namespace Time +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index 4534d230d747..d704020eaedd 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -5,13 +5,14 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Internal -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Date.Unit.Day import Std.Time.Date.Unit.Month namespace Std namespace Time namespace Year +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 3e73a425d812..364546a9b820 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Parsec +import Std.Internal.Parsec import Std.Time.Date import Std.Time.Time import Std.Time.Zoned @@ -13,8 +13,8 @@ import Std.Time.DateTime namespace Std namespace Time open Internal -open Lean.Parsec.String -open Lean.Parsec Lean LocalTime LocalDate TimeZone DateTime +open Std.Internal.Parsec.String +open Std.Internal.Parsec Lean LocalTime LocalDate TimeZone DateTime set_option linter.all true diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 5039147b6041..370ad5a7a289 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -5,12 +5,12 @@ Authors: Sofia Rodrigues -/ prelude import Init.Data -import Lean.Data.Rat +import Std.Internal.Rat namespace Std namespace Time namespace Internal - +open Std.Internal open Lean set_option linter.all true diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index f5b6d185bd64..293aa8e1668f 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Internal import Std.Time.Time.Unit.Minute import Std.Time.Time.Unit.Second @@ -12,6 +12,7 @@ import Std.Time.Time.Unit.Second namespace Std namespace Time namespace Hour +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 1b995a486763..738fe81ccf2c 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Internal namespace Std namespace Time namespace Millisecond +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index ce9fad7176f9..ec36f5d6ed46 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -4,13 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Internal import Std.Time.Time.Unit.Second namespace Std namespace Time namespace Minute +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index de186ff053e4..9b0d1799b416 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Time.Unit.Millisecond namespace Std namespace Time namespace Nanosecond +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index e6e2c5a47000..232f824b637b 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Lean.Data.Rat +import Std.Internal.Rat import Std.Time.Time.Unit.Nanosecond namespace Std namespace Time namespace Second +open Std.Internal open Internal set_option linter.all true diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index 081c7179a144..1f46a12c1c13 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -5,8 +5,8 @@ Authors: Sofia Rodrigues -/ prelude import Init.Data.Range -import Lean.Data.Parsec -import Lean.Data.Parsec.ByteArray +import Std.Internal.Parsec +import Std.Internal.Parsec.ByteArray -- Based on: https://www.rfc-editor.org/rfc/rfc8536.html @@ -14,7 +14,7 @@ namespace Std namespace Time namespace TimeZone namespace TZif -open Lean.Parsec Lean.Parsec.ByteArray +open Std.Internal.Parsec Std.Internal.Parsec.ByteArray set_option linter.all true diff --git a/tests/lean/hpow.lean b/tests/lean/hpow.lean index 6dd24d41c369..398aa14d4d8d 100644 --- a/tests/lean/hpow.lean +++ b/tests/lean/hpow.lean @@ -1,4 +1,5 @@ -import Lean.Data.Rat +import Std.Internal.Rat +open Std.Internal /-! Test the `rightact%` elaborator for `HPow.hPow`, added to address #2854 diff --git a/tests/lean/ppNumericTypes.lean b/tests/lean/ppNumericTypes.lean index 3acd48538b77..2e054457a0b0 100644 --- a/tests/lean/ppNumericTypes.lean +++ b/tests/lean/ppNumericTypes.lean @@ -1,4 +1,5 @@ -import Lean.Data.Rat +import Std.Internal.Rat +open Std.Internal /-! Tests for `pp.numericTypes` and `pp.natLit` diff --git a/tests/lean/rat1.lean b/tests/lean/rat1.lean index 64d2ffd38068..5ddea7524766 100644 --- a/tests/lean/rat1.lean +++ b/tests/lean/rat1.lean @@ -1,4 +1,5 @@ -import Lean.Data.Rat +import Std.Internal.Rat +open Std.Internal open Lean #eval (15 : Rat) / 10 From 6c3cec4923707760d18626189bbff4edeb16d673 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 00:11:32 -0300 Subject: [PATCH 039/118] refactor: changed names add some utilitary functions --- src/Std/Time.lean | 6 +- src/Std/Time/Date.lean | 2 +- .../Date/{LocalDate.lean => PlainDate.lean} | 104 ++++++------ src/Std/Time/Date/Unit/Month.lean | 18 ++ src/Std/Time/Date/Unit/Year.lean | 14 ++ src/Std/Time/Date/WeekDate.lean | 4 +- src/Std/Time/DateTime.lean | 62 +++---- ...{LocalDateTime.lean => PlainDateTime.lean} | 152 ++++++++--------- src/Std/Time/DateTime/Timestamp.lean | 8 +- src/Std/Time/Format.lean | 157 +++++++++++------- src/Std/Time/Format/Basic.lean | 8 +- src/Std/Time/Notation.lean | 12 +- src/Std/Time/Time.lean | 2 +- .../Time/{LocalTime.lean => PlainTime.lean} | 110 ++++++------ src/Std/Time/Time/Unit/Millisecond.lean | 17 ++ src/Std/Time/Time/Unit/Nanosecond.lean | 18 ++ src/Std/Time/Time/Unit/Second.lean | 18 ++ src/Std/Time/Zoned.lean | 60 +++---- src/Std/Time/Zoned/Database.lean | 2 +- src/Std/Time/Zoned/Database/Basic.lean | 22 +-- src/Std/Time/Zoned/Database/TZdb.lean | 6 +- src/Std/Time/Zoned/Database/TzIf.lean | 26 +-- src/Std/Time/Zoned/DateTime.lean | 54 +++--- src/Std/Time/Zoned/ZoneRules.lean | 28 ++-- src/Std/Time/Zoned/ZonedDateTime.lean | 14 +- tests/lean/run/timeFormats.lean | 32 ++-- tests/lean/run/timeLocalDateTime.lean | 20 +-- tests/lean/run/timeParse.lean | 16 +- tests/lean/run/timeTzifParse.lean | 2 +- 29 files changed, 555 insertions(+), 439 deletions(-) rename src/Std/Time/Date/{LocalDate.lean => PlainDate.lean} (69%) rename src/Std/Time/DateTime/{LocalDateTime.lean => PlainDateTime.lean} (57%) rename src/Std/Time/Time/{LocalTime.lean => PlainTime.lean} (59%) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 2cda878c1577..fb5c77a9dac7 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -88,18 +88,18 @@ that are just shifts in dates and times. Types like `Date` are made using of com ## Date These types provide precision down to the day level, useful for representing and manipulating dates. -- **`LocalDate`:** Represents a calendar date in the format `YYYY-MM-DD`. +- **`PlainDate`:** Represents a calendar date in the format `YYYY-MM-DD`. - **`WeekDate`:** Uses the `YYYY-Www` format with week level precision. ## Time These types offer precision down to the nanosecond level, useful for representing and manipulating time of day. -- **`LocalTime`**: Represents a time of day in the format `HH:mm:ss.SSSSSSSSS`. +- **`PlainTime`**: Represents a time of day in the format `HH:mm:ss.SSSSSSSSS`. ## Date and time Combines date and time into a single representation, useful for precise timestamps and scheduling. -- **`LocalDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. +- **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. - **`Timestamp`**: Represents a point in time with second-level precision. It starts on the UNIX epoch and it should be used when you receive or need to send timestamps to another systems. diff --git a/src/Std/Time/Date.lean b/src/Std/Time/Date.lean index b7634f036bac..e3df6fd60a73 100644 --- a/src/Std/Time/Date.lean +++ b/src/Std/Time/Date.lean @@ -6,4 +6,4 @@ Authors: Sofia Rodrigues prelude import Std.Time.Date.Basic import Std.Time.Date.WeekDate -import Std.Time.Date.LocalDate +import Std.Time.Date.PlainDate diff --git a/src/Std/Time/Date/LocalDate.lean b/src/Std/Time/Date/PlainDate.lean similarity index 69% rename from src/Std/Time/Date/LocalDate.lean rename to src/Std/Time/Date/PlainDate.lean index 537f3f141708..baf8b9e43c8a 100644 --- a/src/Std/Time/Date/LocalDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -17,10 +17,10 @@ open Lean set_option linter.all true /-- -`LocalDate` represents a date in the Year-Month-Day (YMD) format. It encapsulates the year, month, +`PlainDate` represents a date in the Year-Month-Day (YMD) format. It encapsulates the year, month, and day components, with validation to ensure the date is valid. -/ -structure LocalDate where +structure PlainDate where /-- The year component of the date. It is represented as an `Offset` type from `Year`. @@ -42,23 +42,23 @@ structure LocalDate where -/ valid : year.Valid month day -instance : BEq LocalDate where +instance : BEq PlainDate where beq x y := x.day == y.day && x.month == y.month && x.year == y.year -namespace LocalDate +namespace PlainDate /-- -Creates a `LocalDate` by clipping the day to ensure validity. This function forces the date to be +Creates a `PlainDate` by clipping the day to ensure validity. This function forces the date to be valid by adjusting the day to fit within the valid range to fit the given month and year. -/ -def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : LocalDate := +def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := let ⟨day, valid⟩ := month.clipDay year.isLeap day - LocalDate.mk year month day valid + PlainDate.mk year month day valid /-- -Creates a `LocalDate` by rolling over the extra days to the next month. +Creates a `PlainDate` by rolling over the extra days to the next month. -/ -def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : LocalDate := by +def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := by let max : Day.Ordinal := month.days year.isLeap have p := month.all_greater_than_27 year.isLeap if h : day.val > max.val then @@ -81,28 +81,28 @@ def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : let h := Int.not_lt.mp h exact ⟨year, month, day, h⟩ -instance : Inhabited LocalDate where +instance : Inhabited PlainDate where default := clip 0 1 1 /-- -Creates a new `LocalDate` from year, month, and day components. +Creates a new `PlainDate` from year, month, and day components. -/ -def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option LocalDate := +def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option PlainDate := if valid : year.Valid month day - then some (LocalDate.mk year month day valid) + then some (PlainDate.mk year month day valid) else none /-- -Creates a `LocalDate` from a year and a day ordinal within that year. +Creates a `PlainDate` from a year and a day ordinal within that year. -/ -def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : LocalDate := +def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : PlainDate := let ⟨⟨month, day⟩, valid⟩ := Month.Ordinal.ofOrdinal ordinal - LocalDate.mk year month day valid + PlainDate.mk year month day valid /-- -Creates a `LocalDate` from the number of days since the UNIX epoch (January 1st, 1970). +Creates a `PlainDate` from the number of days since the UNIX epoch (January 1st, 1970). -/ -def ofDaysSinceUNIXEpoch (day : Day.Offset) : LocalDate := +def ofDaysSinceUNIXEpoch (day : Day.Offset) : PlainDate := let z := day.toInt + 719468 let era := (if z ≥ 0 then z else z - 146096).div 146097 let doe := z - era * 146097 @@ -116,9 +116,9 @@ def ofDaysSinceUNIXEpoch (day : Day.Offset) : LocalDate := .clip y (.clip m (by decide)) (.clip d (by decide)) /-- -Calculates the `Weekday` of a given `LocalDate` using Zeller's Congruence for the Gregorian calendar. +Calculates the `Weekday` of a given `PlainDate` using Zeller's Congruence for the Gregorian calendar. -/ -def weekday (date : LocalDate) : Weekday := +def weekday (date : PlainDate) : Weekday := let q := date.day.toInt let m := date.month.toInt let y := date.year.toInt @@ -135,9 +135,9 @@ def weekday (date : LocalDate) : Weekday := .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ /-- -Converts a `LocalDate` to the number of days since the UNIX epoch. +Converts a `PlainDate` to the number of days since the UNIX epoch. -/ -def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := +def toDaysSinceUNIXEpoch (date : PlainDate) : Day.Offset := let y : Int := if date.month.toInt > 2 then date.year else date.year.toInt - 1 let era : Int := (if y ≥ 0 then y else y - 399).div 400 let yoe : Int := y - era * 400 @@ -149,30 +149,30 @@ def toDaysSinceUNIXEpoch (date : LocalDate) : Day.Offset := .ofInt (era * 146097 + doe - 719468) /-- -Calculates the difference in years between a `LocalDate` and a given year. +Calculates the difference in years between a `PlainDate` and a given year. -/ -def yearsSince (date : LocalDate) (year : Year.Offset) : Year.Offset := +def yearsSince (date : PlainDate) (year : Year.Offset) : Year.Offset := date.year - year /-- -Adds a given number of days to a `LocalDate`. +Adds a given number of days to a `PlainDate`. -/ @[inline] -def addDays (date : LocalDate) (days : Day.Offset) : LocalDate := +def addDays (date : PlainDate) (days : Day.Offset) : PlainDate := let dateDays := date.toDaysSinceUNIXEpoch ofDaysSinceUNIXEpoch (dateDays + days) /-- -Subtracts a given number of days from a `LocalDate`. +Subtracts a given number of days from a `PlainDate`. -/ @[inline] -def subDays (date : LocalDate) (days : Day.Offset) : LocalDate := +def subDays (date : PlainDate) (days : Day.Offset) : PlainDate := addDays date (-days) /-- -Adds a given number of months to a `LocalDate`, clipping the day to the last valid day of the month. +Adds a given number of months to a `PlainDate`, clipping the day to the last valid day of the month. -/ -def addMonthsClip (date : LocalDate) (months : Month.Offset) : LocalDate := +def addMonthsClip (date : PlainDate) (months : Month.Offset) : PlainDate := let yearsOffset := months.div 12 let monthOffset := Bounded.LE.byMod months 12 (by decide) let months := date.month.addBounds monthOffset @@ -187,19 +187,19 @@ def addMonthsClip (date : LocalDate) (months : Month.Offset) : LocalDate := else exact (yearsOffset, months.truncateTop (Int.not_lt.mp h₁) |>.truncateBottom (Int.not_lt.mp h₂)) - LocalDate.clip (date.year.add yearsOffset) months date.day + PlainDate.clip (date.year.add yearsOffset) months date.day /-- -Subtracts `Month.Offset` from a `LocalDate`, it clips the day to the last valid day of that month. +Subtracts `Month.Offset` from a `PlainDate`, it clips the day to the last valid day of that month. -/ @[inline] -def subMonthsClip (date : LocalDate) (months : Month.Offset) : LocalDate := +def subMonthsClip (date : PlainDate) (months : Month.Offset) : PlainDate := addMonthsClip date (-months) /-- -Adds a given number of months to a `LocalDate`, rolling over any excess days into the following month. +Adds a given number of months to a `PlainDate`, rolling over any excess days into the following month. -/ -def addMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := +def addMonthsRollOver (date : PlainDate) (months : Month.Offset) : PlainDate := let yearsOffset := months.div 12 let monthOffset := Bounded.LE.byMod months 12 (by decide) let months := date.month.addBounds monthOffset @@ -221,54 +221,54 @@ def addMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := have p : year.Valid months date.day := by simp_all [Year.Offset.Valid, Month.Ordinal.Valid] exact Int.le_trans h proof - LocalDate.mk year months date.day p + PlainDate.mk year months date.day p else - let roll : Day.Offset := UnitVal.mk (date.day.val - days.toInt) - let date := LocalDate.clip (date.year.add yearsOffset) months date.day + let roll := Day.Offset.ofInt (date.day.val - days.toInt) + let date := PlainDate.clip (date.year.add yearsOffset) months date.day let days := date.toDaysSinceUNIXEpoch + roll - LocalDate.ofDaysSinceUNIXEpoch days + PlainDate.ofDaysSinceUNIXEpoch days /-- -Subtracts `Month.Offset` from a `LocalDate`, rolling over excess days as needed. +Subtracts `Month.Offset` from a `PlainDate`, rolling over excess days as needed. -/ @[inline] -def subMonthsRollOver (date : LocalDate) (months : Month.Offset) : LocalDate := +def subMonthsRollOver (date : PlainDate) (months : Month.Offset) : PlainDate := addMonthsRollOver date (-months) /-- -Adds `Year.Offset` to a `LocalDate`, rolling over excess days to the next month, or next year. +Adds `Year.Offset` to a `PlainDate`, rolling over excess days to the next month, or next year. -/ @[inline] -def addYearsRollOver (date : LocalDate) (years : Year.Offset) : LocalDate := +def addYearsRollOver (date : PlainDate) (years : Year.Offset) : PlainDate := addMonthsRollOver date (years.mul 12) /-- -Subtracts `Year.Offset` from a `LocalDate`, rolling over excess days to the next month. +Subtracts `Year.Offset` from a `PlainDate`, rolling over excess days to the next month. -/ @[inline] -def subYearsRollOver (date : LocalDate) (years : Year.Offset) : LocalDate := +def subYearsRollOver (date : PlainDate) (years : Year.Offset) : PlainDate := addMonthsRollOver date (- years.mul 12) /-- -Adds `Year.Offset` to a `LocalDate`, clipping the day to the last valid day of the month. +Adds `Year.Offset` to a `PlainDate`, clipping the day to the last valid day of the month. -/ @[inline] -def addYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := +def addYearsClip (date : PlainDate) (years : Year.Offset) : PlainDate := addMonthsClip date (years.mul 12) /-- -Subtracts `Year.Offset` from a `LocalDate`, clipping the day to the last valid day of the month. +Subtracts `Year.Offset` from a `PlainDate`, clipping the day to the last valid day of the month. -/ @[inline] -def subYearsClip (date : LocalDate) (years : Year.Offset) : LocalDate := +def subYearsClip (date : PlainDate) (years : Year.Offset) : PlainDate := addMonthsClip date (- years.mul 12) -instance : HAdd LocalDate Day.Offset LocalDate where +instance : HAdd PlainDate Day.Offset PlainDate where hAdd := addDays -instance : HSub LocalDate Day.Offset LocalDate where +instance : HSub PlainDate Day.Offset PlainDate where hSub := subDays -end LocalDate +end PlainDate end Time end Std diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 45ba1e703ae2..dccee95aa178 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -36,6 +36,24 @@ def Offset : Type := Int instance : OfNat Offset n := ⟨Int.ofNat n⟩ +namespace Offset + +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + .ofNat data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + data + +end Offset + namespace Ordinal /-- diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index d704020eaedd..f615d4621735 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -27,6 +27,20 @@ instance : OfNat Offset n := ⟨Int.ofNat n⟩ namespace Offset +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + .ofNat data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + data + /-- Converts the `Year` offset to an `Int`. -/ diff --git a/src/Std/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean index 18ce67ea831e..a8de77043822 100644 --- a/src/Std/Time/Date/WeekDate.lean +++ b/src/Std/Time/Date/WeekDate.lean @@ -6,7 +6,7 @@ Authors: Sofia Rodrigues prelude import Std.Time.Date.Unit.Year import Std.Time.Date.Unit.WeekOfYear -import Std.Time.Date.LocalDate +import Std.Time.Date.PlainDate namespace Std namespace Time @@ -37,7 +37,7 @@ Converts a `WeekDate` to a `Day.Offset`. -/ def toDays (wd : WeekDate) : Day.Offset := let days := wd.year.toInt * 365 + wd.week.val * 7 - UnitVal.mk days + Day.Offset.ofInt days /-- Creates a `WeekDate` from a `Day.Offset`. diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index ae6583f3989e..b5159d0d4276 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -5,7 +5,7 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.DateTime.Timestamp -import Std.Time.DateTime.LocalDateTime +import Std.Time.DateTime.PlainDateTime namespace Std namespace Time @@ -14,83 +14,83 @@ namespace Timestamp set_option linter.all true /-- -Converts a `LocalDateTime` to a `Timestamp` +Converts a `PlainDateTime` to a `Timestamp` -/ @[inline] -def ofLocalDateTime (ldt : LocalDateTime) : Timestamp := - ldt.toLocalTimestamp +def ofPlainDateTime (ldt : PlainDateTime) : Timestamp := + ldt.toPlainTimestamp /-- -Converts a `Timestamp` to a `LocalDateTime` +Converts a `Timestamp` to a `PlainDateTime` -/ @[inline] -def toLocalDateTime (timestamp : Timestamp) : LocalDateTime := - LocalDateTime.ofUTCTimestamp timestamp +def toPlainDateTime (timestamp : Timestamp) : PlainDateTime := + PlainDateTime.ofUTCTimestamp timestamp /-- -Converts a `LocalDate` to a `Timestamp` +Converts a `PlainDate` to a `Timestamp` -/ @[inline] -def ofLocalDate (ld : LocalDate) : Timestamp := +def ofPlainDate (ld : PlainDate) : Timestamp := let days := ld.toDaysSinceUNIXEpoch let secs := days.toSeconds Timestamp.ofSecondsSinceUnixEpoch secs /-- -Converts a `Timestamp` to a `LocalDate` +Converts a `Timestamp` to a `PlainDate` -/ @[inline] -def toLocalDate (timestamp : Timestamp) : LocalDate := +def toPlainDate (timestamp : Timestamp) : PlainDate := let secs := timestamp.toSecondsSinceUnixEpoch let days := Day.Offset.ofSeconds secs - LocalDate.ofDaysSinceUNIXEpoch days + PlainDate.ofDaysSinceUNIXEpoch days /-- -Converts a `LocalTime` to a `Timestamp` +Converts a `PlainTime` to a `Timestamp` -/ @[inline] -def ofLocalTime (lt : LocalTime) : Timestamp := +def ofPlainTime (lt : PlainTime) : Timestamp := let nanos := lt.toNanoseconds Timestamp.ofNanosecondsSinceUnixEpoch nanos /-- -Converts a `Timestamp` to a `LocalTime` +Converts a `Timestamp` to a `PlainTime` -/ @[inline] -def toLocalTime (timestamp : Timestamp) : LocalTime := +def toPlainTime (timestamp : Timestamp) : PlainTime := let nanos := timestamp.toNanosecondsSinceUnixEpoch - LocalTime.ofNanoseconds nanos + PlainTime.ofNanoseconds nanos end Timestamp -namespace LocalDateTime +namespace PlainDateTime /-- -Converts a `LocalDate` to a `Timestamp` +Converts a `PlainDate` to a `Timestamp` -/ @[inline] -def ofLocalDate (ld : LocalDate) : LocalDateTime := - LocalDateTime.ofUTCTimestamp (Timestamp.ofLocalDate ld) +def ofPlainDate (ld : PlainDate) : PlainDateTime := + PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainDate ld) /-- -Converts a `LocalDateTime` to a `LocalDate` +Converts a `PlainDateTime` to a `PlainDate` -/ @[inline] -def toLocalDate (ldt : LocalDateTime) : LocalDate := - Timestamp.toLocalDate ldt.toLocalTimestamp +def toPlainDate (ldt : PlainDateTime) : PlainDate := + Timestamp.toPlainDate ldt.toPlainTimestamp /-- -Converts a `LocalTime` to a `LocalDateTime` +Converts a `PlainTime` to a `PlainDateTime` -/ @[inline] -def ofLocalTime (lt : LocalTime) : LocalDateTime := - LocalDateTime.ofUTCTimestamp (Timestamp.ofLocalTime lt) +def ofPlainTime (lt : PlainTime) : PlainDateTime := + PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainTime lt) /-- -Converts a `LocalDateTime` to a `LocalTime` +Converts a `PlainDateTime` to a `PlainTime` -/ @[inline] -def toLocalTime (ldt : LocalDateTime) : LocalTime := - Timestamp.toLocalTime ldt.toLocalTimestamp +def toPlainTime (ldt : PlainDateTime) : PlainTime := + Timestamp.toPlainTime ldt.toPlainTimestamp -end LocalDateTime +end PlainDateTime diff --git a/src/Std/Time/DateTime/LocalDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean similarity index 57% rename from src/Std/Time/DateTime/LocalDateTime.lean rename to src/Std/Time/DateTime/PlainDateTime.lean index 3f8239e52b65..dd63dbc401c0 100644 --- a/src/Std/Time/DateTime/LocalDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -18,35 +18,35 @@ set_option linter.all true /-- Represents a date and time with components for Year, Month, Day, Hour, Minute, Second, and Nanosecond. -/ -structure LocalDateTime where +structure PlainDateTime where /-- - The `Date` component of a `LocalDate` + The `Date` component of a `PlainDate` -/ - date : LocalDate + date : PlainDate /-- - The `Time` component of a `LocalTime` + The `Time` component of a `PlainTime` -/ - time : LocalTime + time : PlainTime deriving Inhabited, BEq -namespace LocalDateTime +namespace PlainDateTime /-- -Converts a `LocalDateTime` to a `Timestamp` +Converts a `PlainDateTime` to a `Timestamp` -/ -def toLocalTimestamp (dt : LocalDateTime) : Timestamp := +def toLocalTimestamp (dt : PlainDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 let nanos := nanos.val + dt.time.nano.val - Timestamp.ofNanosecondsSinceUnixEpoch (UnitVal.mk nanos) + Timestamp.ofNanosecondsSinceUnixEpoch (Nanosecond.Offset.ofInt nanos) /-- -Converts a UNIX `Timestamp` to a `LocalDateTime`. +Converts a UNIX `Timestamp` to a `PlainDateTime`. -/ -def ofUTCTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do +def ofUTCTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do let leapYearEpoch := 11017 let daysPer400Y := 365 * 400 + 97 let daysPer100Y := 365 * 100 + 24 @@ -122,233 +122,233 @@ def ofUTCTimestamp (stamp : Timestamp) : LocalDateTime := Id.run do let nano : Bounded.LE 0 999999999 := Bounded.LE.byEmod nanos.val 1000000000 (by decide) return { - date := LocalDate.clip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) - time := LocalTime.ofValidHourMinuteSecondsNano (hour.expandTop (by decide)) minute (second.expandTop (by decide)) nano + date := PlainDate.clip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) + time := PlainTime.ofValidHourMinuteSecondsNano (hour.expandTop (by decide)) minute (second.expandTop (by decide)) nano } /-- -Adds a `Day.Offset` to a `LocalDateTime`. +Adds a `Day.Offset` to a `PlainDateTime`. -/ @[inline] -def addDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := +def addDays (dt : PlainDateTime) (days : Day.Offset) : PlainDateTime := { dt with date := dt.date.addDays days } /-- -Subtracts a `Day.Offset` from a `LocalDateTime`. +Subtracts a `Day.Offset` from a `PlainDateTime`. -/ @[inline] -def subDays (dt : LocalDateTime) (days : Day.Offset) : LocalDateTime := +def subDays (dt : PlainDateTime) (days : Day.Offset) : PlainDateTime := { dt with date := dt.date.subDays days } /-- -Adds a `Month.Offset` to a `LocalDateTime`, adjusting the day to the last valid day of the resulting +Adds a `Month.Offset` to a `PlainDateTime`, adjusting the day to the last valid day of the resulting month. -/ -def addMonthsClip (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := +def addMonthsClip (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := { dt with date := dt.date.addMonthsClip months } /-- -Subtracts `Month.Offset` from a `LocalDateTime`, it clips the day to the last valid day of that month. +Subtracts `Month.Offset` from a `PlainDateTime`, it clips the day to the last valid day of that month. -/ @[inline] -def subMonthsClip (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := +def subMonthsClip (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := { dt with date := dt.date.subMonthsClip months } /-- -Adds a `Month.Offset` to a `LocalDateTime`, rolling over excess days to the following month if needed. +Adds a `Month.Offset` to a `PlainDateTime`, rolling over excess days to the following month if needed. -/ -def addMonthsRollOver (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := +def addMonthsRollOver (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := { dt with date := dt.date.addMonthsRollOver months } /-- -Subtracts a `Month.Offset` from a `LocalDateTime`, adjusting the day to the last valid day of the +Subtracts a `Month.Offset` from a `PlainDateTime`, adjusting the day to the last valid day of the resulting month. -/ @[inline] -def subMonthsRollOver (dt : LocalDateTime) (months : Month.Offset) : LocalDateTime := +def subMonthsRollOver (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := { dt with date := dt.date.subMonthsRollOver months } /-- -Adds a `Month.Offset` to a `LocalDateTime`, rolling over excess days to the following month if needed. +Adds a `Month.Offset` to a `PlainDateTime`, rolling over excess days to the following month if needed. -/ @[inline] -def addYearsRollOver (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := +def addYearsRollOver (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := { dt with date := dt.date.addYearsRollOver years } /-- -Subtracts a `Month.Offset` from a `LocalDateTime`, rolling over excess days to the following month if +Subtracts a `Month.Offset` from a `PlainDateTime`, rolling over excess days to the following month if needed. -/ @[inline] -def addYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := +def addYearsClip (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := { dt with date := dt.date.addYearsClip years } /-- -Subtracts a `Year.Offset` from a `LocalDateTime`, this function rolls over any excess days into the +Subtracts a `Year.Offset` from a `PlainDateTime`, this function rolls over any excess days into the following month. -/ @[inline] -def subYearsRollOver (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := +def subYearsRollOver (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := { dt with date := dt.date.subYearsRollOver years } /-- -Subtracts a `Year.Offset` from a `LocalDateTime`, adjusting the day to the last valid day of the +Subtracts a `Year.Offset` from a `PlainDateTime`, adjusting the day to the last valid day of the resulting month. -/ @[inline] -def subYearsClip (dt : LocalDateTime) (years : Year.Offset) : LocalDateTime := +def subYearsClip (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := { dt with date := dt.date.subYearsClip years } /-- -Adds an `Hour.Offset` to a `LocalDateTime`, adjusting the date if the hour overflows. +Adds an `Hour.Offset` to a `PlainDateTime`, adjusting the date if the hour overflows. -/ @[inline] -def addHours (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := +def addHours (dt : PlainDateTime) (hours : Hour.Offset) : PlainDateTime := let totalSeconds := dt.time.toSeconds + hours.toSeconds let days := totalSeconds.ediv 86400 let newTime := dt.time.addSeconds (hours.toSeconds) { dt with date := dt.date.addDays days, time := newTime } /-- -Subtracts an `Hour.Offset` from a `LocalDateTime`, adjusting the date if the hour underflows. +Subtracts an `Hour.Offset` from a `PlainDateTime`, adjusting the date if the hour underflows. -/ @[inline] -def subHours (dt : LocalDateTime) (hours : Hour.Offset) : LocalDateTime := +def subHours (dt : PlainDateTime) (hours : Hour.Offset) : PlainDateTime := addHours dt (-hours) /-- -Adds a `Minute.Offset` to a `LocalDateTime`, adjusting the hour and date if the minutes overflow. +Adds a `Minute.Offset` to a `PlainDateTime`, adjusting the hour and date if the minutes overflow. -/ @[inline] -def addMinutes (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := +def addMinutes (dt : PlainDateTime) (minutes : Minute.Offset) : PlainDateTime := let totalSeconds := dt.time.toSeconds + minutes.toSeconds let days := totalSeconds.ediv 86400 let newTime := dt.time.addSeconds (minutes.toSeconds) { dt with date := dt.date.addDays days, time := newTime } /-- -Subtracts a `Minute.Offset` from a `LocalDateTime`, adjusting the hour and date if the minutes underflow. +Subtracts a `Minute.Offset` from a `PlainDateTime`, adjusting the hour and date if the minutes underflow. -/ @[inline] -def subMinutes (dt : LocalDateTime) (minutes : Minute.Offset) : LocalDateTime := +def subMinutes (dt : PlainDateTime) (minutes : Minute.Offset) : PlainDateTime := addMinutes dt (-minutes) /-- -Adds a `Second.Offset` to a `LocalDateTime`, adjusting the minute, hour, and date if the seconds overflow. +Adds a `Second.Offset` to a `PlainDateTime`, adjusting the minute, hour, and date if the seconds overflow. -/ @[inline] -def addSeconds (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := +def addSeconds (dt : PlainDateTime) (seconds : Second.Offset) : PlainDateTime := let totalSeconds := dt.time.toSeconds + seconds let days := totalSeconds.ediv 86400 let newTime := dt.time.addSeconds seconds { dt with date := dt.date.addDays days, time := newTime } /-- -Subtracts a `Second.Offset` from a `LocalDateTime`, adjusting the minute, hour, and date if the seconds underflow. +Subtracts a `Second.Offset` from a `PlainDateTime`, adjusting the minute, hour, and date if the seconds underflow. -/ @[inline] -def subSeconds (dt : LocalDateTime) (seconds : Second.Offset) : LocalDateTime := +def subSeconds (dt : PlainDateTime) (seconds : Second.Offset) : PlainDateTime := addSeconds dt (-seconds) /-- -Adds a `Nanosecond.Offset` to a `LocalDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds overflow. +Adds a `Nanosecond.Offset` to a `PlainDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds overflow. -/ @[inline] -def addNanoseconds (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := - let nano : Nanosecond.Offset := UnitVal.mk dt.time.nano.val +def addNanoseconds (dt : PlainDateTime) (nanos : Nanosecond.Offset) : PlainDateTime := + let nano := Nanosecond.Offset.ofInt dt.time.nano.val let totalNanos := nano + nanos - let extraSeconds : Second.Offset := totalNanos.ediv 1000000000 + let extraSeconds := totalNanos.ediv 1000000000 let nano := Bounded.LE.byEmod totalNanos.val 1000000000 (by decide) let newTime := dt.time.addSeconds extraSeconds { dt with time := { newTime with nano } } /-- -Subtracts a `Nanosecond.Offset` from a `LocalDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds underflow. +Subtracts a `Nanosecond.Offset` from a `PlainDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds underflow. -/ @[inline] -def subNanoseconds (dt : LocalDateTime) (nanos : Nanosecond.Offset) : LocalDateTime := +def subNanoseconds (dt : PlainDateTime) (nanos : Nanosecond.Offset) : PlainDateTime := addNanoseconds dt (-nanos) /-- -Getter for the `Year` inside of a `LocalDateTime`. +Getter for the `Year` inside of a `PlainDateTime`. -/ @[inline] -def year (dt : LocalDateTime) : Year.Offset := +def year (dt : PlainDateTime) : Year.Offset := dt.date.year /-- -Getter for the `Month` inside of a `LocalDateTime`. +Getter for the `Month` inside of a `PlainDateTime`. -/ @[inline] -def month (dt : LocalDateTime) : Month.Ordinal := +def month (dt : PlainDateTime) : Month.Ordinal := dt.date.month /-- -Getter for the `Day` inside of a `LocalDateTime`. +Getter for the `Day` inside of a `PlainDateTime`. -/ @[inline] -def day (dt : LocalDateTime) : Day.Ordinal := +def day (dt : PlainDateTime) : Day.Ordinal := dt.date.day /-- -Getter for the `Hour` inside of a `LocalDateTime`. +Getter for the `Hour` inside of a `PlainDateTime`. -/ @[inline] -def hour (dt : LocalDateTime) : Hour.Ordinal dt.time.hour.fst := +def hour (dt : PlainDateTime) : Hour.Ordinal dt.time.hour.fst := dt.time.hour.snd /-- -Getter for the `Minute` inside of a `LocalDateTime`. +Getter for the `Minute` inside of a `PlainDateTime`. -/ @[inline] -def minute (dt : LocalDateTime) : Minute.Ordinal := +def minute (dt : PlainDateTime) : Minute.Ordinal := dt.time.minute /-- -Getter for the `Second` inside of a `LocalDateTime`. +Getter for the `Second` inside of a `PlainDateTime`. -/ @[inline] -def second (dt : LocalDateTime) : Second.Ordinal dt.time.second.fst := +def second (dt : PlainDateTime) : Second.Ordinal dt.time.second.fst := dt.time.second.snd /-- -Getter for the `Second` inside of a `LocalDateTime`. +Getter for the `Second` inside of a `PlainDateTime`. -/ @[inline] -def nanosecond (dt : LocalDateTime) : Nanosecond.Ordinal := +def nanosecond (dt : PlainDateTime) : Nanosecond.Ordinal := dt.time.nano -instance : HAdd LocalDateTime Day.Offset LocalDateTime where +instance : HAdd PlainDateTime Day.Offset PlainDateTime where hAdd := addDays -instance : HSub LocalDateTime Day.Offset LocalDateTime where +instance : HSub PlainDateTime Day.Offset PlainDateTime where hSub := subDays -instance : HAdd LocalDateTime Hour.Offset LocalDateTime where +instance : HAdd PlainDateTime Hour.Offset PlainDateTime where hAdd := addHours -instance : HSub LocalDateTime Hour.Offset LocalDateTime where +instance : HSub PlainDateTime Hour.Offset PlainDateTime where hSub := subHours -instance : HAdd LocalDateTime Minute.Offset LocalDateTime where +instance : HAdd PlainDateTime Minute.Offset PlainDateTime where hAdd := addMinutes -instance : HSub LocalDateTime Minute.Offset LocalDateTime where +instance : HSub PlainDateTime Minute.Offset PlainDateTime where hSub := subMinutes -instance : HAdd LocalDateTime Second.Offset LocalDateTime where +instance : HAdd PlainDateTime Second.Offset PlainDateTime where hAdd := addSeconds -instance : HSub LocalDateTime Second.Offset LocalDateTime where +instance : HSub PlainDateTime Second.Offset PlainDateTime where hSub := subSeconds -instance : HAdd LocalDateTime Nanosecond.Offset LocalDateTime where +instance : HAdd PlainDateTime Nanosecond.Offset PlainDateTime where hAdd := addNanoseconds -instance : HSub LocalDateTime Nanosecond.Offset LocalDateTime where +instance : HSub PlainDateTime Nanosecond.Offset PlainDateTime where hSub := subNanoseconds -end LocalDateTime +end PlainDateTime end Time end Std diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 7853374ead90..65a0cba1c96c 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -44,7 +44,7 @@ instance : Inhabited Timestamp where instance : OfNat Timestamp n where ofNat := by - refine ⟨UnitVal.mk n, ⟨0, by decide⟩, ?_⟩ + refine ⟨.ofInt n, ⟨0, by decide⟩, ?_⟩ simp <;> exact Int.le_total s.val 0 |>.symm exact Int.le_total 0 n @@ -98,7 +98,7 @@ protected def add (t₁ t₂ : Timestamp) : Timestamp := by let truncated := diffNano.truncateTop (Int.le_sub_one_of_lt h.right) let nano := truncated.addTop 1000000000 (by decide) let proof₁ : 0 ≤ diffSecs - 1 := Int.le_sub_one_of_lt h.left - refine { second := UnitVal.mk (diffSecs.val - 1), nano, proof := ?_ } + refine { second := .ofInt (diffSecs.val - 1), nano, proof := ?_ } simp [nano, Bounded.LE.addTop] refine (Or.inl (And.intro proof₁ ?_)) let h₃ := (Int.add_le_add_iff_left 1000000000).mpr diffNano.property.left @@ -108,7 +108,7 @@ protected def add (t₁ t₂ : Timestamp) : Timestamp := by let second := diffSecs.val + 1 let truncated := diffNano.truncateBottom h₁.right let nano := truncated.subBottom 1000000000 (by decide) - refine { second := UnitVal.mk second, nano, proof := ?_ } + refine { second := .ofInt second, nano, proof := ?_ } simp [nano, truncated, Bounded.LE.subBottom, Bounded.LE.truncateBottom] refine (Or.inr (And.intro ?_ ?_)) · exact h₁.left @@ -159,7 +159,7 @@ Converts a `Timestamp` from a `Nanosecond.Offset` @[inline] def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := let nanos := tm.toSecondsSinceUnixEpoch.mul 1000000000 - let nanos := nanos + (UnitVal.mk tm.nano.val) + let nanos := nanos + (.ofInt tm.nano.val) nanos /-- diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 131e36164078..dde99b46ff72 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -42,6 +42,18 @@ in a 24-hour clock format. -/ def time24Hour : Format .any := date-spec% "hh:mm:ss" +/-- +The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD hh:mm:ss Z` for +representing date, time, and time zone. +-/ +def dateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD hh:mm:ss" + +/-- +The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD hh:mm:ss Z` +for representing date, time, and time zone. +-/ +def dateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD hh:mm:ss ZZZ" + /-- The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used in SQL databases to represent dates. @@ -74,12 +86,12 @@ def rfc850 : Format .any := date-spec% "EEEE, DD-MMM-YY hh:mm:ss ZZZ" end Formats -namespace LocalDate +namespace PlainDate /-- -Formats a `LocalDate` using a specific format. +Formats a `PlainDate` using a specific format. -/ -def format (date : LocalDate) (format : String) : String := +def format (date : PlainDate) (format : String) : String := let format : Except String (Format .any) := Format.spec format match format with | .error err => s!"error: {err}" @@ -95,50 +107,50 @@ def format (date : LocalDate) (format : String) : String := | none => "invalid time" /-- -Parses a date string in the American format (`MM/DD/YYYY`) and returns a `LocalDate`. +Parses a date string in the American format (`MM/DD/YYYY`) and returns a `PlainDate`. -/ -def fromAmericanDateString (input : String) : Except String LocalDate := do - Formats.americanDate.parseBuilder (λm d y => LocalDate.ofYearMonthDay y m d) input +def fromAmericanDateString (input : String) : Except String PlainDate := do + Formats.americanDate.parseBuilder (λm d y => PlainDate.ofYearMonthDay y m d) input /-- Converts a Date in the American format (`MM/DD/YYYY`) into a `String`. -/ -def toAmericanDateString (input : LocalDate) : String := +def toAmericanDateString (input : PlainDate) : String := Formats.americanDate.formatBuilder input.month input.day input.year /-- Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. -/ -def fromSQLDateString (input : String) : Except String LocalDate := do - Formats.sqlDate.parseBuilder (λy m d => LocalDate.ofYearMonthDay y m d) input +def fromSQLDateString (input : String) : Except String PlainDate := do + Formats.sqlDate.parseBuilder (λy m d => PlainDate.ofYearMonthDay y m d) input /-- Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. -/ -def toSQLDateString (input : LocalDate) : String := +def toSQLDateString (input : PlainDate) : String := Formats.sqlDate.formatBuilder input.year input.month input.day /-- -Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `LocalDate`. +Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `PlainDate`. -/ -def parse (input : String) : Except String LocalDate := +def parse (input : String) : Except String PlainDate := fromAmericanDateString input <|> fromSQLDateString input -instance : ToString LocalDate where +instance : ToString PlainDate where toString := toSQLDateString -instance : Repr LocalDate where +instance : Repr PlainDate where reprPrec data := Repr.addAppParen (toString data) -end LocalDate +end PlainDate -namespace LocalTime +namespace PlainTime /-- -Formats a `LocalTime` using a specific format. +Formats a `PlainTime` using a specific format. -/ -def format (time : LocalTime) (format : String) : String := +def format (time : PlainTime) (format : String) : String := let format : Except String (Format .any) := Format.spec format match format with | .error err => s!"error: {err}" @@ -154,47 +166,47 @@ def format (time : LocalTime) (format : String) : String := | none => "invalid time" /-- -Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `LocalTime`. +Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `PlainTime`. -/ -def fromTime24Hour (input : String) : Except String LocalTime := - Formats.time24Hour.parseBuilder (λh m s => LocalTime.ofHourMinuteSeconds? h.snd m s.snd) input +def fromTime24Hour (input : String) : Except String PlainTime := + Formats.time24Hour.parseBuilder (λh m s => PlainTime.ofHourMinuteSeconds? h.snd m s.snd) input /-- -Formats a `LocalTime` value into a 24-hour format string (`hh:mm:ss`). +Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss`). -/ -def toTime24Hour (input : LocalTime) : String := +def toTime24Hour (input : PlainTime) : String := Formats.time24Hour.formatBuilder input.hour input.minute input.second /-- -Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `LocalTime`. +Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `PlainTime`. -/ -def fromTime12Hour (input : String) : Except String LocalTime := do - let builder h m s a : Option LocalTime := do +def fromTime12Hour (input : String) : Except String PlainTime := do + let builder h m s a : Option PlainTime := do let value ← Internal.Bounded.ofInt? h.snd.val - LocalTime.ofHourMinuteSeconds? (leap₂ := false) (HourMarker.toAbsolute a value) m s.snd + PlainTime.ofHourMinuteSeconds? (leap₂ := false) (HourMarker.toAbsolute a value) m s.snd Formats.time12Hour.parseBuilder builder input /-- -Formats a `LocalTime` value into a 12-hour format string (`hh:mm:ss aa`). +Formats a `PlainTime` value into a 12-hour format string (`hh:mm:ss aa`). -/ -def toTime12Hour (input : LocalTime) : String := +def toTime12Hour (input : PlainTime) : String := Formats.time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.snd.val ≥ 12 then HourMarker.pm else HourMarker.am) /-- -Parses a `String` in the `Time12Hour` or `Time24Hour` format and returns a `LocalTime`. +Parses a `String` in the `Time12Hour` or `Time24Hour` format and returns a `PlainTime`. -/ -def parse (input : String) : Except String LocalTime := +def parse (input : String) : Except String PlainTime := fromTime12Hour input <|> fromTime24Hour input -instance : ToString LocalTime where +instance : ToString PlainTime where toString := toTime24Hour -instance : Repr LocalTime where +instance : Repr PlainTime where reprPrec data := Repr.addAppParen (toString data) -end LocalTime +end PlainTime namespace ZonedDateTime @@ -244,6 +256,18 @@ Formats a `ZonedDateTime` value into an RFC850 format string. def toRFC850String (date : ZonedDateTime) : String := Formats.rfc850.format date.snd +/-- +Parses a `String` in the `DateTimeWithZone` format and returns a `DateTime` object in the GMT time zone. +-/ +def fromDateTimeWithZoneString (input : String) : Except String ZonedDateTime := + Formats.dateTimeWithZone.parse input + +/-- +Formats a `DateTime` value into a `DateTimeWithZone` format string. +-/ +def toDateTimeWithZoneString (ldt : ZonedDateTime) : String := + Formats.dateTimeWithZone.format ldt.snd + /-- Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. -/ @@ -253,19 +277,19 @@ def parse (input : String) : Except String ZonedDateTime := <|> fromRFC850String input instance : ToString ZonedDateTime where - toString := toRFC822String + toString := toDateTimeWithZoneString instance : Repr ZonedDateTime where reprPrec data := Repr.addAppParen (toString data) end ZonedDateTime -namespace LocalDateTime +namespace PlainDateTime /-- -Formats a `LocalDateTime` using a specific format. +Formats a `PlainDateTime` using a specific format. -/ -def format (date : LocalDateTime) (format : String) : String := +def format (date : PlainDateTime) (format : String) : String := let format : Except String (Format .any) := Format.spec format match format with | .error err => s!"error: {err}" @@ -285,45 +309,58 @@ def format (date : LocalDateTime) (format : String) : String := | none => "invalid time" /-- -Parses a `String` in the `AscTime` format and returns a `LocalDateTime` object in the GMT time zone. +Parses a `String` in the `AscTime` format and returns a `PlainDateTime` object in the GMT time zone. -/ -def fromAscTimeString (input : String) : Except String LocalDateTime := +def fromAscTimeString (input : String) : Except String PlainDateTime := Formats.ascTime.parse input - |>.map DateTime.toLocalDateTime + |>.map DateTime.toPlainDateTime /-- -Formats a `LocalDateTime` value into an AscTime format string. +Formats a `PlainDateTime` value into an AscTime format string. -/ -def toAscTimeString (ldt : LocalDateTime) : String := - Formats.ascTime.format (DateTime.ofLocalDateTime ldt .UTC) +def toAscTimeString (ldt : PlainDateTime) : String := + Formats.ascTime.format (DateTime.ofPlainDateTime ldt .UTC) /-- -Parses a `String` in the `LongDateFormat` and returns a `LocalDateTime` object in the GMT time zone. +Parses a `String` in the `LongDateFormat` and returns a `PlainDateTime` object in the GMT time zone. -/ -def fromLongDateFormatString (input : String) : Except String LocalDateTime := +def fromLongDateFormatString (input : String) : Except String PlainDateTime := Formats.longDateFormat.parse input - |>.map DateTime.toLocalDateTime + |>.map DateTime.toPlainDateTime + +/-- +Formats a `PlainDateTime` value into a LongDateFormat string. +-/ +def toLongDateFormatString (ldt : PlainDateTime) : String := + Formats.longDateFormat.format (DateTime.ofPlainDateTime ldt .UTC) + +/-- +Parses a `String` in the `DateTime` format and returns a `PlainDateTime`. +-/ +def fromDateTimeString (input : String) : Except String PlainDateTime := + Formats.dateTime24Hour.parse input + |>.map DateTime.toPlainDateTime /-- -Formats a `LocalDateTime` value into a LongDateFormat string. +Formats a `PlainDateTime` value into a `DateTime` format string. -/ -def toLongDateFormatString (ldt : LocalDateTime) : String := - Formats.longDateFormat.format (DateTime.ofLocalDateTime ldt .UTC) +def toDateTimeString (ldt : PlainDateTime) : String := + Formats.dateTime24Hour.format (DateTime.ofPlainDateTime ldt .UTC) /-- -Parses a `String` in the `AscTime` or `LongDate` format and returns a `LocalDateTime`. +Parses a `String` in the `AscTime` or `LongDate` format and returns a `PlainDateTime`. -/ -def parse (date : String) : Except String LocalDateTime := +def parse (date : String) : Except String PlainDateTime := fromAscTimeString date <|> fromLongDateFormatString date -instance : ToString LocalDateTime where - toString := toLongDateFormatString +instance : ToString PlainDateTime where + toString := toDateTimeString -instance : Repr LocalDateTime where +instance : Repr PlainDateTime where reprPrec data := Repr.addAppParen (toString data) -end LocalDateTime +end PlainDateTime namespace DateTime @@ -378,6 +415,12 @@ Formats a `DateTime` value into an RFC850 format string. def toRFC850String (date : DateTime tz) : String := Formats.rfc850.format date +/-- +Formats a `DateTime` value into a `DateTimeWithZone` format string. +-/ +def toDateTimeWithZoneString (ldt : DateTime tz) : String := + Formats.dateTimeWithZone.format ldt + /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `DateTime`. -/ @@ -386,7 +429,7 @@ def parse (date : String) : Except String (DateTime .GMT) := <|> fromLongDateFormatString date instance : ToString (DateTime tz) where - toString := toRFC822String + toString := toDateTimeWithZoneString instance : Repr (DateTime tz) where reprPrec data := Repr.addAppParen (toString data) diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 364546a9b820..49935360ec69 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -14,7 +14,7 @@ namespace Std namespace Time open Internal open Std.Internal.Parsec.String -open Std.Internal.Parsec Lean LocalTime LocalDate TimeZone DateTime +open Std.Internal.Parsec Lean PlainTime PlainDate TimeZone DateTime set_option linter.all true @@ -583,9 +583,9 @@ private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except pure builder.hour if let .isTrue p := inferInstanceAs (Decidable (ValidTime hour.snd builder.minute builder.second.snd)) then - let build := DateTime.ofLocalDateTime { - date := LocalDate.clip builder.year builder.month builder.day - time := LocalTime.mk hour builder.minute builder.second (.ofMillisecond builder.millisecond) p + let build := DateTime.ofPlainDateTime { + date := PlainDate.clip builder.year builder.month builder.day + time := PlainTime.mk hour builder.minute builder.second (.ofMillisecond builder.millisecond) p } match aw with diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 3a1aac6135b6..19367db841b6 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -63,7 +63,7 @@ syntax date_component noWs "-" noWs date_component noWs "-" noWs date_component private def parseDate : TSyntax `date -> MacroM (TSyntax `term) | `(date|$year:date_component-$month:date_component-$day:date_component) => do - `(Std.Time.LocalDate.mk $(← parseComponent none none year) $(← parseComponent (some 1) (some 12) month) $(← parseComponent (some 1) (some 31) day) (by decide)) + `(Std.Time.PlainDate.mk $(← parseComponent none none year) $(← parseComponent (some 1) (some 12) month) $(← parseComponent (some 1) (some 31) day) (by decide)) | syn => Macro.throwErrorAt syn "unsupported type" /-- @@ -83,9 +83,9 @@ syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component private def parseTime : TSyntax `time -> MacroM (TSyntax `term) | `(time| $hour:date_component:$minute:date_component:$second:date_component) => do - `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ 0 (by decide)) + `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ 0 (by decide)) | `(time| $hour:date_component:$minute:date_component:$second:date_component.$nanos:date_component) => do - `(Std.Time.LocalTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ $(← parseComponent (some 0) (some 999) nanos) (by decide)) + `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ $(← parseComponent (some 0) (some 999) nanos) (by decide)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -105,9 +105,9 @@ syntax date_component : datetime private def parseDateTime : TSyntax `datetime -> MacroM (TSyntax `term) | `(datetime| $date:date:$time:time) => do - `(Std.Time.LocalDateTime.mk $(← parseDate date) $(← parseTime time)) + `(Std.Time.PlainDateTime.mk $(← parseDate date) $(← parseTime time)) | `(datetime|$tm:date_component) => do - `(Std.Time.LocalDateTime.ofUTCTimestamp $(← parseComponent none none tm)) + `(Std.Time.PlainDateTime.ofUTCTimestamp $(← parseComponent none none tm)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -159,7 +159,7 @@ private def parseZoned : TSyntax `zoned -> MacroM (TSyntax `term) | `(zoned| $timestamp:num $zone) => do `(Std.Time.DateTime.ofUTCTimestamp $timestamp $(← parseZone zone)) | `(zoned| $datetime:datetime $zone) => do - `(Std.Time.DateTime.ofLocalDateTime $(← parseDateTime datetime) $(← parseZone zone)) + `(Std.Time.DateTime.ofPlainDateTime $(← parseDateTime datetime) $(← parseZone zone)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- diff --git a/src/Std/Time/Time.lean b/src/Std/Time/Time.lean index 467a4fb35bb9..43590cf600a0 100644 --- a/src/Std/Time/Time.lean +++ b/src/Std/Time/Time.lean @@ -6,4 +6,4 @@ Authors: Sofia Rodrigues prelude import Std.Time.Time.Basic import Std.Time.Time.HourMarker -import Std.Time.Time.LocalTime +import Std.Time.Time.PlainTime diff --git a/src/Std/Time/Time/LocalTime.lean b/src/Std/Time/Time/PlainTime.lean similarity index 59% rename from src/Std/Time/Time/LocalTime.lean rename to src/Std/Time/Time/PlainTime.lean index 073aa2ae437a..9945bd28c55c 100644 --- a/src/Std/Time/Time/LocalTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -15,25 +15,25 @@ set_option linter.all true /-- Represents a specific point in time, including hours, minutes, seconds, and nanoseconds. -/ -structure LocalTime where +structure PlainTime where /-- - `Hour` component of the `LocalTime` + `Hour` component of the `PlainTime` -/ hour : Sigma Hour.Ordinal /-- - `Minute` component of the `LocalTime` + `Minute` component of the `PlainTime` -/ minute : Minute.Ordinal /-- - `Second` component of the `LocalTime` + `Second` component of the `PlainTime` -/ second : Sigma Second.Ordinal /-- - `Nanoseconds` component of the `LocalTime` + `Nanoseconds` component of the `PlainTime` -/ nano : Nanosecond.Ordinal @@ -43,15 +43,15 @@ structure LocalTime where proof : (second.snd.val = 60 → hour.snd.val = 23 ∧ minute.val = 59) ∧ (hour.snd.val = 24 → second.snd.val = 0 ∧ minute.val = 0) -instance : Inhabited LocalTime where +instance : Inhabited PlainTime where default := ⟨Sigma.mk false 0, 0, Sigma.mk false 0, 0, by simp; decide⟩ -instance : BEq LocalTime where +instance : BEq PlainTime where beq x y := x.hour.snd.val == y.hour.snd.val && x.minute == y.minute && x.second.snd.val == y.second.snd.val && x.nano.val == y.nano.val -namespace LocalTime +namespace PlainTime /-- Checks if the hour is valid if it has a leap second or leap hour. @@ -62,77 +62,77 @@ abbrev ValidTime (hour : Hour.Ordinal l) (minute : Minute.Ordinal) (second : Sec ∧ (hour.val = 24 → second.val = 0 ∧ minute.val = 0) /-- -Creates a `LocalTime` value from hours, minutes, and seconds. +Creates a `PlainTime` value from hours, minutes, and seconds. -/ -def ofHourMinuteSeconds (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (proof : ValidTime hour minute second) : LocalTime := +def ofHourMinuteSeconds (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (proof : ValidTime hour minute second) : PlainTime := ⟨Sigma.mk leap₂ hour, minute, Sigma.mk leap second, 0, proof⟩ /-- -Creates a `LocalTime` value from hours, minutes, and seconds. Return `none` if its invalid. +Creates a `PlainTime` value from hours, minutes, and seconds. Return `none` if its invalid. -/ -def ofHourMinuteSeconds? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : Option LocalTime := +def ofHourMinuteSeconds? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : Option PlainTime := if h : ValidTime hour minute second then some <| ofHourMinuteSeconds hour minute second h else none /-- -Creates a `LocalTime` value from hours, minutes, and seconds, setting nanoseconds to zero. +Creates a `PlainTime` value from hours, minutes, and seconds, setting nanoseconds to zero. -/ -def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Ordinal) (second : Second.Ordinal false) (nano : Nanosecond.Ordinal) : LocalTime := by +def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Ordinal) (second : Second.Ordinal false) (nano : Nanosecond.Ordinal) : PlainTime := by refine ⟨Sigma.mk false hour, minute, Sigma.mk false second, nano, ?_⟩ constructor exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr second.property.right)) x) exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) /-- -Converts a `LocalTime` value to the total number of milliseconds. +Converts a `PlainTime` value to the total number of milliseconds. -/ -def toMilliseconds (time : LocalTime) : Millisecond.Offset := +def toMilliseconds (time : PlainTime) : Millisecond.Offset := let secs := time.hour.snd.toOffset.toSeconds + time.minute.toOffset.toSeconds + time.second.snd.toOffset let millis := secs.mul 1000 - UnitVal.mk (millis.val + time.nano.val) + Millisecond.Offset.ofInt (millis.val + time.nano.val) /-- -Converts a `LocalTime` value to the total number of nanoseconds. +Converts a `PlainTime` value to the total number of nanoseconds. -/ -def toNanoseconds (time : LocalTime) : Nanosecond.Offset := +def toNanoseconds (time : PlainTime) : Nanosecond.Offset := let secs := time.hour.snd.toOffset.toSeconds + time.minute.toOffset.toSeconds + time.second.snd.toOffset let nanos := secs.mul 1000000000 - UnitVal.mk (nanos.val + time.nano.val) + Nanosecond.Offset.ofInt (nanos.val + time.nano.val) /-- -Converts a `LocalTime` value to the total number of seconds. +Converts a `PlainTime` value to the total number of seconds. -/ -def toSeconds (time : LocalTime) : Second.Offset := +def toSeconds (time : PlainTime) : Second.Offset := time.hour.snd.toOffset.toSeconds + time.minute.toOffset.toSeconds + time.second.snd.toOffset /-- -Converts a `LocalTime` value to the total number of minutes. +Converts a `PlainTime` value to the total number of minutes. -/ -def toMinutes (time : LocalTime) : Minute.Offset := +def toMinutes (time : PlainTime) : Minute.Offset := time.hour.snd.toOffset.toMinutes + time.minute.toOffset + time.second.snd.toOffset.toMinutes /-- -Converts a `LocalTime` value to the total number of hours. +Converts a `PlainTime` value to the total number of hours. -/ -def toHours (time : LocalTime) : Hour.Offset := +def toHours (time : PlainTime) : Hour.Offset := let hour : Hour.Offset := time.minute.toOffset.ediv 60 time.hour.snd.toOffset + hour + time.second.snd.toOffset.toHours /-- -Creates a `LocalTime` value from a total number of nanoseconds. +Creates a `PlainTime` value from a total number of nanoseconds. -/ -def ofNanoseconds (nanos : Nanosecond.Offset) : LocalTime := +def ofNanoseconds (nanos : Nanosecond.Offset) : PlainTime := have totalSeconds := nanos.ediv 1000000000 have remainingNanos := Bounded.LE.byEmod nanos.val 1000000000 (by decide) have hours := Bounded.LE.byEmod (totalSeconds.val / 3600) 24 (by decide) @@ -142,96 +142,96 @@ def ofNanoseconds (nanos : Nanosecond.Offset) : LocalTime := ofValidHourMinuteSecondsNano hours minutes seconds nanos /-- -Creates a `LocalTime` value from a total number of seconds. +Creates a `PlainTime` value from a total number of seconds. -/ @[inline] -def ofSeconds (secs : Second.Offset) : LocalTime := +def ofSeconds (secs : Second.Offset) : PlainTime := have hours := Bounded.LE.byEmod (secs.val / 3600) 24 (by decide) have minutes := (Bounded.LE.byEmod secs.val 3600 (by decide)).ediv 60 (by decide) have seconds := Bounded.LE.byEmod secs.val 60 (by decide) ofValidHourMinuteSecondsNano hours minutes seconds 0 /-- -Adds seconds to a `LocalTime`. +Adds seconds to a `PlainTime`. -/ @[inline] -def addSeconds (time : LocalTime) (secondsToAdd : Second.Offset) : LocalTime := +def addSeconds (time : PlainTime) (secondsToAdd : Second.Offset) : PlainTime := let totalSeconds := time.toSeconds + secondsToAdd ofSeconds totalSeconds /-- -Subtracts seconds from a `LocalTime`. +Subtracts seconds from a `PlainTime`. -/ @[inline] -def subSeconds (time : LocalTime) (secondsToSub : Second.Offset) : LocalTime := +def subSeconds (time : PlainTime) (secondsToSub : Second.Offset) : PlainTime := addSeconds time (-secondsToSub) /-- -Adds minutes to a `LocalTime`. +Adds minutes to a `PlainTime`. -/ @[inline] -def addMinutes (time : LocalTime) (minutesToAdd : Minute.Offset) : LocalTime := +def addMinutes (time : PlainTime) (minutesToAdd : Minute.Offset) : PlainTime := let total := time.toSeconds + minutesToAdd.toSeconds ofSeconds total /-- -Subtracts minutes from a `LocalTime`. +Subtracts minutes from a `PlainTime`. -/ @[inline] -def subMinutes (time : LocalTime) (minutesToSub : Minute.Offset) : LocalTime := +def subMinutes (time : PlainTime) (minutesToSub : Minute.Offset) : PlainTime := addMinutes time (-minutesToSub) /-- -Adds hours to a `LocalTime`. +Adds hours to a `PlainTime`. -/ -def addHours (time : LocalTime) (hoursToAdd : Hour.Offset) : LocalTime := +def addHours (time : PlainTime) (hoursToAdd : Hour.Offset) : PlainTime := let total := time.toSeconds + hoursToAdd.toSeconds ofSeconds total /-- -Subtracts hours from a `LocalTime`. +Subtracts hours from a `PlainTime`. -/ @[inline] -def subHours (time : LocalTime) (hoursToSub : Hour.Offset) : LocalTime := +def subHours (time : PlainTime) (hoursToSub : Hour.Offset) : PlainTime := addHours time (-hoursToSub) /-- -Adds nanoseconds to a `LocalTime`. +Adds nanoseconds to a `PlainTime`. -/ -def addNanoseconds (time : LocalTime) (nanosToAdd : Nanosecond.Offset) : LocalTime := +def addNanoseconds (time : PlainTime) (nanosToAdd : Nanosecond.Offset) : PlainTime := let total := time.toNanoseconds + nanosToAdd ofNanoseconds total /-- -Subtracts nanoseconds from a `LocalTime`. +Subtracts nanoseconds from a `PlainTime`. -/ -def subNanoseconds (time : LocalTime) (nanosToSub : Nanosecond.Offset) : LocalTime := +def subNanoseconds (time : PlainTime) (nanosToSub : Nanosecond.Offset) : PlainTime := addNanoseconds time (-nanosToSub) -instance : HAdd LocalTime Nanosecond.Offset LocalTime where +instance : HAdd PlainTime Nanosecond.Offset PlainTime where hAdd := addNanoseconds -instance : HSub LocalTime Nanosecond.Offset LocalTime where +instance : HSub PlainTime Nanosecond.Offset PlainTime where hSub := subNanoseconds -instance : HAdd LocalTime Second.Offset LocalTime where +instance : HAdd PlainTime Second.Offset PlainTime where hAdd := addSeconds -instance : HSub LocalTime Second.Offset LocalTime where +instance : HSub PlainTime Second.Offset PlainTime where hSub := subSeconds -instance : HAdd LocalTime Minute.Offset LocalTime where +instance : HAdd PlainTime Minute.Offset PlainTime where hAdd := addMinutes -instance : HSub LocalTime Minute.Offset LocalTime where +instance : HSub PlainTime Minute.Offset PlainTime where hSub := subMinutes -instance : HAdd LocalTime Hour.Offset LocalTime where +instance : HAdd PlainTime Hour.Offset PlainTime where hAdd := addHours -instance : HSub LocalTime Hour.Offset LocalTime where +instance : HSub PlainTime Hour.Offset PlainTime where hSub := subHours -end LocalTime +end PlainTime end Time end Std diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 738fe81ccf2c..3528885dadac 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -37,6 +37,23 @@ def Offset : Type := UnitVal (1 / 1000) instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +namespace Offset + +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + UnitVal.mk data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + UnitVal.mk data + +end Offset end Millisecond end Time end Std diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index 9b0d1799b416..33be8713187f 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -58,6 +58,24 @@ def Offset : Type := UnitVal (1 / 1000000000) instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +namespace Offset + +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + UnitVal.mk data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + UnitVal.mk data + +end Offset + /-- `Span` represents a bounded value for nanoseconds, ranging between -999999999 and 999999999. This can be used for operations that involve differences or adjustments within this range. diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 232f824b637b..036ce5f5b3a0 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -41,6 +41,24 @@ def Offset : Type := UnitVal 1 instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +namespace Offset + +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + UnitVal.mk data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + UnitVal.mk data + +end Offset + namespace Ordinal /-- diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 04742721bd08..810faf8b6ffe 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -14,54 +14,54 @@ import Std.Time.Zoned.Database.Basic namespace Std namespace Time -namespace LocalDateTime +namespace PlainDateTime /-- -Creaates a new `LocalDateTime` out of a `Timestamp` and a `TimeZone`. +Creaates a new `PlainDateTime` out of a `Timestamp` and a `TimeZone`. -/ -def ofTimestamp (stamp : Timestamp) (tz : TimeZone) : LocalDateTime := +def ofTimestamp (stamp : Timestamp) (tz : TimeZone) : PlainDateTime := let stamp := stamp.addSeconds tz.toSeconds - LocalDateTime.ofUTCTimestamp stamp + PlainDateTime.ofUTCTimestamp stamp /-- Get the current monotonic time. -/ @[inline] -def now : IO LocalDateTime := do +def now : IO PlainDateTime := do let tm ← Timestamp.now let tz ← TimeZone.getCurrentTimezone - return LocalDateTime.ofTimestamp tm tz + return PlainDateTime.ofTimestamp tm tz -end LocalDateTime +end PlainDateTime namespace DateTime /-- -Converts a `LocalDate` to a `DateTime` +Converts a `PlainDate` to a `DateTime` -/ @[inline] -def ofLocalDate (ld : LocalDate) (tz : TimeZone) : DateTime tz := - DateTime.ofTimestamp (Timestamp.ofLocalDate ld) tz +def ofPlainDate (ld : PlainDate) (tz : TimeZone) : DateTime tz := + DateTime.ofTimestamp (Timestamp.ofPlainDate ld) tz /-- -Converts a `DateTime` to a `LocalDate` +Converts a `DateTime` to a `PlainDate` -/ @[inline] -def toLocalDate (dt : DateTime tz) : LocalDate := - Timestamp.toLocalDate dt.toTimestamp +def toPlainDate (dt : DateTime tz) : PlainDate := + Timestamp.toPlainDate dt.toTimestamp /-- -Converts a `LocalTime` to a `DateTime` +Converts a `PlainTime` to a `DateTime` -/ @[inline] -def ofLocalTime (lt : LocalTime) (tz : TimeZone) : DateTime tz := - DateTime.ofTimestamp (Timestamp.ofLocalTime lt) tz +def ofPlainTime (lt : PlainTime) (tz : TimeZone) : DateTime tz := + DateTime.ofTimestamp (Timestamp.ofPlainTime lt) tz /-- -Converts a `DateTime` to a `LocalTime` +Converts a `DateTime` to a `PlainTime` -/ @[inline] -def toLocalTime (dt : DateTime tz) : LocalTime := +def toPlainTime (dt : DateTime tz) : PlainTime := dt.date.get.time end DateTime @@ -78,31 +78,31 @@ def now : IO ZonedDateTime := do return ofTimestamp date tz /-- -Converts a `LocalDate` to a `ZonedDateTime` +Converts a `PlainDate` to a `ZonedDateTime` -/ @[inline] -def ofLocalDate (ld : LocalDate) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp (Timestamp.ofLocalDate ld) tz⟩ +def ofPlainDate (ld : PlainDate) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainDate ld) tz⟩ /-- -Converts a `ZonedDateTime` to a `LocalDate` +Converts a `ZonedDateTime` to a `PlainDate` -/ @[inline] -def toLocalDate (dt : ZonedDateTime) : LocalDate := - DateTime.toLocalDate dt.snd +def toPlainDate (dt : ZonedDateTime) : PlainDate := + DateTime.toPlainDate dt.snd /-- -Converts a `LocalTime` to a `ZonedDateTime` +Converts a `PlainTime` to a `ZonedDateTime` -/ @[inline] -def ofLocalTime (lt : LocalTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp (Timestamp.ofLocalTime lt) tz⟩ +def ofPlainTime (lt : PlainTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainTime lt) tz⟩ /-- -Converts a `ZonedDateTime` to a `LocalTime` +Converts a `ZonedDateTime` to a `PlainTime` -/ @[inline] -def toLocalTime (dt : ZonedDateTime) : LocalTime := - DateTime.toLocalTime dt.snd +def toPlainTime (dt : ZonedDateTime) : PlainTime := + DateTime.toPlainTime dt.snd end ZonedDateTime diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index 289247565bad..cb27b13377e7 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -24,7 +24,7 @@ def applyLeapSecondsOnUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO /-- Gets the TimeZone at the local timezone. -/ -def getLocalTimeZoneAt [Database α] (db : α) (tm : Timestamp) : IO TimeZone := do +def getPlainTimeZoneAt [Database α] (db : α) (tm : Timestamp) : IO TimeZone := do (IO.ofExcept <| timezoneAt · tm) =<< Database.localRules db /-- diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index a4b7911d30e6..67a710eca737 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -47,19 +47,19 @@ def convertUt : Bool → UTLocal Converts a `TZif.LeapSecond` structure to a `LeapSecond` structure. -/ def convertLeapSecond (tz : TZif.LeapSecond) : LeapSecond := - { transitionTime := Internal.UnitVal.mk tz.transitionTime, correction := Internal.UnitVal.mk tz.correction } + { transitionTime := .ofInt tz.transitionTime, correction := Second.Offset.ofInt tz.correction } /-- -Converts a LocalTime. +Converts a PlainTime. -/ -def convertLocalTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do - let localType ← tz.localTimeTypes.get? index +def convertPlainTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option PlainTimeType := do + let localType ← tz.PlainTimeTypes.get? index let abbreviation ← tz.abbreviations.getD index "Unknown" let wallflag := convertWall (tz.stdWallIndicators.getD index true) let utLocal := convertUt (tz.utLocalIndicators.getD index true) return { - gmtOffset := Offset.ofSeconds <| Internal.UnitVal.mk localType.gmtOffset + gmtOffset := Offset.ofSeconds <| .ofInt localType.gmtOffset isDst := localType.isDst abbreviation wall := wallflag @@ -70,20 +70,20 @@ def convertLocalTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Op /-- Converts a transition. -/ -def convertTransition (times: Array LocalTimeType) (index : Nat) (tz : TZif.TZifV1) : Option Transition := do +def convertTransition (times: Array PlainTimeType) (index : Nat) (tz : TZif.TZifV1) : Option Transition := do let time := tz.transitionTimes.get! index - let time := Internal.UnitVal.mk time + let time := Second.Offset.ofInt time let indice := tz.transitionIndices.get! index - return { time, localTimeType := times.get! indice.toNat } + return { time, PlainTimeType := times.get! indice.toNat } /-- Converts a `TZif.TZifV1` structure to a `ZoneRules` structure. -/ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := do - let mut times : Array LocalTimeType := #[] + let mut times : Array PlainTimeType := #[] for i in [0:tz.header.typecnt.toNat] do - if let some result := convertLocalTime i tz id + if let some result := convertPlainTime i tz id then times := times.push result else .error s!"cannot convert local time {i} of the file" @@ -95,7 +95,7 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := then transitions := transitions.push result else .error s!"cannot convert transtiion {i} of the file" - .ok { leapSeconds, transitions, localTimes := times } + .ok { leapSeconds, transitions, PlainTimes := times } /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 70ca1f2d8b47..6e78bb3e6b23 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -24,7 +24,7 @@ structure TZdb where The path to the local timezone file. This is typically a symlink to a file within the timezone database that corresponds to the current local time zone. -/ - localPath : System.FilePath := "/etc/localtime" + localPath : System.FilePath := "/etc/PlainTime" /-- The path to the directory containing all available time zone files. These files define various @@ -71,9 +71,9 @@ def idFromPath (path : System.FilePath) : Option String := do Retrieves the timezone rules from the local timezone data file. -/ def localRules (path : System.FilePath) : IO ZoneRules := do - let localtimePath ← IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } + let PlainTimePath ← IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } - if let some id := idFromPath localtimePath + if let some id := idFromPath PlainTimePath then parseTZIfFromDisk path id else throw (IO.userError "cannot read the id of the path.") diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index 1f46a12c1c13..d07471955ad3 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -66,7 +66,7 @@ structure Header where /-- Represents the local time type information, including offset and daylight saving details. -/ -structure LocalTimeType where +structure PlainTimeType where /-- The GMT offset in seconds for this local time type. @@ -123,7 +123,7 @@ structure TZifV1 where /-- The array of local time type structures. -/ - localTimeTypes : Array LocalTimeType + PlainTimeTypes : Array PlainTimeType /-- The array of abbreviation strings used by local time types. @@ -216,8 +216,8 @@ private def parseHeader : Parser Header := <*> pu32 <*> pu32 -private def parseLocalTimeType : Parser LocalTimeType := - LocalTimeType.mk +private def parsePlainTimeType : Parser PlainTimeType := + PlainTimeType.mk <$> pi32 <*> pbool <*> pu8 @@ -233,10 +233,10 @@ private def parseTransitionTimes (size : Parser Int32) (n : UInt32) : Parser (Ar private def parseTransitionIndices (n : UInt32) : Parser (Array UInt8) := manyN (n.toNat) pu8 -private def parseLocalTimeTypes (n : UInt32) : Parser (Array LocalTimeType) := - manyN (n.toNat) parseLocalTimeType +private def parsePlainTimeTypes (n : UInt32) : Parser (Array PlainTimeType) := + manyN (n.toNat) parsePlainTimeType -private def parseAbbreviations (times : Array LocalTimeType) (n : UInt32) : Parser (Array String) := do +private def parseAbbreviations (times : Array PlainTimeType) (n : UInt32) : Parser (Array String) := do let mut strings := #[] let mut current := "" let mut chars ← manyN n.toNat pu8 @@ -264,8 +264,8 @@ private def parseTZifV1 : Parser TZifV1 := do let transitionTimes ← parseTransitionTimes pi32 header.timecnt let transitionIndices ← parseTransitionIndices header.timecnt - let localTimeTypes ← parseLocalTimeTypes header.typecnt - let abbreviations ← parseAbbreviations localTimeTypes header.charcnt + let PlainTimeTypes ← parsePlainTimeTypes header.typecnt + let abbreviations ← parseAbbreviations PlainTimeTypes header.charcnt let leapSeconds ← parseLeapSeconds pi32 header.leapcnt let stdWallIndicators ← parseIndicators header.isstdcnt let utLocalIndicators ← parseIndicators header.isutcnt @@ -274,7 +274,7 @@ private def parseTZifV1 : Parser TZifV1 := do header transitionTimes transitionIndices - localTimeTypes + PlainTimeTypes abbreviations leapSeconds stdWallIndicators @@ -299,8 +299,8 @@ private def parseTZifV2 : Parser (Option TZifV2) := optional do let transitionTimes ← parseTransitionTimes pi64 header.timecnt let transitionIndices ← parseTransitionIndices header.timecnt - let localTimeTypes ← parseLocalTimeTypes header.typecnt - let abbreviations ← parseAbbreviations localTimeTypes header.charcnt + let PlainTimeTypes ← parsePlainTimeTypes header.typecnt + let abbreviations ← parseAbbreviations PlainTimeTypes header.charcnt let leapSeconds ← parseLeapSeconds pi64 header.leapcnt let stdWallIndicators ← parseIndicators header.isstdcnt let utLocalIndicators ← parseIndicators header.isutcnt @@ -309,7 +309,7 @@ private def parseTZifV2 : Parser (Option TZifV2) := optional do header transitionTimes transitionIndices - localTimeTypes + PlainTimeTypes abbreviations leapSeconds stdWallIndicators diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 37a8ebd218b4..eb36b976ff7e 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -13,7 +13,7 @@ namespace Time set_option linter.all true /-- -It stores a `Timestamp`, a `LocalDateTime` and a `TimeZone` +It stores a `Timestamp`, a `PlainDateTime` and a `TimeZone` -/ structure DateTime (tz : TimeZone) where private mk :: @@ -24,10 +24,10 @@ structure DateTime (tz : TimeZone) where timestamp : Timestamp /-- - `Date` is a `Thunk` containing the `LocalDateTime` that represents the local date and time, it's + `Date` is a `Thunk` containing the `PlainDateTime` that represents the local date and time, it's used for accessing data like `day` and `month` without having to recompute the data everytime. -/ - date : Thunk LocalDateTime + date : Thunk PlainDateTime instance : BEq (DateTime tz) where beq x y := x.timestamp == y.timestamp @@ -42,7 +42,7 @@ Creates a new `DateTime` out of a `Timestamp` that is in a `TimeZone`. -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := - DateTime.mk tm (Thunk.mk <| λ_ => (tm.addSeconds tz.toSeconds).toLocalDateTime) + DateTime.mk tm (Thunk.mk <| λ_ => (tm.addSeconds tz.toSeconds).toPlainDateTime) /-- Creates a new zone aware `Timestamp` out of a `DateTime`. @@ -59,11 +59,11 @@ def convertTimeZone (date : DateTime tz) (tz₁ : TimeZone) : DateTime tz₁ := ofTimestamp date.timestamp tz₁ /-- -Creates a new `DateTime` out of a `LocalDateTime` +Creates a new `DateTime` out of a `PlainDateTime` -/ @[inline] -def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : DateTime tz := - let tm := Timestamp.ofLocalDateTime date +def ofPlainDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := + let tm := Timestamp.ofPlainDateTime date DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk <| λ_ => date) /-- @@ -71,90 +71,90 @@ Add `Hour.Offset` to a `DateTime`. -/ @[inline] def addHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addHours hours) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addHours hours) tz /-- Subtract `Hour.Offset` from a `DateTime`. -/ @[inline] def subHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subHours hours) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subHours hours) tz /-- Add `Minute.Offset` to a `DateTime`. -/ @[inline] def addMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addMinutes minutes) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addMinutes minutes) tz /-- Subtract `Minute.Offset` from a `DateTime`. -/ @[inline] def subMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subMinutes minutes) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subMinutes minutes) tz /-- Add `Second.Offset` to a `DateTime`. -/ @[inline] def addSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addSeconds seconds) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addSeconds seconds) tz /-- Subtract `Second.Offset` from a `DateTime`. -/ @[inline] def subSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subSeconds seconds) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subSeconds seconds) tz /-- Add `Nanosecond.Offset` to a `DateTime`. -/ @[inline] def addNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addNanoseconds nanoseconds) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addNanoseconds nanoseconds) tz /-- Subtract `Nanosecond.Offset` from a `DateTime`. -/ @[inline] def subNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subNanoseconds nanoseconds) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subNanoseconds nanoseconds) tz /-- Add `Day.Offset` to a `DateTime`. -/ @[inline] def addDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addDays days) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addDays days) tz /-- Subtracts `Day.Offset` to a `DateTime`. -/ @[inline] def subDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subDays days) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subDays days) tz /-- Add `Month.Offset` to a `DateTime`, it clips the day to the last valid day of that month. -/ def addMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addMonthsClip months) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addMonthsClip months) tz /-- Subtracts `Month.Offset` from a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def subMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subMonthsClip months) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subMonthsClip months) tz /-- Add `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following month. -/ def addMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addMonthsRollOver months) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addMonthsRollOver months) tz /-- Subtract `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following @@ -162,7 +162,7 @@ month. -/ @[inline] def subMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subMonthsRollOver months) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subMonthsRollOver months) tz /-- Add `Year.Offset` to a `DateTime`, this function rolls over any excess days into the following @@ -170,14 +170,14 @@ month. -/ @[inline] def addYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addYearsRollOver years) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addYearsRollOver years) tz /-- Add `Year.Offset` to a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def addYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.addYearsClip years) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.addYearsClip years) tz /-- Subtract `Year.Offset` from a `DateTime`, this function rolls over any excess days into the following @@ -185,20 +185,20 @@ month. -/ @[inline] def subYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subYearsRollOver years) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subYearsRollOver years) tz /-- Subtract `Year.Offset` from to a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def subYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.timestamp.toLocalDateTime.subYearsClip years) tz + ofPlainDateTime (dt.timestamp.toPlainDateTime.subYearsClip years) tz /-- -Converts a `Timestamp` to a `LocalDateTime` +Converts a `Timestamp` to a `PlainDateTime` -/ @[inline] -def toLocalDateTime (dt : DateTime tz) : LocalDateTime := +def toPlainDateTime (dt : DateTime tz) : PlainDateTime := dt.date.get /-- diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index ccfdfb714cc0..d0e578e99852 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -47,7 +47,7 @@ inductive StdWall /-- Represents a type of local time, including offset and daylight saving information. -/ -structure LocalTimeType where +structure PlainTimeType where /-- The offset from GMT for this local time. @@ -80,15 +80,15 @@ structure LocalTimeType where identifier : String deriving Repr, Inhabited -namespace LocalTimeType +namespace PlainTimeType /-- -Gets the `TimeZone` offset from a `LocalTimeType`. +Gets the `TimeZone` offset from a `PlainTimeType`. -/ -def getTimeZone (time : LocalTimeType) : TimeZone := +def getTimeZone (time : PlainTimeType) : TimeZone := ⟨time.gmtOffset, time.abbreviation, time.isDst⟩ -end LocalTimeType +end PlainTimeType /-- Represents a leap second event, including the time of the transition and the correction applied. @@ -119,7 +119,7 @@ structure Transition where /-- The local time type associated with this transition. -/ - localTimeType : LocalTimeType + PlainTimeType : PlainTimeType deriving Repr, Inhabited /-- @@ -130,7 +130,7 @@ structure ZoneRules where /-- The array of local time types for the time zone. -/ - localTimes : Array LocalTimeType + PlainTimes : Array PlainTimeType /-- The array of transitions for the time zone. @@ -149,17 +149,17 @@ namespace Transition Create a TimeZone from a Transition. -/ def createTimeZoneFromTransition (transition : Transition) : TimeZone := - let offset := transition.localTimeType.gmtOffset - let name := transition.localTimeType.abbreviation - TimeZone.mk offset name transition.localTimeType.isDst + let offset := transition.PlainTimeType.gmtOffset + let name := transition.PlainTimeType.abbreviation + TimeZone.mk offset name transition.PlainTimeType.isDst /-- Applies the transition to a Timestamp. -/ def apply (timestamp : Timestamp) (transition : Transition) : Timestamp := - let offsetInSeconds := transition.localTimeType.gmtOffset.hour.mul 3600 |>.add transition.localTimeType.gmtOffset.second - let localTimestamp := timestamp.addSeconds offsetInSeconds - localTimestamp + let offsetInSeconds := transition.PlainTimeType.gmtOffset.hour.mul 3600 |>.add transition.PlainTimeType.gmtOffset.second + let PlainTimestamp := timestamp.addSeconds offsetInSeconds + PlainTimestamp end Transition @@ -192,7 +192,7 @@ def applyLeapSeconds (tm : Timestamp) (leapSeconds : ZoneRules) : Timestamp := I for i in [:leapSeconds.size] do let leapSec := leapSeconds.get! i if currentTime.second.val >= leapSec.transitionTime.val then - currentTime := tm.addSeconds (UnitVal.mk leapSec.correction.val) + currentTime := tm.addSeconds (.ofInt leapSec.correction.val) return currentTime /-- diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 181f83b39c7b..eb03075c4363 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -51,7 +51,7 @@ Creates a new `DateTime` out of a `Timestamp` def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do let transition ← rules.findTransitionForTimestamp tm let tm := rules.applyLeapSeconds tm - return ofTimestamp tm transition.localTimeType.getTimeZone + return ofTimestamp tm transition.PlainTimeType.getTimeZone /-- Changes the `TimeZone` to a new one. @@ -61,18 +61,18 @@ def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := ofTimestamp (date.toTimestamp) tz₁ /-- -Creates a new `ZonedDateTime` out of a `LocalDateTime` +Creates a new `ZonedDateTime` out of a `PlainDateTime` -/ @[inline] -def ofLocalDateTime (date : LocalDateTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofLocalDateTime date tz⟩ +def ofPlainDateTime (date : PlainDateTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofPlainDateTime date tz⟩ /-- -Converts a `ZonedDateTime` to a `LocalDateTime` +Converts a `ZonedDateTime` to a `PlainDateTime` -/ @[inline] -def toLocalDateTime (dt : ZonedDateTime) : LocalDateTime := - DateTime.toLocalDateTime dt.snd +def toPlainDateTime (dt : ZonedDateTime) : PlainDateTime := + DateTime.toPlainDateTime dt.snd /-- Getter for the `Year` inside of a `ZonedDateTime` diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index f40ee9e7366f..b3db9e75e8c2 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -128,19 +128,13 @@ Your time zone: 15 Aguust 2024 11:03:47 GMT-03:00 def localTm : Second.Offset := 1723730627 /-- -This localDate is relative to the local time. +This PlainDate is relative to the local time. -/ -def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) +def PlainDate : PlainDateTime := Timestamp.toPlainDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) -/-- -info: "08/15/2024 14:03:47" --/ -#guard_msgs in -#eval ShortDateTime.formatBuilder localDate.month localDate.day localDate.year localDate.time.hour localDate.minute localDate.time.second - -def dateBR₁ := DateTime.ofLocalDateTime localDate brTZ -def dateJP₁ := DateTime.ofLocalDateTime localDate jpTZ -def dateUTC₁ := DateTime.ofLocalDateTime localDate .UTC +def dateBR₁ := DateTime.ofPlainDateTime PlainDate brTZ +def dateJP₁ := DateTime.ofPlainDateTime PlainDate jpTZ +def dateUTC₁ := DateTime.ofPlainDateTime PlainDate .UTC /-- info: "Thursday, August 15, 2024 14:03:47 -0300" @@ -206,37 +200,37 @@ info: "06/16/2014" info: "0053-06-19" -/ #guard_msgs in -#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-700000⟩) .UTC) +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-700000⟩) .UTC) /-- info: "-0002-09-16" -/ #guard_msgs in -#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-720000⟩) .UTC) +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-720000⟩) .UTC) /-- info: "-0084-07-28" -/ #guard_msgs in -#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-750000⟩) .UTC) +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-750000⟩) .UTC) /-- info: "-0221-09-04" -/ #guard_msgs in -#eval Formats.sqlDate.format (DateTime.ofLocalDate (LocalDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) /-- info: -0221-09-04 -/ #guard_msgs in -#eval (LocalDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) +#eval (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) /-- info: "-0221-09-04" -/ #guard_msgs in -#eval toString (LocalDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) +#eval toString (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) /-- info: 2002-07-14 @@ -251,13 +245,13 @@ info: 14:13:12 #eval time% 14:13:12 /-- -info: Sunday, July 14, 2002 14:13:12 +info: 2002-07-14 14:13:12 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12 /-- -info: Sun, 14 Jul 2002 14:13:12 +0900 +info: 2002-07-14 14:13:12 +0900 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12+09:00 diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 63b475f477d0..7e0d287ac330 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -4,8 +4,8 @@ open Std.Time def ShortDateTime : Format .any := date-spec% "DD/MM/YYYY hh:mm:ss" def ShortDate : Format .any := date-spec% "DD/MM/YYYY" -def format (localDate : LocalDateTime) : String := ShortDateTime.formatBuilder localDate.day localDate.month localDate.year localDate.time.hour localDate.minute localDate.time.second -def format₂ (localDate : LocalDate) : String := ShortDate.formatBuilder localDate.day localDate.month localDate.year +def format (PlainDate : PlainDateTime) : String := ShortDateTime.formatBuilder PlainDate.day PlainDate.month PlainDate.year PlainDate.time.hour PlainDate.minute PlainDate.time.second +def format₂ (PlainDate : PlainDate) : String := ShortDate.formatBuilder PlainDate.day PlainDate.month PlainDate.year def date₁ := date% 1993-11-19:09:08:07 def date₂ := date% 1993-05-09:12:59:59 @@ -31,43 +31,43 @@ info: "09/05/1993 12:59:59" info: 753700087 -/ #guard_msgs in -#eval date₁.toLocalTimestamp.toSeconds +#eval date₁.toLocalTimestamp.toSecondsSinceUnixEpoch /-- info: 736952399 -/ #guard_msgs in -#eval date₂.toLocalTimestamp.toSeconds +#eval date₂.toLocalTimestamp.toSecondsSinceUnixEpoch /-- info: "09/05/1993 12:59:59" -/ #guard_msgs in -#eval LocalDateTime.ofUTCTimestamp 736952399 |> format +#eval PlainDateTime.ofUTCTimestamp 736952399 |> format /-- info: 736952399 -/ #guard_msgs in -#eval LocalDateTime.toLocalTimestamp date₂ |>.toSeconds +#eval PlainDateTime.toPlainTimestamp date₂ |>.toSecondsSinceUnixEpoch /-- info: "16/08/2024" -/ #guard_msgs in -#eval LocalDate.ofDaysSinceUNIXEpoch 19951 |> format₂ +#eval PlainDate.ofDaysSinceUNIXEpoch 19951 |> format₂ /-- info: 19951 -/ #guard_msgs in -#eval LocalDate.toDaysSinceUNIXEpoch date₃ +#eval PlainDate.toDaysSinceUNIXEpoch date₃ /-- info: Std.Time.Weekday.friday -/ #guard_msgs in -#eval LocalDate.weekday date₃ +#eval PlainDate.weekday date₃ /-- info: #[] @@ -78,7 +78,7 @@ info: #[] for i in [0:10000] do let i := Int.ofNat i - 999975 - let date := LocalDate.ofDaysSinceUNIXEpoch (Day.Offset.ofInt i) + let date := PlainDate.ofDaysSinceUNIXEpoch (Day.Offset.ofInt i) let num := date.toDaysSinceUNIXEpoch if i ≠ num.val then res := res.push i diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index e584f126d3d8..26d316042fd5 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -138,19 +138,13 @@ Your time zone: 15 Aguust 2024 11:03:47 GMT-03:00 def localTm : Second.Offset := 1723730627 /-- -This localDate is relative to the local time. +This PlainDate is relative to the local time. -/ -def localDate : LocalDateTime := Timestamp.toLocalDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) +def PlainDate : PlainDateTime := Timestamp.toPlainDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) -/-- -info: "08/15/2024 14:03:47" --/ -#guard_msgs in -#eval ShortDateTime.formatBuilder localDate.month localDate.day localDate.year localDate.time.hour localDate.minute localDate.time.second - -def dateBR₁ := DateTime.ofLocalDateTime localDate brTZ -def dateJP₁ := DateTime.ofLocalDateTime localDate jpTZ -def dateUTC₁ := DateTime.ofLocalDateTime localDate .UTC +def dateBR₁ := DateTime.ofPlainDateTime PlainDate brTZ +def dateJP₁ := DateTime.ofPlainDateTime PlainDate jpTZ +def dateUTC₁ := DateTime.ofPlainDateTime PlainDate .UTC /-- info: "2024-08-15T14:03:47-0300" diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index 48816f44fc35..d1306a9c8636 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -44,7 +44,7 @@ info: 91 info: 3 -/ #guard_msgs in -#eval code.v1.localTimeTypes.size +#eval code.v1.PlainTimeTypes.size /-- info: 0 From c30e4a159632d8d51d4d7b1a27386aa3e8208cf5 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 01:30:06 -0300 Subject: [PATCH 040/118] feat: add conversion functions --- src/Std/Time/Date/PlainDate.lean | 15 ++++++ src/Std/Time/Date/Unit/Year.lean | 14 +++++ src/Std/Time/DateTime.lean | 56 +++++++++++++++----- src/Std/Time/DateTime/PlainDateTime.lean | 15 ++++++ src/Std/Time/DateTime/Timestamp.lean | 13 +++++ src/Std/Time/Duration.lean | 49 ++++++++++++++++++ src/Std/Time/Format.lean | 20 ++++---- src/Std/Time/Zoned.lean | 16 +++--- src/Std/Time/Zoned/DateTime.lean | 18 +++++++ src/Std/Time/Zoned/ZonedDateTime.lean | 19 +++++++ tests/lean/run/timeClassOperations.lean | 32 ++++++------ tests/lean/run/timeLocalDateTime.lean | 2 +- tests/lean/run/timeOperations.lean | 65 ++++++++++++------------ 13 files changed, 252 insertions(+), 82 deletions(-) diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index baf8b9e43c8a..6eb289ee1321 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -134,6 +134,21 @@ def weekday (date : PlainDate) : Weekday := .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ +/-- +Determines the era of the given `PlainDate` based on its year. +-/ +def era (date : PlainDate) : Year.Era := + if date.year.toInt ≥ 0 then + .ce + else + .bce + +/-- +Checks if the `PlainDate` is in a leap year. +-/ +def inLeapYear (date : PlainDate) : Bool := + date.year.isLeap + /-- Converts a `PlainDate` to the number of days since the UNIX epoch. -/ diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index f615d4621735..6136b34a5bcd 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -17,6 +17,20 @@ open Internal set_option linter.all true +/-- +Defines the different eras. +-/ +inductive Era + /-- + The era before the Common Era (BCE), always represents a date before year 0. + -/ + | bce + + /-- + The Common Era (CE), represents dates from year 0 onwards. + -/ + | ce + /-- `Offset` represents a year offset, defined as an `Int`. -/ diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index b5159d0d4276..b6fcfc03db3f 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -17,8 +17,8 @@ set_option linter.all true Converts a `PlainDateTime` to a `Timestamp` -/ @[inline] -def ofPlainDateTime (ldt : PlainDateTime) : Timestamp := - ldt.toPlainTimestamp +def ofPlainDateTime (pdt : PlainDateTime) : Timestamp := + pdt.toLocalTimestamp /-- Converts a `Timestamp` to a `PlainDateTime` @@ -31,8 +31,8 @@ def toPlainDateTime (timestamp : Timestamp) : PlainDateTime := Converts a `PlainDate` to a `Timestamp` -/ @[inline] -def ofPlainDate (ld : PlainDate) : Timestamp := - let days := ld.toDaysSinceUNIXEpoch +def ofPlainDate (pd : PlainDate) : Timestamp := + let days := pd.toDaysSinceUNIXEpoch let secs := days.toSeconds Timestamp.ofSecondsSinceUnixEpoch secs @@ -49,8 +49,8 @@ def toPlainDate (timestamp : Timestamp) : PlainDate := Converts a `PlainTime` to a `Timestamp` -/ @[inline] -def ofPlainTime (lt : PlainTime) : Timestamp := - let nanos := lt.toNanoseconds +def ofPlainTime (pt : PlainTime) : Timestamp := + let nanos := pt.toNanoseconds Timestamp.ofNanosecondsSinceUnixEpoch nanos /-- @@ -69,28 +69,56 @@ namespace PlainDateTime Converts a `PlainDate` to a `Timestamp` -/ @[inline] -def ofPlainDate (ld : PlainDate) : PlainDateTime := - PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainDate ld) +def ofPlainDate (pd : PlainDate) : PlainDateTime := + PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainDate pd) /-- Converts a `PlainDateTime` to a `PlainDate` -/ @[inline] -def toPlainDate (ldt : PlainDateTime) : PlainDate := - Timestamp.toPlainDate ldt.toPlainTimestamp +def toPlainDate (pdt : PlainDateTime) : PlainDate := + Timestamp.toPlainDate pdt.toLocalTimestamp /-- Converts a `PlainTime` to a `PlainDateTime` -/ @[inline] -def ofPlainTime (lt : PlainTime) : PlainDateTime := - PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainTime lt) +def ofPlainTime (pt : PlainTime) : PlainDateTime := + PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainTime pt) /-- Converts a `PlainDateTime` to a `PlainTime` -/ @[inline] -def toPlainTime (ldt : PlainDateTime) : PlainTime := - Timestamp.toPlainTime ldt.toPlainTimestamp +def toPlainTime (pdt : PlainDateTime) : PlainTime := + Timestamp.toPlainTime pdt.toLocalTimestamp + +instance : ToTimestamp PlainDateTime where + toTimestamp := Timestamp.ofPlainDateTime + +instance : ToTimestamp PlainDate where + toTimestamp := Timestamp.ofPlainDate + +end PlainDateTime + +namespace PlainDate + +/-- +Converts a `PlainDate` to a `Timestamp` +-/ +@[inline] +def toTimestamp (pdt : PlainDate) : Timestamp := + Timestamp.ofPlainDate pdt + +end PlainDate + +namespace PlainDateTime + +/-- +Converts a `PlainDateTime` to a `Timestamp` +-/ +@[inline] +def toTimestamp (pdt : PlainDateTime) : Timestamp := + Timestamp.ofPlainDateTime pdt end PlainDateTime diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index dd63dbc401c0..6dc3e8db742c 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -312,6 +312,21 @@ Getter for the `Second` inside of a `PlainDateTime`. def second (dt : PlainDateTime) : Second.Ordinal dt.time.second.fst := dt.time.second.snd +/-- +Determines the era of the given `PlainDateTime` based on its year. +-/ +def era (date : PlainDateTime) : Year.Era := + if date.year.toInt ≥ 0 then + .ce + else + .bce + +/-- +Checks if the `PlainDateTime` is in a leap year. +-/ +def inLeapYear (date : PlainDateTime) : Bool := + date.year.isLeap + /-- Getter for the `Second` inside of a `PlainDateTime`. -/ diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 65a0cba1c96c..2798f5207c86 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -36,6 +36,19 @@ structure Timestamp where proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) deriving Repr +/-- +Type class to transform to `Timestamp`. It's used internally to generate durations out of +some absolute date types. +-/ +class ToTimestamp (α : Type) where + /-- + Transforms into a `Timestamp`. + -/ + toTimestamp : α → Timestamp + +instance : ToTimestamp Timestamp where + toTimestamp := id + instance : BEq Timestamp where beq x y := x.second == y.second && y.nano == x.nano diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 26ce8e0e91eb..5da3c12c3df3 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -5,6 +5,7 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.DateTime +import Std.Time.Zoned namespace Std namespace Time @@ -64,3 +65,51 @@ instance : HAdd Timestamp Duration Timestamp where instance : HSub Timestamp Timestamp Duration where hSub := Std.Time.Timestamp.sub + +namespace PlainDate + +/-- +Calculates the duration between a given `PlainDate` and a specified date. +-/ +def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Timestamp.sub since date + +end PlainDate + +namespace PlainDateTime + +/-- +Calculates the duration between a given `PlainDateTime` and a specified date. +-/ +def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Timestamp.sub since date + +end PlainDateTime + +namespace DateTime + +/-- +Calculates the duration between a given `DateTime` and a specified date. +-/ +def since [ToTimestamp α] (date : DateTime tz) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Timestamp.sub since date + +end DateTime + +namespace ZonedDateTime + +/-- +Calculates the duration between a given `ZonedDateTime` and a specified date. +-/ +def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Timestamp.sub since date + +end ZonedDateTime diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index dde99b46ff72..d5a6838c2e73 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -265,8 +265,8 @@ def fromDateTimeWithZoneString (input : String) : Except String ZonedDateTime := /-- Formats a `DateTime` value into a `DateTimeWithZone` format string. -/ -def toDateTimeWithZoneString (ldt : ZonedDateTime) : String := - Formats.dateTimeWithZone.format ldt.snd +def toDateTimeWithZoneString (pdt : ZonedDateTime) : String := + Formats.dateTimeWithZone.format pdt.snd /-- Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. @@ -318,8 +318,8 @@ def fromAscTimeString (input : String) : Except String PlainDateTime := /-- Formats a `PlainDateTime` value into an AscTime format string. -/ -def toAscTimeString (ldt : PlainDateTime) : String := - Formats.ascTime.format (DateTime.ofPlainDateTime ldt .UTC) +def toAscTimeString (pdt : PlainDateTime) : String := + Formats.ascTime.format (DateTime.ofPlainDateTime pdt .UTC) /-- Parses a `String` in the `LongDateFormat` and returns a `PlainDateTime` object in the GMT time zone. @@ -331,8 +331,8 @@ def fromLongDateFormatString (input : String) : Except String PlainDateTime := /-- Formats a `PlainDateTime` value into a LongDateFormat string. -/ -def toLongDateFormatString (ldt : PlainDateTime) : String := - Formats.longDateFormat.format (DateTime.ofPlainDateTime ldt .UTC) +def toLongDateFormatString (pdt : PlainDateTime) : String := + Formats.longDateFormat.format (DateTime.ofPlainDateTime pdt .UTC) /-- Parses a `String` in the `DateTime` format and returns a `PlainDateTime`. @@ -344,8 +344,8 @@ def fromDateTimeString (input : String) : Except String PlainDateTime := /-- Formats a `PlainDateTime` value into a `DateTime` format string. -/ -def toDateTimeString (ldt : PlainDateTime) : String := - Formats.dateTime24Hour.format (DateTime.ofPlainDateTime ldt .UTC) +def toDateTimeString (pdt : PlainDateTime) : String := + Formats.dateTime24Hour.format (DateTime.ofPlainDateTime pdt .UTC) /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `PlainDateTime`. @@ -418,8 +418,8 @@ def toRFC850String (date : DateTime tz) : String := /-- Formats a `DateTime` value into a `DateTimeWithZone` format string. -/ -def toDateTimeWithZoneString (ldt : DateTime tz) : String := - Formats.dateTimeWithZone.format ldt +def toDateTimeWithZoneString (pdt : DateTime tz) : String := + Formats.dateTimeWithZone.format pdt /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `DateTime`. diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 810faf8b6ffe..6753300f9bcd 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -40,8 +40,8 @@ namespace DateTime Converts a `PlainDate` to a `DateTime` -/ @[inline] -def ofPlainDate (ld : PlainDate) (tz : TimeZone) : DateTime tz := - DateTime.ofTimestamp (Timestamp.ofPlainDate ld) tz +def ofPlainDate (pd : PlainDate) (tz : TimeZone) : DateTime tz := + DateTime.ofTimestamp (Timestamp.ofPlainDate pd) tz /-- Converts a `DateTime` to a `PlainDate` @@ -54,8 +54,8 @@ def toPlainDate (dt : DateTime tz) : PlainDate := Converts a `PlainTime` to a `DateTime` -/ @[inline] -def ofPlainTime (lt : PlainTime) (tz : TimeZone) : DateTime tz := - DateTime.ofTimestamp (Timestamp.ofPlainTime lt) tz +def ofPlainTime (pt : PlainTime) (tz : TimeZone) : DateTime tz := + DateTime.ofTimestamp (Timestamp.ofPlainTime pt) tz /-- Converts a `DateTime` to a `PlainTime` @@ -81,8 +81,8 @@ def now : IO ZonedDateTime := do Converts a `PlainDate` to a `ZonedDateTime` -/ @[inline] -def ofPlainDate (ld : PlainDate) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainDate ld) tz⟩ +def ofPlainDate (pd : PlainDate) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainDate pd) tz⟩ /-- Converts a `ZonedDateTime` to a `PlainDate` @@ -95,8 +95,8 @@ def toPlainDate (dt : ZonedDateTime) : PlainDate := Converts a `PlainTime` to a `ZonedDateTime` -/ @[inline] -def ofPlainTime (lt : PlainTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainTime lt) tz⟩ +def ofPlainTime (pt : PlainTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainTime pt) tz⟩ /-- Converts a `ZonedDateTime` to a `PlainTime` diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index eb36b976ff7e..ea8b5b93a02d 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -257,6 +257,24 @@ Gets the `Weekday` of a DateTime. def weekday (dt : DateTime tz) : Weekday := dt.date.get.date.weekday +/-- +Determines the era of the given `DateTime` based on its year. +-/ +def era (date : DateTime tz) : Year.Era := + if date.year.toInt ≥ 0 then + .ce + else + .bce + +/-- +Checks if the `DateTime` is in a leap year. +-/ +def inLeapYear (date : DateTime tz) : Bool := + date.year.isLeap + +instance : ToTimestamp (DateTime tz) where + toTimestamp dt := dt.toTimestamp + instance : HAdd (DateTime tz) (Day.Offset) (DateTime tz) where hAdd := addDays diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index eb03075c4363..8234d73cd6db 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -232,6 +232,25 @@ Subtract `Second.Offset` from a `ZonedDateTime`, adjusting the date if necessary def subSeconds (dt : ZonedDateTime) (seconds : Second.Offset) : ZonedDateTime := Sigma.mk dt.fst (dt.snd.subSeconds seconds) +/-- +Determines the era of the given `PlainDate` based on its year. +-/ +@[inline] +def era (date : ZonedDateTime) : Year.Era := + if date.year.toInt ≥ 0 then + .ce + else + .bce + +/-- +Checks if the `DateTime` is in a leap year. +-/ +def inLeapYear (date : ZonedDateTime) : Bool := + date.year.isLeap + +instance : ToTimestamp ZonedDateTime where + toTimestamp dt := dt.toTimestamp + instance : HAdd ZonedDateTime (Day.Offset) ZonedDateTime where hAdd := addDays diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index ca1fcbbd00f5..d9c7bf088212 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -18,49 +18,49 @@ info: 1970-01-08 def datetime := date% 2000-01-20:03:02:01 /-- -info: Thursday, January 20, 2000 04:02:01 +info: 2000-01-20 04:02:01 -/ #guard_msgs in #eval datetime + (1 : Hour.Offset) /-- -info: Thursday, January 20, 2000 02:02:01 +info: 2000-01-20 02:02:01 -/ #guard_msgs in #eval datetime - (1 : Hour.Offset) /-- -info: Thursday, January 20, 2000 03:12:01 +info: 2000-01-20 03:12:01 -/ #guard_msgs in #eval datetime + (10 : Minute.Offset) /-- -info: Thursday, January 20, 2000 02:52:01 +info: 2000-01-20 02:52:01 -/ #guard_msgs in #eval datetime - (10 : Minute.Offset) /-- -info: Thursday, January 20, 2000 03:03:01 +info: 2000-01-20 03:03:01 -/ #guard_msgs in #eval datetime + (60 : Second.Offset) /-- -info: Thursday, January 20, 2000 03:01:01 +info: 2000-01-20 03:01:01 -/ #guard_msgs in #eval datetime - (60 : Second.Offset) /-- -info: Friday, January 21, 2000 03:02:01 +info: 2000-01-21 03:02:01 -/ #guard_msgs in #eval datetime + (1 : Day.Offset) /-- -info: Wednesday, January 19, 2000 03:02:01 +info: 2000-01-19 03:02:01 -/ #guard_msgs in #eval datetime - (1 : Day.Offset) @@ -106,49 +106,49 @@ info: 13:01:01 def datetimetz := date% 2000-01-20:03:02:01-03:00 /-- -info: Sat, 22 Jan 2000 06:02:01 -0300 +info: 2000-01-22 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz + (2 : Day.Offset) /-- -info: Wed, 19 Jan 2000 06:02:01 -0300 +info: 2000-01-19 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz - (1 : Day.Offset) /-- -info: Thu, 20 Jan 2000 07:02:01 -0300 +info: 2000-01-20 07:02:01 -0300 -/ #guard_msgs in #eval datetimetz + (1 : Hour.Offset) /-- -info: Thu, 20 Jan 2000 05:02:01 -0300 +info: 2000-01-20 05:02:01 -0300 -/ #guard_msgs in #eval datetimetz - (1 : Hour.Offset) /-- -info: Thu, 20 Jan 2000 06:12:01 -0300 +info: 2000-01-20 06:12:01 -0300 -/ #guard_msgs in #eval datetimetz + (10 : Minute.Offset) /-- -info: Thu, 20 Jan 2000 05:52:01 -0300 +info: 2000-01-20 05:52:01 -0300 -/ #guard_msgs in #eval datetimetz - (10 : Minute.Offset) /-- -info: Thu, 20 Jan 2000 06:03:01 -0300 +info: 2000-01-20 06:03:01 -0300 -/ #guard_msgs in #eval datetimetz + (60 : Second.Offset) /-- -info: Thu, 20 Jan 2000 06:01:01 -0300 +info: 2000-01-20 06:01:01 -0300 -/ #guard_msgs in #eval datetimetz - (60 : Second.Offset) diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 7e0d287ac330..654c5107f42d 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -49,7 +49,7 @@ info: "09/05/1993 12:59:59" info: 736952399 -/ #guard_msgs in -#eval PlainDateTime.toPlainTimestamp date₂ |>.toSecondsSinceUnixEpoch +#eval PlainDateTime.toLocalTimestamp date₂ |>.toSecondsSinceUnixEpoch /-- info: "16/08/2024" diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index f7536f030c04..46fc714bf366 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -36,97 +36,97 @@ info: 1969-12-20 def datetime := date% 2000-01-20:03:02:01 /-- -info: Thursday, January 20, 2000 04:02:01 +info: 2000-01-20 04:02:01 -/ #guard_msgs in #eval datetime.addHours 1 /-- -info: Thursday, January 20, 2000 02:02:01 +info: 2000-01-20 02:02:01 -/ #guard_msgs in #eval datetime.subHours 1 /-- -info: Thursday, January 20, 2000 03:12:01 +info: 2000-01-20 03:12:01 -/ #guard_msgs in #eval datetime.addMinutes 10 /-- -info: Thursday, January 20, 2000 02:52:01 +info: 2000-01-20 02:52:01 -/ #guard_msgs in #eval datetime.subMinutes 10 /-- -info: Thursday, January 20, 2000 03:03:01 +info: 2000-01-20 03:03:01 -/ #guard_msgs in #eval datetime.addSeconds 60 /-- -info: Thursday, January 20, 2000 03:01:01 +info: 2000-01-20 03:01:01 -/ #guard_msgs in #eval datetime.subSeconds 60 /-- -info: Friday, January 21, 2000 03:02:01 +info: 2000-01-21 03:02:01 -/ #guard_msgs in #eval datetime.addDays 1 /-- -info: Wednesday, January 19, 2000 03:02:01 +info: 2000-01-19 03:02:01 -/ #guard_msgs in #eval datetime.subDays 1 /-- -info: Friday, February 20, 2000 03:02:01 +info: 2000-02-20 03:02:01 -/ #guard_msgs in #eval datetime.addMonthsClip 1 /-- -info: Monday, December 20, 1999 03:02:01 +info: 1999-12-20 03:02:01 -/ #guard_msgs in #eval datetime.subMonthsClip 1 /-- -info: Friday, February 20, 2000 03:02:01 +info: 2000-02-20 03:02:01 -/ #guard_msgs in #eval datetime.addMonthsRollOver 1 /-- -info: Monday, December 20, 1999 03:02:01 +info: 1999-12-20 03:02:01 -/ #guard_msgs in #eval datetime.subMonthsRollOver 1 /-- -info: Saturday, January 20, 2001 03:02:01 +info: 2001-01-20 03:02:01 -/ #guard_msgs in #eval datetime.addYearsClip 1 /-- -info: Wednesday, January 20, 1999 03:02:01 +info: 1999-01-20 03:02:01 -/ #guard_msgs in #eval datetime.subYearsClip 1 /-- -info: Saturday, January 20, 2001 03:02:01 +info: 2001-01-20 03:02:01 -/ #guard_msgs in #eval datetime.addYearsRollOver 1 /-- -info: Wednesday, January 20, 1999 03:02:01 +info: 1999-01-20 03:02:01 -/ #guard_msgs in #eval datetime.subYearsRollOver 1 @@ -170,99 +170,98 @@ info: 13:01:01 #eval time.subSeconds 60 def datetimetz := date% 2000-01-20:03:02:01-03:00 - /-- -info: Sat, 22 Jan 2000 06:02:01 -0300 +info: 2000-01-22 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.addDays 2 /-- -info: Wed, 19 Jan 2000 06:02:01 -0300 +info: 2000-01-19 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.subDays 1 /-- -info: Fri, 20 Feb 2000 06:02:01 -0300 +info: 2000-02-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.addMonthsClip 1 /-- -info: Mon, 20 Dec 1999 06:02:01 -0300 +info: 1999-12-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.subMonthsClip 1 /-- -info: Fri, 20 Feb 2000 06:02:01 -0300 +info: 2000-02-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.addMonthsRollOver 1 /-- -info: Mon, 20 Dec 1999 06:02:01 -0300 +info: 1999-12-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.subMonthsRollOver 1 /-- -info: Sat, 20 Jan 2001 06:02:01 -0300 +info: 2001-01-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: Sat, 20 Jan 2001 06:02:01 -0300 +info: 2001-01-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: Sat, 20 Jan 2001 06:02:01 -0300 +info: 2001-01-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.addYearsRollOver 1 /-- -info: Wed, 20 Jan 1999 06:02:01 -0300 +info: 1999-01-20 06:02:01 -0300 -/ #guard_msgs in #eval datetimetz.subYearsRollOver 1 /-- -info: Thu, 20 Jan 2000 07:02:01 -0300 +info: 2000-01-20 07:02:01 -0300 -/ #guard_msgs in #eval datetimetz.addHours 1 /-- -info: Thu, 20 Jan 2000 05:02:01 -0300 +info: 2000-01-20 05:02:01 -0300 -/ #guard_msgs in #eval datetimetz.subHours 1 /-- -info: Thu, 20 Jan 2000 06:12:01 -0300 +info: 2000-01-20 06:12:01 -0300 -/ #guard_msgs in #eval datetimetz.addMinutes 10 /-- -info: Thu, 20 Jan 2000 05:52:01 -0300 +info: 2000-01-20 05:52:01 -0300 -/ #guard_msgs in #eval datetimetz.subMinutes 10 /-- -info: Thu, 20 Jan 2000 06:03:01 -0300 +info: 2000-01-20 06:03:01 -0300 -/ #guard_msgs in #eval datetimetz.addSeconds 60 /-- -info: Thu, 20 Jan 2000 06:01:01 -0300 +info: 2000-01-20 06:01:01 -0300 -/ #guard_msgs in #eval datetimetz.subSeconds 60 From ada5670e51fa10addd7bba04f89b71023b418687 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 01:38:47 -0300 Subject: [PATCH 041/118] fix: problems with since functions --- src/Std/Time/Duration.lean | 13 +++++++++---- tests/lean/run/timeOperations.lean | 6 ++++++ 2 files changed, 15 insertions(+), 4 deletions(-) diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 5da3c12c3df3..c61bf70d9dcf 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -27,6 +27,9 @@ instance : ToString Duration where else if s.nano.val < 0 then ("-", -s.second.val, -s.nano.val) else ("", s.second.val, s.nano.val) sign ++ toString secs ++ "." ++ toString nanos ++ "s" +instance : Repr Duration where + reprPrec s := reprPrec (toString s) + namespace Duration /-- @@ -66,6 +69,8 @@ instance : HAdd Timestamp Duration Timestamp where instance : HSub Timestamp Timestamp Duration where hSub := Std.Time.Timestamp.sub +end Duration + namespace PlainDate /-- @@ -74,7 +79,7 @@ Calculates the duration between a given `PlainDate` and a specified date. def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := let date := date.toTimestamp let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub since date + Std.Time.Timestamp.sub date since end PlainDate @@ -86,7 +91,7 @@ Calculates the duration between a given `PlainDateTime` and a specified date. def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := let date := date.toTimestamp let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub since date + Std.Time.Timestamp.sub date since end PlainDateTime @@ -98,7 +103,7 @@ Calculates the duration between a given `DateTime` and a specified date. def since [ToTimestamp α] (date : DateTime tz) (since : α) : Duration := let date := date.toTimestamp let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub since date + Std.Time.Timestamp.sub date since end DateTime @@ -110,6 +115,6 @@ Calculates the duration between a given `ZonedDateTime` and a specified date. def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := let date := date.toTimestamp let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub since date + Std.Time.Timestamp.sub date since end ZonedDateTime diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index 46fc714bf366..1a76271a045c 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -265,3 +265,9 @@ info: 2000-01-20 06:01:01 -0300 -/ #guard_msgs in #eval datetimetz.subSeconds 60 + +/-- +info: "86400.0s" +-/ +#guard_msgs in +#eval (date% 2000-1-20).since (date% 2000-1-19) From 408fc65b6d0af246a56bc772860d799c29e560f8 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 01:39:23 -0300 Subject: [PATCH 042/118] tests: add more tests --- tests/lean/run/timeOperations.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index 1a76271a045c..4d94d6125e47 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -271,3 +271,15 @@ info: "86400.0s" -/ #guard_msgs in #eval (date% 2000-1-20).since (date% 2000-1-19) + +/-- +info: "86399.0s" +-/ +#guard_msgs in +#eval (date% 2000-1-20:0:0:0).since (date% 2000-1-19:0:0:1) + +/-- +info: "86399.0s" +-/ +#guard_msgs in +#eval (date% 2000-1-20:0:0:0UTC).since (date% 2000-1-19:0:0:1UTC) From 3f9899ba425029c58278bda791da6c5d44901636 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 01:46:20 -0300 Subject: [PATCH 043/118] feat: add more conversion functions for timestamp --- src/Std/Time/DateTime.lean | 4 ++-- src/Std/Time/DateTime/PlainDateTime.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 24 ++++++++++++++++++------ src/Std/Time/Duration.lean | 16 ++++++++++++++-- 4 files changed, 35 insertions(+), 11 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index b6fcfc03db3f..01a10ac59be7 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -34,7 +34,7 @@ Converts a `PlainDate` to a `Timestamp` def ofPlainDate (pd : PlainDate) : Timestamp := let days := pd.toDaysSinceUNIXEpoch let secs := days.toSeconds - Timestamp.ofSecondsSinceUnixEpoch secs + Timestamp.ofSeconds secs /-- Converts a `Timestamp` to a `PlainDate` @@ -51,7 +51,7 @@ Converts a `PlainTime` to a `Timestamp` @[inline] def ofPlainTime (pt : PlainTime) : Timestamp := let nanos := pt.toNanoseconds - Timestamp.ofNanosecondsSinceUnixEpoch nanos + Timestamp.ofNanoseconds nanos /-- Converts a `Timestamp` to a `PlainTime` diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 6dc3e8db742c..f43b0073adab 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -41,7 +41,7 @@ def toLocalTimestamp (dt : PlainDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 let nanos := nanos.val + dt.time.nano.val - Timestamp.ofNanosecondsSinceUnixEpoch (Nanosecond.Offset.ofInt nanos) + Timestamp.ofNanoseconds (Nanosecond.Offset.ofInt nanos) /-- Converts a UNIX `Timestamp` to a `PlainDateTime`. diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 2798f5207c86..e5888354122d 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -148,7 +148,7 @@ protected def sub (t₁ t₂ : Timestamp) : Timestamp := Creates a new `Timestamp` out of `Second.Offset`. -/ @[inline] -def ofSecondsSinceUnixEpoch (s : Second.Offset) : Timestamp := by +def ofSeconds (s : Second.Offset) : Timestamp := by refine ⟨s, ⟨0, by decide⟩, ?_⟩ simp <;> exact Int.le_total s.val 0 |>.symm @@ -156,7 +156,7 @@ def ofSecondsSinceUnixEpoch (s : Second.Offset) : Timestamp := by Creates a new `Timestamp` out of `Second.Offset`. -/ @[inline] -def ofNanosecondsSinceUnixEpoch (s : Nanosecond.Offset) : Timestamp := by +def ofNanoseconds (s : Nanosecond.Offset) : Timestamp := by refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ cases Int.le_total s.val 0 next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) @@ -175,33 +175,45 @@ def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := let nanos := nanos + (.ofInt tm.nano.val) nanos +/-- +Converts a `Timestamp` to a `Minute.Offset` +-/ +def toMinutes (tm : Timestamp) : Minute.Offset := + tm.second.ediv 60 + +/-- +Converts a `Timestamp` to a `Day.Offset` +-/ +def toDays (tm : Timestamp) : Day.Offset := + tm.second.ediv 86400 + /-- Adds a `Nanosecond.Offset` to a `Timestamp` -/ @[inline] def addNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := - t.add (ofNanosecondsSinceUnixEpoch s) + t.add (ofNanoseconds s) /-- Adds a `Nanosecond.Offset` to a `Timestamp` -/ @[inline] def subNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := - t.sub (ofNanosecondsSinceUnixEpoch s) + t.sub (ofNanoseconds s) /-- Adds a `Second.Offset` to a `Timestamp` -/ @[inline] def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.add (ofSecondsSinceUnixEpoch s) + t.add (ofSeconds s) /-- Subtracts a `Second.Offset` from a `Timestamp` -/ @[inline] def subSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.sub (ofSecondsSinceUnixEpoch s) + t.sub (ofSeconds s) /-- Adds a `Minute.Offset` to a `Timestamp` diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index c61bf70d9dcf..a03e37550a61 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -32,17 +32,29 @@ instance : Repr Duration where namespace Duration +/-- +Converts a `Duration` to a `Minute.Offset` +-/ +def toMinutes (tm : Duration) : Minute.Offset := + tm.second.ediv 60 + +/-- +Converts a `Duration` to a `Day.Offset` +-/ +def toDays (tm : Duration) : Day.Offset := + tm.second.ediv 86400 + /-- Creates a new `Duration` out of `Second.Offset`. -/ def ofSeconds (secs : Second.Offset) : Duration := - Timestamp.ofSecondsSinceUnixEpoch secs + Timestamp.ofSeconds secs /-- Creates a new `Duration` out of `Nanosecond.Offset`. -/ def ofNanosecond (secs : Nanosecond.Offset) : Duration := - Timestamp.ofNanosecondsSinceUnixEpoch secs + Timestamp.ofNanoseconds secs /-- Calculates a `Duration` out of two `Timestamp`s. From ebadb4f8778cbd261dde75e2168fb3f1df03bee2 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 09:04:54 -0300 Subject: [PATCH 044/118] feat: remove notation import --- src/Std/Time/Notation.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 19367db841b6..d2b8d978ae53 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -9,11 +9,10 @@ import Std.Time.Time import Std.Time.Zoned import Std.Time.DateTime import Std.Time.Format.Basic -import Lean.Elab.Command namespace Std namespace Time -open Lean.Elab Term Lean Parser Command Std +open Lean Parser Command Std /-- Category of units that are valid inside a date. From 8e0794f1b86739bdf276531af0cbd94b342ec18f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 09:36:38 -0300 Subject: [PATCH 045/118] fix: problems with tests and function names --- src/Std/Time/DateTime.lean | 4 ++-- src/Std/Time/DateTime/PlainDateTime.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 12 ++++++------ src/Std/Time/Duration.lean | 4 ++-- 4 files changed, 11 insertions(+), 11 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 01a10ac59be7..b6fcfc03db3f 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -34,7 +34,7 @@ Converts a `PlainDate` to a `Timestamp` def ofPlainDate (pd : PlainDate) : Timestamp := let days := pd.toDaysSinceUNIXEpoch let secs := days.toSeconds - Timestamp.ofSeconds secs + Timestamp.ofSecondsSinceUnixEpoch secs /-- Converts a `Timestamp` to a `PlainDate` @@ -51,7 +51,7 @@ Converts a `PlainTime` to a `Timestamp` @[inline] def ofPlainTime (pt : PlainTime) : Timestamp := let nanos := pt.toNanoseconds - Timestamp.ofNanoseconds nanos + Timestamp.ofNanosecondsSinceUnixEpoch nanos /-- Converts a `Timestamp` to a `PlainTime` diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index f43b0073adab..6dc3e8db742c 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -41,7 +41,7 @@ def toLocalTimestamp (dt : PlainDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 let nanos := nanos.val + dt.time.nano.val - Timestamp.ofNanoseconds (Nanosecond.Offset.ofInt nanos) + Timestamp.ofNanosecondsSinceUnixEpoch (Nanosecond.Offset.ofInt nanos) /-- Converts a UNIX `Timestamp` to a `PlainDateTime`. diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index e5888354122d..cadf4c600c39 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -148,7 +148,7 @@ protected def sub (t₁ t₂ : Timestamp) : Timestamp := Creates a new `Timestamp` out of `Second.Offset`. -/ @[inline] -def ofSeconds (s : Second.Offset) : Timestamp := by +def ofSecondsSinceUnixEpoch (s : Second.Offset) : Timestamp := by refine ⟨s, ⟨0, by decide⟩, ?_⟩ simp <;> exact Int.le_total s.val 0 |>.symm @@ -156,7 +156,7 @@ def ofSeconds (s : Second.Offset) : Timestamp := by Creates a new `Timestamp` out of `Second.Offset`. -/ @[inline] -def ofNanoseconds (s : Nanosecond.Offset) : Timestamp := by +def ofNanosecondsSinceUnixEpoch (s : Nanosecond.Offset) : Timestamp := by refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ cases Int.le_total s.val 0 next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) @@ -192,28 +192,28 @@ Adds a `Nanosecond.Offset` to a `Timestamp` -/ @[inline] def addNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := - t.add (ofNanoseconds s) + t.add (ofNanosecondsSinceUnixEpoch s) /-- Adds a `Nanosecond.Offset` to a `Timestamp` -/ @[inline] def subNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := - t.sub (ofNanoseconds s) + t.sub (ofNanosecondsSinceUnixEpoch s) /-- Adds a `Second.Offset` to a `Timestamp` -/ @[inline] def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.add (ofSeconds s) + t.add (ofSecondsSinceUnixEpoch s) /-- Subtracts a `Second.Offset` from a `Timestamp` -/ @[inline] def subSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.sub (ofSeconds s) + t.sub (ofSecondsSinceUnixEpoch s) /-- Adds a `Minute.Offset` to a `Timestamp` diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index a03e37550a61..e89fa01d1fbe 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -48,13 +48,13 @@ def toDays (tm : Duration) : Day.Offset := Creates a new `Duration` out of `Second.Offset`. -/ def ofSeconds (secs : Second.Offset) : Duration := - Timestamp.ofSeconds secs + Timestamp.ofSecondsSinceUnixEpoch secs /-- Creates a new `Duration` out of `Nanosecond.Offset`. -/ def ofNanosecond (secs : Nanosecond.Offset) : Duration := - Timestamp.ofNanoseconds secs + Timestamp.ofNanosecondsSinceUnixEpoch secs /-- Calculates a `Duration` out of two `Timestamp`s. From 0f755ec534634d858420daa25de4344f6093a9e6 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 09:39:22 -0300 Subject: [PATCH 046/118] fix: tests using Rat --- tests/lean/hpow.lean | 2 +- tests/lean/ppNumericTypes.lean | 2 +- tests/lean/run/1780.lean | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/tests/lean/hpow.lean b/tests/lean/hpow.lean index 398aa14d4d8d..fe896373faed 100644 --- a/tests/lean/hpow.lean +++ b/tests/lean/hpow.lean @@ -15,7 +15,7 @@ variable (n : Nat) (m : Int) (q : Rat) #check (n ^ 2 + (1 : Nat) : Int) instance instNatPowRat : NatPow Rat where - pow q n := Lean.mkRat (q.num ^ n) (q.den ^ n) + pow q n := Std.Internal.mkRat (q.num ^ n) (q.den ^ n) instance instPowRatInt : Pow Rat Int where pow q m := if 0 ≤ m then q ^ (m.toNat : Nat) else (1/q) ^ ((-m).toNat) diff --git a/tests/lean/ppNumericTypes.lean b/tests/lean/ppNumericTypes.lean index 2e054457a0b0..a52dcda22f4c 100644 --- a/tests/lean/ppNumericTypes.lean +++ b/tests/lean/ppNumericTypes.lean @@ -7,7 +7,7 @@ Tests for `pp.numericTypes` and `pp.natLit` RFC #3021 -/ -open Lean (Rat) +open Lean section diff --git a/tests/lean/run/1780.lean b/tests/lean/run/1780.lean index d3345288469d..0712fd9a1ff9 100644 --- a/tests/lean/run/1780.lean +++ b/tests/lean/run/1780.lean @@ -6,6 +6,6 @@ open Meta -- ok end section -open Lean hiding Rat +open Lean hiding AttributeImplCore open Meta -- ok end From 751b85257f3b0a106d6fe9b8f9ab5f4dddb61015 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 09:42:32 -0300 Subject: [PATCH 047/118] fix: update src/Std/Time/Duration.lean Co-authored-by: Markus Himmel --- src/Std/Time/Duration.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index e89fa01d1fbe..c26b3f4aaa78 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -70,7 +70,7 @@ def add (f : Timestamp) (d : Duration) : Timestamp := Std.Time.Timestamp.add f d /-- -Checks if the duration is zero seconds ands zero nanoseconds. +Checks if the duration is zero seconds and zero nanoseconds. -/ def isZero (d : Duration) : Bool := d.second.val = 0 ∧ d.nano.val = 0 From ec3968cabb6cf028ef7df46bdd913870f54caeca Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 09:42:49 -0300 Subject: [PATCH 048/118] fix: update src/Std/Time.lean message Co-authored-by: Markus Himmel --- src/Std/Time.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index fb5c77a9dac7..ccda79bd1e4b 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -101,7 +101,7 @@ Combines date and time into a single representation, useful for precise timestam - **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. - **`Timestamp`**: Represents a point in time with second-level precision. It starts on the UNIX -epoch and it should be used when you receive or need to send timestamps to another systems. +epoch and it should be used when you receive or need to send timestamps to other systems. ## Zoned date and times. Combines date, time and time zones. From 86fec58fe1f88739f175408bafd2666a11f90fc8 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 09:43:03 -0300 Subject: [PATCH 049/118] style: remove whitspace. Co-authored-by: Markus Himmel --- src/Std/Time/Date/Unit/Month.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index dccee95aa178..648cc5651d12 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -192,7 +192,7 @@ def daysWithoutProof (leap : Bool) (month : Ordinal) : Day.Ordinal := rw [← p] at r exact months.get r -theorem all_greater_than_27 (leap : Bool) (i: Month.Ordinal) : daysWithoutProof leap i > 27 := by +theorem all_greater_than_27 (leap : Bool) (i : Month.Ordinal) : daysWithoutProof leap i > 27 := by simp [daysWithoutProof, monthSizesNonLeap, Bounded.LE.sub, Bounded.LE.add, Bounded.LE.toFin] match i with | ⟨2, _⟩ => split <;> (simp; try split); all_goals decide From 3dabdcd024ddb919aadf61d0e1345606fae619c1 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 11:42:03 -0300 Subject: [PATCH 050/118] refactor: change Timestamp to depend on Duration --- src/Std/Time/DateTime.lean | 47 ++-- src/Std/Time/DateTime/PlainDateTime.lean | 2 +- src/Std/Time/DateTime/Timestamp.lean | 186 +++++---------- src/Std/Time/Duration.lean | 289 ++++++++++++++++++----- src/Std/Time/Zoned.lean | 18 +- src/Std/Time/Zoned/ZoneRules.lean | 2 +- tests/lean/run/timeLocalDateTime.lean | 6 +- tests/lean/run/timeTzifParse.lean | 4 +- 8 files changed, 334 insertions(+), 220 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index b6fcfc03db3f..e45360bb4011 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -9,6 +9,7 @@ import Std.Time.DateTime.PlainDateTime namespace Std namespace Time + namespace Timestamp set_option linter.all true @@ -18,7 +19,7 @@ Converts a `PlainDateTime` to a `Timestamp` -/ @[inline] def ofPlainDateTime (pdt : PlainDateTime) : Timestamp := - pdt.toLocalTimestamp + pdt.toTimestampAssumingUTC /-- Converts a `Timestamp` to a `PlainDateTime` @@ -62,7 +63,24 @@ def toPlainTime (timestamp : Timestamp) : PlainTime := PlainTime.ofNanoseconds nanos end Timestamp +namespace PlainDate + +/-- +Converts a `PlainDate` to a `Timestamp` +-/ +@[inline] +def toTimestamp (pdt : PlainDate) : Timestamp := + Timestamp.ofPlainDate pdt + +/-- +Calculates the duration between a given `PlainDate` and a specified date. +-/ +def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch +end PlainDate namespace PlainDateTime /-- @@ -77,7 +95,7 @@ Converts a `PlainDateTime` to a `PlainDate` -/ @[inline] def toPlainDate (pdt : PlainDateTime) : PlainDate := - Timestamp.toPlainDate pdt.toLocalTimestamp + Timestamp.toPlainDate pdt.toTimestampAssumingUTC /-- Converts a `PlainTime` to a `PlainDateTime` @@ -91,7 +109,7 @@ Converts a `PlainDateTime` to a `PlainTime` -/ @[inline] def toPlainTime (pdt : PlainDateTime) : PlainTime := - Timestamp.toPlainTime pdt.toLocalTimestamp + Timestamp.toPlainTime pdt.toTimestampAssumingUTC instance : ToTimestamp PlainDateTime where toTimestamp := Timestamp.ofPlainDateTime @@ -99,21 +117,6 @@ instance : ToTimestamp PlainDateTime where instance : ToTimestamp PlainDate where toTimestamp := Timestamp.ofPlainDate -end PlainDateTime - -namespace PlainDate - -/-- -Converts a `PlainDate` to a `Timestamp` --/ -@[inline] -def toTimestamp (pdt : PlainDate) : Timestamp := - Timestamp.ofPlainDate pdt - -end PlainDate - -namespace PlainDateTime - /-- Converts a `PlainDateTime` to a `Timestamp` -/ @@ -121,4 +124,12 @@ Converts a `PlainDateTime` to a `Timestamp` def toTimestamp (pdt : PlainDateTime) : Timestamp := Timestamp.ofPlainDateTime pdt +/-- +Calculates the duration between a given `PlainDateTime` and a specified date. +-/ +def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch + end PlainDateTime diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 6dc3e8db742c..6f8c3f0407a6 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -37,7 +37,7 @@ namespace PlainDateTime /-- Converts a `PlainDateTime` to a `Timestamp` -/ -def toLocalTimestamp (dt : PlainDateTime) : Timestamp := +def toTimestampAssumingUTC (dt : PlainDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 let nanos := nanos.val + dt.time.nano.val diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index cadf4c600c39..33ba3a1b775c 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -8,6 +8,7 @@ import Std.Time.Internal import Init.Data.Int import Std.Time.Time import Std.Time.Date +import Std.Time.Duration namespace Std namespace Time @@ -16,28 +17,32 @@ open Internal set_option linter.all true /-- -Represents a point in time relative to the UNIX Epoch, with nanoseconds precision. +Timestamp is just a period between two timestamps. -/ structure Timestamp where /-- - Second offset of the timestamp. + Duration since the unix epoch. -/ - second : Second.Offset + val : Duration + deriving Repr, BEq, Inhabited - /-- - Nanosecond span that ranges from -999999999 and 999999999 - -/ - nano : Nanosecond.Span +instance : OfNat Timestamp n where + ofNat := ⟨OfNat.ofNat n⟩ - /-- - Proof that the timestamp is valid, ensuring that the `second` and `nano` values are correctly related. - -/ - proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) - deriving Repr +instance : ToString Timestamp where + toString s := + let (sign, secs, nanos) := + if s.val.second.val > 0 then ("" ,s.val.second.val, s.val.nano.val) + else if s.val.second.val < 0 then ("-", -s.val.second.val, -s.val.nano.val) + else if s.val.nano.val < 0 then ("-", -s.val.second.val, -s.val.nano.val) else ("", s.val.second.val, s.val.nano.val) + sign ++ toString secs ++ "." ++ toString nanos ++ "s" + +instance : Repr Timestamp where + reprPrec s := reprPrec (toString s) /-- -Type class to transform to `Timestamp`. It's used internally to generate durations out of +Type class to transform to `Timestamp`. It's used internally to generate timestamps out of some absolute date types. -/ class ToTimestamp (α : Type) where @@ -49,122 +54,44 @@ class ToTimestamp (α : Type) where instance : ToTimestamp Timestamp where toTimestamp := id -instance : BEq Timestamp where - beq x y := x.second == y.second && y.nano == x.nano - -instance : Inhabited Timestamp where - default := ⟨0, Bounded.LE.mk 0 (by decide), by decide⟩ - -instance : OfNat Timestamp n where - ofNat := by - refine ⟨.ofInt n, ⟨0, by decide⟩, ?_⟩ - simp <;> exact Int.le_total s.val 0 |>.symm - exact Int.le_total 0 n - namespace Timestamp /-- -Fetches the current timestamp from the system. +Fetches the current duration from the system. -/ @[extern "lean_get_current_time"] opaque now : IO Timestamp /-- -Converts a `Timestamp` into its equivalent `Second.Offset`. --/ -@[inline] -def toSecondsSinceUnixEpoch (t : Timestamp) : Second.Offset := - t.second - -/-- -Negates a `Timestamp`, flipping its second and nanosecond values. +Converts a `Timestamp` to a `Minute.Offset` -/ -@[inline] -protected def neg (t : Timestamp) : Timestamp := by - refine ⟨-t.second, t.nano.neg, ?_⟩ - cases t.proof with - | inl n => exact Or.inr (n.imp Int.neg_le_neg Int.neg_le_neg) - | inr n => exact Or.inl (n.imp Int.neg_le_neg Int.neg_le_neg) +def toMinutes (tm : Timestamp) : Minute.Offset := + tm.val.second.ediv 60 /-- -Adds two timestamps together, handling any carry-over in nanoseconds. It should not be used for `Timestamp`. -The subtraction of two `Timestamp` returns a duration but the addition does not make sense at all. +Converts a `Timestamp` to a `Day.Offset` -/ -protected def add (t₁ t₂ : Timestamp) : Timestamp := by - let diffSecs := t₁.second + t₂.second - let diffNano := t₁.nano.addBounds t₂.nano - - let (diffSecs, diffNano) : Second.Offset × Nanosecond.Span := - if h₀ : diffNano.val > 999999999 then - have diffNano := diffNano.truncateBottom h₀ |>.sub 999999999 - (diffSecs + 1, diffNano.expandBottom (by decide)) - else if h₁ : diffNano.val < -999999999 then - have diffNano := diffNano.truncateTop (Int.le_sub_one_of_lt h₁) |>.add 999999999 - (diffSecs - 1, diffNano.expandTop (by decide)) - else by - have h₀ := Int.le_sub_one_of_lt <| Int.not_le.mp h₀ - have h₁ := Int.le_sub_one_of_lt <| Int.not_le.mp h₁ - simp_all [Int.add_sub_cancel] - exact ⟨diffSecs, Bounded.mk diffNano.val (And.intro h₁ h₀)⟩ - - if h : diffSecs.val > 0 ∧ diffNano.val < 0 then - let truncated := diffNano.truncateTop (Int.le_sub_one_of_lt h.right) - let nano := truncated.addTop 1000000000 (by decide) - let proof₁ : 0 ≤ diffSecs - 1 := Int.le_sub_one_of_lt h.left - refine { second := .ofInt (diffSecs.val - 1), nano, proof := ?_ } - simp [nano, Bounded.LE.addTop] - refine (Or.inl (And.intro proof₁ ?_)) - let h₃ := (Int.add_le_add_iff_left 1000000000).mpr diffNano.property.left - rw [Int.add_comm] - exact Int.le_trans (by decide) h₃ - else if h₁ : diffSecs.val < 0 ∧ diffNano.val > 0 then - let second := diffSecs.val + 1 - let truncated := diffNano.truncateBottom h₁.right - let nano := truncated.subBottom 1000000000 (by decide) - refine { second := .ofInt second, nano, proof := ?_ } - simp [nano, truncated, Bounded.LE.subBottom, Bounded.LE.truncateBottom] - refine (Or.inr (And.intro ?_ ?_)) - · exact h₁.left - · let h₃ := Int.sub_le_sub_right diffNano.property.right 1000000000 - simp at h₃ - exact Int.le_trans h₃ (by decide) - else - refine ⟨diffSecs, diffNano, ?_⟩ - if h₂ : diffSecs.val > 0 then - exact Or.inl (And.intro (Int.le_of_lt h₂) (Int.not_lt.mp (not_and.mp h h₂))) - else if h₃ : diffSecs.val < 0 then - exact Or.inr (And.intro (Int.le_of_lt h₃) (Int.not_lt.mp (not_and.mp h₁ h₃))) - else - exact Int.le_total diffNano.val 0 |>.symm.imp (And.intro (Int.not_lt.mp h₃)) (And.intro (Int.not_lt.mp h₂)) +def toDays (tm : Timestamp) : Day.Offset := + tm.val.second.ediv 86400 /-- -Subtracts one `Timestamp` from another. +Creates a new `Timestamp` out of `Second.Offset`. -/ -@[inline] -protected def sub (t₁ t₂ : Timestamp) : Timestamp := - t₁.add t₂.neg +def ofSecondsSinceUnixEpoch (secs : Second.Offset) : Timestamp := + ⟨Duration.ofSeconds secs⟩ /-- -Creates a new `Timestamp` out of `Second.Offset`. +Creates a new `Timestamp` out of `Nanosecond.Offset`. -/ -@[inline] -def ofSecondsSinceUnixEpoch (s : Second.Offset) : Timestamp := by - refine ⟨s, ⟨0, by decide⟩, ?_⟩ - simp <;> exact Int.le_total s.val 0 |>.symm +def ofNanosecondsSinceUnixEpoch (secs : Nanosecond.Offset) : Timestamp := + ⟨Duration.ofNanoseconds secs⟩ /-- -Creates a new `Timestamp` out of `Second.Offset`. +Converts a `Timestamp` into its equivalent `Second.Offset`. -/ @[inline] -def ofNanosecondsSinceUnixEpoch (s : Nanosecond.Offset) : Timestamp := by - refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ - cases Int.le_total s.val 0 - next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) - next n => exact Or.inl (And.intro (Int.ediv_nonneg n (by decide)) (Int.mod_nonneg 1000000000 n)) - where - mod_nonpos : ∀ {a : Int} (b : Int), (a ≤ 0) → (b ≥ 0) → 0 ≥ a.mod b - | .negSucc m, .ofNat n, _, _ => Int.neg_le_neg (Int.mod_nonneg (↑n) (Int.ofNat_le.mpr (Nat.zero_le (m + 1)))) - | 0, n, _, _ => Int.eq_iff_le_and_ge.mp (Int.zero_mod n) |>.left +def toSecondsSinceUnixEpoch (t : Timestamp) : Second.Offset := + t.val.second /-- Converts a `Timestamp` from a `Nanosecond.Offset` @@ -172,96 +99,93 @@ Converts a `Timestamp` from a `Nanosecond.Offset` @[inline] def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := let nanos := tm.toSecondsSinceUnixEpoch.mul 1000000000 - let nanos := nanos + (.ofInt tm.nano.val) + let nanos := nanos + (.ofInt tm.val.nano.val) nanos /-- -Converts a `Timestamp` to a `Minute.Offset` +Calculates a `Timestamp` out of two `Timestamp`s. -/ -def toMinutes (tm : Timestamp) : Minute.Offset := - tm.second.ediv 60 +@[inline] +def since (f : Timestamp) : IO Duration := do + let cur ← Timestamp.now + return Std.Time.Duration.sub cur.val f.val /-- -Converts a `Timestamp` to a `Day.Offset` +Creates a duration out of the `Timestamp` since the unix epoch. -/ -def toDays (tm : Timestamp) : Day.Offset := - tm.second.ediv 86400 +@[inline] +def toDurationSinceUnixEpoch (tm : Timestamp) : Duration := + tm.val /-- Adds a `Nanosecond.Offset` to a `Timestamp` -/ @[inline] def addNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := - t.add (ofNanosecondsSinceUnixEpoch s) + ⟨t.val + s⟩ /-- Adds a `Nanosecond.Offset` to a `Timestamp` -/ @[inline] def subNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := - t.sub (ofNanosecondsSinceUnixEpoch s) + ⟨t.val - s⟩ /-- Adds a `Second.Offset` to a `Timestamp` -/ @[inline] def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.add (ofSecondsSinceUnixEpoch s) + ⟨t.val + s⟩ /-- Subtracts a `Second.Offset` from a `Timestamp` -/ @[inline] def subSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := - t.sub (ofSecondsSinceUnixEpoch s) + ⟨t.val - s⟩ /-- Adds a `Minute.Offset` to a `Timestamp` -/ @[inline] def addMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := - let seconds := m.mul 60 - t.addSeconds seconds + ⟨t.val + m⟩ /-- Subtracts a `Minute.Offset` from a `Timestamp` -/ @[inline] def subMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := - let seconds := m.mul 60 - t.subSeconds seconds + ⟨t.val - m⟩ /-- Adds an `Hour.Offset` to a `Timestamp` -/ @[inline] def addHours (t : Timestamp) (h : Hour.Offset) : Timestamp := - let seconds := h.mul 3600 - t.addSeconds seconds + ⟨t.val + h⟩ /-- Subtracts an `Hour.Offset` from a `Timestamp` -/ @[inline] def subHours (t : Timestamp) (h : Hour.Offset) : Timestamp := - let seconds := h.mul 3600 - t.subSeconds seconds + ⟨t.val - h⟩ /-- Adds a `Day.Offset` to a `Timestamp` -/ @[inline] def addDays (t : Timestamp) (d : Day.Offset) : Timestamp := - let seconds := d.mul 86400 - t.addSeconds seconds + ⟨t.val + d⟩ /-- Subtracts a `Day.Offset` from a `Timestamp` -/ @[inline] def subDays (t : Timestamp) (d : Day.Offset) : Timestamp := - let seconds := d.mul 86400 - t.subSeconds seconds + ⟨t.val - d⟩ instance : HAdd Timestamp Day.Offset Timestamp where hAdd := addDays @@ -294,5 +218,3 @@ instance : HSub Timestamp Nanosecond.Offset Timestamp where hSub := subNanoseconds end Timestamp -end Time -end Std diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index c26b3f4aaa78..f9b9e669a9bc 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.DateTime -import Std.Time.Zoned +import Std.Time.Date +import Std.Time.Time namespace Std namespace Time @@ -14,119 +14,284 @@ open Internal set_option linter.all true /-- -Duration is just a period between two timestamps. +Represents a point in time relative to the UNIX Epoch, with nanoseconds precision. -/ -def Duration := Timestamp - deriving Repr, BEq +structure Duration where + + /-- + Second offset of the duration. + -/ + second : Second.Offset + + /-- + Nanosecond span that ranges from -999999999 and 999999999 + -/ + nano : Nanosecond.Span + + /-- + Proof that the duration is valid, ensuring that the `second` and `nano` values are correctly related. + -/ + proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) + deriving Repr instance : ToString Duration where toString s := let (sign, secs, nanos) := - if s.second.val > 0 then ("" ,s.second.val, s.nano.val) - else if s.second.val < 0 then ("-", -s.second.val, -s.nano.val) - else if s.nano.val < 0 then ("-", -s.second.val, -s.nano.val) else ("", s.second.val, s.nano.val) + if s.second.val > 0 then ("" ,s.second, s.nano.val) + else if s.second.val < 0 then ("-", -s.second, -s.nano.val) + else if s.nano.val < 0 then ("-", -s.second, -s.nano.val) else ("", s.second, s.nano.val) sign ++ toString secs ++ "." ++ toString nanos ++ "s" instance : Repr Duration where reprPrec s := reprPrec (toString s) +instance : BEq Duration where + beq x y := x.second == y.second && y.nano == x.nano + +instance : Inhabited Duration where + default := ⟨0, Bounded.LE.mk 0 (by decide), by decide⟩ + +instance : OfNat Duration n where + ofNat := by + refine ⟨.ofInt n, ⟨0, by decide⟩, ?_⟩ + simp <;> exact Int.le_total s.val 0 |>.symm + exact Int.le_total 0 n + namespace Duration /-- -Converts a `Duration` to a `Minute.Offset` +Negates a `Duration`, flipping its second and nanosecond values. -/ -def toMinutes (tm : Duration) : Minute.Offset := - tm.second.ediv 60 +@[inline] +protected def neg (duration : Duration) : Duration := by + refine ⟨-duration.second, duration.nano.neg, ?_⟩ + cases duration.proof with + | inl n => exact Or.inr (n.imp Int.neg_le_neg Int.neg_le_neg) + | inr n => exact Or.inl (n.imp Int.neg_le_neg Int.neg_le_neg) /-- -Converts a `Duration` to a `Day.Offset` +Adds two durations together, handling any carry-over in nanoseconds. It should not be used for `Duration`. +The subtraction of two `Duration` returns a duration but the addition does not make sense at all. -/ -def toDays (tm : Duration) : Day.Offset := - tm.second.ediv 86400 +def add (t₁ t₂ : Duration) : Duration := by + let diffSecs := t₁.second + t₂.second + let diffNano := t₁.nano.addBounds t₂.nano -/-- -Creates a new `Duration` out of `Second.Offset`. --/ -def ofSeconds (secs : Second.Offset) : Duration := - Timestamp.ofSecondsSinceUnixEpoch secs + let (diffSecs, diffNano) : Second.Offset × Nanosecond.Span := + if h₀ : diffNano.val > 999999999 then + have diffNano := diffNano.truncateBottom h₀ |>.sub 999999999 + (diffSecs + 1, diffNano.expandBottom (by decide)) + else if h₁ : diffNano.val < -999999999 then + have diffNano := diffNano.truncateTop (Int.le_sub_one_of_lt h₁) |>.add 999999999 + (diffSecs - 1, diffNano.expandTop (by decide)) + else by + have h₀ := Int.le_sub_one_of_lt <| Int.not_le.mp h₀ + have h₁ := Int.le_sub_one_of_lt <| Int.not_le.mp h₁ + simp_all [Int.add_sub_cancel] + exact ⟨diffSecs, Bounded.mk diffNano.val (And.intro h₁ h₀)⟩ + + if h : diffSecs.val > 0 ∧ diffNano.val < 0 then + let truncated := diffNano.truncateTop (Int.le_sub_one_of_lt h.right) + let nano := truncated.addTop 1000000000 (by decide) + let proof₁ : 0 ≤ diffSecs - 1 := Int.le_sub_one_of_lt h.left + refine { second := .ofInt (diffSecs.val - 1), nano, proof := ?_ } + simp [nano, Bounded.LE.addTop] + refine (Or.inl (And.intro proof₁ ?_)) + let h₃ := (Int.add_le_add_iff_left 1000000000).mpr diffNano.property.left + rw [Int.add_comm] + exact Int.le_trans (by decide) h₃ + else if h₁ : diffSecs.val < 0 ∧ diffNano.val > 0 then + let second := diffSecs.val + 1 + let truncated := diffNano.truncateBottom h₁.right + let nano := truncated.subBottom 1000000000 (by decide) + refine { second := .ofInt second, nano, proof := ?_ } + simp [nano, truncated, Bounded.LE.subBottom, Bounded.LE.truncateBottom] + refine (Or.inr (And.intro ?_ ?_)) + · exact h₁.left + · let h₃ := Int.sub_le_sub_right diffNano.property.right 1000000000 + simp at h₃ + exact Int.le_trans h₃ (by decide) + else + refine ⟨diffSecs, diffNano, ?_⟩ + if h₂ : diffSecs.val > 0 then + exact Or.inl (And.intro (Int.le_of_lt h₂) (Int.not_lt.mp (not_and.mp h h₂))) + else if h₃ : diffSecs.val < 0 then + exact Or.inr (And.intro (Int.le_of_lt h₃) (Int.not_lt.mp (not_and.mp h₁ h₃))) + else + exact Int.le_total diffNano.val 0 |>.symm.imp (And.intro (Int.not_lt.mp h₃)) (And.intro (Int.not_lt.mp h₂)) /-- -Creates a new `Duration` out of `Nanosecond.Offset`. +Subtracts one `Duration` from another. -/ -def ofNanosecond (secs : Nanosecond.Offset) : Duration := - Timestamp.ofNanosecondsSinceUnixEpoch secs +@[inline] +def sub (t₁ t₂ : Duration) : Duration := + t₁.add t₂.neg /-- -Calculates a `Duration` out of two `Timestamp`s. +Creates a new `Duration` out of `Second.Offset`. -/ -def since (f : Timestamp) : IO Duration := do - let cur ← Timestamp.now - return Std.Time.Timestamp.sub cur f +@[inline] +def ofSeconds (s : Second.Offset) : Duration := by + refine ⟨s, ⟨0, by decide⟩, ?_⟩ + simp <;> exact Int.le_total s.val 0 |>.symm /-- -Adds a `Duration` to a `Timestamp`. +Creates a new `Duration` out of `Nanosecond.Offset`. -/ -def add (f : Timestamp) (d : Duration) : Timestamp := - Std.Time.Timestamp.add f d +@[inline] +def ofNanoseconds (s : Nanosecond.Offset) : Duration := by + refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ + cases Int.le_total s.val 0 + next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) + next n => exact Or.inl (And.intro (Int.ediv_nonneg n (by decide)) (Int.mod_nonneg 1000000000 n)) + where + mod_nonpos : ∀ {a : Int} (b : Int), (a ≤ 0) → (b ≥ 0) → 0 ≥ a.mod b + | .negSucc m, .ofNat n, _, _ => Int.neg_le_neg (Int.mod_nonneg (↑n) (Int.ofNat_le.mpr (Nat.zero_le (m + 1)))) + | 0, n, _, _ => Int.eq_iff_le_and_ge.mp (Int.zero_mod n) |>.left /-- Checks if the duration is zero seconds and zero nanoseconds. -/ +@[inline] def isZero (d : Duration) : Bool := d.second.val = 0 ∧ d.nano.val = 0 -instance : HAdd Timestamp Duration Timestamp where - hAdd := add +/-- +Converts a `Duration` from a `Nanosecond.Offset` +-/ +@[inline] +def toSeconds (duration : Duration) : Second.Offset := + duration.second -instance : HSub Timestamp Timestamp Duration where - hSub := Std.Time.Timestamp.sub +/-- +Converts a `Duration` from a `Nanosecond.Offset` +-/ +@[inline] +def toNanoseconds (duration : Duration) : Nanosecond.Offset := + let nanos := duration.second.mul 1000000000 + let nanos := nanos + (.ofInt duration.nano.val) + nanos -end Duration +/-- +Converts a `Duration` to a `Minute.Offset` +-/ +@[inline] +def toMinutes (tm : Duration) : Minute.Offset := + tm.second.ediv 60 -namespace PlainDate +/-- +Converts a `Duration` to a `Day.Offset` +-/ +@[inline] +def toDays (tm : Duration) : Day.Offset := + tm.second.ediv 86400 /-- -Calculates the duration between a given `PlainDate` and a specified date. +Adds a `Nanosecond.Offset` to a `Duration` -/ -def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := - let date := date.toTimestamp - let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub date since +@[inline] +def addNanoseconds (t : Duration) (s : Nanosecond.Offset) : Duration := + t.add (ofNanoseconds s) -end PlainDate +/-- +Adds a `Nanosecond.Offset` to a `Duration` +-/ +@[inline] +def subNanoseconds (t : Duration) (s : Nanosecond.Offset) : Duration := + t.sub (ofNanoseconds s) -namespace PlainDateTime +/-- +Adds a `Second.Offset` to a `Duration` +-/ +@[inline] +def addSeconds (t : Duration) (s : Second.Offset) : Duration := + t.add (ofSeconds s) /-- -Calculates the duration between a given `PlainDateTime` and a specified date. +Subtracts a `Second.Offset` from a `Duration` -/ -def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := - let date := date.toTimestamp - let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub date since +@[inline] +def subSeconds (t : Duration) (s : Second.Offset) : Duration := + t.sub (ofSeconds s) -end PlainDateTime +/-- +Adds a `Minute.Offset` to a `Duration` +-/ +@[inline] +def addMinutes (t : Duration) (m : Minute.Offset) : Duration := + let seconds := m.mul 60 + t.addSeconds seconds -namespace DateTime +/-- +Subtracts a `Minute.Offset` from a `Duration` +-/ +@[inline] +def subMinutes (t : Duration) (m : Minute.Offset) : Duration := + let seconds := m.mul 60 + t.subSeconds seconds /-- -Calculates the duration between a given `DateTime` and a specified date. +Adds an `Hour.Offset` to a `Duration` -/ -def since [ToTimestamp α] (date : DateTime tz) (since : α) : Duration := - let date := date.toTimestamp - let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub date since +@[inline] +def addHours (t : Duration) (h : Hour.Offset) : Duration := + let seconds := h.mul 3600 + t.addSeconds seconds -end DateTime +/-- +Subtracts an `Hour.Offset` from a `Duration` +-/ +@[inline] +def subHours (t : Duration) (h : Hour.Offset) : Duration := + let seconds := h.mul 3600 + t.subSeconds seconds -namespace ZonedDateTime +/-- +Adds a `Day.Offset` to a `Duration` +-/ +@[inline] +def addDays (t : Duration) (d : Day.Offset) : Duration := + let seconds := d.mul 86400 + t.addSeconds seconds /-- -Calculates the duration between a given `ZonedDateTime` and a specified date. +Subtracts a `Day.Offset` from a `Duration` -/ -def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := - let date := date.toTimestamp - let since := ToTimestamp.toTimestamp since - Std.Time.Timestamp.sub date since +@[inline] +def subDays (t : Duration) (d : Day.Offset) : Duration := + let seconds := d.mul 86400 + t.subSeconds seconds + +instance : HAdd Duration Day.Offset Duration where + hAdd := addDays + +instance : HSub Duration Day.Offset Duration where + hSub := subDays + +instance : HAdd Duration Hour.Offset Duration where + hAdd := addHours + +instance : HSub Duration Hour.Offset Duration where + hSub := subHours -end ZonedDateTime +instance : HAdd Duration Minute.Offset Duration where + hAdd := addMinutes + +instance : HSub Duration Minute.Offset Duration where + hSub := subMinutes + +instance : HAdd Duration Second.Offset Duration where + hAdd := addSeconds + +instance : HSub Duration Second.Offset Duration where + hSub := subSeconds + +instance : HAdd Duration Nanosecond.Offset Duration where + hAdd := addNanoseconds + +instance : HSub Duration Nanosecond.Offset Duration where + hSub := subNanoseconds + +end Duration +end Time +end Std diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 6753300f9bcd..924b8f661be4 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -64,8 +64,15 @@ Converts a `DateTime` to a `PlainTime` def toPlainTime (dt : DateTime tz) : PlainTime := dt.date.get.time -end DateTime +/-- +Calculates the duration between a given `DateTime` and a specified date. +-/ +def since [ToTimestamp α] (date : DateTime tz) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch +end DateTime namespace ZonedDateTime /-- @@ -105,4 +112,13 @@ Converts a `ZonedDateTime` to a `PlainTime` def toPlainTime (dt : ZonedDateTime) : PlainTime := DateTime.toPlainTime dt.snd +/-- +Calculates the duration between a given `ZonedDateTime` and a specified date. +-/ +@[inline] +def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := + let date := date.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch + end ZonedDateTime diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index d0e578e99852..0a1cfb75b43c 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -191,7 +191,7 @@ def applyLeapSeconds (tm : Timestamp) (leapSeconds : ZoneRules) : Timestamp := I let leapSeconds := leapSeconds.leapSeconds for i in [:leapSeconds.size] do let leapSec := leapSeconds.get! i - if currentTime.second.val >= leapSec.transitionTime.val then + if currentTime.toSecondsSinceUnixEpoch.val >= leapSec.transitionTime.val then currentTime := tm.addSeconds (.ofInt leapSec.correction.val) return currentTime diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 654c5107f42d..0d6ee0584fb9 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -31,13 +31,13 @@ info: "09/05/1993 12:59:59" info: 753700087 -/ #guard_msgs in -#eval date₁.toLocalTimestamp.toSecondsSinceUnixEpoch +#eval date₁.toTimestampAssumingUTC.toSecondsSinceUnixEpoch /-- info: 736952399 -/ #guard_msgs in -#eval date₂.toLocalTimestamp.toSecondsSinceUnixEpoch +#eval date₂.toTimestampAssumingUTC.toSecondsSinceUnixEpoch /-- info: "09/05/1993 12:59:59" @@ -49,7 +49,7 @@ info: "09/05/1993 12:59:59" info: 736952399 -/ #guard_msgs in -#eval PlainDateTime.toLocalTimestamp date₂ |>.toSecondsSinceUnixEpoch +#eval PlainDateTime.toTimestampAssumingUTC date₂ |>.toSecondsSinceUnixEpoch /-- info: "16/08/2024" diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index d1306a9c8636..c826927def7d 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -59,13 +59,13 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: some (Tue, 30 Dec 1969 21:00:00 -0300) +info: some (1969-12-30 21:00:00 -0300) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules /-- -info: some (Mon, 10 Dec 2012 00:35:47 -0200) +info: some (2012-12-10 00:35:47 -0200) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From 1cccd6576658fe8d1bec6c433d512ad640b4d05f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 11:43:57 -0300 Subject: [PATCH 051/118] refactor: fix some function names --- src/Std/Time/DateTime.lean | 13 +++---------- 1 file changed, 3 insertions(+), 10 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index e45360bb4011..d9b33b73359a 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -69,14 +69,14 @@ namespace PlainDate Converts a `PlainDate` to a `Timestamp` -/ @[inline] -def toTimestamp (pdt : PlainDate) : Timestamp := +def toTimestampAssumingUTC (pdt : PlainDate) : Timestamp := Timestamp.ofPlainDate pdt /-- Calculates the duration between a given `PlainDate` and a specified date. -/ def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := - let date := date.toTimestamp + let date := date.toTimestampAssumingUTC let since := ToTimestamp.toTimestamp since Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch @@ -117,18 +117,11 @@ instance : ToTimestamp PlainDateTime where instance : ToTimestamp PlainDate where toTimestamp := Timestamp.ofPlainDate -/-- -Converts a `PlainDateTime` to a `Timestamp` --/ -@[inline] -def toTimestamp (pdt : PlainDateTime) : Timestamp := - Timestamp.ofPlainDateTime pdt - /-- Calculates the duration between a given `PlainDateTime` and a specified date. -/ def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := - let date := date.toTimestamp + let date := date.toTimestampAssumingUTC let since := ToTimestamp.toTimestamp since Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch From b1da2da2c975642af7ba9621481cada39f5c4a0c Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 11:57:09 -0300 Subject: [PATCH 052/118] feat: add some class operations that result in durations --- src/Std/Time.lean | 4 ++-- src/Std/Time/DateTime.lean | 6 ++++++ src/Std/Time/DateTime/Timestamp.lean | 3 +++ src/Std/Time/Duration.lean | 6 ++++++ src/Std/Time/Zoned/DateTime.lean | 3 +++ src/Std/Time/Zoned/ZonedDateTime.lean | 3 +++ tests/lean/run/timeClassOperations.lean | 6 ++++++ 7 files changed, 29 insertions(+), 2 deletions(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index ccda79bd1e4b..e777bcb54aff 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -100,7 +100,7 @@ These types offer precision down to the nanosecond level, useful for representin Combines date and time into a single representation, useful for precise timestamps and scheduling. - **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. -- **`Timestamp`**: Represents a point in time with second-level precision. It starts on the UNIX +- **`Timestamp`**: Represents a point in time with nanosecond-level precision. It starts on the UNIX epoch and it should be used when you receive or need to send timestamps to other systems. ## Zoned date and times. @@ -161,10 +161,10 @@ The supported formats include: In order to help the user build dates easily, there are a lot of macros available for creating dates. - **`date% YYYY-MM-DD`**: Defines a date in the `YYYY-MM-DD` format. -- **`date% HH:mm:ss`**: Defines a time in the `HH:mm:ss` format. - **`date% HH:mm:ss.sssssss`**: Defines a time in the `HH:mm:ss.sssssss` format, including fractional seconds. - **`date% YYYY-MM-DD:HH:mm:ss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss` format. - **`date% YYYY-MM-DDTHH:mm:ssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ssZ` format. +- **`time% HH:mm:ss`**: Defines a time in the `HH:mm:ss` format. - **`offset% +HH:mm`**: Defines a timezone offset in the format `+HH:mm`. - **`timezone% NAME/ID offset% +HH:mm`**: Defines a timezone with a name and an offset. - **`date-spec% "format"`**: Defines a date specification format at compile time using the provided format string. diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index d9b33b73359a..de0fa5dd619c 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -80,6 +80,9 @@ def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := let since := ToTimestamp.toTimestamp since Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch +instance : HSub PlainDate PlainDate Duration where + hSub x y := x.toTimestampAssumingUTC - y.toTimestampAssumingUTC + end PlainDate namespace PlainDateTime @@ -125,4 +128,7 @@ def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := let since := ToTimestamp.toTimestamp since Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch +instance : HSub PlainDateTime PlainDateTime Duration where + hSub x y := x.toTimestampAssumingUTC - y.toTimestampAssumingUTC + end PlainDateTime diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 33ba3a1b775c..d196bef7a71e 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -217,4 +217,7 @@ instance : HAdd Timestamp Nanosecond.Offset Timestamp where instance : HSub Timestamp Nanosecond.Offset Timestamp where hSub := subNanoseconds +instance : HSub Timestamp Timestamp Duration where + hSub x y := x.val - y.val + end Timestamp diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index f9b9e669a9bc..c41a37063553 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -292,6 +292,12 @@ instance : HAdd Duration Nanosecond.Offset Duration where instance : HSub Duration Nanosecond.Offset Duration where hSub := subNanoseconds +instance : HSub Duration Duration Duration where + hSub := sub + +instance : HAdd Duration Duration Duration where + hAdd := add + end Duration end Time end Std diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index ea8b5b93a02d..53dbd9009791 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -305,6 +305,9 @@ instance : HAdd (DateTime tz) (Nanosecond.Offset) (DateTime tz) where instance : HSub (DateTime tz) (Nanosecond.Offset) (DateTime tz) where hSub := subNanoseconds +instance : HSub (DateTime tz) (DateTime tz₁) Duration where + hSub x y := x.toTimestamp - y.toTimestamp + end DateTime end Time end Std diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 8234d73cd6db..0ac31c987acd 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -275,6 +275,9 @@ instance : HAdd ZonedDateTime (Second.Offset) ZonedDateTime where instance : HSub ZonedDateTime (Second.Offset) ZonedDateTime where hSub := subSeconds +instance : HSub ZonedDateTime ZonedDateTime Duration where + hSub x y := x.toTimestamp - y.toTimestamp + end ZonedDateTime end Time end Std diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index d9c7bf088212..e058854acd7f 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -152,3 +152,9 @@ info: 2000-01-20 06:01:01 -0300 -/ #guard_msgs in #eval datetimetz - (60 : Second.Offset) + +/-- +info: "3600.0s" +-/ +#guard_msgs in +#eval (date% 2000-12-20:00:0:00-03:00) - (date% 2000-12-20:00:00:00-02:00) From 7bc1a7cf7e4d20b93ee1038af5836cf06285b064 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 13:50:09 -0300 Subject: [PATCH 053/118] feat: adjust pretty printing and notation for dates --- src/Std/Time.lean | 14 +++-- src/Std/Time/DateTime/Timestamp.lean | 7 +-- src/Std/Time/Duration.lean | 4 +- src/Std/Time/Format.lean | 19 +++--- src/Std/Time/Format/Basic.lean | 17 ++++- src/Std/Time/Notation.lean | 9 +-- src/Std/Time/Time/PlainTime.lean | 14 ++++- src/Std/Time/Zoned/DateTime.lean | 7 +++ src/Std/Time/Zoned/ZonedDateTime.lean | 19 ++++++ tests/lean/run/timeClassOperations.lean | 68 +++++++++++++------- tests/lean/run/timeFormats.lean | 6 +- tests/lean/run/timeOperations.lean | 82 ++++++++++++------------- tests/lean/run/timeTzifParse.lean | 4 +- 13 files changed, 170 insertions(+), 100 deletions(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index e777bcb54aff..db39299eb649 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -94,12 +94,12 @@ These types provide precision down to the day level, useful for representing and ## Time These types offer precision down to the nanosecond level, useful for representing and manipulating time of day. -- **`PlainTime`**: Represents a time of day in the format `HH:mm:ss.SSSSSSSSS`. +- **`PlainTime`**: Represents a time of day in the format `HH:mm:ss.sssssssss`. ## Date and time Combines date and time into a single representation, useful for precise timestamps and scheduling. -- **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.SSSSSSSSS`. +- **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.sssssssss`. - **`Timestamp`**: Represents a point in time with nanosecond-level precision. It starts on the UNIX epoch and it should be used when you receive or need to send timestamps to other systems. @@ -146,6 +146,7 @@ The supported formats include: - `mm`: Two-digit minute (e.g., 01, 02) - `m`: One or two-digit minute (e.g., 1, 2) - **Second:** + - `sssssssss` : Nine-digit nanosseconds (e.g., 000000009) - `sss`: Three-digit milliseconds (e.g., 001, 202) - `ss`: Two-digit second (e.g., 01, 02) - `s`: One or two-digit second (e.g., 1, 2) @@ -159,12 +160,13 @@ The supported formats include: # Macros In order to help the user build dates easily, there are a lot of macros available for creating dates. +The `.sssssssss` can be ommited in most cases. - **`date% YYYY-MM-DD`**: Defines a date in the `YYYY-MM-DD` format. -- **`date% HH:mm:ss.sssssss`**: Defines a time in the `HH:mm:ss.sssssss` format, including fractional seconds. -- **`date% YYYY-MM-DD:HH:mm:ss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss` format. -- **`date% YYYY-MM-DDTHH:mm:ssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ssZ` format. -- **`time% HH:mm:ss`**: Defines a time in the `HH:mm:ss` format. +- **`date% HH:mm:ss:sssssssss`**: Defines a time in the `HH:mm:ss:sssssssss` format, including fractional seconds. +- **`date% YYYY-MM-DD:HH:mm:ss.sssssssss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. +- **`date% YYYY-MM-DDTHH:mm:ss:sssssssssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ss:sssssssssZ` format. +- **`time% HH:mm:ss:sssssssss`**: Defines a time in the `HH:mm:ss` format. - **`offset% +HH:mm`**: Defines a timezone offset in the format `+HH:mm`. - **`timezone% NAME/ID offset% +HH:mm`**: Defines a timezone with a name and an offset. - **`date-spec% "format"`**: Defines a date specification format at compile time using the provided format string. diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index d196bef7a71e..56670abd90cd 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -31,12 +31,7 @@ instance : OfNat Timestamp n where ofNat := ⟨OfNat.ofNat n⟩ instance : ToString Timestamp where - toString s := - let (sign, secs, nanos) := - if s.val.second.val > 0 then ("" ,s.val.second.val, s.val.nano.val) - else if s.val.second.val < 0 then ("-", -s.val.second.val, -s.val.nano.val) - else if s.val.nano.val < 0 then ("-", -s.val.second.val, -s.val.nano.val) else ("", s.val.second.val, s.val.nano.val) - sign ++ toString secs ++ "." ++ toString nanos ++ "s" + toString s := toString s.val instance : Repr Timestamp where reprPrec s := reprPrec (toString s) diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index c41a37063553..2d77ff04dfb4 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -40,7 +40,9 @@ instance : ToString Duration where if s.second.val > 0 then ("" ,s.second, s.nano.val) else if s.second.val < 0 then ("-", -s.second, -s.nano.val) else if s.nano.val < 0 then ("-", -s.second, -s.nano.val) else ("", s.second, s.nano.val) - sign ++ toString secs ++ "." ++ toString nanos ++ "s" + sign ++ toString secs ++ (if s.nano.val == 0 then "" else "." ++ (leftPad 9 <| toString nanos)) ++ "s" + where + leftPad n s := "".pushn '0' (n - s.length) ++ s instance : Repr Duration where reprPrec s := reprPrec (toString s) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index d5a6838c2e73..8bc34fd75b3e 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -40,19 +40,19 @@ def time12Hour : Format .any := date-spec% "HH:mm:ss AA" The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time in a 24-hour clock format. -/ -def time24Hour : Format .any := date-spec% "hh:mm:ss" +def time24Hour : Format .any := date-spec% "hh:mm:ss.sssssssss" /-- -The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD hh:mm:ss Z` for +The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD hh:mm:ss.sssssssss` for representing date, time, and time zone. -/ -def dateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD hh:mm:ss" +def dateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD hh:mm:ss.sssssssss" /-- -The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD hh:mm:ss Z` +The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD hh:mm:ss.sssssssss` for representing date, time, and time zone. -/ -def dateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD hh:mm:ss ZZZ" +def dateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD hh:mm:ss.sssssssss ZZZ" /-- The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used @@ -169,13 +169,13 @@ def format (time : PlainTime) (format : String) : String := Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `PlainTime`. -/ def fromTime24Hour (input : String) : Except String PlainTime := - Formats.time24Hour.parseBuilder (λh m s => PlainTime.ofHourMinuteSeconds? h.snd m s.snd) input + Formats.time24Hour.parseBuilder (λh m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input /-- Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss`). -/ def toTime24Hour (input : PlainTime) : String := - Formats.time24Hour.formatBuilder input.hour input.minute input.second + Formats.time24Hour.formatBuilder input.hour input.minute input.second input.nano /-- Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `PlainTime`. @@ -300,8 +300,11 @@ def format (date : PlainDateTime) (format : String) : String := | .DD | .D | .d => some date.day | .EEEE | .EEE => some date.date.weekday | .HH | .H => some date.time.hour + | .hh | .h => some date.time.hour + | .aa | .AA => some (if date.time.hour.snd.val > 12 then HourMarker.pm else HourMarker.am) | .mm | .m => some date.time.minute - | .sss => some (Internal.Bounded.LE.ofNat 0 (by decide)) + | .sssssssss => some date.time.nano + | .sss => some date.time.nano.toMillisecond | .ss | .s => some date.time.second | _ => none match res with diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 49935360ec69..2cfcfd936187 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -118,6 +118,11 @@ inductive Modifier -/ | m + /-- + `sssssssss` Nine-digit nanosecond (e.g., 000000001). + -/ + | sssssssss + /-- `sss`: Three-digit milliseconds (e.g., 001, 202). -/ @@ -246,6 +251,7 @@ private def parseModifier : Parser Modifier <|> pstring "aa" *> pure .aa <|> pstring "mm" *> pure .mm <|> pstring "m" *> pure .m + <|> pstring "sssssssss" *> pure .sssssssss <|> pstring "sss" *> pure .sss <|> pstring "ss" *> pure .ss <|> pstring "s" *> pure .s @@ -354,6 +360,7 @@ private def formatWithDate (date : DateTime tz) : Modifier → String | .aa => if date.hour.toInt < 12 then "am" else "pm" | .mm => s!"{leftPadNum 2 <| date.minute.toInt}" | .m => s!"{date.minute.toInt}" + | .sssssssss => s!"{leftPadNum 9 <| date.nanoseconds.toInt}" | .sss => s!"{leftPadNum 3 <| date.milliseconds.toInt}" | .ss => s!"{leftPadNum 2 <| date.second.toInt}" | .s => s!"{date.second.toInt}" @@ -378,6 +385,7 @@ private def SingleFormatType : Modifier → Type | .hh | .h | .HH | .H => Sigma Hour.Ordinal | .AA | .aa => HourMarker | .mm | .m => Minute.Ordinal + | .sssssssss => Nanosecond.Ordinal | .sss => Millisecond.Ordinal | .ss | .s => Sigma Second.Ordinal | .ZZZZZ | .ZZZZ | .ZZZ | .Z => Offset @@ -404,6 +412,7 @@ private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) | .aa => match data with | .am => "am" | .pm => "pm" | .mm => s!"{leftPadNum 2 data.toInt}" | .m => s!"{data.toInt}" + | .sssssssss => s!"{leftPadNum 9 data.toInt}" | .sss => s!"{leftPadNum 3 data.toInt}" | .ss => s!"{leftPadNum 2 data.snd.toInt}" | .s => s!"{data.snd.toInt}" @@ -551,6 +560,7 @@ private def parserWithFormat : (typ: Modifier) → Parser (SingleFormatType typ) | .aa => parserLowerHourMarker | .mm => transform Bounded.LE.ofInt twoDigit | .m => transform Bounded.LE.ofInt number + | .sssssssss => transform Bounded.LE.ofInt threeDigit | .sss => transform Bounded.LE.ofInt threeDigit | .ss => Sigma.mk true <$> transform Bounded.LE.ofInt twoDigit | .s => Sigma.mk true <$> transform Bounded.LE.ofInt number @@ -569,7 +579,7 @@ private structure DateBuilder where hour : Sigma Hour.Ordinal := ⟨true, 0⟩ minute : Minute.Ordinal := 0 second : Sigma Second.Ordinal := ⟨true, 0⟩ - millisecond : Millisecond.Ordinal := 0 + nanoseconds : Nanosecond.Ordinal := 0 marker : Option HourMarker := none private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except String aw.type := do @@ -585,7 +595,7 @@ private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except if let .isTrue p := inferInstanceAs (Decidable (ValidTime hour.snd builder.minute builder.second.snd)) then let build := DateTime.ofPlainDateTime { date := PlainDate.clip builder.year builder.month builder.day - time := PlainTime.mk hour builder.minute builder.second (.ofMillisecond builder.millisecond) p + time := PlainTime.mk hour builder.minute builder.second builder.nanoseconds p } match aw with @@ -605,7 +615,8 @@ private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : Sin | .EEEE | .EEE => .ok data | .hh | .h | .HH | .H => .ok { data with hour := value } | .mm | .m => .ok { data with minute := value } - | .sss => .ok { data with millisecond := value } + | .sssssssss => .ok { data with nanoseconds := value } + | .sss => .ok { data with nanoseconds := Nanosecond.Ordinal.ofMillisecond value } | .ss | .s => .ok { data with second := value } | .ZZZZZ | .ZZZZ | .ZZZ | .Z => .ok { data with tz := value } diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index d2b8d978ae53..06a9ebec86cf 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -76,15 +76,15 @@ Date in `HH-mm-ss` format. syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component : time /-- -Date in `HH-mm-ss.sssssss` format. +Date in `HH-mm-ss.sssssssss` format. -/ -syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component "." noWs date_component : time +syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component noWs ":" noWs date_component : time private def parseTime : TSyntax `time -> MacroM (TSyntax `term) + | `(time| $hour:date_component:$minute:date_component:$second:date_component:$nanos:date_component) => do + `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ $(← parseComponent (some 0) (some 999999999) nanos) (by decide)) | `(time| $hour:date_component:$minute:date_component:$second:date_component) => do `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ 0 (by decide)) - | `(time| $hour:date_component:$minute:date_component:$second:date_component.$nanos:date_component) => do - `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ $(← parseComponent (some 0) (some 999) nanos) (by decide)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- @@ -234,6 +234,7 @@ private def convertModifier : Modifier → MacroM (TSyntax `term) | .aa => `(Std.Time.Modifier.aa) | .mm => `(Std.Time.Modifier.mm) | .m => `(Std.Time.Modifier.m) + | .sssssssss => `(Std.Time.Modifier.sssssssss) | .sss => `(Std.Time.Modifier.sss) | .ss => `(Std.Time.Modifier.ss) | .s => `(Std.Time.Modifier.s) diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 9945bd28c55c..0c12aa0ce1e6 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -64,15 +64,15 @@ abbrev ValidTime (hour : Hour.Ordinal l) (minute : Minute.Ordinal) (second : Sec /-- Creates a `PlainTime` value from hours, minutes, and seconds. -/ -def ofHourMinuteSeconds (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (proof : ValidTime hour minute second) : PlainTime := - ⟨Sigma.mk leap₂ hour, minute, Sigma.mk leap second, 0, proof⟩ +def ofHourMinuteSecondsNano (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (nano : Nanosecond.Ordinal) (proof : ValidTime hour minute second) : PlainTime := + ⟨Sigma.mk leap₂ hour, minute, Sigma.mk leap second, nano, proof⟩ /-- Creates a `PlainTime` value from hours, minutes, and seconds. Return `none` if its invalid. -/ def ofHourMinuteSeconds? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : Option PlainTime := if h : ValidTime hour minute second - then some <| ofHourMinuteSeconds hour minute second h + then some <| ofHourMinuteSecondsNano hour minute second 0 h else none /-- @@ -84,6 +84,14 @@ def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Or exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr second.property.right)) x) exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) +/-- +Creates a `PlainTime` value from hours, minutes, and seconds. Return `none` if its invalid. +-/ +def ofHourMinuteSecondsNano? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (nano : Nanosecond.Ordinal) : Option PlainTime := + if h : ValidTime hour minute second + then some <| ofHourMinuteSecondsNano hour minute second nano h + else none + /-- Converts a `PlainTime` value to the total number of milliseconds. -/ diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 53dbd9009791..bb8df5c0df26 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -250,6 +250,13 @@ Getter for the `Milliseconds` inside of a `DateTime` def milliseconds (dt : DateTime tz) : Millisecond.Ordinal := dt.date.get.time.nano.toMillisecond +/-- +Getter for the `Nanosecond` inside of a `DateTime` +-/ +@[inline] +def nanoseconds (dt : DateTime tz) : Nanosecond.Ordinal := + dt.date.get.time.nano + /-- Gets the `Weekday` of a DateTime. -/ diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 0ac31c987acd..a08886e2f7a8 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -232,6 +232,19 @@ Subtract `Second.Offset` from a `ZonedDateTime`, adjusting the date if necessary def subSeconds (dt : ZonedDateTime) (seconds : Second.Offset) : ZonedDateTime := Sigma.mk dt.fst (dt.snd.subSeconds seconds) +/-- +Add `Nanosecond.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +-/ +def addNanoseconds (dt : ZonedDateTime) (nanoseconds : Nanosecond.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addNanoseconds nanoseconds) + +/-- +Subtract `Nanosecond.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +-/ +@[inline] +def subNanoseconds (dt : ZonedDateTime) (nanoseconds : Nanosecond.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subNanoseconds nanoseconds) + /-- Determines the era of the given `PlainDate` based on its year. -/ @@ -275,6 +288,12 @@ instance : HAdd ZonedDateTime (Second.Offset) ZonedDateTime where instance : HSub ZonedDateTime (Second.Offset) ZonedDateTime where hSub := subSeconds +instance : HAdd ZonedDateTime (Nanosecond.Offset) ZonedDateTime where + hAdd := addNanoseconds + +instance : HSub ZonedDateTime (Nanosecond.Offset) ZonedDateTime where + hSub := subNanoseconds + instance : HSub ZonedDateTime ZonedDateTime Duration where hSub x y := x.toTimestamp - y.toTimestamp diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index e058854acd7f..067deb7354c2 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -18,49 +18,49 @@ info: 1970-01-08 def datetime := date% 2000-01-20:03:02:01 /-- -info: 2000-01-20 04:02:01 +info: 2000-01-20 04:02:01.000000000 -/ #guard_msgs in #eval datetime + (1 : Hour.Offset) /-- -info: 2000-01-20 02:02:01 +info: 2000-01-20 02:02:01.000000000 -/ #guard_msgs in #eval datetime - (1 : Hour.Offset) /-- -info: 2000-01-20 03:12:01 +info: 2000-01-20 03:12:01.000000000 -/ #guard_msgs in #eval datetime + (10 : Minute.Offset) /-- -info: 2000-01-20 02:52:01 +info: 2000-01-20 02:52:01.000000000 -/ #guard_msgs in #eval datetime - (10 : Minute.Offset) /-- -info: 2000-01-20 03:03:01 +info: 2000-01-20 03:03:01.000000000 -/ #guard_msgs in #eval datetime + (60 : Second.Offset) /-- -info: 2000-01-20 03:01:01 +info: 2000-01-20 03:01:01.000000000 -/ #guard_msgs in #eval datetime - (60 : Second.Offset) /-- -info: 2000-01-21 03:02:01 +info: 2000-01-21 03:02:01.000000000 -/ #guard_msgs in #eval datetime + (1 : Day.Offset) /-- -info: 2000-01-19 03:02:01 +info: 2000-01-19 03:02:01.000000000 -/ #guard_msgs in #eval datetime - (1 : Day.Offset) @@ -68,37 +68,37 @@ info: 2000-01-19 03:02:01 def time := time% 13:02:01 /-- -info: 14:02:01 +info: 14:02:01.000000000 -/ #guard_msgs in #eval time + (1 : Hour.Offset) /-- -info: 12:02:01 +info: 12:02:01.000000000 -/ #guard_msgs in #eval time - (1 : Hour.Offset) /-- -info: 13:12:01 +info: 13:12:01.000000000 -/ #guard_msgs in #eval time + (10 : Minute.Offset) /-- -info: 12:52:01 +info: 12:52:01.000000000 -/ #guard_msgs in #eval time - (10 : Minute.Offset) /-- -info: 13:03:01 +info: 13:03:01.000000000 -/ #guard_msgs in #eval time + (60 : Second.Offset) /-- -info: 13:01:01 +info: 13:01:01.000000000 -/ #guard_msgs in #eval time - (60 : Second.Offset) @@ -106,55 +106,77 @@ info: 13:01:01 def datetimetz := date% 2000-01-20:03:02:01-03:00 /-- -info: 2000-01-22 06:02:01 -0300 +info: 2000-01-22 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz + (2 : Day.Offset) /-- -info: 2000-01-19 06:02:01 -0300 +info: 2000-01-19 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz - (1 : Day.Offset) /-- -info: 2000-01-20 07:02:01 -0300 +info: 2000-01-20 07:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz + (1 : Hour.Offset) /-- -info: 2000-01-20 05:02:01 -0300 +info: 2000-01-20 05:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz - (1 : Hour.Offset) /-- -info: 2000-01-20 06:12:01 -0300 +info: 2000-01-20 06:12:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz + (10 : Minute.Offset) /-- -info: 2000-01-20 05:52:01 -0300 +info: 2000-01-20 05:52:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz - (10 : Minute.Offset) /-- -info: 2000-01-20 06:03:01 -0300 +info: 2000-01-20 06:03:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz + (60 : Second.Offset) /-- -info: 2000-01-20 06:01:01 -0300 +info: 2000-01-20 06:01:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz - (60 : Second.Offset) /-- -info: "3600.0s" +info: "3600s" -/ #guard_msgs in #eval (date% 2000-12-20:00:0:00-03:00) - (date% 2000-12-20:00:00:00-02:00) + +def now := PlainDateTime.ofUTCTimestamp ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ +def after := PlainDateTime.ofUTCTimestamp ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ +def duration := after.since now + +/-- +info: "03:40:38.907328169" +-/ +#guard_msgs in +#eval now.format "HH:mm:ss.sssssssss" + +/-- +info: "03.40:39.907641224" +-/ +#guard_msgs in +#eval after.format "HH.mm:ss.sssssssss" + +/-- +info: "1.000313055s" +-/ +#guard_msgs in +#eval duration diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index b3db9e75e8c2..b1f6d884cb93 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -239,19 +239,19 @@ info: 2002-07-14 #eval date% 2002-07-14 /-- -info: 14:13:12 +info: 14:13:12.000000000 -/ #guard_msgs in #eval time% 14:13:12 /-- -info: 2002-07-14 14:13:12 +info: 2002-07-14 14:13:12.000000000 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12 /-- -info: 2002-07-14 14:13:12 +0900 +info: 2002-07-14 14:13:12.000000000 +0900 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12+09:00 diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index 4d94d6125e47..635aa472e868 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -36,97 +36,97 @@ info: 1969-12-20 def datetime := date% 2000-01-20:03:02:01 /-- -info: 2000-01-20 04:02:01 +info: 2000-01-20 04:02:01.000000000 -/ #guard_msgs in #eval datetime.addHours 1 /-- -info: 2000-01-20 02:02:01 +info: 2000-01-20 02:02:01.000000000 -/ #guard_msgs in #eval datetime.subHours 1 /-- -info: 2000-01-20 03:12:01 +info: 2000-01-20 03:12:01.000000000 -/ #guard_msgs in #eval datetime.addMinutes 10 /-- -info: 2000-01-20 02:52:01 +info: 2000-01-20 02:52:01.000000000 -/ #guard_msgs in #eval datetime.subMinutes 10 /-- -info: 2000-01-20 03:03:01 +info: 2000-01-20 03:03:01.000000000 -/ #guard_msgs in #eval datetime.addSeconds 60 /-- -info: 2000-01-20 03:01:01 +info: 2000-01-20 03:01:01.000000000 -/ #guard_msgs in #eval datetime.subSeconds 60 /-- -info: 2000-01-21 03:02:01 +info: 2000-01-21 03:02:01.000000000 -/ #guard_msgs in #eval datetime.addDays 1 /-- -info: 2000-01-19 03:02:01 +info: 2000-01-19 03:02:01.000000000 -/ #guard_msgs in #eval datetime.subDays 1 /-- -info: 2000-02-20 03:02:01 +info: 2000-02-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.addMonthsClip 1 /-- -info: 1999-12-20 03:02:01 +info: 1999-12-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.subMonthsClip 1 /-- -info: 2000-02-20 03:02:01 +info: 2000-02-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.addMonthsRollOver 1 /-- -info: 1999-12-20 03:02:01 +info: 1999-12-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.subMonthsRollOver 1 /-- -info: 2001-01-20 03:02:01 +info: 2001-01-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.addYearsClip 1 /-- -info: 1999-01-20 03:02:01 +info: 1999-01-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.subYearsClip 1 /-- -info: 2001-01-20 03:02:01 +info: 2001-01-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.addYearsRollOver 1 /-- -info: 1999-01-20 03:02:01 +info: 1999-01-20 03:02:01.000000000 -/ #guard_msgs in #eval datetime.subYearsRollOver 1 @@ -134,152 +134,152 @@ info: 1999-01-20 03:02:01 def time := time% 13:02:01 /-- -info: 14:02:01 +info: 14:02:01.000000000 -/ #guard_msgs in #eval time.addHours 1 /-- -info: 12:02:01 +info: 12:02:01.000000000 -/ #guard_msgs in #eval time.subHours 1 /-- -info: 13:12:01 +info: 13:12:01.000000000 -/ #guard_msgs in #eval time.addMinutes 10 /-- -info: 12:52:01 +info: 12:52:01.000000000 -/ #guard_msgs in #eval time.subMinutes 10 /-- -info: 13:03:01 +info: 13:03:01.000000000 -/ #guard_msgs in #eval time.addSeconds 60 /-- -info: 13:01:01 +info: 13:01:01.000000000 -/ #guard_msgs in #eval time.subSeconds 60 def datetimetz := date% 2000-01-20:03:02:01-03:00 /-- -info: 2000-01-22 06:02:01 -0300 +info: 2000-01-22 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addDays 2 /-- -info: 2000-01-19 06:02:01 -0300 +info: 2000-01-19 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.subDays 1 /-- -info: 2000-02-20 06:02:01 -0300 +info: 2000-02-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addMonthsClip 1 /-- -info: 1999-12-20 06:02:01 -0300 +info: 1999-12-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.subMonthsClip 1 /-- -info: 2000-02-20 06:02:01 -0300 +info: 2000-02-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addMonthsRollOver 1 /-- -info: 1999-12-20 06:02:01 -0300 +info: 1999-12-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.subMonthsRollOver 1 /-- -info: 2001-01-20 06:02:01 -0300 +info: 2001-01-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: 2001-01-20 06:02:01 -0300 +info: 2001-01-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: 2001-01-20 06:02:01 -0300 +info: 2001-01-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addYearsRollOver 1 /-- -info: 1999-01-20 06:02:01 -0300 +info: 1999-01-20 06:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.subYearsRollOver 1 /-- -info: 2000-01-20 07:02:01 -0300 +info: 2000-01-20 07:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addHours 1 /-- -info: 2000-01-20 05:02:01 -0300 +info: 2000-01-20 05:02:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.subHours 1 /-- -info: 2000-01-20 06:12:01 -0300 +info: 2000-01-20 06:12:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addMinutes 10 /-- -info: 2000-01-20 05:52:01 -0300 +info: 2000-01-20 05:52:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.subMinutes 10 /-- -info: 2000-01-20 06:03:01 -0300 +info: 2000-01-20 06:03:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.addSeconds 60 /-- -info: 2000-01-20 06:01:01 -0300 +info: 2000-01-20 06:01:01.000000000 -0300 -/ #guard_msgs in #eval datetimetz.subSeconds 60 /-- -info: "86400.0s" +info: "86400s" -/ #guard_msgs in #eval (date% 2000-1-20).since (date% 2000-1-19) /-- -info: "86399.0s" +info: "86399s" -/ #guard_msgs in #eval (date% 2000-1-20:0:0:0).since (date% 2000-1-19:0:0:1) /-- -info: "86399.0s" +info: "86399s" -/ #guard_msgs in #eval (date% 2000-1-20:0:0:0UTC).since (date% 2000-1-19:0:0:1UTC) diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index c826927def7d..f0213be67a41 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -59,13 +59,13 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: some (1969-12-30 21:00:00 -0300) +info: some (1969-12-30 21:00:00.000000000 -0300) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules /-- -info: some (2012-12-10 00:35:47 -0200) +info: some (2012-12-10 00:35:47.000000000 -0200) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From fe480070d3abdca8fb80b9ed87f5268a2dd4df09 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 28 Aug 2024 16:24:05 -0300 Subject: [PATCH 054/118] refactor: modified each formatter to use the a default format that can be parsed by the date% macro --- src/Std/Time/DateTime.lean | 20 ++++++ src/Std/Time/Format.lean | 85 ++++++++++++++++++++++--- src/Std/Time/Zoned.lean | 19 ++++++ tests/lean/run/timeClassOperations.lean | 44 ++++++------- tests/lean/run/timeFormats.lean | 6 +- tests/lean/run/timeOperations.lean | 76 +++++++++++----------- tests/lean/run/timeTzifParse.lean | 4 +- 7 files changed, 179 insertions(+), 75 deletions(-) diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index de0fa5dd619c..6cfea6996854 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -74,7 +74,17 @@ def toTimestampAssumingUTC (pdt : PlainDate) : Timestamp := /-- Calculates the duration between a given `PlainDate` and a specified date. + +## Example + +```lean +def example : Duration := + let startDate := date% 2023-1-1 + let endDate := date% 2023-3-15 + endDate.since startDate +``` -/ +@[inline] def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := let date := date.toTimestampAssumingUTC let since := ToTimestamp.toTimestamp since @@ -122,7 +132,17 @@ instance : ToTimestamp PlainDate where /-- Calculates the duration between a given `PlainDateTime` and a specified date. + +## Example + +```lean +example : Duration := + let startDate := date% 2023-1-1:05:10:20 + let endDate := date% 2023-3-15:05:10:20 + endDate.since startDate +``` -/ +@[inline] def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := let date := date.toTimestampAssumingUTC let since := ToTimestamp.toTimestamp since diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 8bc34fd75b3e..c1629a7f0f79 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -40,19 +40,40 @@ def time12Hour : Format .any := date-spec% "HH:mm:ss AA" The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time in a 24-hour clock format. -/ -def time24Hour : Format .any := date-spec% "hh:mm:ss.sssssssss" +def time24Hour : Format .any := date-spec% "hh:mm:ss:sssssssss" /-- The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD hh:mm:ss.sssssssss` for representing date, time, and time zone. -/ -def dateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD hh:mm:ss.sssssssss" +def dateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD:hh:mm:ss.sssssssss" /-- The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD hh:mm:ss.sssssssss` for representing date, time, and time zone. -/ -def dateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD hh:mm:ss.sssssssss ZZZ" +def dateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD:hh:mm:ss.sssssssssZZZ" + +/-- +The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time +in a 24-hour clock format. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanTime24Hour : Format .any := date-spec% "hh:mm:ss:sssssssss" + +/-- +The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD hh:mm:ss:sssssssss` for +representing date, time, and time zone. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanDateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD:hh:mm:ss:sssssssss" + +/-- +The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD hh:mm:ss:sssssssss` +for representing date, time, and time zone. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanDateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD:hh:mm:ss:sssssssssZZZZZ" /-- The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used @@ -166,17 +187,29 @@ def format (time : PlainTime) (format : String) : String := | none => "invalid time" /-- -Parses a time string in the 24-hour format (`hh:mm:ss`) and returns a `PlainTime`. +Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. -/ def fromTime24Hour (input : String) : Except String PlainTime := Formats.time24Hour.parseBuilder (λh m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input /-- -Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss`). +Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). -/ def toTime24Hour (input : PlainTime) : String := Formats.time24Hour.formatBuilder input.hour input.minute input.second input.nano +/-- +Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. +-/ +def fromLeanTime24Hour (input : String) : Except String PlainTime := + Formats.leanTime24Hour.parseBuilder (λh m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input + +/-- +Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). +-/ +def toLeanTime24Hour (input : PlainTime) : String := + Formats.leanTime24Hour.formatBuilder input.hour input.minute input.second input.nano + /-- Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `PlainTime`. -/ @@ -201,7 +234,7 @@ def parse (input : String) : Except String PlainTime := <|> fromTime24Hour input instance : ToString PlainTime where - toString := toTime24Hour + toString := toLeanTime24Hour instance : Repr PlainTime where reprPrec data := Repr.addAppParen (toString data) @@ -263,11 +296,23 @@ def fromDateTimeWithZoneString (input : String) : Except String ZonedDateTime := Formats.dateTimeWithZone.parse input /-- -Formats a `DateTime` value into a `DateTimeWithZone` format string. +Formats a `DateTime` value into a simple date time with timezone string. -/ def toDateTimeWithZoneString (pdt : ZonedDateTime) : String := Formats.dateTimeWithZone.format pdt.snd +/-- +Parses a `String` in the `DateTimeWithZone` format and returns a `DateTime` object in the GMT time zone. +-/ +def fromLeanDateTimeWithZoneString (input : String) : Except String ZonedDateTime := + Formats.leanDateTimeWithZone.parse input + +/-- +Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notationg. +-/ +def toLeanDateTimeWithZoneString (pdt : ZonedDateTime) : String := + Formats.leanDateTimeWithZone.format pdt.snd + /-- Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. -/ @@ -277,7 +322,7 @@ def parse (input : String) : Except String ZonedDateTime := <|> fromRFC850String input instance : ToString ZonedDateTime where - toString := toDateTimeWithZoneString + toString := toLeanDateTimeWithZoneString instance : Repr ZonedDateTime where reprPrec data := Repr.addAppParen (toString data) @@ -350,15 +395,29 @@ Formats a `PlainDateTime` value into a `DateTime` format string. def toDateTimeString (pdt : PlainDateTime) : String := Formats.dateTime24Hour.format (DateTime.ofPlainDateTime pdt .UTC) +/-- +Parses a `String` in the `DateTime` format and returns a `PlainDateTime`. +-/ +def fromLeanDateTimeString (input : String) : Except String PlainDateTime := + Formats.leanDateTime24Hour.parse input + |>.map DateTime.toPlainDateTime + +/-- +Formats a `PlainDateTime` value into a `DateTime` format string. +-/ +def toLeanDateTimeString (pdt : PlainDateTime) : String := + Formats.leanDateTime24Hour.format (DateTime.ofPlainDateTime pdt .UTC) + /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `PlainDateTime`. -/ def parse (date : String) : Except String PlainDateTime := fromAscTimeString date <|> fromLongDateFormatString date + <|> fromDateTimeString date instance : ToString PlainDateTime where - toString := toDateTimeString + toString := toLeanDateTimeString instance : Repr PlainDateTime where reprPrec data := Repr.addAppParen (toString data) @@ -424,6 +483,12 @@ Formats a `DateTime` value into a `DateTimeWithZone` format string. def toDateTimeWithZoneString (pdt : DateTime tz) : String := Formats.dateTimeWithZone.format pdt +/-- +Formats a `DateTime` value into a `DateTimeWithZone` format string that can be parsed by `date%`. +-/ +def toLeanDateTimeWithZoneString (pdt : DateTime tz) : String := + Formats.leanDateTimeWithZone.format pdt + /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `DateTime`. -/ @@ -432,7 +497,7 @@ def parse (date : String) : Except String (DateTime .GMT) := <|> fromLongDateFormatString date instance : ToString (DateTime tz) where - toString := toDateTimeWithZoneString + toString := toLeanDateTimeWithZoneString instance : Repr (DateTime tz) where reprPrec data := Repr.addAppParen (toString data) diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 924b8f661be4..73b01d561af5 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -66,7 +66,17 @@ def toPlainTime (dt : DateTime tz) : PlainTime := /-- Calculates the duration between a given `DateTime` and a specified date. + +## Example + +```lean +example : Duration := + let startDate := date% 2023-1-1:05:10:20UTC + let endDate := date% 2023-3-15:05:10:20UTC + endDate.since startDate +``` -/ +@[inline] def since [ToTimestamp α] (date : DateTime tz) (since : α) : Duration := let date := date.toTimestamp let since := ToTimestamp.toTimestamp since @@ -114,6 +124,15 @@ def toPlainTime (dt : ZonedDateTime) : PlainTime := /-- Calculates the duration between a given `ZonedDateTime` and a specified date. + +## Example + +```lean +def example : Duration := + let startDate := date% 2023-1-1:05:10:20UTC + let endDate := date% 2023-3-15:05:10:20UTC + endDate.since startDate +``` -/ @[inline] def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index 067deb7354c2..4132ee84c555 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -18,49 +18,49 @@ info: 1970-01-08 def datetime := date% 2000-01-20:03:02:01 /-- -info: 2000-01-20 04:02:01.000000000 +info: 2000-01-20:04:02:01:000000000 -/ #guard_msgs in #eval datetime + (1 : Hour.Offset) /-- -info: 2000-01-20 02:02:01.000000000 +info: 2000-01-20:02:02:01:000000000 -/ #guard_msgs in #eval datetime - (1 : Hour.Offset) /-- -info: 2000-01-20 03:12:01.000000000 +info: 2000-01-20:03:12:01:000000000 -/ #guard_msgs in #eval datetime + (10 : Minute.Offset) /-- -info: 2000-01-20 02:52:01.000000000 +info: 2000-01-20:02:52:01:000000000 -/ #guard_msgs in #eval datetime - (10 : Minute.Offset) /-- -info: 2000-01-20 03:03:01.000000000 +info: 2000-01-20:03:03:01:000000000 -/ #guard_msgs in #eval datetime + (60 : Second.Offset) /-- -info: 2000-01-20 03:01:01.000000000 +info: 2000-01-20:03:01:01:000000000 -/ #guard_msgs in #eval datetime - (60 : Second.Offset) /-- -info: 2000-01-21 03:02:01.000000000 +info: 2000-01-21:03:02:01:000000000 -/ #guard_msgs in #eval datetime + (1 : Day.Offset) /-- -info: 2000-01-19 03:02:01.000000000 +info: 2000-01-19:03:02:01:000000000 -/ #guard_msgs in #eval datetime - (1 : Day.Offset) @@ -68,37 +68,37 @@ info: 2000-01-19 03:02:01.000000000 def time := time% 13:02:01 /-- -info: 14:02:01.000000000 +info: 14:02:01:000000000 -/ #guard_msgs in #eval time + (1 : Hour.Offset) /-- -info: 12:02:01.000000000 +info: 12:02:01:000000000 -/ #guard_msgs in #eval time - (1 : Hour.Offset) /-- -info: 13:12:01.000000000 +info: 13:12:01:000000000 -/ #guard_msgs in #eval time + (10 : Minute.Offset) /-- -info: 12:52:01.000000000 +info: 12:52:01:000000000 -/ #guard_msgs in #eval time - (10 : Minute.Offset) /-- -info: 13:03:01.000000000 +info: 13:03:01:000000000 -/ #guard_msgs in #eval time + (60 : Second.Offset) /-- -info: 13:01:01.000000000 +info: 13:01:01:000000000 -/ #guard_msgs in #eval time - (60 : Second.Offset) @@ -106,49 +106,49 @@ info: 13:01:01.000000000 def datetimetz := date% 2000-01-20:03:02:01-03:00 /-- -info: 2000-01-22 06:02:01.000000000 -0300 +info: 2000-01-22:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (2 : Day.Offset) /-- -info: 2000-01-19 06:02:01.000000000 -0300 +info: 2000-01-19:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (1 : Day.Offset) /-- -info: 2000-01-20 07:02:01.000000000 -0300 +info: 2000-01-20:07:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (1 : Hour.Offset) /-- -info: 2000-01-20 05:02:01.000000000 -0300 +info: 2000-01-20:05:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (1 : Hour.Offset) /-- -info: 2000-01-20 06:12:01.000000000 -0300 +info: 2000-01-20:06:12:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (10 : Minute.Offset) /-- -info: 2000-01-20 05:52:01.000000000 -0300 +info: 2000-01-20:05:52:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (10 : Minute.Offset) /-- -info: 2000-01-20 06:03:01.000000000 -0300 +info: 2000-01-20:06:03:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (60 : Second.Offset) /-- -info: 2000-01-20 06:01:01.000000000 -0300 +info: 2000-01-20:06:01:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (60 : Second.Offset) diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index b1f6d884cb93..8cc2584022a8 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -239,19 +239,19 @@ info: 2002-07-14 #eval date% 2002-07-14 /-- -info: 14:13:12.000000000 +info: 14:13:12:000000000 -/ #guard_msgs in #eval time% 14:13:12 /-- -info: 2002-07-14 14:13:12.000000000 +info: 2002-07-14:14:13:12:000000000 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12 /-- -info: 2002-07-14 14:13:12.000000000 +0900 +info: 2002-07-14:14:13:12:000000000+09:00 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12+09:00 diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index 635aa472e868..dbef85f37d99 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -36,97 +36,97 @@ info: 1969-12-20 def datetime := date% 2000-01-20:03:02:01 /-- -info: 2000-01-20 04:02:01.000000000 +info: 2000-01-20:04:02:01:000000000 -/ #guard_msgs in #eval datetime.addHours 1 /-- -info: 2000-01-20 02:02:01.000000000 +info: 2000-01-20:02:02:01:000000000 -/ #guard_msgs in #eval datetime.subHours 1 /-- -info: 2000-01-20 03:12:01.000000000 +info: 2000-01-20:03:12:01:000000000 -/ #guard_msgs in #eval datetime.addMinutes 10 /-- -info: 2000-01-20 02:52:01.000000000 +info: 2000-01-20:02:52:01:000000000 -/ #guard_msgs in #eval datetime.subMinutes 10 /-- -info: 2000-01-20 03:03:01.000000000 +info: 2000-01-20:03:03:01:000000000 -/ #guard_msgs in #eval datetime.addSeconds 60 /-- -info: 2000-01-20 03:01:01.000000000 +info: 2000-01-20:03:01:01:000000000 -/ #guard_msgs in #eval datetime.subSeconds 60 /-- -info: 2000-01-21 03:02:01.000000000 +info: 2000-01-21:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addDays 1 /-- -info: 2000-01-19 03:02:01.000000000 +info: 2000-01-19:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subDays 1 /-- -info: 2000-02-20 03:02:01.000000000 +info: 2000-02-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addMonthsClip 1 /-- -info: 1999-12-20 03:02:01.000000000 +info: 1999-12-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subMonthsClip 1 /-- -info: 2000-02-20 03:02:01.000000000 +info: 2000-02-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addMonthsRollOver 1 /-- -info: 1999-12-20 03:02:01.000000000 +info: 1999-12-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subMonthsRollOver 1 /-- -info: 2001-01-20 03:02:01.000000000 +info: 2001-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addYearsClip 1 /-- -info: 1999-01-20 03:02:01.000000000 +info: 1999-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subYearsClip 1 /-- -info: 2001-01-20 03:02:01.000000000 +info: 2001-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addYearsRollOver 1 /-- -info: 1999-01-20 03:02:01.000000000 +info: 1999-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subYearsRollOver 1 @@ -134,134 +134,134 @@ info: 1999-01-20 03:02:01.000000000 def time := time% 13:02:01 /-- -info: 14:02:01.000000000 +info: 14:02:01:000000000 -/ #guard_msgs in #eval time.addHours 1 /-- -info: 12:02:01.000000000 +info: 12:02:01:000000000 -/ #guard_msgs in #eval time.subHours 1 /-- -info: 13:12:01.000000000 +info: 13:12:01:000000000 -/ #guard_msgs in #eval time.addMinutes 10 /-- -info: 12:52:01.000000000 +info: 12:52:01:000000000 -/ #guard_msgs in #eval time.subMinutes 10 /-- -info: 13:03:01.000000000 +info: 13:03:01:000000000 -/ #guard_msgs in #eval time.addSeconds 60 /-- -info: 13:01:01.000000000 +info: 13:01:01:000000000 -/ #guard_msgs in #eval time.subSeconds 60 def datetimetz := date% 2000-01-20:03:02:01-03:00 /-- -info: 2000-01-22 06:02:01.000000000 -0300 +info: 2000-01-22:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addDays 2 /-- -info: 2000-01-19 06:02:01.000000000 -0300 +info: 2000-01-19:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subDays 1 /-- -info: 2000-02-20 06:02:01.000000000 -0300 +info: 2000-02-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMonthsClip 1 /-- -info: 1999-12-20 06:02:01.000000000 -0300 +info: 1999-12-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMonthsClip 1 /-- -info: 2000-02-20 06:02:01.000000000 -0300 +info: 2000-02-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMonthsRollOver 1 /-- -info: 1999-12-20 06:02:01.000000000 -0300 +info: 1999-12-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMonthsRollOver 1 /-- -info: 2001-01-20 06:02:01.000000000 -0300 +info: 2001-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: 2001-01-20 06:02:01.000000000 -0300 +info: 2001-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: 2001-01-20 06:02:01.000000000 -0300 +info: 2001-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsRollOver 1 /-- -info: 1999-01-20 06:02:01.000000000 -0300 +info: 1999-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subYearsRollOver 1 /-- -info: 2000-01-20 07:02:01.000000000 -0300 +info: 2000-01-20:07:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addHours 1 /-- -info: 2000-01-20 05:02:01.000000000 -0300 +info: 2000-01-20:05:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subHours 1 /-- -info: 2000-01-20 06:12:01.000000000 -0300 +info: 2000-01-20:06:12:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMinutes 10 /-- -info: 2000-01-20 05:52:01.000000000 -0300 +info: 2000-01-20:05:52:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMinutes 10 /-- -info: 2000-01-20 06:03:01.000000000 -0300 +info: 2000-01-20:06:03:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addSeconds 60 /-- -info: 2000-01-20 06:01:01.000000000 -0300 +info: 2000-01-20:06:01:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subSeconds 60 diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index f0213be67a41..98e7e7fe5ff0 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -59,13 +59,13 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: some (1969-12-30 21:00:00.000000000 -0300) +info: some (1969-12-30:21:00:00:000000000-03:00) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules /-- -info: some (2012-12-10 00:35:47.000000000 -0200) +info: some (2012-12-10:00:35:47:000000000-02:00) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From f85c0cf147b11c87e3a557acd287f6593e3051fb Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 29 Aug 2024 17:42:03 -0300 Subject: [PATCH 055/118] feat: add a bunch of conversion functions and fixes --- src/Std/Time.lean | 4 +- src/Std/Time/Date.lean | 1 - src/Std/Time/Date/PlainDate.lean | 23 ++ src/Std/Time/Date/Unit/Basic.lean | 42 ++- src/Std/Time/Date/Unit/Day.lean | 49 +++- src/Std/Time/Date/Unit/Month.lean | 3 + .../Date/Unit/{WeekOfYear.lean => Week.lean} | 63 +++- src/Std/Time/Date/WeekDate.lean | 53 ---- src/Std/Time/DateTime/PlainDateTime.lean | 20 ++ src/Std/Time/DateTime/Timestamp.lean | 22 +- src/Std/Time/Duration.lean | 101 +++---- src/Std/Time/Format.lean | 34 ++- src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Notation.lean | 2 +- src/Std/Time/Time/HourMarker.lean | 1 + src/Std/Time/Time/PlainTime.lean | 1 - src/Std/Time/Time/Unit/Basic.lean | 274 +++++++++++++++++- src/Std/Time/Time/Unit/Hour.lean | 23 +- src/Std/Time/Time/Unit/Millisecond.lean | 3 + src/Std/Time/Time/Unit/Minute.lean | 14 +- src/Std/Time/Time/Unit/Nanosecond.lean | 12 + src/Std/Time/Time/Unit/Second.lean | 5 +- src/Std/Time/Zoned.lean | 4 +- src/Std/Time/Zoned/Database/Basic.lean | 6 +- src/Std/Time/Zoned/DateTime.lean | 70 +++-- src/Std/Time/Zoned/TimeZone.lean | 2 +- src/Std/Time/Zoned/ZonedDateTime.lean | 47 ++- src/runtime/io.cpp | 2 +- tests/lean/run/timeClassOperations.lean | 56 ++-- tests/lean/run/timeFormats.lean | 16 +- tests/lean/run/timeOperations.lean | 111 ++++--- tests/lean/run/timeParse.lean | 6 +- tests/lean/run/timeTzifParse.lean | 4 +- 33 files changed, 779 insertions(+), 297 deletions(-) rename src/Std/Time/Date/Unit/{WeekOfYear.lean => Week.lean} (54%) delete mode 100644 src/Std/Time/Date/WeekDate.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index db39299eb649..0a18451a6067 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -17,7 +17,7 @@ namespace Time /-! # Time -The Lean4 API for date, time, and duration functionalities, following the ISO standards. +The Lean4 API for date, time, and duration functionalities, following the ISO8601 standards. # Overview @@ -164,7 +164,7 @@ The `.sssssssss` can be ommited in most cases. - **`date% YYYY-MM-DD`**: Defines a date in the `YYYY-MM-DD` format. - **`date% HH:mm:ss:sssssssss`**: Defines a time in the `HH:mm:ss:sssssssss` format, including fractional seconds. -- **`date% YYYY-MM-DD:HH:mm:ss.sssssssss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. +- **`date% YYYY-MM-DD:HH:mm:ss:sssssssss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. - **`date% YYYY-MM-DDTHH:mm:ss:sssssssssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ss:sssssssssZ` format. - **`time% HH:mm:ss:sssssssss`**: Defines a time in the `HH:mm:ss` format. - **`offset% +HH:mm`**: Defines a timezone offset in the format `+HH:mm`. diff --git a/src/Std/Time/Date.lean b/src/Std/Time/Date.lean index e3df6fd60a73..8caabbe8e108 100644 --- a/src/Std/Time/Date.lean +++ b/src/Std/Time/Date.lean @@ -5,5 +5,4 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Date.Basic -import Std.Time.Date.WeekDate import Std.Time.Date.PlainDate diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 6eb289ee1321..d0e289c5fd71 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -184,6 +184,23 @@ Subtracts a given number of days from a `PlainDate`. def subDays (date : PlainDate) (days : Day.Offset) : PlainDate := addDays date (-days) +/-- +Adds a given number of weeks to a `PlainDate`. +-/ + +@[inline] +def addWeeks (date : PlainDate) (weeks : Week.Offset) : PlainDate := + let dateDays := date.toDaysSinceUNIXEpoch + let daysToAdd := weeks.toDays + ofDaysSinceUNIXEpoch (dateDays + daysToAdd) + +/-- +Subtracts a given number of weeks from a `PlainDate`. +-/ +@[inline] +def subWeeks (date : PlainDate) (weeks : Week.Offset) : PlainDate := + addWeeks date (-weeks) + /-- Adds a given number of months to a `PlainDate`, clipping the day to the last valid day of the month. -/ @@ -284,6 +301,12 @@ instance : HAdd PlainDate Day.Offset PlainDate where instance : HSub PlainDate Day.Offset PlainDate where hSub := subDays +instance : HAdd PlainDate Week.Offset PlainDate where + hAdd := addWeeks + +instance : HSub PlainDate Week.Offset PlainDate where + hSub := subWeeks + end PlainDate end Time end Std diff --git a/src/Std/Time/Date/Unit/Basic.lean b/src/Std/Time/Date/Unit/Basic.lean index d64675eddd1f..b16bb3219082 100644 --- a/src/Std/Time/Date/Unit/Basic.lean +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -8,7 +8,7 @@ import Std.Time.Date.Unit.Day import Std.Time.Date.Unit.Month import Std.Time.Date.Unit.Year import Std.Time.Date.Unit.Weekday -import Std.Time.Date.Unit.WeekOfYear +import Std.Time.Date.Unit.Week /-! This module defines various units used for measuring, counting, and converting between days, months, @@ -17,3 +17,43 @@ years, weekdays, and weeks of the year. The units are organized into types representing these time-related concepts, with operations provided to facilitate conversions and manipulations between them. -/ + +namespace Std +namespace Time +open Internal + +namespace Day.Offset + +/-- +Convert `Week.Offset` into `Day.Offset`. +-/ +@[inline] +def ofWeeks (week : Week.Offset) : Day.Offset := + week.mul 7 + +/-- +Convert `Day.Offset` into `Week.Offset`. +-/ +@[inline] +def toWeeks (day : Day.Offset) : Week.Offset := + day.ediv 7 + +end Day.Offset + +namespace Week.Offset + +/-- +Convert `Week.Offset` into `Day.Offset`. +-/ +@[inline] +def toDays (week : Week.Offset) : Day.Offset := + week.mul 7 + +/-- +Convert `Day.Offset` into `Week.Offset`. +-/ +@[inline] +def ofDays (day : Day.Offset) : Week.Offset := + day.ediv 7 + +end Week.Offset diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 51ada9fc337f..95e3a21eb005 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -19,6 +19,9 @@ set_option linter.all true def Ordinal := Bounded.LE 1 31 deriving Repr, BEq, LE, LT +instance : ToString Ordinal where + toString x := toString x.val + instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (30 : Nat))) n) @@ -101,25 +104,39 @@ def ofInt (data : Int) : Offset := UnitVal.mk data /-- -Convert `Day.Offset` into `Second.Offset`. +Convert `Day.Offset` into `Nanosecond.Offset`. -/ @[inline] -def toSeconds (days : Offset) : Second.Offset := - days.mul 86400 +def toNanoseconds (days : Offset) : Nanosecond.Offset := + days.mul 86400000000000 /-- -Convert `Day.Offset` into `Minute.Offset`. +Convert `Nanosecond.Offset` into `Day.Offset`. -/ @[inline] -def toMinutes (days : Offset) : Minute.Offset := - days.mul 1440 +def ofNanoseconds (ns : Nanosecond.Offset) : Offset := + ns.ediv 86400000000000 /-- -Convert `Day.Offset` into `Hour.Offset`. +Convert `Day.Offset` into `Millisecond.Offset`. -/ @[inline] -def toHours (days : Offset) : Hour.Offset := - days.mul 24 +def toMilliseconds (days : Offset) : Millisecond.Offset := + days.mul 86400000 + +/-- +Convert `Millisecond.Offset` into `Day.Offset`. +-/ +@[inline] +def ofMilliseconds (ms : Millisecond.Offset) : Offset := + ms.ediv 86400000 + +/-- +Convert `Day.Offset` into `Second.Offset`. +-/ +@[inline] +def toSeconds (days : Offset) : Second.Offset := + days.mul 86400 /-- Convert `Second.Offset` into `Day.Offset`. @@ -128,6 +145,13 @@ Convert `Second.Offset` into `Day.Offset`. def ofSeconds (secs : Second.Offset) : Offset := secs.ediv 86400 +/-- +Convert `Day.Offset` into `Minute.Offset`. +-/ +@[inline] +def toMinutes (days : Offset) : Minute.Offset := + days.mul 1440 + /-- Convert `Minute.Offset` into `Day.Offset`. -/ @@ -135,6 +159,13 @@ Convert `Minute.Offset` into `Day.Offset`. def ofMinutes (minutes : Minute.Offset) : Offset := minutes.ediv 1440 +/-- +Convert `Day.Offset` into `Hour.Offset`. +-/ +@[inline] +def toHours (days : Offset) : Hour.Offset := + days.mul 24 + /-- Convert `Hour.Offset` into `Day.Offset`. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 648cc5651d12..3b7c36bf30fd 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -21,6 +21,9 @@ set_option linter.all true def Ordinal := Bounded.LE 1 12 deriving Repr, BEq, LE +instance : ToString Ordinal where + toString x := toString x.val + instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (11 : Nat))) n) diff --git a/src/Std/Time/Date/Unit/WeekOfYear.lean b/src/Std/Time/Date/Unit/Week.lean similarity index 54% rename from src/Std/Time/Date/Unit/WeekOfYear.lean rename to src/Std/Time/Date/Unit/Week.lean index 0df5a2db38a0..61d63f850cb5 100644 --- a/src/Std/Time/Date/Unit/WeekOfYear.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -9,7 +9,7 @@ import Std.Time.Date.Unit.Day namespace Std namespace Time -namespace WeekOfYear +namespace Week open Std.Internal open Internal @@ -21,6 +21,9 @@ set_option linter.all true def Ordinal := Bounded.LE 1 53 deriving Repr, BEq, LE, LT +instance : ToString Ordinal where + toString x := toString x.val + instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (52 : Nat))) n) @@ -55,8 +58,64 @@ Converts an `Ordinal` to an `Offset`. def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val +/-- +Convert `Week.Offset` into `Second.Offset`. +-/ +@[inline] +def toSeconds (weeks : Week.Offset) : Second.Offset := + weeks.mul 604800 + +/-- +Convert `Second.Offset` into `Week.Offset`. +-/ +@[inline] +def ofSeconds (secs : Second.Offset) : Week.Offset := + secs.ediv 604800 + +/-- +Convert `Week.Offset` into `Minute.Offset`. +-/ +@[inline] +def toMinutes (weeks : Week.Offset) : Minute.Offset := + weeks.mul 10080 + +/-- +Convert `Minute.Offset` into `Week.Offset`. +-/ +@[inline] +def ofMinutes (minutes : Minute.Offset) : Week.Offset := + minutes.ediv 10080 + +/-- +Convert `Week.Offset` into `Hour.Offset`. +-/ +@[inline] +def toHours (weeks : Week.Offset) : Hour.Offset := + weeks.mul 168 + +/-- +Convert `Hour.Offset` into `Week.Offset`. +-/ +@[inline] +def ofHours (hours : Hour.Offset) : Week.Offset := + hours.ediv 168 + +/-- +Convert `Week.Offset` into `Day.Offset`. +-/ +@[inline] +def toDays (weeks : Week.Offset) : Day.Offset := + weeks.mul 7 + +/-- +Convert `Day.Offset` into `Week.Offset`. +-/ +@[inline] +def ofDays (hours : Day.Offset) : Week.Offset := + hours.ediv 7 + end Ordinal -end WeekOfYear +end Week end Time end Std diff --git a/src/Std/Time/Date/WeekDate.lean b/src/Std/Time/Date/WeekDate.lean deleted file mode 100644 index a8de77043822..000000000000 --- a/src/Std/Time/Date/WeekDate.lean +++ /dev/null @@ -1,53 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Std.Time.Date.Unit.Year -import Std.Time.Date.Unit.WeekOfYear -import Std.Time.Date.PlainDate - -namespace Std -namespace Time -open Internal - -set_option linter.all true - -/-- -`WeekDate` represents a date using a combination of a week of the year and the year. --/ -structure WeekDate where - - /-- - The year component of the date. It is represented as an `Offset` type from `Year`. - -/ - year : Year.Offset - - /-- - The week of the year component. It is represented as an `Ordinal` type from `WeekOfYear`. - -/ - week : WeekOfYear.Ordinal - deriving Repr, BEq, Inhabited - -namespace WeekDate - -/-- -Converts a `WeekDate` to a `Day.Offset`. --/ -def toDays (wd : WeekDate) : Day.Offset := - let days := wd.year.toInt * 365 + wd.week.val * 7 - Day.Offset.ofInt days - -/-- -Creates a `WeekDate` from a `Day.Offset`. --/ -def fromDays (scalar : Day.Offset) : WeekDate := - let totalDays := scalar.val - let year := totalDays / 365 - let week := Bounded.LE.byEmod totalDays 365 (by decide) |>.ediv 7 (by decide) |>.add 1 - { year := year, week := week } - -end WeekDate -end Time -end Std diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 6f8c3f0407a6..75e932174d84 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -140,6 +140,20 @@ Subtracts a `Day.Offset` from a `PlainDateTime`. def subDays (dt : PlainDateTime) (days : Day.Offset) : PlainDateTime := { dt with date := dt.date.subDays days } +/-- +Adds a `Week.Offset` to a `PlainDateTime`. +-/ +@[inline] +def addWeeks (dt : PlainDateTime) (weeks : Week.Offset) : PlainDateTime := + { dt with date := dt.date.addWeeks weeks } + +/-- +Subtracts a `Week.Offset` from a `PlainDateTime`. +-/ +@[inline] +def subWeeks (dt : PlainDateTime) (weeks : Week.Offset) : PlainDateTime := + { dt with date := dt.date.subWeeks weeks } + /-- Adds a `Month.Offset` to a `PlainDateTime`, adjusting the day to the last valid day of the resulting month. @@ -340,6 +354,12 @@ instance : HAdd PlainDateTime Day.Offset PlainDateTime where instance : HSub PlainDateTime Day.Offset PlainDateTime where hSub := subDays +instance : HAdd PlainDateTime Week.Offset PlainDateTime where + hAdd := addWeeks + +instance : HSub PlainDateTime Week.Offset PlainDateTime where + hSub := subWeeks + instance : HAdd PlainDateTime Hour.Offset PlainDateTime where hAdd := addHours diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 56670abd90cd..43d39522b264 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -17,7 +17,7 @@ open Internal set_option linter.all true /-- -Timestamp is just a period between two timestamps. +`Timestamp` represents a specific point in time since the UNIX Epoch. -/ structure Timestamp where @@ -182,12 +182,32 @@ Subtracts a `Day.Offset` from a `Timestamp` def subDays (t : Timestamp) (d : Day.Offset) : Timestamp := ⟨t.val - d⟩ +/-- +Adds a `Week.Offset` to a `Timestamp` +-/ +@[inline] +def addWeeks (t : Timestamp) (d : Week.Offset) : Timestamp := + ⟨t.val + d⟩ + +/-- +Subtracts a `Week.Offset` from a `Timestamp` +-/ +@[inline] +def subWeeks (t : Timestamp) (d : Week.Offset) : Timestamp := + ⟨t.val - d⟩ + instance : HAdd Timestamp Day.Offset Timestamp where hAdd := addDays instance : HSub Timestamp Day.Offset Timestamp where hSub := subDays +instance : HAdd Timestamp Week.Offset Timestamp where + hAdd := addWeeks + +instance : HSub Timestamp Week.Offset Timestamp where + hSub := subWeeks + instance : HAdd Timestamp Hour.Offset Timestamp where hAdd := addHours diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 2d77ff04dfb4..16f1d1770d18 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -71,64 +71,6 @@ protected def neg (duration : Duration) : Duration := by | inl n => exact Or.inr (n.imp Int.neg_le_neg Int.neg_le_neg) | inr n => exact Or.inl (n.imp Int.neg_le_neg Int.neg_le_neg) -/-- -Adds two durations together, handling any carry-over in nanoseconds. It should not be used for `Duration`. -The subtraction of two `Duration` returns a duration but the addition does not make sense at all. --/ -def add (t₁ t₂ : Duration) : Duration := by - let diffSecs := t₁.second + t₂.second - let diffNano := t₁.nano.addBounds t₂.nano - - let (diffSecs, diffNano) : Second.Offset × Nanosecond.Span := - if h₀ : diffNano.val > 999999999 then - have diffNano := diffNano.truncateBottom h₀ |>.sub 999999999 - (diffSecs + 1, diffNano.expandBottom (by decide)) - else if h₁ : diffNano.val < -999999999 then - have diffNano := diffNano.truncateTop (Int.le_sub_one_of_lt h₁) |>.add 999999999 - (diffSecs - 1, diffNano.expandTop (by decide)) - else by - have h₀ := Int.le_sub_one_of_lt <| Int.not_le.mp h₀ - have h₁ := Int.le_sub_one_of_lt <| Int.not_le.mp h₁ - simp_all [Int.add_sub_cancel] - exact ⟨diffSecs, Bounded.mk diffNano.val (And.intro h₁ h₀)⟩ - - if h : diffSecs.val > 0 ∧ diffNano.val < 0 then - let truncated := diffNano.truncateTop (Int.le_sub_one_of_lt h.right) - let nano := truncated.addTop 1000000000 (by decide) - let proof₁ : 0 ≤ diffSecs - 1 := Int.le_sub_one_of_lt h.left - refine { second := .ofInt (diffSecs.val - 1), nano, proof := ?_ } - simp [nano, Bounded.LE.addTop] - refine (Or.inl (And.intro proof₁ ?_)) - let h₃ := (Int.add_le_add_iff_left 1000000000).mpr diffNano.property.left - rw [Int.add_comm] - exact Int.le_trans (by decide) h₃ - else if h₁ : diffSecs.val < 0 ∧ diffNano.val > 0 then - let second := diffSecs.val + 1 - let truncated := diffNano.truncateBottom h₁.right - let nano := truncated.subBottom 1000000000 (by decide) - refine { second := .ofInt second, nano, proof := ?_ } - simp [nano, truncated, Bounded.LE.subBottom, Bounded.LE.truncateBottom] - refine (Or.inr (And.intro ?_ ?_)) - · exact h₁.left - · let h₃ := Int.sub_le_sub_right diffNano.property.right 1000000000 - simp at h₃ - exact Int.le_trans h₃ (by decide) - else - refine ⟨diffSecs, diffNano, ?_⟩ - if h₂ : diffSecs.val > 0 then - exact Or.inl (And.intro (Int.le_of_lt h₂) (Int.not_lt.mp (not_and.mp h h₂))) - else if h₃ : diffSecs.val < 0 then - exact Or.inr (And.intro (Int.le_of_lt h₃) (Int.not_lt.mp (not_and.mp h₁ h₃))) - else - exact Int.le_total diffNano.val 0 |>.symm.imp (And.intro (Int.not_lt.mp h₃)) (And.intro (Int.not_lt.mp h₂)) - -/-- -Subtracts one `Duration` from another. --/ -@[inline] -def sub (t₁ t₂ : Duration) : Duration := - t₁.add t₂.neg - /-- Creates a new `Duration` out of `Second.Offset`. -/ @@ -188,6 +130,27 @@ Converts a `Duration` to a `Day.Offset` def toDays (tm : Duration) : Day.Offset := tm.second.ediv 86400 +/-- +Normalizes `Second.Offset` and `NanoSecond.span` in order to build a new `Duration` out of it. +-/ +@[inline] +def fromComponents (secs : Second.Offset) (nanos : Nanosecond.Span) : Duration := + ofNanoseconds (secs.toNanoseconds + nanos.toOffset) + +/-- +Adds two durations together, handling any carry-over in nanoseconds. +-/ +@[inline] +def add (t₁ t₂ : Duration) : Duration := + ofNanoseconds (toNanoseconds t₁ + toNanoseconds t₂) + +/-- +Subtracts one `Duration` from another. +-/ +@[inline] +def sub (t₁ t₂ : Duration) : Duration := + t₁.add t₂.neg + /-- Adds a `Nanosecond.Offset` to a `Duration` -/ @@ -264,12 +227,34 @@ def subDays (t : Duration) (d : Day.Offset) : Duration := let seconds := d.mul 86400 t.subSeconds seconds +/-- +Adds a `Week.Offset` to a `Duration` +-/ +@[inline] +def addWeeks (t : Duration) (w : Week.Offset) : Duration := + let seconds := w.mul 604800 + t.addSeconds seconds + +/-- +Subtracts a `Week.Offset` from a `Duration` +-/ +@[inline] +def subWeeks (t : Duration) (w : Week.Offset) : Duration := + let seconds := w.mul 604800 + t.subSeconds seconds + instance : HAdd Duration Day.Offset Duration where hAdd := addDays instance : HSub Duration Day.Offset Duration where hSub := subDays +instance : HAdd Duration Week.Offset Duration where + hAdd := addWeeks + +instance : HSub Duration Week.Offset Duration where + hSub := subWeeks + instance : HAdd Duration Hour.Offset Duration where hAdd := addHours diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index c1629a7f0f79..73664030ad91 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -55,25 +55,31 @@ for representing date, time, and time zone. def dateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD:hh:mm:ss.sssssssssZZZ" /-- -The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time +The Time24Hour format, which follows the pattern `date% hh:mm:ss` for representing time in a 24-hour clock format. It uses the default value that can be parsed with the notation of dates. -/ -def leanTime24Hour : Format .any := date-spec% "hh:mm:ss:sssssssss" +def leanTime24Hour : Format .any := date-spec% "'time% 'hh:mm:ss:sssssssss" /-- -The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD hh:mm:ss:sssssssss` for +The DateTimeZone24Hour format, which follows the pattern `date% YYYY-MM-DD hh:mm:ss:sssssssss` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD:hh:mm:ss:sssssssss" +def leanDateTime24Hour : Format (.only .GMT) := date-spec% "'date% 'YYYY-MM-DD:hh:mm:ss:sssssssss" /-- -The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD hh:mm:ss:sssssssss` +The DateTimeWithZone format, which follows the pattern `date% YYYY-MM-DD hh:mm:ss:sssssssss` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD:hh:mm:ss:sssssssssZZZZZ" +def leanDateTimeWithZone : Format .any := date-spec% "'date% 'YYYY-MM-DD:hh:mm:ss:sssssssssZZZZZ" + +/-- +The Lean Date format, which follows the pattern `date% YYYY-MM-DD`. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanDate : Format .any := date-spec% "'date% 'YYYY-MM-DD" /-- The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used @@ -151,6 +157,12 @@ Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. def toSQLDateString (input : PlainDate) : String := Formats.sqlDate.formatBuilder input.year input.month input.day +/-- +Converts a Date in the Lean format (`YYYY-MM-DD`) into a `String` with the format `date% YYY-MM-DD`. +-/ +def toLeanDateString (input : PlainDate) : String := + Formats.leanDate.formatBuilder input.year input.month input.day + /-- Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `PlainDate`. -/ @@ -162,7 +174,7 @@ instance : ToString PlainDate where toString := toSQLDateString instance : Repr PlainDate where - reprPrec data := Repr.addAppParen (toString data) + reprPrec data := Repr.addAppParen (toLeanDateString data) end PlainDate @@ -367,7 +379,7 @@ def fromAscTimeString (input : String) : Except String PlainDateTime := Formats a `PlainDateTime` value into an AscTime format string. -/ def toAscTimeString (pdt : PlainDateTime) : String := - Formats.ascTime.format (DateTime.ofPlainDateTime pdt .UTC) + Formats.ascTime.format (DateTime.ofPlainDateTimeAssumingUTC pdt .UTC) /-- Parses a `String` in the `LongDateFormat` and returns a `PlainDateTime` object in the GMT time zone. @@ -380,7 +392,7 @@ def fromLongDateFormatString (input : String) : Except String PlainDateTime := Formats a `PlainDateTime` value into a LongDateFormat string. -/ def toLongDateFormatString (pdt : PlainDateTime) : String := - Formats.longDateFormat.format (DateTime.ofPlainDateTime pdt .UTC) + Formats.longDateFormat.format (DateTime.ofPlainDateTimeAssumingUTC pdt .UTC) /-- Parses a `String` in the `DateTime` format and returns a `PlainDateTime`. @@ -393,7 +405,7 @@ def fromDateTimeString (input : String) : Except String PlainDateTime := Formats a `PlainDateTime` value into a `DateTime` format string. -/ def toDateTimeString (pdt : PlainDateTime) : String := - Formats.dateTime24Hour.format (DateTime.ofPlainDateTime pdt .UTC) + Formats.dateTime24Hour.format (DateTime.ofPlainDateTimeAssumingUTC pdt .UTC) /-- Parses a `String` in the `DateTime` format and returns a `PlainDateTime`. @@ -406,7 +418,7 @@ def fromLeanDateTimeString (input : String) : Except String PlainDateTime := Formats a `PlainDateTime` value into a `DateTime` format string. -/ def toLeanDateTimeString (pdt : PlainDateTime) : String := - Formats.leanDateTime24Hour.format (DateTime.ofPlainDateTime pdt .UTC) + Formats.leanDateTime24Hour.format (DateTime.ofPlainDateTimeAssumingUTC pdt .UTC) /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `PlainDateTime`. diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 2cfcfd936187..cc8a02fcfe74 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -593,7 +593,7 @@ private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except pure builder.hour if let .isTrue p := inferInstanceAs (Decidable (ValidTime hour.snd builder.minute builder.second.snd)) then - let build := DateTime.ofPlainDateTime { + let build := DateTime.ofLocalDateTime { date := PlainDate.clip builder.year builder.month builder.day time := PlainTime.mk hour builder.minute builder.second builder.nanoseconds p } diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 06a9ebec86cf..b432dd2c85b4 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -158,7 +158,7 @@ private def parseZoned : TSyntax `zoned -> MacroM (TSyntax `term) | `(zoned| $timestamp:num $zone) => do `(Std.Time.DateTime.ofUTCTimestamp $timestamp $(← parseZone zone)) | `(zoned| $datetime:datetime $zone) => do - `(Std.Time.DateTime.ofPlainDateTime $(← parseDateTime datetime) $(← parseZone zone)) + `(Std.Time.DateTime.ofLocalDateTime $(← parseDateTime datetime) $(← parseZone zone)) | syn => Macro.throwErrorAt syn "unsupported syntax" /-- diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 3f87610f4bf5..449ae283a5ca 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -21,6 +21,7 @@ inductive HourMarker | am /-- Post meridiem. -/ | pm + deriving Repr, BEq /-- Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker`. diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 0c12aa0ce1e6..9508b060b24a 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -46,7 +46,6 @@ structure PlainTime where instance : Inhabited PlainTime where default := ⟨Sigma.mk false 0, 0, Sigma.mk false 0, 0, by simp; decide⟩ - instance : BEq PlainTime where beq x y := x.hour.snd.val == y.hour.snd.val && x.minute == y.minute && x.second.snd.val == y.second.snd.val && x.nano.val == y.nano.val diff --git a/src/Std/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean index 47e8cf276c27..e699a3670a96 100644 --- a/src/Std/Time/Time/Unit/Basic.lean +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -21,11 +21,160 @@ to facilitate conversions and manipulations between them. namespace Std namespace Time -namespace Second.Offset open Internal set_option linter.all true +namespace Nanosecond.Offset + +/-- +Converts a `Nanosecond.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def toMilliseconds (offset : Nanosecond.Offset) : Millisecond.Offset := + offset.ediv 1000000 + +/-- +Converts a `Millisecond.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def ofMilliseconds (offset : Millisecond.Offset) : Nanosecond.Offset := + offset.mul 1000000 + +/-- +Converts a `Nanosecond.Offset` to a `Second.Offset`. +-/ +@[inline] +def toSeconds (offset : Nanosecond.Offset) : Second.Offset := + offset.ediv 1000000000 + +/-- +Converts a `Second.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def ofSeconds (offset : Second.Offset) : Nanosecond.Offset := + offset.mul 1000000000 + +/-- +Converts a `Nanosecond.Offset` to a `Minute.Offset`. +-/ +@[inline] +def toMinutes (offset : Nanosecond.Offset) : Minute.Offset := + offset.ediv 60000000000 + +/-- +Converts a `Minute.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def ofMinutes (offset : Minute.Offset) : Nanosecond.Offset := + offset.mul 60000000000 + +/-- +Converts a `Nanosecond.Offset` to an `Hour.Offset`. +-/ +@[inline] +def toHours (offset : Nanosecond.Offset) : Hour.Offset := + offset.ediv 3600000000000 + +/-- +Converts an `Hour.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def ofHours (offset : Hour.Offset) : Nanosecond.Offset := + offset.mul 3600000000000 + +end Nanosecond.Offset + +namespace Millisecond.Offset + +/-- +Converts a `Millisecond.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def toNanoseconds (offset : Millisecond.Offset) : Nanosecond.Offset := + offset.mul 1000000 + +/-- +Converts a `Nanosecond.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def ofNanoseconds (offset : Nanosecond.Offset) : Millisecond.Offset := + offset.ediv 1000000 + +/-- +Converts a `Millisecond.Offset` to a `Second.Offset`. +-/ +@[inline] +def toSeconds (offset : Millisecond.Offset) : Second.Offset := + offset.ediv 1000 + +/-- +Converts a `Second.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def ofSeconds (offset : Second.Offset) : Millisecond.Offset := + offset.mul 1000 + +/-- +Converts a `Millisecond.Offset` to a `Minute.Offset`. +-/ +@[inline] +def toMinutes (offset : Millisecond.Offset) : Minute.Offset := + offset.ediv 60000 + +/-- +Converts a `Minute.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def ofMinutes (offset : Minute.Offset) : Millisecond.Offset := + offset.mul 60000 + +/-- +Converts a `Millisecond.Offset` to an `Hour.Offset`. +-/ +@[inline] +def toHours (offset : Millisecond.Offset) : Hour.Offset := + offset.ediv 3600000 + +/-- +Converts an `Hour.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def ofHours (offset : Hour.Offset) : Millisecond.Offset := + offset.mul 3600000 + +end Millisecond.Offset + +namespace Second.Offset + +/-- +Converts a `Second.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def toNanoseconds (offset : Second.Offset) : Nanosecond.Offset := + offset.mul 1000000000 + +/-- +Converts a `Nanosecond.Offset` to a `Second.Offset`. +-/ +@[inline] +def ofNanoseconds (offset : Nanosecond.Offset) : Second.Offset := + offset.ediv 1000000000 + +/-- +Converts a `Second.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def toMilliseconds (offset : Second.Offset) : Millisecond.Offset := + offset.mul 1000 + +/-- +Converts a `Millisecond.Offset` to a `Second.Offset`. +-/ +@[inline] +def ofMilliseconds (offset : Millisecond.Offset) : Second.Offset := + offset.ediv 1000 + /-- Converts a `Second.Offset` to a `Minute.Offset`. -/ @@ -33,6 +182,13 @@ Converts a `Second.Offset` to a `Minute.Offset`. def toMinutes (offset : Second.Offset) : Minute.Offset := offset.ediv 60 +/-- +Converts a `Minute.Offset` to a `Second.Offset`. +-/ +@[inline] +def ofMinutes (offset : Minute.Offset) : Second.Offset := + offset.mul 60 + /-- Converts a `Second.Offset` to an `Hour.Offset`. -/ @@ -40,10 +196,59 @@ Converts a `Second.Offset` to an `Hour.Offset`. def toHours (offset : Second.Offset) : Hour.Offset := offset.ediv 3600 +/-- +Converts an `Hour.Offset` to a `Second.Offset`. +-/ +@[inline] +def ofHours (offset : Hour.Offset) : Second.Offset := + offset.mul 3600 + end Second.Offset namespace Minute.Offset +/-- +Converts a `Minute.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def toNanoseconds (offset : Minute.Offset) : Nanosecond.Offset := + offset.mul 60000000000 + +/-- +Converts a `Nanosecond.Offset` to a `Minute.Offset`. +-/ +@[inline] +def ofNanoseconds (offset : Nanosecond.Offset) : Minute.Offset := + offset.ediv 60000000000 + +/-- +Converts a `Minute.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def toMilliseconds (offset : Minute.Offset) : Millisecond.Offset := + offset.mul 60000 + +/-- +Converts a `Millisecond.Offset` to a `Minute.Offset`. +-/ +@[inline] +def ofMilliseconds (offset : Millisecond.Offset) : Minute.Offset := + offset.ediv 60000 + +/-- +Converts a `Minute.Offset` to a `Second.Offset`. +-/ +@[inline] +def toSeconds (offset : Minute.Offset) : Second.Offset := + offset.mul 60 + +/-- +Converts a `Second.Offset` to a `Minute.Offset`. +-/ +@[inline] +def ofSeconds (offset : Second.Offset) : Minute.Offset := + offset.ediv 60 + /-- Converts a `Minute.Offset` to an `Hour.Offset`. -/ @@ -51,7 +256,74 @@ Converts a `Minute.Offset` to an `Hour.Offset`. def toHours (offset : Minute.Offset) : Hour.Offset := offset.ediv 60 +/-- +Converts an `Hour.Offset` to a `Minute.Offset`. +-/ +@[inline] +def ofHours (offset : Hour.Offset) : Minute.Offset := + offset.mul 60 + end Minute.Offset +namespace Hour.Offset + +/-- +Converts an `Hour.Offset` to a `Nanosecond.Offset`. +-/ +@[inline] +def toNanoseconds (offset : Hour.Offset) : Nanosecond.Offset := + offset.mul 3600000000000 + +/-- +Converts a `Nanosecond.Offset` to an `Hour.Offset`. +-/ +@[inline] +def ofNanoseconds (offset : Nanosecond.Offset) : Hour.Offset := + offset.ediv 3600000000000 + +/-- +Converts an `Hour.Offset` to a `Millisecond.Offset`. +-/ +@[inline] +def toMilliseconds (offset : Hour.Offset) : Millisecond.Offset := + offset.mul 3600000 + +/-- +Converts a `Millisecond.Offset` to an `Hour.Offset`. +-/ +@[inline] +def ofMilliseconds (offset : Millisecond.Offset) : Hour.Offset := + offset.ediv 3600000 + +/-- +Converts an `Hour.Offset` to a `Second.Offset`. +-/ +@[inline] +def toSeconds (offset : Hour.Offset) : Second.Offset := + offset.mul 3600 + +/-- +Converts a `Second.Offset` to an `Hour.Offset`. +-/ +@[inline] +def ofSeconds (offset : Second.Offset) : Hour.Offset := + offset.ediv 3600 + +/-- +Converts an `Hour.Offset` to a `Minute.Offset`. +-/ +@[inline] +def toMinutes (offset : Hour.Offset) : Minute.Offset := + offset.mul 60 + +/-- +Converts a `Minute.Offset` to an `Hour.Offset`. +-/ +@[inline] +def ofMinutes (offset : Minute.Offset) : Hour.Offset := + offset.ediv 60 + +end Hour.Offset + end Time end Std diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 293aa8e1668f..4454e4e663ae 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -23,6 +23,9 @@ account for valid timestamps like 24:00:00 with leap seconds. -/ def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 24 else 23)) +instance : ToString (Ordinal leap) where + toString x := toString x.val + instance : Repr (Ordinal l) where reprPrec r l := reprPrec r.val l @@ -32,9 +35,6 @@ instance : OfNat (Ordinal leap) n := by · exact inst · exact ⟨inst.ofNat.expandTop (by decide)⟩ -instance : OfNat (Ordinal true) 24 where - ofNat := Bounded.LE.mk (Int.ofNat 24) (by decide) - /-- `Offset` represents an offset in hours, defined as an `Int`. This can be used to express durations or differences in hours. @@ -72,23 +72,6 @@ def toOffset (ordinal : Ordinal leap) : Offset := UnitVal.ofInt ordinal.val end Ordinal -namespace Offset - -/-- -Converts an `Hour.Offset` to a `Second.Offset`. --/ -@[inline] -def toSeconds (val : Offset) : Second.Offset := - val.mul 3600 - -/-- -Converts an `Hour.Offset` to a `Minute.Offset`. --/ -@[inline] -def toMinutes (val : Offset) : Minute.Offset := - val.mul 60 - -end Offset end Hour end Time end Std diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 3528885dadac..389473cc8e19 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -21,6 +21,9 @@ set_option linter.all true def Ordinal := Bounded.LE 0 999 deriving Repr, BEq, LE, LT +instance : ToString Ordinal where + toString x := toString x.val + instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (999 : Nat))) n) diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index ec36f5d6ed46..346e5eff5988 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -22,6 +22,9 @@ set_option linter.all true def Ordinal := Bounded.LE 0 59 deriving Repr, BEq, LE +instance : ToString Ordinal where + toString x := toString x.val + instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) n) @@ -62,17 +65,6 @@ def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val end Ordinal - -namespace Offset - -/-- -Converts a `Minute.Offset` to `Second.Offset`. --/ -@[inline] -def toSeconds (val : Offset) : Second.Offset := - val.mul 60 - -end Offset end Minute end Time end Std diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index 33be8713187f..2cfe052c313d 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -21,6 +21,9 @@ set_option linter.all true def Ordinal := Bounded.LE 0 999999999 deriving Repr, BEq, LE, LT +instance : ToString Ordinal where + toString x := toString x.val + instance : OfNat Ordinal n where ofNat := Bounded.LE.ofFin (Fin.ofNat n) @@ -85,6 +88,15 @@ def Span := Bounded.LE (-999999999) 999999999 instance : Inhabited Span where default := Bounded.LE.mk 0 (by decide) +namespace Span + +/-- +Creates a new `Offset` out of a `Span`. +-/ +def toOffset (span : Span) : Offset := + UnitVal.mk span.val + +end Span end Nanosecond end Time end Std diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 036ce5f5b3a0..b8b5d0093183 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -21,8 +21,11 @@ for potential leap second. -/ def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 60 else 59)) +instance : ToString (Ordinal leap) where + toString x := toString x.val + instance : Repr (Ordinal l) where - reprPrec r l := reprPrec r.val l + reprPrec r := reprPrec r.val instance : OfNat (Ordinal leap) n := by have inst := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) n) diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 73b01d561af5..06a110acc27c 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -29,7 +29,7 @@ Get the current monotonic time. @[inline] def now : IO PlainDateTime := do let tm ← Timestamp.now - let tz ← TimeZone.getCurrentTimezone + let tz ← TimeZone.getSystemTimezone return PlainDateTime.ofTimestamp tm tz end PlainDateTime @@ -91,7 +91,7 @@ Gets the current `DateTime`. @[inline] def now : IO ZonedDateTime := do let date ← Timestamp.now - let tz ← TimeZone.getCurrentTimezone + let tz ← TimeZone.getSystemTimezone return ofTimestamp date tz /-- diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 67a710eca737..fc4d848c6166 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -13,17 +13,17 @@ namespace Time set_option linter.all true /-- -A timezone database that we can read the `ZoneRules` of some area by it's id. +A timezone database from which we can read the `ZoneRules` of some area by it's id. -/ class Database (α : Type) where /-- - Loads a `ZoneRules` by it's id. + Loads a `ZoneRules` by its id. -/ load : α → String → IO TimeZone.ZoneRules /-- - Loads a `ZoneRules` that is setted by the local computer. + Loads a `ZoneRules` that is set by the local computer. -/ localRules : α → IO TimeZone.ZoneRules diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index bb8df5c0df26..dbea6fbcf2f9 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -59,10 +59,20 @@ def convertTimeZone (date : DateTime tz) (tz₁ : TimeZone) : DateTime tz₁ := ofTimestamp date.timestamp tz₁ /-- -Creates a new `DateTime` out of a `PlainDateTime` +Creates a new `DateTime` out of a `PlainDateTime`. It assumes that the `PlainDateTime` is relative +to UTC. If you're using hte PlainDateTime -/ @[inline] -def ofPlainDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := +def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : DateTime tz := + let tm := Timestamp.ofPlainDateTime date + DateTime.mk tm (Thunk.mk <| λ_ => date.addSeconds tz.toSeconds) + +/-- +Creates a new `DateTime` out of a `PlainDateTime`. It assumes that the `PlainDateTime` it's relative +to the +-/ +@[inline] +def ofLocalDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := let tm := Timestamp.ofPlainDateTime date DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk <| λ_ => date) @@ -71,90 +81,104 @@ Add `Hour.Offset` to a `DateTime`. -/ @[inline] def addHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addHours hours) tz + ofLocalDateTime (dt.date.get.addHours hours) tz /-- Subtract `Hour.Offset` from a `DateTime`. -/ @[inline] def subHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subHours hours) tz + ofLocalDateTime (dt.date.get.subHours hours) tz /-- Add `Minute.Offset` to a `DateTime`. -/ @[inline] def addMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addMinutes minutes) tz + ofLocalDateTime (dt.date.get.addMinutes minutes) tz /-- Subtract `Minute.Offset` from a `DateTime`. -/ @[inline] def subMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subMinutes minutes) tz + ofLocalDateTime (dt.date.get.subMinutes minutes) tz /-- Add `Second.Offset` to a `DateTime`. -/ @[inline] def addSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addSeconds seconds) tz + ofLocalDateTime (dt.date.get.addSeconds seconds) tz /-- Subtract `Second.Offset` from a `DateTime`. -/ @[inline] def subSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subSeconds seconds) tz + ofLocalDateTime (dt.date.get.subSeconds seconds) tz /-- Add `Nanosecond.Offset` to a `DateTime`. -/ @[inline] def addNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addNanoseconds nanoseconds) tz + ofLocalDateTime (dt.date.get.addNanoseconds nanoseconds) tz /-- Subtract `Nanosecond.Offset` from a `DateTime`. -/ @[inline] def subNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subNanoseconds nanoseconds) tz + ofLocalDateTime (dt.date.get.subNanoseconds nanoseconds) tz /-- Add `Day.Offset` to a `DateTime`. -/ @[inline] def addDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addDays days) tz + ofLocalDateTime (dt.date.get.addDays days) tz /-- Subtracts `Day.Offset` to a `DateTime`. -/ @[inline] def subDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subDays days) tz + ofLocalDateTime (dt.date.get.subDays days) tz + +/-- +Add `Week.Offset` to a `DateTime`. +-/ +@[inline] +def addWeeks (dt : DateTime tz) (weeks : Week.Offset) : DateTime tz := + ofLocalDateTime (dt.date.get.addWeeks weeks) tz + +/-- +Subtracts `Week.Offset` to a `DateTime`. +-/ +@[inline] +def subWeeks (dt : DateTime tz) (weeks : Week.Offset) : DateTime tz := + ofLocalDateTime (dt.date.get.subWeeks weeks) tz /-- Add `Month.Offset` to a `DateTime`, it clips the day to the last valid day of that month. -/ def addMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addMonthsClip months) tz + ofLocalDateTime (dt.date.get.addMonthsClip months) tz /-- Subtracts `Month.Offset` from a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def subMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subMonthsClip months) tz + ofLocalDateTime (dt.date.get.subMonthsClip months) tz /-- Add `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following month. -/ def addMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addMonthsRollOver months) tz + ofLocalDateTime (dt.date.get.addMonthsRollOver months) tz /-- Subtract `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following @@ -162,7 +186,7 @@ month. -/ @[inline] def subMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subMonthsRollOver months) tz + ofLocalDateTime (dt.date.get.subMonthsRollOver months) tz /-- Add `Year.Offset` to a `DateTime`, this function rolls over any excess days into the following @@ -170,14 +194,14 @@ month. -/ @[inline] def addYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addYearsRollOver years) tz + ofLocalDateTime (dt.date.get.addYearsRollOver years) tz /-- Add `Year.Offset` to a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def addYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.addYearsClip years) tz + ofLocalDateTime (dt.date.get.addYearsClip years) tz /-- Subtract `Year.Offset` from a `DateTime`, this function rolls over any excess days into the following @@ -185,14 +209,14 @@ month. -/ @[inline] def subYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subYearsRollOver years) tz + ofLocalDateTime (dt.date.get.subYearsRollOver years) tz /-- Subtract `Year.Offset` from to a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def subYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofPlainDateTime (dt.timestamp.toPlainDateTime.subYearsClip years) tz + ofLocalDateTime (dt.date.get.subYearsClip years) tz /-- Converts a `Timestamp` to a `PlainDateTime` @@ -288,6 +312,12 @@ instance : HAdd (DateTime tz) (Day.Offset) (DateTime tz) where instance : HSub (DateTime tz) (Day.Offset) (DateTime tz) where hSub := subDays +instance : HAdd (DateTime tz) (Week.Offset) (DateTime tz) where + hAdd := addWeeks + +instance : HSub (DateTime tz) (Week.Offset) (DateTime tz) where + hSub := subWeeks + instance : HAdd (DateTime tz) (Hour.Offset) (DateTime tz) where hAdd := addHours diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index ffa6bfd3df23..f46a0fe2045f 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -38,7 +38,7 @@ namespace TimeZone Fetches the current timestamp from the system. -/ @[extern "lean_get_timezone_offset"] -opaque getCurrentTimezone : IO TimeZone +opaque getSystemTimezone : IO TimeZone /-- A zeroed `Timezone` representing UTC (no offset). diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index a08886e2f7a8..05c08e5faedb 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -45,7 +45,7 @@ def toTimestamp (date : ZonedDateTime) : Timestamp := date.snd.toTimestamp /-- -Creates a new `DateTime` out of a `Timestamp` +Creates a new `ZonedDateTime` out of a `Timestamp` -/ @[inline] def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do @@ -58,14 +58,14 @@ Changes the `TimeZone` to a new one. -/ @[inline] def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := - ofTimestamp (date.toTimestamp) tz₁ + ofTimestamp date.toTimestamp tz₁ /-- Creates a new `ZonedDateTime` out of a `PlainDateTime` -/ @[inline] -def ofPlainDateTime (date : PlainDateTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofPlainDateTime date tz⟩ +def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofPlainDateTimeAssumingUTC date tz⟩ /-- Converts a `ZonedDateTime` to a `PlainDateTime` @@ -150,6 +150,19 @@ Subtracts `Day.Offset` to a `ZonedDateTime`. def subDays (dt : ZonedDateTime) (days : Day.Offset) : ZonedDateTime := Sigma.mk dt.fst (dt.snd.subDays days) +/-- +Add `Week.Offset` to a `ZonedDateTime`. +-/ +def addWeeks (dt : ZonedDateTime) (weeks : Week.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.addWeeks weeks) + +/-- +Subtracts `Week.Offset` to a `ZonedDateTime`. +-/ +@[inline] +def subWeeks (dt : ZonedDateTime) (weeks : Week.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subWeeks weeks) + /-- Add `Month.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. -/ @@ -264,34 +277,40 @@ def inLeapYear (date : ZonedDateTime) : Bool := instance : ToTimestamp ZonedDateTime where toTimestamp dt := dt.toTimestamp -instance : HAdd ZonedDateTime (Day.Offset) ZonedDateTime where +instance : HAdd ZonedDateTime Day.Offset ZonedDateTime where hAdd := addDays -instance : HSub ZonedDateTime (Day.Offset) ZonedDateTime where +instance : HSub ZonedDateTime Day.Offset ZonedDateTime where hSub := subDays -instance : HAdd ZonedDateTime (Hour.Offset) ZonedDateTime where +instance : HAdd ZonedDateTime Week.Offset ZonedDateTime where + hAdd := addWeeks + +instance : HSub ZonedDateTime Week.Offset ZonedDateTime where + hSub := subWeeks + +instance : HAdd ZonedDateTime Hour.Offset ZonedDateTime where hAdd := addHours -instance : HSub ZonedDateTime (Hour.Offset) ZonedDateTime where +instance : HSub ZonedDateTime Hour.Offset ZonedDateTime where hSub := subHours -instance : HAdd ZonedDateTime (Minute.Offset) ZonedDateTime where +instance : HAdd ZonedDateTime Minute.Offset ZonedDateTime where hAdd := addMinutes -instance : HSub ZonedDateTime (Minute.Offset) ZonedDateTime where +instance : HSub ZonedDateTime Minute.Offset ZonedDateTime where hSub := subMinutes -instance : HAdd ZonedDateTime (Second.Offset) ZonedDateTime where +instance : HAdd ZonedDateTime Second.Offset ZonedDateTime where hAdd := addSeconds -instance : HSub ZonedDateTime (Second.Offset) ZonedDateTime where +instance : HSub ZonedDateTime Second.Offset ZonedDateTime where hSub := subSeconds -instance : HAdd ZonedDateTime (Nanosecond.Offset) ZonedDateTime where +instance : HAdd ZonedDateTime Nanosecond.Offset ZonedDateTime where hAdd := addNanoseconds -instance : HSub ZonedDateTime (Nanosecond.Offset) ZonedDateTime where +instance : HSub ZonedDateTime Nanosecond.Offset ZonedDateTime where hSub := subNanoseconds instance : HSub ZonedDateTime ZonedDateTime Duration where diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 9b21934014e1..4a1ac3729407 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -546,7 +546,7 @@ extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { auto now_time_t = system_clock::to_time_t(now); std::tm tm_info; -#if defined(_MSC_VER) +#if defined(LEAN_WINDOWS) localtime_s(&tm_info, &now_time_t); #else localtime_r(&now_time_t, &tm_info); diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index 4132ee84c555..542928002ce2 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -4,13 +4,13 @@ open Std.Time def date := date% 1970-01-20 /-- -info: 1970-02-01 +info: date% 1970-02-01 -/ #guard_msgs in #eval date + (12 : Day.Offset) /-- -info: 1970-01-08 +info: date% 1970-01-08 -/ #guard_msgs in #eval date - (12 : Day.Offset) @@ -18,49 +18,49 @@ info: 1970-01-08 def datetime := date% 2000-01-20:03:02:01 /-- -info: 2000-01-20:04:02:01:000000000 +info: date% 2000-01-20:04:02:01:000000000 -/ #guard_msgs in #eval datetime + (1 : Hour.Offset) /-- -info: 2000-01-20:02:02:01:000000000 +info: date% 2000-01-20:02:02:01:000000000 -/ #guard_msgs in #eval datetime - (1 : Hour.Offset) /-- -info: 2000-01-20:03:12:01:000000000 +info: date% 2000-01-20:03:12:01:000000000 -/ #guard_msgs in #eval datetime + (10 : Minute.Offset) /-- -info: 2000-01-20:02:52:01:000000000 +info: date% 2000-01-20:02:52:01:000000000 -/ #guard_msgs in #eval datetime - (10 : Minute.Offset) /-- -info: 2000-01-20:03:03:01:000000000 +info: date% 2000-01-20:03:03:01:000000000 -/ #guard_msgs in #eval datetime + (60 : Second.Offset) /-- -info: 2000-01-20:03:01:01:000000000 +info: date% 2000-01-20:03:01:01:000000000 -/ #guard_msgs in #eval datetime - (60 : Second.Offset) /-- -info: 2000-01-21:03:02:01:000000000 +info: date% 2000-01-21:03:02:01:000000000 -/ #guard_msgs in #eval datetime + (1 : Day.Offset) /-- -info: 2000-01-19:03:02:01:000000000 +info: date% 2000-01-19:03:02:01:000000000 -/ #guard_msgs in #eval datetime - (1 : Day.Offset) @@ -68,87 +68,93 @@ info: 2000-01-19:03:02:01:000000000 def time := time% 13:02:01 /-- -info: 14:02:01:000000000 +info: time% 14:02:01:000000000 -/ #guard_msgs in #eval time + (1 : Hour.Offset) /-- -info: 12:02:01:000000000 +info: time% 12:02:01:000000000 -/ #guard_msgs in #eval time - (1 : Hour.Offset) /-- -info: 13:12:01:000000000 +info: time% 13:12:01:000000000 -/ #guard_msgs in #eval time + (10 : Minute.Offset) /-- -info: 12:52:01:000000000 +info: time% 12:52:01:000000000 -/ #guard_msgs in #eval time - (10 : Minute.Offset) /-- -info: 13:03:01:000000000 +info: time% 13:03:01:000000000 -/ #guard_msgs in #eval time + (60 : Second.Offset) /-- -info: 13:01:01:000000000 +info: time% 13:01:01:000000000 -/ #guard_msgs in #eval time - (60 : Second.Offset) -def datetimetz := date% 2000-01-20:03:02:01-03:00 +def datetimetz := date% 2000-01-20:06:02:01-03:00 /-- -info: 2000-01-22:06:02:01:000000000-03:00 +info: date% 2000-01-20:06:02:01:000000000-03:00 +-/ +#guard_msgs in +#eval datetimetz + +/-- +info: date% 2000-01-22:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (2 : Day.Offset) /-- -info: 2000-01-19:06:02:01:000000000-03:00 +info: date% 2000-01-19:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (1 : Day.Offset) /-- -info: 2000-01-20:07:02:01:000000000-03:00 +info: date% 2000-01-20:07:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (1 : Hour.Offset) /-- -info: 2000-01-20:05:02:01:000000000-03:00 +info: date% 2000-01-20:05:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (1 : Hour.Offset) /-- -info: 2000-01-20:06:12:01:000000000-03:00 +info: date% 2000-01-20:06:12:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (10 : Minute.Offset) /-- -info: 2000-01-20:05:52:01:000000000-03:00 +info: date% 2000-01-20:05:52:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (10 : Minute.Offset) /-- -info: 2000-01-20:06:03:01:000000000-03:00 +info: date% 2000-01-20:06:03:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz + (60 : Second.Offset) /-- -info: 2000-01-20:06:01:01:000000000-03:00 +info: date% 2000-01-20:06:01:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz - (60 : Second.Offset) diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 8cc2584022a8..28a66b5d5782 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -132,9 +132,9 @@ This PlainDate is relative to the local time. -/ def PlainDate : PlainDateTime := Timestamp.toPlainDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) -def dateBR₁ := DateTime.ofPlainDateTime PlainDate brTZ -def dateJP₁ := DateTime.ofPlainDateTime PlainDate jpTZ -def dateUTC₁ := DateTime.ofPlainDateTime PlainDate .UTC +def dateBR₁ := DateTime.ofLocalDateTime PlainDate brTZ +def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ +def dateUTC₁ := DateTime.ofLocalDateTime PlainDate .UTC /-- info: "Thursday, August 15, 2024 14:03:47 -0300" @@ -221,7 +221,7 @@ info: "-0221-09-04" #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) /-- -info: -0221-09-04 +info: date% -0221-09-04 -/ #guard_msgs in #eval (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) @@ -233,25 +233,25 @@ info: "-0221-09-04" #eval toString (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) /-- -info: 2002-07-14 +info: date% 2002-07-14 -/ #guard_msgs in #eval date% 2002-07-14 /-- -info: 14:13:12:000000000 +info: time% 14:13:12:000000000 -/ #guard_msgs in #eval time% 14:13:12 /-- -info: 2002-07-14:14:13:12:000000000 +info: date% 2002-07-14:14:13:12:000000000 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12 /-- -info: 2002-07-14:14:13:12:000000000+09:00 +info: date% 2002-07-14:14:13:12:000000000+09:00 -/ #guard_msgs in #eval date% 2002-07-14:14:13:12+09:00 diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index dbef85f37d99..61b00aa8c5de 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -4,31 +4,31 @@ open Std.Time def date := date% 1970-01-20 /-- -info: 1970-02-01 +info: date% 1970-02-01 -/ #guard_msgs in #eval date.addDays 12 /-- -info: 1970-02-20 +info: date% 1970-02-20 -/ #guard_msgs in #eval date.addMonthsClip 1 /-- -info: 1971-01-20 +info: date% 1971-01-20 -/ #guard_msgs in #eval date.addYearsRollOver 1 /-- -info: 1969-01-20 +info: date% 1969-01-20 -/ #guard_msgs in #eval date.subYearsClip 1 /-- -info: 1969-12-20 +info: date% 1969-12-20 -/ #guard_msgs in #eval date.subMonthsClip 1 @@ -36,97 +36,97 @@ info: 1969-12-20 def datetime := date% 2000-01-20:03:02:01 /-- -info: 2000-01-20:04:02:01:000000000 +info: date% 2000-01-20:04:02:01:000000000 -/ #guard_msgs in #eval datetime.addHours 1 /-- -info: 2000-01-20:02:02:01:000000000 +info: date% 2000-01-20:02:02:01:000000000 -/ #guard_msgs in #eval datetime.subHours 1 /-- -info: 2000-01-20:03:12:01:000000000 +info: date% 2000-01-20:03:12:01:000000000 -/ #guard_msgs in #eval datetime.addMinutes 10 /-- -info: 2000-01-20:02:52:01:000000000 +info: date% 2000-01-20:02:52:01:000000000 -/ #guard_msgs in #eval datetime.subMinutes 10 /-- -info: 2000-01-20:03:03:01:000000000 +info: date% 2000-01-20:03:03:01:000000000 -/ #guard_msgs in #eval datetime.addSeconds 60 /-- -info: 2000-01-20:03:01:01:000000000 +info: date% 2000-01-20:03:01:01:000000000 -/ #guard_msgs in #eval datetime.subSeconds 60 /-- -info: 2000-01-21:03:02:01:000000000 +info: date% 2000-01-21:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addDays 1 /-- -info: 2000-01-19:03:02:01:000000000 +info: date% 2000-01-19:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subDays 1 /-- -info: 2000-02-20:03:02:01:000000000 +info: date% 2000-02-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addMonthsClip 1 /-- -info: 1999-12-20:03:02:01:000000000 +info: date% 1999-12-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subMonthsClip 1 /-- -info: 2000-02-20:03:02:01:000000000 +info: date% 2000-02-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addMonthsRollOver 1 /-- -info: 1999-12-20:03:02:01:000000000 +info: date% 1999-12-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subMonthsRollOver 1 /-- -info: 2001-01-20:03:02:01:000000000 +info: date% 2001-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addYearsClip 1 /-- -info: 1999-01-20:03:02:01:000000000 +info: date% 1999-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subYearsClip 1 /-- -info: 2001-01-20:03:02:01:000000000 +info: date% 2001-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.addYearsRollOver 1 /-- -info: 1999-01-20:03:02:01:000000000 +info: date% 1999-01-20:03:02:01:000000000 -/ #guard_msgs in #eval datetime.subYearsRollOver 1 @@ -134,134 +134,142 @@ info: 1999-01-20:03:02:01:000000000 def time := time% 13:02:01 /-- -info: 14:02:01:000000000 +info: time% 14:02:01:000000000 -/ #guard_msgs in #eval time.addHours 1 /-- -info: 12:02:01:000000000 +info: time% 12:02:01:000000000 -/ #guard_msgs in #eval time.subHours 1 /-- -info: 13:12:01:000000000 +info: time% 13:12:01:000000000 -/ #guard_msgs in #eval time.addMinutes 10 /-- -info: 12:52:01:000000000 +info: time% 12:52:01:000000000 -/ #guard_msgs in #eval time.subMinutes 10 /-- -info: 13:03:01:000000000 +info: time% 13:03:01:000000000 -/ #guard_msgs in #eval time.addSeconds 60 /-- -info: 13:01:01:000000000 +info: time% 13:01:01:000000000 -/ #guard_msgs in #eval time.subSeconds 60 -def datetimetz := date% 2000-01-20:03:02:01-03:00 +def datetimetz := date% 2000-01-20:06:02:01-03:00 + +/-- +info: date% 2000-01-20:06:02:01:000000000-03:00 + +-/ +#guard_msgs in +#eval datetimetz + /-- -info: 2000-01-22:06:02:01:000000000-03:00 +info: date% 2000-01-22:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addDays 2 /-- -info: 2000-01-19:06:02:01:000000000-03:00 +info: date% 2000-01-19:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subDays 1 /-- -info: 2000-02-20:06:02:01:000000000-03:00 +info: date% 2000-02-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMonthsClip 1 /-- -info: 1999-12-20:06:02:01:000000000-03:00 +info: date% 1999-12-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMonthsClip 1 /-- -info: 2000-02-20:06:02:01:000000000-03:00 +info: date% 2000-02-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMonthsRollOver 1 /-- -info: 1999-12-20:06:02:01:000000000-03:00 +info: date% 1999-12-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMonthsRollOver 1 /-- -info: 2001-01-20:06:02:01:000000000-03:00 +info: date% 2001-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: 2001-01-20:06:02:01:000000000-03:00 +info: date% 2001-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: 2001-01-20:06:02:01:000000000-03:00 +info: date% 2001-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsRollOver 1 /-- -info: 1999-01-20:06:02:01:000000000-03:00 +info: date% 1999-01-20:06:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subYearsRollOver 1 /-- -info: 2000-01-20:07:02:01:000000000-03:00 +info: date% 2000-01-20:07:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addHours 1 /-- -info: 2000-01-20:05:02:01:000000000-03:00 +info: date% 2000-01-20:05:02:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subHours 1 /-- -info: 2000-01-20:06:12:01:000000000-03:00 +info: date% 2000-01-20:06:12:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMinutes 10 /-- -info: 2000-01-20:05:52:01:000000000-03:00 +info: date% 2000-01-20:05:52:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMinutes 10 /-- -info: 2000-01-20:06:03:01:000000000-03:00 +info: date% 2000-01-20:06:03:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.addSeconds 60 /-- -info: 2000-01-20:06:01:01:000000000-03:00 +info: date% 2000-01-20:06:01:01:000000000-03:00 -/ #guard_msgs in #eval datetimetz.subSeconds 60 @@ -283,3 +291,18 @@ info: "86399s" -/ #guard_msgs in #eval (date% 2000-1-20:0:0:0UTC).since (date% 2000-1-19:0:0:1UTC) + +def now := date% 2024-08-29:10:56:43:276801081+02:00 + + +/-- +info: date% 2024-08-29:10:56:43:276801081+02:00 +-/ +#guard_msgs in +#eval now + +/-- +info: date% 2024-08-30:10:56:43:276801081+02:00 +-/ +#guard_msgs in +#eval now.addDays 1 diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 26d316042fd5..5dde92fb237b 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -142,9 +142,9 @@ This PlainDate is relative to the local time. -/ def PlainDate : PlainDateTime := Timestamp.toPlainDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) -def dateBR₁ := DateTime.ofPlainDateTime PlainDate brTZ -def dateJP₁ := DateTime.ofPlainDateTime PlainDate jpTZ -def dateUTC₁ := DateTime.ofPlainDateTime PlainDate .UTC +def dateBR₁ := DateTime.ofLocalDateTime PlainDate brTZ +def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ +def dateUTC₁ := DateTime.ofLocalDateTime PlainDate .UTC /-- info: "2024-08-15T14:03:47-0300" diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index 98e7e7fe5ff0..69d60fc24a25 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -59,13 +59,13 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: some (1969-12-30:21:00:00:000000000-03:00) +info: some (date% 1969-12-30:21:00:00:000000000-03:00) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules /-- -info: some (2012-12-10:00:35:47:000000000-02:00) +info: some (date% 2012-12-10:00:35:47:000000000-02:00) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From 79c90c9e63134da4d25cd88c99ebc84a7632617f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 30 Aug 2024 01:08:03 -0300 Subject: [PATCH 056/118] refactor: change notation syntax, add error handling for the get_timezone_offset function --- src/Std/Time.lean | 12 ++-- src/Std/Time/Date/PlainDate.lean | 23 ++++--- src/Std/Time/Date/Unit/Month.lean | 2 +- src/Std/Time/Date/Unit/Weekday.lean | 76 +++++++++++--------- src/Std/Time/Date/Unit/Year.lean | 2 +- src/Std/Time/Format.lean | 6 +- src/Std/Time/Notation.lean | 15 ++-- src/runtime/io.cpp | 12 +++- tests/lean/run/timeClassOperations.lean | 52 +++++++------- tests/lean/run/timeFormats.lean | 18 ++--- tests/lean/run/timeLocalDateTime.lean | 4 +- tests/lean/run/timeOperations.lean | 92 ++++++++++++------------- tests/lean/run/timeParse.lean | 36 +++++----- tests/lean/run/timeTzifParse.lean | 4 +- 14 files changed, 188 insertions(+), 166 deletions(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 0a18451a6067..964d652f48c1 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -94,12 +94,12 @@ These types provide precision down to the day level, useful for representing and ## Time These types offer precision down to the nanosecond level, useful for representing and manipulating time of day. -- **`PlainTime`**: Represents a time of day in the format `HH:mm:ss.sssssssss`. +- **`PlainTime`**: Represents a time of day in the format `HH:mm:ss,sssssssss`. ## Date and time Combines date and time into a single representation, useful for precise timestamps and scheduling. -- **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss.sssssssss`. +- **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss,sssssssss`. - **`Timestamp`**: Represents a point in time with nanosecond-level precision. It starts on the UNIX epoch and it should be used when you receive or need to send timestamps to other systems. @@ -163,10 +163,10 @@ In order to help the user build dates easily, there are a lot of macros availabl The `.sssssssss` can be ommited in most cases. - **`date% YYYY-MM-DD`**: Defines a date in the `YYYY-MM-DD` format. -- **`date% HH:mm:ss:sssssssss`**: Defines a time in the `HH:mm:ss:sssssssss` format, including fractional seconds. -- **`date% YYYY-MM-DD:HH:mm:ss:sssssssss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. -- **`date% YYYY-MM-DDTHH:mm:ss:sssssssssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ss:sssssssssZ` format. -- **`time% HH:mm:ss:sssssssss`**: Defines a time in the `HH:mm:ss` format. +- **`date% HH:mm:ss,sssssssss`**: Defines a time in the `HH:mm:ss:sssssssss` format, including fractional seconds. +- **`date% YYYY-MM-DD T HH:mm:ss,sssssssss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. +- **`date% YYYY-MM-DD T HH:mm:ss,sssssssssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ss:sssssssssZ` format. +- **`time% HH:mm:ss,sssssssss`**: Defines a time in the `HH:mm:ss` format. - **`offset% +HH:mm`**: Defines a timezone offset in the format `+HH:mm`. - **`timezone% NAME/ID offset% +HH:mm`**: Defines a timezone with a name and an offset. - **`date-spec% "format"`**: Defines a date specification format at compile time using the provided format string. diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index d0e289c5fd71..ba3a6e5b5579 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -119,20 +119,23 @@ def ofDaysSinceUNIXEpoch (day : Day.Offset) : PlainDate := Calculates the `Weekday` of a given `PlainDate` using Zeller's Congruence for the Gregorian calendar. -/ def weekday (date : PlainDate) : Weekday := - let q := date.day.toInt - let m := date.month.toInt + let q := date.day + let m := date.month let y := date.year.toInt - let y := if m < 2 then y - 1 else y - let m := if m < 2 then m + 12 else m + let y := if m.val < 2 then y - 1 else y - let k := y % 100 - let j := y.div 100 - let part := q + (13 * (m + 1)).div 5 + k + (k.div 4) - let h := part + (j.div 4) - 2*j - let d := (h + 5) % 7 + let m : Bounded.LE 1 13 := if h : m.val ≤ 1 + then (m.truncateTop h |>.add 12 |>.expandBottom (by decide)) + else m.expandTop (by decide) - .ofFin ⟨d.toNat % 7, Nat.mod_lt d.toNat (by decide)⟩ + let k := Bounded.LE.byEmod y 100 (by decide) + let j : Bounded.LE (-10) 9 := (Bounded.LE.byMod y 1000 (by decide)).ediv 100 (by decide) + let part : Bounded.LE 6 190 := q.addBounds (((m.add 1).mul_pos 13 (by decide)).ediv 5 (by decide)) |>.addBounds k |>.addBounds (k.ediv 4 (by decide)) + let h : Bounded.LE (-15) 212 := part.addBounds ((j.ediv 4 (by decide)).addBounds (j.mul_pos 2 (by decide)).neg) + let d := (h.add 5).emod 7 (by decide) + + .ofOrdinal (d.add 1) /-- Determines the era of the given `PlainDate` based on its year. diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 3b7c36bf30fd..6fd980da57a9 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -19,7 +19,7 @@ set_option linter.all true `Ordinal` represents a bounded value for months, which ranges between 1 and 12. -/ def Ordinal := Bounded.LE 1 12 - deriving Repr, BEq, LE + deriving Repr, BEq, LE, LT instance : ToString Ordinal where toString x := toString x.val diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index c86f3495aca9..c8d95ea24ec7 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -16,6 +16,7 @@ set_option linter.all true /-- Defines the enumeration for days of the week. Each variant corresponds to a day of the week, from Monday to Sunday. +As defined in the ISO 8601, It starts on monday. -/ inductive Weekday /-- Monday. -/ @@ -37,53 +38,62 @@ inductive Weekday namespace Weekday /-- -Converts a `Fin 7` representing a day index into a corresponding `Weekday`. This function is useful +`Ordinal` represents a bounded value for weekdays, which ranges between 1 and 7.. +-/ +def Ordinal := Bounded.LE 1 7 + +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 1 (1 + (6 : Nat))) n) + + +/-- +Converts a `Ordinal` representing a day index into a corresponding `Weekday`. This function is useful for mapping numerical representations to days of the week. -/ -def ofFin : Fin 7 → Weekday - | 0 => .monday - | 1 => .tuesday - | 2 => .wednesday - | 3 => .thursday - | 4 => .friday - | 5 => .saturday - | 6 => .sunday +def ofOrdinal : Ordinal → Weekday + | 1 => .monday + | 2 => .tuesday + | 3 => .wednesday + | 4 => .thursday + | 5 => .friday + | 6 => .saturday + | 7 => .sunday /-- Converts a `Weekday` to a `Nat`. -/ -def toNat : Weekday → Nat - | .monday => 0 - | .tuesday => 1 - | .wednesday => 2 - | .thursday => 3 - | .friday => 4 - | .saturday => 5 - | .sunday => 6 +def toOrdinal : Weekday → Ordinal + | .monday => 1 + | .tuesday => 2 + | .wednesday => 3 + | .thursday => 4 + | .friday => 5 + | .saturday => 6 + | .sunday => 7 /-- -Converts a `Weekday` to a `Fin`. +Converts a `Weekday` to a `Nat`. -/ -def toFin : Weekday → Nat - | .monday => 0 - | .tuesday => 1 - | .wednesday => 2 - | .thursday => 3 - | .friday => 4 - | .saturday => 5 - | .sunday => 6 +def toNat : Weekday → Nat + | .monday => 1 + | .tuesday => 2 + | .wednesday => 3 + | .thursday => 4 + | .friday => 5 + | .saturday => 6 + | .sunday => 7 /-- Converts a `Nat` to an `Option Weekday`. -/ def ofNat? : Nat → Option Weekday - | 0 => some .monday - | 1 => some .tuesday - | 2 => some .wednesday - | 3 => some .thursday - | 4 => some .friday - | 5 => some .saturday - | 6 => some .sunday + | 1 => some .monday + | 2 => some .tuesday + | 3 => some .wednesday + | 4 => some .thursday + | 5 => some .friday + | 6 => some .saturday + | 7 => some .sunday | _ => none /-- diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index 6136b34a5bcd..a747b5a3a3c5 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -70,7 +70,7 @@ def toMonths (val : Offset) : Month.Offset := val.mul 12 /-- -Determines if a year is a Gregorian Leap Year. +Determines if a year is a leap year in the proleptic gregorian calendar. -/ @[inline] def isLeap (y : Offset) : Bool := diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 73664030ad91..96f405f16b40 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -59,21 +59,21 @@ The Time24Hour format, which follows the pattern `date% hh:mm:ss` for representi in a 24-hour clock format. It uses the default value that can be parsed with the notation of dates. -/ -def leanTime24Hour : Format .any := date-spec% "'time% 'hh:mm:ss:sssssssss" +def leanTime24Hour : Format .any := date-spec% "'time% 'hh:mm:ss,sssssssss" /-- The DateTimeZone24Hour format, which follows the pattern `date% YYYY-MM-DD hh:mm:ss:sssssssss` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTime24Hour : Format (.only .GMT) := date-spec% "'date% 'YYYY-MM-DD:hh:mm:ss:sssssssss" +def leanDateTime24Hour : Format (.only .GMT) := date-spec% "'date% 'YYYY-MM-DD'T'hh:mm:ss,sssssssss" /-- The DateTimeWithZone format, which follows the pattern `date% YYYY-MM-DD hh:mm:ss:sssssssss` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithZone : Format .any := date-spec% "'date% 'YYYY-MM-DD:hh:mm:ss:sssssssssZZZZZ" +def leanDateTimeWithZone : Format .any := date-spec% "'date% 'YYYY-MM-DD'T'hh:mm:ss,sssssssssZZZZZ" /-- The Lean Date format, which follows the pattern `date% YYYY-MM-DD`. It uses the default value that can be parsed with the diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index b432dd2c85b4..e9770fe8644e 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -9,10 +9,11 @@ import Std.Time.Time import Std.Time.Zoned import Std.Time.DateTime import Std.Time.Format.Basic +import Lean.Parser namespace Std namespace Time -open Lean Parser Command Std +open Lean Parser Command Std Lean.Parser /-- Category of units that are valid inside a date. @@ -76,12 +77,12 @@ Date in `HH-mm-ss` format. syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component : time /-- -Date in `HH-mm-ss.sssssssss` format. +Date in `HH-mm-ss,sssssssss` format. -/ -syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component noWs ":" noWs date_component : time +syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component noWs "," noWs date_component : time private def parseTime : TSyntax `time -> MacroM (TSyntax `term) - | `(time| $hour:date_component:$minute:date_component:$second:date_component:$nanos:date_component) => do + | `(time| $hour:date_component:$minute:date_component:$second:date_component,$nanos:date_component) => do `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ $(← parseComponent (some 0) (some 999999999) nanos) (by decide)) | `(time| $hour:date_component:$minute:date_component:$second:date_component) => do `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ 0 (by decide)) @@ -93,9 +94,9 @@ Category of date and time together. declare_syntax_cat datetime /-- -Date and time in `YYYY-MM-DD:HH:mm:ss` format. +Date and time in `YYYY-MM-DDTHH:mm:ss,nano` format. -/ -syntax date noWs ":" noWs time : datetime +syntax date &"T" time : datetime /-- Date but using timestamp. @@ -103,7 +104,7 @@ Date but using timestamp. syntax date_component : datetime private def parseDateTime : TSyntax `datetime -> MacroM (TSyntax `term) - | `(datetime| $date:date:$time:time) => do + | `(datetime| $date:date T $time:time) => do `(Std.Time.PlainDateTime.mk $(← parseDate date) $(← parseTime time)) | `(datetime|$tm:date_component) => do `(Std.Time.PlainDateTime.ofUTCTimestamp $(← parseComponent none none tm)) diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 4a1ac3729407..80f475a76c43 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -547,9 +547,17 @@ extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { std::tm tm_info; #if defined(LEAN_WINDOWS) - localtime_s(&tm_info, &now_time_t); + errno_t err = localtime_s(&tm_info, &now_time_t); + + if (err != 0) { + return lean_io_result_mk_error(lean_decode_io_error(err, mk_string(""))); + } #else - localtime_r(&now_time_t, &tm_info); + struct tm *err = localtime_r(&now_time_t, &tm_info); + + if (err == NULL) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string(""))); + } #endif int offset_hour = tm_info.tm_gmtoff / 3600; diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index 542928002ce2..80d10ac8d69a 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -15,52 +15,52 @@ info: date% 1970-01-08 #guard_msgs in #eval date - (12 : Day.Offset) -def datetime := date% 2000-01-20:03:02:01 +def datetime := date% 2000-01-20 T 03:02:01 /-- -info: date% 2000-01-20:04:02:01:000000000 +info: date% 2000-01-20T04:02:01,000000000 -/ #guard_msgs in #eval datetime + (1 : Hour.Offset) /-- -info: date% 2000-01-20:02:02:01:000000000 +info: date% 2000-01-20T02:02:01,000000000 -/ #guard_msgs in #eval datetime - (1 : Hour.Offset) /-- -info: date% 2000-01-20:03:12:01:000000000 +info: date% 2000-01-20T03:12:01,000000000 -/ #guard_msgs in #eval datetime + (10 : Minute.Offset) /-- -info: date% 2000-01-20:02:52:01:000000000 +info: date% 2000-01-20T02:52:01,000000000 -/ #guard_msgs in #eval datetime - (10 : Minute.Offset) /-- -info: date% 2000-01-20:03:03:01:000000000 +info: date% 2000-01-20T03:03:01,000000000 -/ #guard_msgs in #eval datetime + (60 : Second.Offset) /-- -info: date% 2000-01-20:03:01:01:000000000 +info: date% 2000-01-20T03:01:01,000000000 -/ #guard_msgs in #eval datetime - (60 : Second.Offset) /-- -info: date% 2000-01-21:03:02:01:000000000 +info: date% 2000-01-21T03:02:01,000000000 -/ #guard_msgs in #eval datetime + (1 : Day.Offset) /-- -info: date% 2000-01-19:03:02:01:000000000 +info: date% 2000-01-19T03:02:01,000000000 -/ #guard_msgs in #eval datetime - (1 : Day.Offset) @@ -68,93 +68,93 @@ info: date% 2000-01-19:03:02:01:000000000 def time := time% 13:02:01 /-- -info: time% 14:02:01:000000000 +info: time% 14:02:01,000000000 -/ #guard_msgs in #eval time + (1 : Hour.Offset) /-- -info: time% 12:02:01:000000000 +info: time% 12:02:01,000000000 -/ #guard_msgs in #eval time - (1 : Hour.Offset) /-- -info: time% 13:12:01:000000000 +info: time% 13:12:01,000000000 -/ #guard_msgs in #eval time + (10 : Minute.Offset) /-- -info: time% 12:52:01:000000000 +info: time% 12:52:01,000000000 -/ #guard_msgs in #eval time - (10 : Minute.Offset) /-- -info: time% 13:03:01:000000000 +info: time% 13:03:01,000000000 -/ #guard_msgs in #eval time + (60 : Second.Offset) /-- -info: time% 13:01:01:000000000 +info: time% 13:01:01,000000000 -/ #guard_msgs in #eval time - (60 : Second.Offset) -def datetimetz := date% 2000-01-20:06:02:01-03:00 +def datetimetz := date% 2000-01-20 T 06:02:01-03:00 /-- -info: date% 2000-01-20:06:02:01:000000000-03:00 +info: date% 2000-01-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz /-- -info: date% 2000-01-22:06:02:01:000000000-03:00 +info: date% 2000-01-22T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz + (2 : Day.Offset) /-- -info: date% 2000-01-19:06:02:01:000000000-03:00 +info: date% 2000-01-19T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz - (1 : Day.Offset) /-- -info: date% 2000-01-20:07:02:01:000000000-03:00 +info: date% 2000-01-20T07:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz + (1 : Hour.Offset) /-- -info: date% 2000-01-20:05:02:01:000000000-03:00 +info: date% 2000-01-20T05:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz - (1 : Hour.Offset) /-- -info: date% 2000-01-20:06:12:01:000000000-03:00 +info: date% 2000-01-20T06:12:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz + (10 : Minute.Offset) /-- -info: date% 2000-01-20:05:52:01:000000000-03:00 +info: date% 2000-01-20T05:52:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz - (10 : Minute.Offset) /-- -info: date% 2000-01-20:06:03:01:000000000-03:00 +info: date% 2000-01-20T06:03:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz + (60 : Second.Offset) /-- -info: date% 2000-01-20:06:01:01:000000000-03:00 +info: date% 2000-01-20T06:01:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz - (60 : Second.Offset) @@ -163,7 +163,7 @@ info: date% 2000-01-20:06:01:01:000000000-03:00 info: "3600s" -/ #guard_msgs in -#eval (date% 2000-12-20:00:0:00-03:00) - (date% 2000-12-20:00:00:00-02:00) +#eval (date% 2000-12-20 T 00:0:00-03:00) - (date% 2000-12-20 T 00:00:00-02:00) def now := PlainDateTime.ofUTCTimestamp ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ def after := PlainDateTime.ofUTCTimestamp ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 28a66b5d5782..7f8db5c2abbd 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -1,7 +1,7 @@ import Std.Time open Std.Time -def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ssZZZ" +def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" def RFC1123 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" def ShortDate : Format .any := date-spec% "MM/DD/YYYY" def LongDate : Format .any := date-spec% "MMMM D, YYYY" @@ -17,7 +17,7 @@ def CustomDayTime : Format .any := date-spec% "EEE D MMM YYYY hh:mm" def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 -def date₁ := date% 2014-06-16:03:03:03(brTZ) +def date₁ := date% 2014-06-16 T 03:03:03(brTZ) def time₁ := time% 14:11:01 def time₂ := time% 03:11:01 @@ -239,31 +239,31 @@ info: date% 2002-07-14 #eval date% 2002-07-14 /-- -info: time% 14:13:12:000000000 +info: time% 14:13:12,000000000 -/ #guard_msgs in #eval time% 14:13:12 /-- -info: date% 2002-07-14:14:13:12:000000000 +info: date% 2002-07-14T14:13:12,000000000 -/ #guard_msgs in -#eval date% 2002-07-14:14:13:12 +#eval date% 2002-07-14 T 14:13:12 /-- -info: date% 2002-07-14:14:13:12:000000000+09:00 +info: date% 2002-07-14T14:13:12,000000000+09:00 -/ #guard_msgs in -#eval date% 2002-07-14:14:13:12+09:00 +#eval date% 2002-07-14 T 14:13:12+09:00 /-- info: "2002-07-14" -/ #guard_msgs in -#eval (date% 2002-07-14:14:13:12+09:00).format "YYYY-MM-DD" +#eval (date% 2002-07-14 T 14:13:12+09:00).format "YYYY-MM-DD" /-- info: "14-13-12" -/ #guard_msgs in -#eval (date% 2002-07-14:14:13:12+09:00).format "hh-mm-ss" +#eval (date% 2002-07-14 T 14:13:12+09:00).format "hh-mm-ss" diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 0d6ee0584fb9..aee24cc90e49 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -7,8 +7,8 @@ def ShortDate : Format .any := date-spec% "DD/MM/YYYY" def format (PlainDate : PlainDateTime) : String := ShortDateTime.formatBuilder PlainDate.day PlainDate.month PlainDate.year PlainDate.time.hour PlainDate.minute PlainDate.time.second def format₂ (PlainDate : PlainDate) : String := ShortDate.formatBuilder PlainDate.day PlainDate.month PlainDate.year -def date₁ := date% 1993-11-19:09:08:07 -def date₂ := date% 1993-05-09:12:59:59 +def date₁ := date% 1993-11-19 T 09:08:07 +def date₂ := date% 1993-05-09 T 12:59:59 def date₃ := date% 2024-08-16 def date₄ := date% 1500-08-16 diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index 61b00aa8c5de..8a003a652dc6 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -33,100 +33,100 @@ info: date% 1969-12-20 #guard_msgs in #eval date.subMonthsClip 1 -def datetime := date% 2000-01-20:03:02:01 +def datetime := date% 2000-01-20 T 03:02:01 /-- -info: date% 2000-01-20:04:02:01:000000000 +info: date% 2000-01-20T04:02:01,000000000 -/ #guard_msgs in #eval datetime.addHours 1 /-- -info: date% 2000-01-20:02:02:01:000000000 +info: date% 2000-01-20T02:02:01,000000000 -/ #guard_msgs in #eval datetime.subHours 1 /-- -info: date% 2000-01-20:03:12:01:000000000 +info: date% 2000-01-20T03:12:01,000000000 -/ #guard_msgs in #eval datetime.addMinutes 10 /-- -info: date% 2000-01-20:02:52:01:000000000 +info: date% 2000-01-20T02:52:01,000000000 -/ #guard_msgs in #eval datetime.subMinutes 10 /-- -info: date% 2000-01-20:03:03:01:000000000 +info: date% 2000-01-20T03:03:01,000000000 -/ #guard_msgs in #eval datetime.addSeconds 60 /-- -info: date% 2000-01-20:03:01:01:000000000 +info: date% 2000-01-20T03:01:01,000000000 -/ #guard_msgs in #eval datetime.subSeconds 60 /-- -info: date% 2000-01-21:03:02:01:000000000 +info: date% 2000-01-21T03:02:01,000000000 -/ #guard_msgs in #eval datetime.addDays 1 /-- -info: date% 2000-01-19:03:02:01:000000000 +info: date% 2000-01-19T03:02:01,000000000 -/ #guard_msgs in #eval datetime.subDays 1 /-- -info: date% 2000-02-20:03:02:01:000000000 +info: date% 2000-02-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.addMonthsClip 1 /-- -info: date% 1999-12-20:03:02:01:000000000 +info: date% 1999-12-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.subMonthsClip 1 /-- -info: date% 2000-02-20:03:02:01:000000000 +info: date% 2000-02-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.addMonthsRollOver 1 /-- -info: date% 1999-12-20:03:02:01:000000000 +info: date% 1999-12-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.subMonthsRollOver 1 /-- -info: date% 2001-01-20:03:02:01:000000000 +info: date% 2001-01-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.addYearsClip 1 /-- -info: date% 1999-01-20:03:02:01:000000000 +info: date% 1999-01-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.subYearsClip 1 /-- -info: date% 2001-01-20:03:02:01:000000000 +info: date% 2001-01-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.addYearsRollOver 1 /-- -info: date% 1999-01-20:03:02:01:000000000 +info: date% 1999-01-20T03:02:01,000000000 -/ #guard_msgs in #eval datetime.subYearsRollOver 1 @@ -134,142 +134,142 @@ info: date% 1999-01-20:03:02:01:000000000 def time := time% 13:02:01 /-- -info: time% 14:02:01:000000000 +info: time% 14:02:01,000000000 -/ #guard_msgs in #eval time.addHours 1 /-- -info: time% 12:02:01:000000000 +info: time% 12:02:01,000000000 -/ #guard_msgs in #eval time.subHours 1 /-- -info: time% 13:12:01:000000000 +info: time% 13:12:01,000000000 -/ #guard_msgs in #eval time.addMinutes 10 /-- -info: time% 12:52:01:000000000 +info: time% 12:52:01,000000000 -/ #guard_msgs in #eval time.subMinutes 10 /-- -info: time% 13:03:01:000000000 +info: time% 13:03:01,000000000 -/ #guard_msgs in #eval time.addSeconds 60 /-- -info: time% 13:01:01:000000000 +info: time% 13:01:01,000000000 -/ #guard_msgs in #eval time.subSeconds 60 -def datetimetz := date% 2000-01-20:06:02:01-03:00 +def datetimetz := date% 2000-01-20 T 06:02:01-03:00 /-- -info: date% 2000-01-20:06:02:01:000000000-03:00 +info: date% 2000-01-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz /-- -info: date% 2000-01-22:06:02:01:000000000-03:00 +info: date% 2000-01-22T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addDays 2 /-- -info: date% 2000-01-19:06:02:01:000000000-03:00 +info: date% 2000-01-19T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.subDays 1 /-- -info: date% 2000-02-20:06:02:01:000000000-03:00 +info: date% 2000-02-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMonthsClip 1 /-- -info: date% 1999-12-20:06:02:01:000000000-03:00 +info: date% 1999-12-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMonthsClip 1 /-- -info: date% 2000-02-20:06:02:01:000000000-03:00 +info: date% 2000-02-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMonthsRollOver 1 /-- -info: date% 1999-12-20:06:02:01:000000000-03:00 +info: date% 1999-12-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMonthsRollOver 1 /-- -info: date% 2001-01-20:06:02:01:000000000-03:00 +info: date% 2001-01-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: date% 2001-01-20:06:02:01:000000000-03:00 +info: date% 2001-01-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: date% 2001-01-20:06:02:01:000000000-03:00 +info: date% 2001-01-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addYearsRollOver 1 /-- -info: date% 1999-01-20:06:02:01:000000000-03:00 +info: date% 1999-01-20T06:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.subYearsRollOver 1 /-- -info: date% 2000-01-20:07:02:01:000000000-03:00 +info: date% 2000-01-20T07:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addHours 1 /-- -info: date% 2000-01-20:05:02:01:000000000-03:00 +info: date% 2000-01-20T05:02:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.subHours 1 /-- -info: date% 2000-01-20:06:12:01:000000000-03:00 +info: date% 2000-01-20T06:12:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addMinutes 10 /-- -info: date% 2000-01-20:05:52:01:000000000-03:00 +info: date% 2000-01-20T05:52:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.subMinutes 10 /-- -info: date% 2000-01-20:06:03:01:000000000-03:00 +info: date% 2000-01-20T06:03:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.addSeconds 60 /-- -info: date% 2000-01-20:06:01:01:000000000-03:00 +info: date% 2000-01-20T06:01:01,000000000-03:00 -/ #guard_msgs in #eval datetimetz.subSeconds 60 @@ -284,25 +284,25 @@ info: "86400s" info: "86399s" -/ #guard_msgs in -#eval (date% 2000-1-20:0:0:0).since (date% 2000-1-19:0:0:1) +#eval (date% 2000-1-20 T 0:0:0).since (date% 2000-1-19 T 0:0:1) /-- info: "86399s" -/ #guard_msgs in -#eval (date% 2000-1-20:0:0:0UTC).since (date% 2000-1-19:0:0:1UTC) +#eval (date% 2000-1-20 T 0:0:0UTC).since (date% 2000-1-19 T 0:0:1UTC) -def now := date% 2024-08-29:10:56:43:276801081+02:00 +def now := date% 2024-08-29 T 10:56:43,276801081+02:00 /-- -info: date% 2024-08-29:10:56:43:276801081+02:00 +info: date% 2024-08-29T10:56:43,276801081+02:00 -/ #guard_msgs in #eval now /-- -info: date% 2024-08-30:10:56:43:276801081+02:00 +info: date% 2024-08-30T10:56:43,276801081+02:00 -/ #guard_msgs in #eval now.addDays 1 diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 5dde92fb237b..33961dc215e1 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -1,7 +1,7 @@ import Std.Time open Std.Time -def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ssZZZ" +def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" def RFC1123 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" def ShortDate : Format .any := date-spec% "MM/DD/YYYY" def LongDate : Format .any := date-spec% "MMMM D, YYYY" @@ -19,23 +19,23 @@ def Full12HourWrong : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss aa Z" def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 -def date₁ := date% 2014-06-16:03:03:03(brTZ) +def date₁ := date% 2014-06-16 T 03:03:03(brTZ) def time₁ := time% 14:11:01 def time₂ := time% 03:11:01 /-- -info: "2014-06-16T03:03:03-0300" +info: "2014-06-16T03:03:03.000-03:00" -/ #guard_msgs in #eval - let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03-0300" + let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03.000-03:00" ISO8601UTC.format t.snd def tm := date₁.timestamp def date₂ := DateTime.ofTimestamp tm brTZ /-- -info: "2014-06-16T03:03:03-0300" +info: "2014-06-16T03:03:03.000-03:00" -/ #guard_msgs in #eval @@ -46,7 +46,7 @@ def tm₃ := date₁.toTimestamp def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- -info: "2014-06-16T00:00:00UTC" +info: "2014-06-16T00:00:00.000Z" -/ #guard_msgs in #eval @@ -67,7 +67,7 @@ def dateJP := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpT def dateUTC := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC /-- -info: "2024-08-15T13:28:12-0300" +info: "2024-08-15T13:28:12.000-03:00" -/ #guard_msgs in #eval @@ -75,7 +75,7 @@ info: "2024-08-15T13:28:12-0300" ISO8601UTC.format t.snd /-- -info: "2024-08-16T01:28:00UTC" +info: "2024-08-16T01:28:00.000Z" -/ #guard_msgs in #eval @@ -83,7 +83,7 @@ info: "2024-08-16T01:28:00UTC" ISO8601UTC.format t.snd /-- -info: "-0001-12-31T22:28:12+0900" +info: "-0001-12-31T22:28:12.000+09:00" -/ #guard_msgs in #eval @@ -91,7 +91,7 @@ info: "-0001-12-31T22:28:12+0900" ISO8601UTC.format (t.snd.convertTimeZone jpTZ) /-- -info: "-0001-12-31T09:28:12-0300" +info: "-0001-12-31T09:28:12.000-03:00" -/ #guard_msgs in #eval @@ -107,7 +107,7 @@ info: "Thu 15 Aug 2024 16:28" CustomDayTime.format t2.snd /-- -info: "2024-08-16T13:28:00UTC" +info: "2024-08-16T13:28:00.000Z" -/ #guard_msgs in #eval @@ -115,7 +115,7 @@ info: "2024-08-16T13:28:00UTC" ISO8601UTC.format t5.snd /-- -info: "2024-08-16T01:28:12+0900" +info: "2024-08-16T01:28:12.000+09:00" -/ #guard_msgs in #eval @@ -123,7 +123,7 @@ info: "2024-08-16T01:28:12+0900" ISO8601UTC.format (t6.snd.convertTimeZone jpTZ) /-- -info: "2024-08-16T01:28:12+0900" +info: "2024-08-16T01:28:12.000+09:00" -/ #guard_msgs in #eval @@ -147,7 +147,7 @@ def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ def dateUTC₁ := DateTime.ofLocalDateTime PlainDate .UTC /-- -info: "2024-08-15T14:03:47-0300" +info: "2024-08-15T14:03:47.000-03:00" -/ #guard_msgs in #eval @@ -155,7 +155,7 @@ info: "2024-08-15T14:03:47-0300" ISO8601UTC.format t.snd /-- -info: "2024-08-15T14:03:47+0900" +info: "2024-08-15T14:03:47.000+09:00" -/ #guard_msgs in #eval @@ -163,7 +163,7 @@ info: "2024-08-15T14:03:47+0900" ISO8601UTC.format t1.snd /-- -info: "2014-06-16T03:03:03-0300" +info: "2014-06-16T03:03:03.000-03:00" -/ #guard_msgs in #eval @@ -171,7 +171,7 @@ info: "2014-06-16T03:03:03-0300" ISO8601UTC.format t2.snd /-- -info: Except.ok "1993-05-10T10:30:23+0300" +info: Except.ok "1993-05-10T10:30:23.000+03:00" -/ #guard_msgs in #eval @@ -179,7 +179,7 @@ info: Except.ok "1993-05-10T10:30:23+0300" (ISO8601UTC.format ·.snd) <$> t2 /-- -info: Except.ok "1993-05-10T22:30:23+0300" +info: Except.ok "1993-05-10T22:30:23.000+03:00" -/ #guard_msgs in #eval diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index 69d60fc24a25..efd25062834d 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -59,13 +59,13 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: some (date% 1969-12-30:21:00:00:000000000-03:00) +info: some (date% 1969-12-30T21:00:00,000000000-03:00) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules /-- -info: some (date% 2012-12-10:00:35:47:000000000-02:00) +info: some (date% 2012-12-10T00:35:47,000000000-02:00) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From b32c8a6390fc6d692bd8c65957d2f981d027129b Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 30 Aug 2024 18:12:21 -0300 Subject: [PATCH 057/118] fix: avoid creating new tokens Co-authored-by: Mac Malone --- src/Std/Time/Notation.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index e9770fe8644e..cf36393ce394 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -135,7 +135,7 @@ declare_syntax_cat zone /-- Timezone in `+HHmm`, or `-HHmm` format. -/ -syntax (offset <|> "UTC" <|> "GMT" <|> "Z" <|> ("(" term ")")) : zone +syntax (offset <|> &"UTC" <|> &"GMT" <|> &"Z" <|> ("(" term ")")) : zone private def parseZone : TSyntax `zone -> MacroM (TSyntax `term) | `(zone| ($term)) => do `($term) From 60f5b923bb3c330d09bfb1df423ae3c93355e952 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 2 Sep 2024 11:42:04 -0300 Subject: [PATCH 058/118] refactor: rename err to tm_ptr in io.cpp --- src/runtime/io.cpp | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 80f475a76c43..f2742eff00b9 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -553,9 +553,9 @@ extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { return lean_io_result_mk_error(lean_decode_io_error(err, mk_string(""))); } #else - struct tm *err = localtime_r(&now_time_t, &tm_info); + struct tm *tm_ptr = localtime_r(&now_time_t, &tm_info); - if (err == NULL) { + if (tm_ptr == NULL) { return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string(""))); } #endif @@ -564,8 +564,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { int offset_seconds = tm_info.tm_gmtoff; lean_object *lean_offset = lean_alloc_ctor(0, 2, 0); - lean_ctor_set(lean_offset, 0, lean_int_to_int(static_cast(offset_hour))); - lean_ctor_set(lean_offset, 1, lean_int_to_int(static_cast(offset_seconds))); + lean_ctor_set(lean_offset, 0, lean_int_to_int(offset_hour)); + lean_ctor_set(lean_offset, 1, lean_int_to_int(offset_seconds)); lean_object *lean_tz = lean_alloc_ctor(0, 2, 1); lean_ctor_set(lean_tz, 0, lean_offset); From 27e9893ff5747af15b5c5a93824863d80481fa34 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 2 Sep 2024 11:43:52 -0300 Subject: [PATCH 059/118] refactor: fix names of things that were changed with search-and-replace incorrectly --- src/Std/Time/Zoned/ZoneRules.lean | 22 +++++++++++----------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 0a1cfb75b43c..7a1e4f287535 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -47,7 +47,7 @@ inductive StdWall /-- Represents a type of local time, including offset and daylight saving information. -/ -structure PlainTimeType where +structure LocalTimeType where /-- The offset from GMT for this local time. @@ -80,15 +80,15 @@ structure PlainTimeType where identifier : String deriving Repr, Inhabited -namespace PlainTimeType +namespace LocalTimeType /-- -Gets the `TimeZone` offset from a `PlainTimeType`. +Gets the `TimeZone` offset from a `LocalTimeType`. -/ -def getTimeZone (time : PlainTimeType) : TimeZone := +def getTimeZone (time : LocalTimeType) : TimeZone := ⟨time.gmtOffset, time.abbreviation, time.isDst⟩ -end PlainTimeType +end LocalTimeType /-- Represents a leap second event, including the time of the transition and the correction applied. @@ -119,7 +119,7 @@ structure Transition where /-- The local time type associated with this transition. -/ - PlainTimeType : PlainTimeType + localTimeType : LocalTimeType deriving Repr, Inhabited /-- @@ -130,7 +130,7 @@ structure ZoneRules where /-- The array of local time types for the time zone. -/ - PlainTimes : Array PlainTimeType + localTimes : Array LocalTimeType /-- The array of transitions for the time zone. @@ -149,15 +149,15 @@ namespace Transition Create a TimeZone from a Transition. -/ def createTimeZoneFromTransition (transition : Transition) : TimeZone := - let offset := transition.PlainTimeType.gmtOffset - let name := transition.PlainTimeType.abbreviation - TimeZone.mk offset name transition.PlainTimeType.isDst + let offset := transition.localTimeType.gmtOffset + let name := transition.localTimeType.abbreviation + TimeZone.mk offset name transition.localTimeType.isDst /-- Applies the transition to a Timestamp. -/ def apply (timestamp : Timestamp) (transition : Transition) : Timestamp := - let offsetInSeconds := transition.PlainTimeType.gmtOffset.hour.mul 3600 |>.add transition.PlainTimeType.gmtOffset.second + let offsetInSeconds := transition.localTimeType.gmtOffset.hour.mul 3600 |>.add transition.localTimeType.gmtOffset.second let PlainTimestamp := timestamp.addSeconds offsetInSeconds PlainTimestamp From 8d23eb014e3249cd6d0d27a094c798710f14140a Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 2 Sep 2024 11:44:44 -0300 Subject: [PATCH 060/118] fix: name of the localPath that is wrong --- src/Std/Time/Zoned/Database/TZdb.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 6e78bb3e6b23..e1c19ce999ff 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -24,7 +24,7 @@ structure TZdb where The path to the local timezone file. This is typically a symlink to a file within the timezone database that corresponds to the current local time zone. -/ - localPath : System.FilePath := "/etc/PlainTime" + localPath : System.FilePath := "/etc/localtime" /-- The path to the directory containing all available time zone files. These files define various From b84a0e23e3b28281ff40fc5c3407219d370019de Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 2 Sep 2024 11:48:09 -0300 Subject: [PATCH 061/118] style: changes all the lambdas for fun --- src/Std/Time/Format.lean | 8 ++++---- src/Std/Time/Time/PlainTime.lean | 8 ++++---- src/Std/Time/Zoned/Database/TzIf.lean | 2 +- src/Std/Time/Zoned/DateTime.lean | 8 ++++---- src/Std/Time/Zoned/ZoneRules.lean | 2 +- 5 files changed, 14 insertions(+), 14 deletions(-) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 96f405f16b40..5c37537e429b 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -137,7 +137,7 @@ def format (date : PlainDate) (format : String) : String := Parses a date string in the American format (`MM/DD/YYYY`) and returns a `PlainDate`. -/ def fromAmericanDateString (input : String) : Except String PlainDate := do - Formats.americanDate.parseBuilder (λm d y => PlainDate.ofYearMonthDay y m d) input + Formats.americanDate.parseBuilder (fun m d y => PlainDate.ofYearMonthDay y m d) input /-- Converts a Date in the American format (`MM/DD/YYYY`) into a `String`. @@ -149,7 +149,7 @@ def toAmericanDateString (input : PlainDate) : String := Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. -/ def fromSQLDateString (input : String) : Except String PlainDate := do - Formats.sqlDate.parseBuilder (λy m d => PlainDate.ofYearMonthDay y m d) input + Formats.sqlDate.parseBuilder (PlainDate.ofYearMonthDay) input /-- Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. @@ -202,7 +202,7 @@ def format (time : PlainTime) (format : String) : String := Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. -/ def fromTime24Hour (input : String) : Except String PlainTime := - Formats.time24Hour.parseBuilder (λh m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input + Formats.time24Hour.parseBuilder (fun h m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input /-- Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). @@ -214,7 +214,7 @@ def toTime24Hour (input : PlainTime) : String := Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. -/ def fromLeanTime24Hour (input : String) : Except String PlainTime := - Formats.leanTime24Hour.parseBuilder (λh m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input + Formats.leanTime24Hour.parseBuilder (fun h m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input /-- Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 9508b060b24a..b8663cdc712b 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -48,7 +48,7 @@ instance : Inhabited PlainTime where instance : BEq PlainTime where beq x y := x.hour.snd.val == y.hour.snd.val && x.minute == y.minute - && x.second.snd.val == y.second.snd.val && x.nano.val == y.nano.val + && x.second.snd.val == y.second.snd.val && x.nano == y.nano namespace PlainTime @@ -75,13 +75,13 @@ def ofHourMinuteSeconds? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) else none /-- -Creates a `PlainTime` value from hours, minutes, and seconds, setting nanoseconds to zero. +Creates a `PlainTime` value from hours, minutes, seconds and nanoseconds. -/ def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Ordinal) (second : Second.Ordinal false) (nano : Nanosecond.Ordinal) : PlainTime := by refine ⟨Sigma.mk false hour, minute, Sigma.mk false second, nano, ?_⟩ constructor - exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr second.property.right)) x) - exact λx => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) + exact fun x => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr second.property.right)) x) + exact fun x => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) /-- Creates a `PlainTime` value from hours, minutes, and seconds. Return `none` if its invalid. diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index d07471955ad3..e54b70479256 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -286,7 +286,7 @@ private def parseFooter : Parser (Option String) := do if char = 0x0A then pure () else return none - let tzString ← many (satisfy (λ c => c ≠ 0x0A)) + let tzString ← many (satisfy (· ≠ 0x0A)) let mut str := "" for byte in tzString do diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index dbea6fbcf2f9..f23f11342a09 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -33,7 +33,7 @@ instance : BEq (DateTime tz) where beq x y := x.timestamp == y.timestamp instance : Inhabited (DateTime tz) where - default := ⟨Inhabited.default, Thunk.mk λ_ => Inhabited.default⟩ + default := ⟨Inhabited.default, Thunk.mk fun _ => Inhabited.default⟩ namespace DateTime @@ -42,7 +42,7 @@ Creates a new `DateTime` out of a `Timestamp` that is in a `TimeZone`. -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := - DateTime.mk tm (Thunk.mk <| λ_ => (tm.addSeconds tz.toSeconds).toPlainDateTime) + DateTime.mk tm (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds).toPlainDateTime) /-- Creates a new zone aware `Timestamp` out of a `DateTime`. @@ -65,7 +65,7 @@ to UTC. If you're using hte PlainDateTime @[inline] def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : DateTime tz := let tm := Timestamp.ofPlainDateTime date - DateTime.mk tm (Thunk.mk <| λ_ => date.addSeconds tz.toSeconds) + DateTime.mk tm (Thunk.mk fun _ => date.addSeconds tz.toSeconds) /-- Creates a new `DateTime` out of a `PlainDateTime`. It assumes that the `PlainDateTime` it's relative @@ -74,7 +74,7 @@ to the @[inline] def ofLocalDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := let tm := Timestamp.ofPlainDateTime date - DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk <| λ_ => date) + DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk fun _ => date) /-- Add `Hour.Offset` to a `DateTime`. diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 7a1e4f287535..6962f85c62b0 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -171,7 +171,7 @@ If the timestamp falls between two transitions, it returns the most recent trans -/ def findTransitionForTimestamp (zoneRules : ZoneRules) (timestamp : Timestamp) : Option Transition := let value := timestamp.toSecondsSinceUnixEpoch - if let some idx := zoneRules.transitions.findIdx? (λ t => t.time.val > value.val) + if let some idx := zoneRules.transitions.findIdx? (fun t => t.time.val > value.val) then zoneRules.transitions.get? (idx - 1) else zoneRules.transitions.back? From 3031f2650ec1bc498a200a02991e14807dd7992b Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 2 Sep 2024 12:00:55 -0300 Subject: [PATCH 062/118] fix: wrong name of function --- src/Std/Time/Zoned/ZonedDateTime.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 05c08e5faedb..bf9ae9857716 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -51,7 +51,7 @@ Creates a new `ZonedDateTime` out of a `Timestamp` def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do let transition ← rules.findTransitionForTimestamp tm let tm := rules.applyLeapSeconds tm - return ofTimestamp tm transition.PlainTimeType.getTimeZone + return ofTimestamp tm transition.localTimeType.getTimeZone /-- Changes the `TimeZone` to a new one. From 16b4a7271b0ba5d6cea5c47ac8c50f1709c8e213 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 2 Sep 2024 12:23:40 -0300 Subject: [PATCH 063/118] fix: function names, ofOrdinal and others --- src/Std/Time/Date/Unit/Month.lean | 8 +++----- src/Std/Time/Format/Basic.lean | 10 +++++----- src/Std/Time/Zoned/Database/Basic.lean | 12 ++++++------ src/Std/Time/Zoned/Database/TzIf.lean | 16 ++++++++-------- 4 files changed, 22 insertions(+), 24 deletions(-) diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 6fd980da57a9..6c2865ac15f6 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -238,10 +238,7 @@ def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day if h : cumulative.val < ordinal.val ∧ ordinal.val ≤ cumulative.val + days.val then let bounded := Bounded.LE.mk ordinal.val h |>.sub cumulative - - let bounded : Bounded.LE 1 days.val := bounded.cast - (by simp [Int.add_comm _ 1, Int.add_assoc, ←Int.sub_eq_add_neg];) - (by simp [Int.add_comm _ days.val, Int.add_assoc, ←Int.sub_eq_add_neg];) + let bounded : Bounded.LE 1 days.val := bounded.cast (by omega) (by omega) let ⟨left, right⟩ := bounded.property let days₁ : Day.Ordinal := ⟨bounded.val, And.intro left (Int.le_trans right days.property.right)⟩ @@ -252,7 +249,8 @@ def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day let ⟨day, valid⟩ := clipDay leap 1 1 ⟨⟨1, day⟩, valid⟩ else - go ⟨idx.val + 1, Nat.succ_le_succ (Nat.not_le.mp h)⟩ cumulative + go ⟨idx.val + 1, Nat.succ_le_succ (Nat.not_le.mp h)⟩ (cumulative + (Fin.ofNat days.val.toNat)) + termination_by 12 - idx.val go 0 0 diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index cc8a02fcfe74..5061fbb0dc20 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -265,7 +265,7 @@ private def isFormatStart : Char → Bool := Char.isAlpha private def pnumber : Parser Nat := do let numbers ← manyChars digit - return numbers.foldl (λacc char => acc * 10 + (char.toNat - 48)) 0 + return numbers.foldl (fun acc char => acc * 10 + (char.toNat - 48)) 0 private def parseFormatPart : Parser FormatPart := (.modifier <$> parseModifier) @@ -428,7 +428,7 @@ private def FormatType (result : Type) : FormatString → Type | .string _ :: xs => (FormatType result xs) | [] => result -private def position : Parser Nat := λs => (ParseResult.success s (s.pos.byteIdx)) +private def position : Parser Nat := fun s => (ParseResult.success s (s.pos.byteIdx)) private def size (data : Parser α) : Parser (α × Nat) := do let st ← position @@ -545,7 +545,7 @@ private def parserWithFormat : (typ: Modifier) → Parser (SingleFormatType typ) | .M => transform Bounded.LE.ofInt number | .DD => transform Bounded.LE.ofInt twoDigit | .D => transform Bounded.LE.ofInt number - | .d => transform Bounded.LE.ofInt (orElse twoDigit (λ_ => pchar ' ' *> (singleDigit))) + | .d => transform Bounded.LE.ofInt (orElse twoDigit (fun _ => pchar ' ' *> (singleDigit))) | .EEEE => parseWeekdayUnnabrev | .EEE => parseWeekday | .hh => Sigma.mk true <$> transform Bounded.LE.ofInt twoDigit @@ -568,7 +568,7 @@ private def parserWithFormat : (typ: Modifier) → Parser (SingleFormatType typ) | .ZZZZ => timeOffset false | .ZZZ => timeOrUTC "UTC" false | .Z => timeOrUTC "Z" true - | .z => many1Chars (satisfy (λc => c == ' ' || c.isAlpha)) + | .z => many1Chars (satisfy (fun c => c == ' ' || c.isAlpha)) private structure DateBuilder where tzName : String := "Greenwich Mean Time" @@ -665,7 +665,7 @@ Formats the date using the format into a String. -/ def formatBuilder (format : Format aw) : FormatType String format.string := let rec go (data : String) : (format : FormatString) → FormatType String format - | .modifier x :: xs => λres => go (data ++ formatPart x res) xs + | .modifier x :: xs => fun res => go (data ++ formatPart x res) xs | .string x :: xs => go (data ++ x) xs | [] => data go "" format.string diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index fc4d848c6166..e3e4ee8a0229 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -52,8 +52,8 @@ def convertLeapSecond (tz : TZif.LeapSecond) : LeapSecond := /-- Converts a PlainTime. -/ -def convertPlainTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option PlainTimeType := do - let localType ← tz.PlainTimeTypes.get? index +def convertPlainTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do + let localType ← tz.localTimeTypes.get? index let abbreviation ← tz.abbreviations.getD index "Unknown" let wallflag := convertWall (tz.stdWallIndicators.getD index true) let utLocal := convertUt (tz.utLocalIndicators.getD index true) @@ -70,17 +70,17 @@ def convertPlainTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Op /-- Converts a transition. -/ -def convertTransition (times: Array PlainTimeType) (index : Nat) (tz : TZif.TZifV1) : Option Transition := do +def convertTransition (times: Array LocalTimeType) (index : Nat) (tz : TZif.TZifV1) : Option Transition := do let time := tz.transitionTimes.get! index let time := Second.Offset.ofInt time let indice := tz.transitionIndices.get! index - return { time, PlainTimeType := times.get! indice.toNat } + return { time, localTimeType := times.get! indice.toNat } /-- Converts a `TZif.TZifV1` structure to a `ZoneRules` structure. -/ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := do - let mut times : Array PlainTimeType := #[] + let mut times : Array LocalTimeType := #[] for i in [0:tz.header.typecnt.toNat] do if let some result := convertPlainTime i tz id @@ -95,7 +95,7 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := then transitions := transitions.push result else .error s!"cannot convert transtiion {i} of the file" - .ok { leapSeconds, transitions, PlainTimes := times } + .ok { leapSeconds, transitions, localTimes := times } /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index e54b70479256..0b36e1f1e22c 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -123,7 +123,7 @@ structure TZifV1 where /-- The array of local time type structures. -/ - PlainTimeTypes : Array PlainTimeType + localTimeTypes : Array PlainTimeType /-- The array of abbreviation strings used by local time types. @@ -233,7 +233,7 @@ private def parseTransitionTimes (size : Parser Int32) (n : UInt32) : Parser (Ar private def parseTransitionIndices (n : UInt32) : Parser (Array UInt8) := manyN (n.toNat) pu8 -private def parsePlainTimeTypes (n : UInt32) : Parser (Array PlainTimeType) := +private def parseLocalTimeType (n : UInt32) : Parser (Array PlainTimeType) := manyN (n.toNat) parsePlainTimeType private def parseAbbreviations (times : Array PlainTimeType) (n : UInt32) : Parser (Array String) := do @@ -264,8 +264,8 @@ private def parseTZifV1 : Parser TZifV1 := do let transitionTimes ← parseTransitionTimes pi32 header.timecnt let transitionIndices ← parseTransitionIndices header.timecnt - let PlainTimeTypes ← parsePlainTimeTypes header.typecnt - let abbreviations ← parseAbbreviations PlainTimeTypes header.charcnt + let localTimeTypes ← parseLocalTimeType header.typecnt + let abbreviations ← parseAbbreviations localTimeTypes header.charcnt let leapSeconds ← parseLeapSeconds pi32 header.leapcnt let stdWallIndicators ← parseIndicators header.isstdcnt let utLocalIndicators ← parseIndicators header.isutcnt @@ -274,7 +274,7 @@ private def parseTZifV1 : Parser TZifV1 := do header transitionTimes transitionIndices - PlainTimeTypes + localTimeTypes abbreviations leapSeconds stdWallIndicators @@ -299,8 +299,8 @@ private def parseTZifV2 : Parser (Option TZifV2) := optional do let transitionTimes ← parseTransitionTimes pi64 header.timecnt let transitionIndices ← parseTransitionIndices header.timecnt - let PlainTimeTypes ← parsePlainTimeTypes header.typecnt - let abbreviations ← parseAbbreviations PlainTimeTypes header.charcnt + let localTimeTypes ← parseLocalTimeType header.typecnt + let abbreviations ← parseAbbreviations localTimeTypes header.charcnt let leapSeconds ← parseLeapSeconds pi64 header.leapcnt let stdWallIndicators ← parseIndicators header.isstdcnt let utLocalIndicators ← parseIndicators header.isutcnt @@ -309,7 +309,7 @@ private def parseTZifV2 : Parser (Option TZifV2) := optional do header transitionTimes transitionIndices - PlainTimeTypes + localTimeTypes abbreviations leapSeconds stdWallIndicators From cab1802de884a9ecd96351102c00ee94698b82fa Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 20 Sep 2024 00:34:36 -0300 Subject: [PATCH 064/118] reafactor: notation and format --- src/Std/Time.lean | 129 +- src/Std/Time/Date/PlainDate.lean | 141 +- src/Std/Time/Date/Unit/Day.lean | 112 +- src/Std/Time/Date/Unit/Month.lean | 195 ++- src/Std/Time/Date/Unit/Week.lean | 35 + src/Std/Time/Date/Unit/Weekday.lean | 64 +- src/Std/Time/Date/Unit/Year.lean | 16 +- src/Std/Time/DateTime.lean | 18 +- src/Std/Time/DateTime/PlainDateTime.lean | 14 +- src/Std/Time/Format.lean | 187 ++- src/Std/Time/Format/Basic.lean | 1681 ++++++++++++++++------ src/Std/Time/Internal/Bounded.lean | 31 +- src/Std/Time/Internal/UnitVal.lean | 7 + src/Std/Time/Notation.lean | 358 ++--- src/Std/Time/Notation/Spec.lean | 114 ++ src/Std/Time/Time/HourMarker.lean | 27 +- src/Std/Time/Time/PlainTime.lean | 114 +- src/Std/Time/Time/Unit/Basic.lean | 40 +- src/Std/Time/Time/Unit/Hour.lean | 44 +- src/Std/Time/Time/Unit/Millisecond.lean | 38 + src/Std/Time/Time/Unit/Minute.lean | 24 +- src/Std/Time/Time/Unit/Nanosecond.lean | 68 +- src/Std/Time/Time/Unit/Second.lean | 26 +- src/Std/Time/Zoned.lean | 17 +- src/Std/Time/Zoned/DateTime.lean | 62 +- src/Std/Time/Zoned/Offset.lean | 1 - src/Std/Time/Zoned/TimeZone.lean | 17 +- src/Std/Time/Zoned/ZoneRules.lean | 7 +- src/Std/Time/Zoned/ZonedDateTime.lean | 32 +- tests/lean/run/timeClassOperations.lean | 68 +- tests/lean/run/timeFormats.lean | 74 +- tests/lean/run/timeLocalDateTime.lean | 13 +- tests/lean/run/timeOperations.lean | 109 +- tests/lean/run/timeParse.lean | 68 +- 34 files changed, 2613 insertions(+), 1338 deletions(-) create mode 100644 src/Std/Time/Notation/Spec.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 964d652f48c1..ff922c5c1381 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -116,58 +116,93 @@ Represents spans of time and the difference between two points in time. # Formats -Format strings are used to convert between `String` representations and date/time types, e.g, `YYYY-MM-DD'T'hh:mm:ss.sssZ`. -The supported formats include: +Format strings are used to convert between `String` representations and date/time types, such as `yyyy-MM-dd'T'HH:mm:ss.sssZ`. +The table below outlines the available format specifiers. Some specifiers can be modified by repeating characters to +adjust truncation and offsets. Some of them when a character is repeated `n` times, it truncates the corresponding value to +`n` characters, usually when not specified a quantity. -- **Year:** - - `YYYY`: Four-digit year (e.g., 2024) - - `YY`: Two-digit year (e.g., 24 for 2024) -- **Month:** - - `MMMM`: Full month name (e.g., January, February) - - `MMM`: Abbreviated month name (e.g., Jan, Feb) - - `MM`: Two-digit month (e.g., 01 for January) - - `M`: One or two-digit month (e.g., 1 for January, 10 for October) -- **Day:** - - `DD`: Two-digit day of the month (e.g., 01, 02) - - `D`: One or two-digit day of the month (e.g., 1, 2) - - `d`: One or two-digit day of the month with space padding (e.g., " 1", "12") -- **Day of Week:** - - `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday) - - `EEE`: Abbreviated day of the week (e.g., Mon, Tue) -- **Hour:** - - `hh`: Two-digit hour in 24-hour format (e.g., 13, 14) - - `h`: One or two-digit hour in 24-hour format (e.g., 1, 2) - - `HH`: Two-digit hour in 12-hour format (e.g., 01, 02) - - `H`: One or two-digit hour in 12-hour format (e.g., 1, 2) -- **AM/PM Indicator:** - - `AA`: Uppercase AM/PM (e.g., AM, PM) - - `aa`: Lowercase am/pm (e.g., am, pm) -- **Minute:** - - `mm`: Two-digit minute (e.g., 01, 02) - - `m`: One or two-digit minute (e.g., 1, 2) -- **Second:** - - `sssssssss` : Nine-digit nanosseconds (e.g., 000000009) - - `sss`: Three-digit milliseconds (e.g., 001, 202) - - `ss`: Two-digit second (e.g., 01, 02) - - `s`: One or two-digit second (e.g., 1, 2) -- **Timezone:** - - `ZZZZZ`: Full timezone offset with hours and minutes (e.g., +03:00) - - `ZZZZ`: Timezone offset without the colon (e.g., +0300) - - `ZZZ`: Like `ZZZZ`, but with "UTC" for UTC - - `Z`: Like `ZZZZZ`, but with "Z" for UTC - - `z`: Timezone name (e.g., Brasilia Standard Time) +The supported formats include: +- `G`: Represents the era, such as AD (Anno Domini) or BC (Before Christ). + - `G`, `GG`, `GGG` (short): Displays the era in a short format (e.g., "AD"). + - `GGGG` (full): Displays the era in a full format (e.g., "Anno Domini"). + - `GGGGG` (narrow): Displays the era in a narrow format (e.g., "A"). +- `y`: Represents the year of the era. + - `yy`: Displays the year in a two-digit format, showing the last two digits (e.g., "04" for 2004). + - `yyyy`: Displays the year in a four-digit format (e.g., "2004"). + - `yyyy+`: Extended format for years with more than four digits. +- `D`: Represents the day of the year. +- `M`: Represents the month of the year, displayed as either a number or text. + - `M`, `MM`: Displays the month as a number, with `MM` zero-padded (e.g., "7" for July, "07" for July with padding). + - `MMM`: Displays the abbreviated month name (e.g., "Jul"). + - `MMMM`: Displays the full month name (e.g., "July"). + - `MMMMM`: Displays the month in a narrow form (e.g., "J" for July). +- `d`: Represents the day of the month. +- `Q`: Represents the quarter of the year. + - `Q`, `QQ`: Displays the quarter as a number (e.g., "3", "03"). + - `QQQ` (short): Displays the quarter as an abbreviated text (e.g., "Q3"). + - `QQQQ` (full): Displays the full quarter text (e.g., "3rd quarter"). + - `QQQQQ` (narrow): Displays the full quarter text (e.g., "3rd quarter"). +- `w`: Represents the week of the week-based year (e.g., "27"). +- `W`: Represents the week of the month (e.g., "4"). +- `E`: Represents the day of the week as text. + - `E`, `EE`, `EEE`: Displays the abbreviated day name (e.g., "Tue"). + - `EEEE`: Displays the full day name (e.g., "Tuesday"). + - `EEEEE`: Displays the narrow day name (e.g., "T" for Tuesday). +- `e`: Represents the localized day of the week. + - `e`, `ee`: Displays the day of the week as a number, starting from 1 (Monday) to 7 (Sunday). + - `eee`, `eeee`, `eeeee`: Displays the localized day of the week as text (same format as `E`). +- `F`: Represents the aligned week of the month (e.g., "3"). +- `a`: Represents the AM or PM designation of the day. + - `a`, `aa`, `aaa`: Displays AM or PM in a concise format (e.g., "PM"). + - `aaaa`: Displays the full AM/PM designation (e.g., "PM"). +- `B`: Represents the period of the day (e.g., "in the morning", "at night"). + - `b`, `bb`, `bbb` (short): Displays the abbreviated period of the day (e.g., "morning"). + - `bbbb` (full): Displays the full period of the day (e.g., "in the morning"). + - `bbbbb` (narrow): Displays a narrow version of the period of the day (e.g., "m" for morning). +- `h`: Represents the hour of the AM/PM clock (1-12) (e.g., "12"). +- `K`: Represents the hour of the AM/PM clock (0-11) (e.g., "0"). +- `k`: Represents the hour of the day in a 1-24 format (e.g., "24"). +- `H`: Represents the hour of the day in a 0-23 format (e.g., "0"). +- `m`: Represents the minute of the hour (e.g., "30"). +- `s`: Represents the second of the minute (e.g., "55"). +- `S`: Represents a fraction of a second, typically displayed as a decimal number (e.g., "978" for milliseconds). +- `A`: Represents the millisecond of the day (e.g., "1234"). +- `n`: Represents the nanosecond of the second (e.g., "987654321"). +- `N`: Represents the nanosecond of the day (e.g., "1234000000"). +- `V`: Represents the time zone ID, which could be a city-based zone (e.g., "America/Los_Angeles"), a UTC marker (`"Z"`), or a specific offset (e.g., "-08:30"). +- `z`: Represents the time zone name. + - `z`, `zz`, `zzz`: Shows an abbreviated time zone name (e.g., "PST" for Pacific Standard Time). + - `zzzz`: Displays the full time zone name (e.g., "Pacific Standard Time"). +- `O`: Represents the localized zone offset in the format "GMT" followed by the time difference from UTC. + - `O`: Displays the GMT offset in a simple format (e.g., "GMT+8"). + - `O`: Displays the full GMT offset, including hours and minutes (e.g., "GMT+08:00"). +- `X`: Represents the zone offset, using 'Z' for UTC and specifying the offset. + - `X`: Displays the hour offset (e.g., "-08"). + - `XX`: Displays the hour and minute offset without a colon (e.g., "-0830"). + - `XXX`: Displays the hour and minute offset with a colon (e.g., "-08:30"). + - `XXXX`: Displays the hour, minute, and second offset without a colon (e.g., "-083045"). + - `XXXXX`: Displays the hour, minute, and second offset with a colon (e.g., "-08:30:45"). + - It also uses `Z` to represent UTC without any offset. +- `x`: Represents the zone offset without using 'Z' for zero offsets. + - `x`: Displays the hour offset (e.g., "+08"). + - `xx`: Displays the hour and minute offset without a colon (e.g., "+0830"). + - `xxx`: Displays the hour and minute offset with a colon (e.g., "+08:30"). + - `xxxx`: Displays the hour, minute, and second offset without a colon (e.g., "+083045"). + - `xxxxx`: Displays the hour, minute, and second offset with a colon (e.g., "+08:30:45"). +- `Z`: Represents the zone offset, with 'Z' for UTC and an optional time offset. + - `Z`: Displays the hour and minute offset without a colon (e.g., "+0800"). + - `ZZ`: Displays "GMT" followed by the time offset (e.g., "GMT+08:00" or Z). + - `ZZZ`: Displays the full hour, minute, and second offset with a colon (e.g., "+08:30:45" or Z). # Macros In order to help the user build dates easily, there are a lot of macros available for creating dates. The `.sssssssss` can be ommited in most cases. -- **`date% YYYY-MM-DD`**: Defines a date in the `YYYY-MM-DD` format. -- **`date% HH:mm:ss,sssssssss`**: Defines a time in the `HH:mm:ss:sssssssss` format, including fractional seconds. -- **`date% YYYY-MM-DD T HH:mm:ss,sssssssss`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. -- **`date% YYYY-MM-DD T HH:mm:ss,sssssssssZ`**: Defines a datetime with a timezone in the `YYYY-MM-DDTHH:mm:ss:sssssssssZ` format. -- **`time% HH:mm:ss,sssssssss`**: Defines a time in the `HH:mm:ss` format. -- **`offset% +HH:mm`**: Defines a timezone offset in the format `+HH:mm`. -- **`timezone% NAME/ID offset% +HH:mm`**: Defines a timezone with a name and an offset. -- **`date-spec% "format"`**: Defines a date specification format at compile time using the provided format string. +- **`date(yyyy-MM-dd)`**: Defines a date in the `YYYY-MM-DD` format. +- **`time(HH:mm:ss.sssssssss)`**: Defines a time in the `HH:mm:ss:sssssssss` format, including fractional seconds. +- **`datetime("yyy-MM-ddTHH:mm:ss.sssssssss")`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. +- **`offset("+HH:mm")`**: Defines a timezone offset in the format `+HH:mm`. +- **`timezone("NAME/ID ZZZ")`**: Defines a timezone with a name and an offset. +- **`datespec("format")`**: Defines a date specification format at compile time using the provided format string. -/ diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index ba3a6e5b5579..102c3cd6aa63 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -52,37 +52,11 @@ Creates a `PlainDate` by clipping the day to ensure validity. This function forc valid by adjusting the day to fit within the valid range to fit the given month and year. -/ def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := - let ⟨day, valid⟩ := month.clipDay year.isLeap day - PlainDate.mk year month day valid - -/-- -Creates a `PlainDate` by rolling over the extra days to the next month. --/ -def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := by - let max : Day.Ordinal := month.days year.isLeap - have p := month.all_greater_than_27 year.isLeap - if h : day.val > max.val then - if h₁ : month.val > 11 then - let eq : month.val = 12 := Int.eq_iff_le_and_ge.mpr (And.intro month.property.right h₁) - let h : max.val = 31 := by simp [max, Month.Ordinal.days, Month.Ordinal.daysWithoutProof, Bounded.LE.sub , Bounded.LE.add, Bounded.LE.toFin, eq]; rfl - let h₂ := Int.le_trans day.property.right (Int.eq_iff_le_and_ge.mp h |>.right) |> Int.not_lt.mpr - contradiction - else - let max₂ : Bounded.LE 28 31 := max.truncateBottom p - let sub := Int.sub_nonneg_of_le h - simp [←Int.sub_sub] at sub - let roll := day.addBounds (max₂.neg) |>.truncateBottom (Int.add_le_of_le_sub_left sub) - let day : Day.Ordinal := roll.expandTop (by decide) - let h₂ : roll.val ≤ 27 := Int.le_trans roll.property.right (by decide) - let month := month.truncateTop (Int.not_lt.mp h₁) |>.addTop 1 (by decide) - refine ⟨year, month, day, ?_⟩ - exact Int.le_of_lt (Int.le_trans (Int.add_le_add_right h₂ 1) (Month.Ordinal.all_greater_than_27 year.isLeap month)) - else - let h := Int.not_lt.mp h - exact ⟨year, month, day, h⟩ + let day := month.clipDay year.isLeap day + PlainDate.mk year month day Month.Ordinal.clipDay_valid instance : Inhabited PlainDate where - default := clip 0 1 1 + default := mk 0 1 1 (by decide) /-- Creates a new `PlainDate` from year, month, and day components. @@ -96,8 +70,8 @@ def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordin Creates a `PlainDate` from a year and a day ordinal within that year. -/ def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : PlainDate := - let ⟨⟨month, day⟩, valid⟩ := Month.Ordinal.ofOrdinal ordinal - PlainDate.mk year month day valid + let ⟨⟨month, day⟩, proof⟩ := Month.Ordinal.ofOrdinal ordinal + ⟨year, month, day, proof⟩ /-- Creates a `PlainDate` from the number of days since the UNIX epoch (January 1st, 1970). @@ -133,7 +107,7 @@ def weekday (date : PlainDate) : Weekday := let j : Bounded.LE (-10) 9 := (Bounded.LE.byMod y 1000 (by decide)).ediv 100 (by decide) let part : Bounded.LE 6 190 := q.addBounds (((m.add 1).mul_pos 13 (by decide)).ediv 5 (by decide)) |>.addBounds k |>.addBounds (k.ediv 4 (by decide)) let h : Bounded.LE (-15) 212 := part.addBounds ((j.ediv 4 (by decide)).addBounds (j.mul_pos 2 (by decide)).neg) - let d := (h.add 5).emod 7 (by decide) + let d := (h.add 6).emod 7 (by decide) .ofOrdinal (d.add 1) @@ -203,26 +177,15 @@ Subtracts a given number of weeks from a `PlainDate`. @[inline] def subWeeks (date : PlainDate) (weeks : Week.Offset) : PlainDate := addWeeks date (-weeks) - /-- Adds a given number of months to a `PlainDate`, clipping the day to the last valid day of the month. -/ def addMonthsClip (date : PlainDate) (months : Month.Offset) : PlainDate := - let yearsOffset := months.div 12 - let monthOffset := Bounded.LE.byMod months 12 (by decide) - let months := date.month.addBounds monthOffset - - let (yearsOffset, months) : Year.Offset × Month.Ordinal := by - if h₁ : months.val > 12 then - let months := months |>.truncateBottom h₁ |>.sub 12 - exact (yearsOffset.add 1, months.expandTop (by decide)) - else if h₂ : months.val < 1 then - let months := months |>.truncateTop (Int.le_sub_one_of_lt h₂) |>.add 12 - exact (yearsOffset.sub 1, months.expandBottom (by decide)) - else - exact (yearsOffset, months.truncateTop (Int.not_lt.mp h₁) |>.truncateBottom (Int.not_lt.mp h₂)) - - PlainDate.clip (date.year.add yearsOffset) months date.day + let totalMonths := (date.month.toOffset - 1) + months + let totalMonths : Int := totalMonths + let wrappedMonths := Bounded.LE.byEmod totalMonths 12 (by decide) |>.add 1 + let yearsOffset := totalMonths / 12 + PlainDate.clip (date.year.add yearsOffset) wrappedMonths date.day /-- Subtracts `Month.Offset` from a `PlainDate`, it clips the day to the last valid day of that month. @@ -231,37 +194,65 @@ Subtracts `Month.Offset` from a `PlainDate`, it clips the day to the last valid def subMonthsClip (date : PlainDate) (months : Month.Offset) : PlainDate := addMonthsClip date (-months) +/-- +Creates a `PlainDate` by rolling over the extra days to the next month. +-/ +def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := + clip year month 1 |>.addDays (day.toOffset - 1) + +/-- +Creates a new `PlainDate` by adjusting the day of the month to the given `days` value, with any +out-of-range days clipped to the nearest valid date. +-/ +@[inline] +def withDaysClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := + clip dt.year dt.month days + +/-- +Creates a new `PlainDate` by adjusting the day of the month to the given `days` value, with any +out-of-range days rolled over to the next month or year as needed. +-/ +@[inline] +def withDaysRollOver (dt : PlainDate) (days : Day.Ordinal) : PlainDate := + rollOver dt.year dt.month days + +/-- +Creates a new `PlainDate` by adjusting the month to the given `month` value. +The day remains unchanged, and any invalid days for the new month will be handled according to the `clip` behavior. +-/ +@[inline] +def withMonthClip (dt : PlainDate) (month : Month.Ordinal) : PlainDate := + clip dt.year month dt.day + +/-- +Creates a new `PlainDate` by adjusting the month to the given `month` value. +The day is rolled over to the next valid month if necessary. +-/ +@[inline] +def withMonthRollOver (dt : PlainDate) (month : Month.Ordinal) : PlainDate := + rollOver dt.year month dt.day + +/-- +Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day remain unchanged, +and any invalid days for the new year will be handled according to the `clip` behavior. +-/ +@[inline] +def withYearClip (dt : PlainDate) (year : Year.Offset) : PlainDate := + clip year dt.month dt.day + +/-- +Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day are rolled +over to the next valid month and day if necessary. +-/ +@[inline] +def withYearRollOver (dt : PlainDate) (year : Year.Offset) : PlainDate := + rollOver year dt.month dt.day + /-- Adds a given number of months to a `PlainDate`, rolling over any excess days into the following month. -/ def addMonthsRollOver (date : PlainDate) (months : Month.Offset) : PlainDate := - let yearsOffset := months.div 12 - let monthOffset := Bounded.LE.byMod months 12 (by decide) - let months := date.month.addBounds monthOffset - - let (yearsOffset, months) : Year.Offset × Month.Ordinal := by - if h₁ : months.val > 12 then - let months := months |>.truncateBottom h₁ |>.sub 12 - exact (yearsOffset.add 1, months.expandTop (by decide)) - else if h₂ : months.val < 1 then - let months := months |>.truncateTop (Int.le_sub_one_of_lt h₂) |>.add 12 - exact (yearsOffset.sub 1, months.expandBottom (by decide)) - else - exact (yearsOffset, months.truncateTop (Int.not_lt.mp h₁) |>.truncateBottom (Int.not_lt.mp h₂)) - - let year : Year.Offset := date.year.add yearsOffset - let ⟨days, proof, _⟩ := Month.Ordinal.days year.isLeap months - - if h : days.val ≥ date.day.val then - have p : year.Valid months date.day := by - simp_all [Year.Offset.Valid, Month.Ordinal.Valid] - exact Int.le_trans h proof - PlainDate.mk year months date.day p - else - let roll := Day.Offset.ofInt (date.day.val - days.toInt) - let date := PlainDate.clip (date.year.add yearsOffset) months date.day - let days := date.toDaysSinceUNIXEpoch + roll - PlainDate.ofDaysSinceUNIXEpoch days + addMonthsClip (clip date.year date.month 1) months |>.addDays (date.day.toOffset - 1) /-- Subtracts `Month.Offset` from a `PlainDate`, rolling over excess days as needed. diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 95e3a21eb005..79f6b4af3572 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -34,45 +34,93 @@ instance {x y : Ordinal} : Decidable (x < y) := instance : Inhabited Ordinal where default := 1 /-- -`Ordinal.OfYear` represents the number of days in a year, accounting for leap years. It ensures that -the day is within the correct bounds—either 1 to 365 for regular years or 1 to 366 for leap years. +`Offset` represents an offset in days. It is defined as an `Int` with a base unit of 86400 (the number of seconds in a day). +This type supports arithmetic operations like addition, subtraction, multiplication, and division, and also comparisons like less than or equal. -/ -def Ordinal.OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) +def Offset : Type := UnitVal 86400 + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString -instance : OfNat (Ordinal.OfYear leap) n := by - have inst := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (364 : Nat))) n) - cases leap - · exact inst - · exact ⟨inst.ofNat.expandTop (by decide)⟩ +/-- +Provides an instance for creating an `Offset` from a natural number (`OfNat`), converting the input to the base unit (days). +-/ +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ -instance : OfNat (Ordinal.OfYear true) 366 where - ofNat := Bounded.LE.mk (Int.ofNat 366) (by decide) +/-- +Provides a decidable instance to check if one `Offset` is less than or equal to another. +-/ +instance {x y : Offset} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) -instance : Inhabited (Ordinal.OfYear leap) where - default := by - refine ⟨1, And.intro (by decide) ?_⟩ - split <;> simp +/-- +Provides a decidable instance to check if one `Offset` is strictly less than another. +-/ +instance {x y : Offset} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + +namespace Ordinal /-- -`Offset` represents an offset in days. It is defined as an `Int`. +`OfYear` represents the day ordinal within a year, which can be bounded between 1 and 365 or 366, +depending on whether it's a leap year. -/ -def Offset : Type := UnitVal 86400 - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString +def OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) -instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +namespace OfYear -namespace Ordinal +/-- +Creates an ordinal for a specific day within the year, ensuring that the provided day (`data`) +is within the valid range for the year, which can be 1 to 365 or 366 for leap years. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ (if leap then 366 else 365) := by decide) : OfYear leap := + Bounded.LE.ofNat' data h + +end OfYear /-- -Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +`Period` is an enumeration representing different times of the day: morning, afternoon, evening, and night. +-/ +inductive Period + /-- Represents the morning period. -/ + | morning + + /-- Represents the afternoon period. -/ + | afternoon + + /-- Represents the evening period. -/ + | evening + + /-- Represents the night period. -/ + | night + +/-- +Instance to allow creation of an `Ordinal.OfYear` from a natural number, ensuring the value is +within the bounds of the year, which depends on whether it's a leap year or not. +-/ +instance : OfNat (Ordinal.OfYear leap) n := + match leap with + | true => inferInstanceAs (OfNat (Bounded.LE 1 (1 + (365 : Nat))) n) + | false => inferInstanceAs (OfNat (Bounded.LE 1 (1 + (364 : Nat))) n) + +/-- +Provides a default value for `Ordinal.OfYear`, defaulting to day 1. +-/ +instance : Inhabited (Ordinal.OfYear leap) where + default := by + refine ⟨1, And.intro (by decide) ?_⟩ + split <;> simp + +/-- +Creates an ordinal from a natural number, ensuring the number is within the valid range +for days of a month (1 to 31). -/ @[inline] def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 31 := by decide) : Ordinal := Bounded.LE.ofNat' data h /-- -Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds, if its 0 then its converted -to 1. +Creates an ordinal from a `Fin` value, ensuring it is within the valid range for days of the month (1 to 31). +If the `Fin` value is 0, it is converted to 1. -/ @[inline] def ofFin (data : Fin 32) : Ordinal := @@ -85,10 +133,30 @@ Converts an `Ordinal` to an `Offset`. def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val +namespace OfYear + +/-- +Converts an `OfYear` ordinal to a `Offset`. +-/ +def toOffset (of: OfYear leap) : Offset := + UnitVal.mk of.val + +end OfYear end Ordinal namespace Offset +/-- +Converts an `Ordinal` to an `Offset`. +-/ +@[inline] +def toOrdinal (off : Offset) (h : off.val ≥ 1 ∧ off.val ≤ 31) : Ordinal := + Bounded.LE.mk off.val h + +theorem toOffset_toOrdinal {d : Ordinal} : ∃h, d.toOffset.toOrdinal h = d := by + simp [Ordinal.toOffset, toOrdinal, Bounded.LE.mk, UnitVal.ofInt] + exists d.property + /-- Creates an `Offset` from a natural number. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 6c2865ac15f6..be9ac02cec6a 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -34,13 +34,30 @@ instance : Inhabited Ordinal where `Offset` represents an offset in months. It is defined as an `Int`. -/ def Offset : Type := Int - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString, LT, LE, DecidableEq instance : OfNat Offset n := ⟨Int.ofNat n⟩ +namespace Ordinal + +/-- +`OfYear` represents the number of days in a year, accounting for leap years. It ensures that +the day is within the correct bounds—either 1 to 365 for regular years or 1 to 366 for leap years. +-/ +def Quarter := Bounded.LE 1 4 + +end Ordinal namespace Offset +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds, if its 0 then its converted +to 1. +-/ +@[inline] +def ofFin (data : Fin 13) : Ordinal := + Bounded.LE.ofFin' data (by decide) + /-- Creates an `Offset` from a natural number. -/ @@ -119,6 +136,13 @@ The ordinal value representing the month of December. -/ @[inline] def december : Ordinal := 12 +/-- +Converts a `Ordinal` into a `Offset`. +-/ +@[inline] +def toOffset (month : Ordinal) : Offset := + month.val + /-- Creates an `Ordinal` from a `Nat`, ensuring the value is within bounds. -/ @@ -130,8 +154,10 @@ def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 12 := by decide) : Ordinal : Converts a `Ordinal` into a `Nat`. -/ @[inline] -def toNat (month : Ordinal) : Nat := - Bounded.LE.toNat month +def toNat (month : Ordinal) : Nat := by + match month with + | ⟨.ofNat s, _⟩ => exact s + | ⟨.negSucc s, h⟩ => nomatch h.left /-- Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds, if its 0 then its converted @@ -180,79 +206,152 @@ Size in days of each month if the year is not a leap year. -/ @[inline] def monthSizesNonLeap : { val : Array Day.Ordinal // val.size = 12 } := - ⟨#[31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31], by simp⟩ + ⟨#[31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31], by decide⟩ /-- -Gets the number of days in a month without requiring a proof of the validity of the ordinal in a -month and year. +Size in days of each month if the year is not a leap year. -/ -def daysWithoutProof (leap : Bool) (month : Ordinal) : Day.Ordinal := +@[inline] +def cumulativeSizes : { val : Array Day.Offset // val.size = 12 } := + ⟨#[0, 31, 59, 90, 120, 151, 181, 212, 243, 273, 304, 334], by decide⟩ + +/-- +Gets the number of days in a month. +-/ +def days (leap : Bool) (month : Ordinal) : Day.Ordinal := if month.val = 2 then if leap then 29 else 28 else by let ⟨months, p⟩ := monthSizesNonLeap - let r : Fin 12 := (month.sub 1).toFin (by decide) (by decide) - rw [← p] at r - exact months.get r + let index : Fin 12 := (month.sub 1).toFin (by decide) (by decide) + rw [← p] at index + exact months.get index -theorem all_greater_than_27 (leap : Bool) (i : Month.Ordinal) : daysWithoutProof leap i > 27 := by - simp [daysWithoutProof, monthSizesNonLeap, Bounded.LE.sub, Bounded.LE.add, Bounded.LE.toFin] +theorem days_gt_27 (leap : Bool) (i : Month.Ordinal) : days leap i > 27 := by match i with - | ⟨2, _⟩ => split <;> (simp; try split); all_goals decide + | ⟨2, _⟩ => + simp [days] + split <;> decide | ⟨1, _⟩ | ⟨3, _⟩ | ⟨4, _⟩ | ⟨5, _⟩ | ⟨6, _⟩ | ⟨7, _⟩ - | ⟨8, _⟩ | ⟨9, _⟩ | ⟨10, _⟩ | ⟨11, _⟩ | ⟨12, _⟩ => simp; decide + | ⟨8, _⟩ | ⟨9, _⟩ | ⟨10, _⟩ | ⟨11, _⟩ | ⟨12, _⟩ => + simp [days, monthSizesNonLeap, Bounded.LE.sub, Bounded.LE.add, Bounded.LE.toFin] + decide + +/-- +Gets the number of days in a month without requiring a proof of the validity of the ordinal in a +month and year. +-/ +def cumulativeDays (leap : Bool) (month : Ordinal) : Day.Offset := by + let ⟨months, p⟩ := cumulativeSizes + let index : Fin 12 := (month.sub 1).toFin (by decide) (by decide) + rw [← p] at index + let res := months.get index + exact res + (if leap ∧ month.val > 2 then 1 else 0) + +theorem cumulativeDays_le_335 (leap : Bool) (month : Month.Ordinal) : cumulativeDays leap month ≥ 0 ∧ cumulativeDays leap month ≤ 334 + (if leap then 1 else 0) := by + match month with + | ⟨1, _⟩ | ⟨2, _⟩ | ⟨3, _⟩ | ⟨4, _⟩ | ⟨5, _⟩ | ⟨6, _⟩ | ⟨7, _⟩ | ⟨8, _⟩ | ⟨9, _⟩ | ⟨10, _⟩ | ⟨11, _⟩ | ⟨12, _⟩ => + simp [cumulativeSizes, Bounded.LE.sub, Bounded.LE.add, Bounded.LE.toFin, cumulativeDays] + try split + all_goals decide + +theorem difference_eq (p : month.val ≤ 11) : + let next := month.truncateTop p |>.addTop 1 (by decide) + (cumulativeDays leap next).val = (cumulativeDays leap month).val + (days leap month).val := by + match month with + | ⟨1, _⟩ | ⟨2, _⟩ | ⟨3, _⟩ | ⟨4, _⟩ | ⟨5, _⟩ | ⟨6, _⟩ | ⟨7, _⟩ | ⟨8, _⟩ | ⟨9, _⟩ | ⟨10, _⟩ | ⟨11, _⟩ => + simp [cumulativeDays, Bounded.LE.addTop, days, monthSizesNonLeap]; + try split <;> rfl + try rfl + | ⟨12, _⟩ => contradiction /-- Check if the day is valid in a month and a leap year. -/ abbrev Valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := - day.val ≤ (daysWithoutProof leap month).val + day.val ≤ (days leap month).val -instance : Decidable (Valid leap month day) := - dite (day ≤ daysWithoutProof leap month) isTrue isFalse +/-- +Type for dates and months that are valid within a leap year. +-/ +def ValidDate (leap : Bool) := { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } + +instance : Inhabited (ValidDate l) where + default := ⟨⟨1, 1⟩, (by cases l <;> decide)⟩ /-- -Gets the number of days in a month along with a proof of its validity. +Transforms a tuple of a `Month` and a `Day` into a `Day.Ordinal.OfYear`. -/ -@[inline] -def days (leap : Bool) (month : Ordinal) : { day : Day.Ordinal // Valid leap month day ∧ day.val > 27 } := by - refine ⟨daysWithoutProof leap month, ⟨Int.le_refl ((daysWithoutProof leap month).val), all_greater_than_27 leap month⟩⟩ +def toOrdinal (ordinal : ValidDate leap) : Day.Ordinal.OfYear leap := + let days := cumulativeDays leap ordinal.val.fst + let proof := cumulativeDays_le_335 leap ordinal.val.fst + let bounded := Bounded.LE.mk days.toInt proof |>.addBounds ordinal.val.snd + match leap, bounded with + | true, bounded => bounded + | false, bounded => bounded + +/-- +Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. +-/ +def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : ValidDate leap := + let rec go (idx : Month.Ordinal) (acc : Int) (h : ordinal.val > acc) (p : acc = (cumulativeDays leap idx).val) : ValidDate leap := + let monthDays := days leap idx + if h₁ : ordinal.val ≤ acc + monthDays.val then + let bounded := Bounded.LE.mk ordinal.val (And.intro h h₁) |>.sub acc + let bounded : Bounded.LE 1 monthDays.val := bounded.cast (by omega) (by omega) + let days₁ : Day.Ordinal := ⟨bounded.val, And.intro bounded.property.left (Int.le_trans bounded.property.right monthDays.property.right)⟩ + ⟨⟨idx, days₁⟩, Int.le_trans bounded.property.right (by simp)⟩ + else + let h₂ := Int.not_le.mp h₁ + if h₃ : idx.val > 11 then by + have h₅ := ordinal.property.right + let eq := Int.eq_iff_le_and_ge.mpr (And.intro idx.property.right h₃) + simp [monthDays, days, eq] at h₂ + simp [cumulativeDays, eq] at p + simp [p] at h₂ + cases leap + all_goals (simp at h₂; simp_all) + · have h₂ : 365 < ordinal.val := h₂ + omega + · have h₂ : 366 < ordinal.val := h₂ + omega + else by + let h₃ := Int.not_le.mp h₃ + let idx₂ := idx.truncateTop (Int.le_sub_one_of_lt h₃) |>.addTop 1 (by decide) + refine go idx₂ (acc + monthDays.val) h₂ ?_ + simp [monthDays, p] + rw [difference_eq (Int.le_of_lt_add_one h₃)] + termination_by 12 - idx.val.toNat + decreasing_by + simp_wf + simp [Bounded.LE.addTop] + let gt0 : idx.val ≥ 0 := Int.le_trans (by decide) idx.property.left + refine Nat.sub_lt_sub_left (Int.toNat_lt gt0 |>.mpr h₃) ?_ + let toNat_lt_lt {n z : Int} (h : 0 ≤ z) (h₁ : 0 ≤ n) : z.toNat < n.toNat ↔ z < n := by + rw [← Int.not_le, ← Nat.not_le, ← Int.ofNat_le, Int.toNat_of_nonneg h, Int.toNat_of_nonneg h₁] + rw [toNat_lt_lt (by omega) (by omega)] + omega + + go 1 0 (Int.le_trans (by decide) ordinal.property.left) (by cases leap <;> decide) /-- Clips the day to be within the valid range. -/ @[inline] -def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : { day : Day.Ordinal // Valid leap month day } := +def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Day.Ordinal := let max : Day.Ordinal := month.days leap - if h : day.val > max.val - then ⟨max, Int.le_refl max.val⟩ - else ⟨day, Int.not_lt.mp h⟩ + if day.val > max.val + then max + else day /-- -Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. +Proves that every value provided by a clipDay is a valid day in a year. -/ -def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } := Id.run do - let rec go (idx : Fin 12) (cumulative : Fin 366) := - let month := Month.Ordinal.ofFin idx.succ - let ⟨days, valid, _⟩ := days leap month - - if h : cumulative.val < ordinal.val ∧ ordinal.val ≤ cumulative.val + days.val then - let bounded := Bounded.LE.mk ordinal.val h |>.sub cumulative - let bounded : Bounded.LE 1 days.val := bounded.cast (by omega) (by omega) - - let ⟨left, right⟩ := bounded.property - let days₁ : Day.Ordinal := ⟨bounded.val, And.intro left (Int.le_trans right days.property.right)⟩ - ⟨⟨month, days₁⟩, Int.le_trans right valid⟩ - else - if h : idx.val ≥ 11 then - -- Need to remove this in the future. - let ⟨day, valid⟩ := clipDay leap 1 1 - ⟨⟨1, day⟩, valid⟩ - else - go ⟨idx.val + 1, Nat.succ_le_succ (Nat.not_le.mp h)⟩ (cumulative + (Fin.ofNat days.val.toNat)) - - termination_by 12 - idx.val - go 0 0 +theorem clipDay_valid : Valid leap month (clipDay leap month day) := by + simp [Valid, clipDay] + split + exact Int.le_refl (days leap month).val + next h => exact Int.not_lt.mp h end Ordinal end Month diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 61d63f850cb5..5e3bb35294da 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -40,6 +40,13 @@ instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ namespace Ordinal +/-- +`OfMonth` represents the number of weeks within a month. It ensures that the week is within the +correct bounds—either 1 to 6, representing the possible weeks in a month. +-/ +def OfMonth := Bounded.LE 1 6 + + /-- Creates an `Ordinal` from a natural number, ensuring the value is within bounds. -/ @@ -58,6 +65,34 @@ Converts an `Ordinal` to an `Offset`. def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val +/-- +Convert `Week.Offset` into `Millisecond.Offset`. +-/ +@[inline] +def toMilliseconds (weeks : Week.Offset) : Millisecond.Offset := + weeks.mul 604800000 + +/-- +Convert `Millisecond.Offset` into `Week.Offset`. +-/ +@[inline] +def ofMilliseconds (millis : Millisecond.Offset) : Week.Offset := + millis.ediv 604800000 + +/-- +Convert `Week.Offset` into `Nanosecond.Offset`. +-/ +@[inline] +def toNanoseconds (weeks : Week.Offset) : Nanosecond.Offset := + weeks.mul 604800000000000 + +/-- +Convert `Nanosecond.Offset` into `Week.Offset`. +-/ +@[inline] +def ofNanoseconds (nanos : Nanosecond.Offset) : Week.Offset := + nanos.ediv 604800000000000 + /-- Convert `Week.Offset` into `Second.Offset`. -/ diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index c8d95ea24ec7..976fd02dfeed 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -16,9 +16,11 @@ set_option linter.all true /-- Defines the enumeration for days of the week. Each variant corresponds to a day of the week, from Monday to Sunday. -As defined in the ISO 8601, It starts on monday. +It starts on the sunday because the library is based on the Gregorian Calendar. -/ inductive Weekday + /-- Sunday. -/ + | sunday /-- Monday. -/ | monday /-- Tuesday. -/ @@ -31,8 +33,6 @@ inductive Weekday | friday /-- Saturday. -/ | saturday - /-- Sunday. -/ - | sunday deriving Repr, Inhabited, BEq namespace Weekday @@ -45,55 +45,54 @@ def Ordinal := Bounded.LE 1 7 instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (6 : Nat))) n) - /-- Converts a `Ordinal` representing a day index into a corresponding `Weekday`. This function is useful for mapping numerical representations to days of the week. -/ def ofOrdinal : Ordinal → Weekday - | 1 => .monday - | 2 => .tuesday - | 3 => .wednesday - | 4 => .thursday - | 5 => .friday - | 6 => .saturday - | 7 => .sunday + | 1 => .sunday + | 2 => .monday + | 3 => .tuesday + | 4 => .wednesday + | 5 => .thursday + | 6 => .friday + | 7 => .saturday /-- Converts a `Weekday` to a `Nat`. -/ def toOrdinal : Weekday → Ordinal - | .monday => 1 - | .tuesday => 2 - | .wednesday => 3 - | .thursday => 4 - | .friday => 5 - | .saturday => 6 - | .sunday => 7 + | .sunday => 1 + | .monday => 2 + | .tuesday => 3 + | .wednesday => 4 + | .thursday => 5 + | .friday => 6 + | .saturday => 7 /-- Converts a `Weekday` to a `Nat`. -/ def toNat : Weekday → Nat - | .monday => 1 - | .tuesday => 2 - | .wednesday => 3 - | .thursday => 4 - | .friday => 5 - | .saturday => 6 - | .sunday => 7 + | .sunday => 1 + | .monday => 2 + | .tuesday => 3 + | .wednesday => 4 + | .thursday => 5 + | .friday => 6 + | .saturday => 7 /-- Converts a `Nat` to an `Option Weekday`. -/ def ofNat? : Nat → Option Weekday - | 1 => some .monday - | 2 => some .tuesday - | 3 => some .wednesday - | 4 => some .thursday - | 5 => some .friday - | 6 => some .saturday - | 7 => some .sunday + | 1 => some .sunday + | 2 => some .monday + | 3 => some .tuesday + | 4 => some .wednesday + | 5 => some .thursday + | 6 => some .friday + | 7 => some .saturday | _ => none /-- @@ -116,6 +115,7 @@ def next : Weekday → Weekday | .friday => .saturday | .saturday => .sunday | .sunday => .monday + /-- Check if it's a Weekend. -/ diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index a747b5a3a3c5..4cf634a8281e 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -35,7 +35,7 @@ inductive Era `Offset` represents a year offset, defined as an `Int`. -/ def Offset : Type := Int - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString, LT, LE instance : OfNat Offset n := ⟨Int.ofNat n⟩ @@ -70,21 +70,29 @@ def toMonths (val : Offset) : Month.Offset := val.mul 12 /-- -Determines if a year is a leap year in the proleptic gregorian calendar. +Determines if a year is a leap year in the proleptic Gregorian calendar. -/ @[inline] def isLeap (y : Offset) : Bool := y.toInt.mod 4 = 0 ∧ (y.toInt.mod 100 ≠ 0 ∨ y.toInt.mod 400 = 0) +/-- +Returns the `Era` of the `Year`. +-/ +def era (year : Offset) : Era := + if year.toInt ≥ 1 + then .ce + else .bce + /-- Checks if the given date is valid for the specified year, month, and day. -/ @[inline] abbrev Valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := - month.Valid year.isLeap day + day ≤ month.days year.isLeap instance : Decidable (Valid year month day) := - dite (month.Valid year.isLeap day) isTrue isFalse + dite (day ≤ month.days year.isLeap) isTrue isFalse end Offset end Year diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 6cfea6996854..1da4b00f1482 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -18,21 +18,21 @@ set_option linter.all true Converts a `PlainDateTime` to a `Timestamp` -/ @[inline] -def ofPlainDateTime (pdt : PlainDateTime) : Timestamp := +def ofPlainDateTimeAssumingUTC (pdt : PlainDateTime) : Timestamp := pdt.toTimestampAssumingUTC /-- Converts a `Timestamp` to a `PlainDateTime` -/ @[inline] -def toPlainDateTime (timestamp : Timestamp) : PlainDateTime := +def toPlainDateTimeAssumingUTC (timestamp : Timestamp) : PlainDateTime := PlainDateTime.ofUTCTimestamp timestamp /-- Converts a `PlainDate` to a `Timestamp` -/ @[inline] -def ofPlainDate (pd : PlainDate) : Timestamp := +def ofPlainDateAssumingUTC (pd : PlainDate) : Timestamp := let days := pd.toDaysSinceUNIXEpoch let secs := days.toSeconds Timestamp.ofSecondsSinceUnixEpoch secs @@ -41,7 +41,7 @@ def ofPlainDate (pd : PlainDate) : Timestamp := Converts a `Timestamp` to a `PlainDate` -/ @[inline] -def toPlainDate (timestamp : Timestamp) : PlainDate := +def toPlainDateAssumingUTC (timestamp : Timestamp) : PlainDate := let secs := timestamp.toSecondsSinceUnixEpoch let days := Day.Offset.ofSeconds secs PlainDate.ofDaysSinceUNIXEpoch days @@ -70,7 +70,7 @@ Converts a `PlainDate` to a `Timestamp` -/ @[inline] def toTimestampAssumingUTC (pdt : PlainDate) : Timestamp := - Timestamp.ofPlainDate pdt + Timestamp.ofPlainDateAssumingUTC pdt /-- Calculates the duration between a given `PlainDate` and a specified date. @@ -101,14 +101,14 @@ Converts a `PlainDate` to a `Timestamp` -/ @[inline] def ofPlainDate (pd : PlainDate) : PlainDateTime := - PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainDate pd) + PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainDateAssumingUTC pd) /-- Converts a `PlainDateTime` to a `PlainDate` -/ @[inline] def toPlainDate (pdt : PlainDateTime) : PlainDate := - Timestamp.toPlainDate pdt.toTimestampAssumingUTC + Timestamp.toPlainDateAssumingUTC pdt.toTimestampAssumingUTC /-- Converts a `PlainTime` to a `PlainDateTime` @@ -125,10 +125,10 @@ def toPlainTime (pdt : PlainDateTime) : PlainTime := Timestamp.toPlainTime pdt.toTimestampAssumingUTC instance : ToTimestamp PlainDateTime where - toTimestamp := Timestamp.ofPlainDateTime + toTimestamp := Timestamp.ofPlainDateTimeAssumingUTC instance : ToTimestamp PlainDate where - toTimestamp := Timestamp.ofPlainDate + toTimestamp := Timestamp.ofPlainDateAssumingUTC /-- Calculates the duration between a given `PlainDateTime` and a specified date. diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 75e932174d84..71a659b6f1b7 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -123,7 +123,7 @@ def ofUTCTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do return { date := PlainDate.clip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) - time := PlainTime.ofValidHourMinuteSecondsNano (hour.expandTop (by decide)) minute (second.expandTop (by decide)) nano + time := PlainTime.ofHourMinuteSecondsNano (leap := false) (hour.expandTop (by decide)) minute second nano } /-- @@ -140,6 +140,14 @@ Subtracts a `Day.Offset` from a `PlainDateTime`. def subDays (dt : PlainDateTime) (days : Day.Offset) : PlainDateTime := { dt with date := dt.date.subDays days } +/-- +Sets the `Day.Offset` from a `PlainDateTime`. +-/ +@[inline] +def withDaysClip (dt : PlainDateTime) (days : Day.Ordinal) : PlainDateTime := + { dt with date := dt.date.withDaysClip days } + + /-- Adds a `Week.Offset` to a `PlainDateTime`. -/ @@ -309,8 +317,8 @@ def day (dt : PlainDateTime) : Day.Ordinal := Getter for the `Hour` inside of a `PlainDateTime`. -/ @[inline] -def hour (dt : PlainDateTime) : Hour.Ordinal dt.time.hour.fst := - dt.time.hour.snd +def hour (dt : PlainDateTime) : Hour.Ordinal := + dt.time.hour /-- Getter for the `Minute` inside of a `PlainDateTime`. diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 5c37537e429b..c7f5b4c95c0c 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -4,115 +4,158 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ prelude -import Std.Time.Notation +import Std.Time.Notation.Spec import Std.Time.Format.Basic import Std.Time.Internal.Bounded namespace Std namespace Time namespace Formats +open Internal set_option linter.all true /-- The ISO8601 format, which is always 24 or 27 characters long, used for representing date and time in -a standardized format. The format follows the pattern `YYYY-MM-DD'T'hh:mm:ss'Z'`. +a standardized format. The format follows the pattern `yyyy-MM-dd:HH::mm:ss.SSSZ`. -/ -def iso8601 : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" +def iso8601 : Format .any := datespec("yyyy-MM-dd:HH::mm:ss.SSSZ") /-- -The americanDate format, which follows the pattern `MM/DD/YYYY`. +The americanDate format, which follows the pattern `MM-dd-yyyy`. -/ -def americanDate : Format .any := date-spec% "MM-DD-YYYY" +def americanDate : Format .any := datespec("MM-dd-yyyy") /-- -The europeanDate format, which follows the pattern `DD-MM-YYYY`. +The europeanDate format, which follows the pattern `dd-MM-yyyy`. -/ -def europeanDate : Format .any := date-spec% "DD-MM-YYYY" +def europeanDate : Format .any := datespec("dd-MM-yyyy") /-- -The time12Hour format, which follows the pattern `HH:mm:ss AA` for representing time +The time12Hour format, which follows the pattern `hh:mm:ss aa` for representing time in a 12-hour clock format with an upper case AM/PM marker. -/ -def time12Hour : Format .any := date-spec% "HH:mm:ss AA" +def time12Hour : Format .any := datespec("hh:mm:ss aa") /-- -The Time24Hour format, which follows the pattern `hh:mm:ss` for representing time +The Time24Hour format, which follows the pattern `HH:mm:ss:SSSSSSSSS` for representing time in a 24-hour clock format. -/ -def time24Hour : Format .any := date-spec% "hh:mm:ss:sssssssss" +def time24Hour : Format .any := datespec("HH:mm:ss.SSSSSSSSS") /-- -The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD hh:mm:ss.sssssssss` for +The DateTimeZone24Hour format, which follows the pattern `yyyy-MM-dd:HH:mm:ss.SSSSSSSSS` for representing date, time, and time zone. -/ -def dateTime24Hour : Format (.only .GMT) := date-spec% "YYYY-MM-DD:hh:mm:ss.sssssssss" +def dateTime24Hour : Format (.only .GMT) := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSS") /-- -The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD hh:mm:ss.sssssssss` +The DateTimeWithZone format, which follows the pattern `yyyy-MM-dd:HH:mm:ss.SSSSSSSSSZZZ` for representing date, time, and time zone. -/ -def dateTimeWithZone : Format .any := date-spec% "YYYY-MM-DD:hh:mm:ss.sssssssssZZZ" +def dateTimeWithZone : Format .any := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSSZZZ") /-- -The Time24Hour format, which follows the pattern `date% hh:mm:ss` for representing time +The Time24Hour format, which follows the pattern `HH:mm:ss` for representing time in a 24-hour clock format. It uses the default value that can be parsed with the notation of dates. -/ -def leanTime24Hour : Format .any := date-spec% "'time% 'hh:mm:ss,sssssssss" +def leanTime24Hour : Format .any := datespec("HH:mm:ss.SSSSSSSSS") /-- -The DateTimeZone24Hour format, which follows the pattern `date% YYYY-MM-DD hh:mm:ss:sssssssss` for +The Time24Hour format, which follows the pattern `HH:mm:ss` for representing time +in a 24-hour clock format. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanTime24HourNoNanos : Format .any := datespec("HH:mm:ss") + +/-- +The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD HH:mm:ss:sssssssss` for +representing date, time, and time zone. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanDateTime24Hour : Format (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSS") + +/-- +The DateTime24HourNoNanos format, which follows the pattern `YYYY-MM-DD HH:mm:ss` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTime24Hour : Format (.only .GMT) := date-spec% "'date% 'YYYY-MM-DD'T'hh:mm:ss,sssssssss" +def leanDateTime24HourNoNanos : Format (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss") + +/-- +The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD HH:mm:ss.SSSSSSSSS` +for representing date, time, and time zone. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanDateTimeWithZone : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ") /-- -The DateTimeWithZone format, which follows the pattern `date% YYYY-MM-DD hh:mm:ss:sssssssss` +The DateTimeWithZoneNoNanos format, which follows the pattern `YYYY-MM-DD HH:mm:ssZZZZZ` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithZone : Format .any := date-spec% "'date% 'YYYY-MM-DD'T'hh:mm:ss,sssssssssZZZZZ" +def leanDateTimeWithZoneNoNanos : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ssZZZZZ") /-- -The Lean Date format, which follows the pattern `date% YYYY-MM-DD`. It uses the default value that can be parsed with the +The Lean Date format, which follows the pattern `yyyy-MM-DD`. It uses the default value that can be parsed with the notation of dates. -/ -def leanDate : Format .any := date-spec% "'date% 'YYYY-MM-DD" +def leanDate : Format .any := datespec("yyyy-MM-dd") /-- -The SQLDate format, which follows the pattern `YYYY-MM-DD` and is commonly used +The SQLDate format, which follows the pattern `yyyy-MM-DD` and is commonly used in SQL databases to represent dates. -/ -def sqlDate : Format .any := date-spec% "YYYY-MM-DD" +def sqlDate : Format .any := datespec("yyyy-MM-dd") /-- -The LongDateFormat, which follows the pattern `EEEE, MMMM D, YYYY hh:mm:ss` for +The LongDateFormat, which follows the pattern `EEEE, MMMM D, yyyy HH:mm:ss` for representing a full date and time with the day of the week and month name. -/ -def longDateFormat : Format (.only .GMT) := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss" +def longDateFormat : Format (.only .GMT) := datespec("EEEE, MMMM D, yyyy HH:mm:ss") /-- -The AscTime format, which follows the pattern `EEE MMM d hh:mm:ss YYYY`. This format +The AscTime format, which follows the pattern `EEE MMM d HH:mm:ss yyyy`. This format is often used in older systems for logging and time-stamping events. -/ -def ascTime : Format (.only .GMT) := date-spec% "EEE MMM d hh:mm:ss YYYY" +def ascTime : Format (.only .GMT) := datespec("EEE MMM d HH:mm:ss yyyy") /-- -The RFC822 format, which follows the pattern `EEE, DD MMM YYYY hh:mm:ss ZZZ`. +The RFC822 format, which follows the pattern `eee, dd MMM yyyy HH:mm:ss ZZZ`. This format is used in email headers and HTTP headers. -/ -def rfc822 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" +def rfc822 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") /-- -The RFC850 format, which follows the pattern `EEEE, DD-MMM-YY hh:mm:ss ZZZ`. +The RFC850 format, which follows the pattern `eee, dd-MMM-YY HH:mm:ss ZZZ`. This format is an older standard for representing date and time in headers. -/ -def rfc850 : Format .any := date-spec% "EEEE, DD-MMM-YY hh:mm:ss ZZZ" +def rfc850 : Format .any := datespec("eee, dd-MM-yyyy HH:mm:ss ZZZ") end Formats +namespace TimeZone + +/-- +Parses a string into a `TimeZone` object. The input string must be in the format `"VV ZZZZZ"`. +-/ +def fromTimeZone (input : String) : Except String TimeZone := do + let spec : Format .any := datespec("VV ZZZZZ") + spec.parseBuilder (λid off => some (TimeZone.mk off id "Unknown" false)) input + +namespace Offset + +/-- +Parses a string representing an offset into an `Offset` object. The input string must follow the `"xxx"` format. +-/ +def fromOffset (input : String) : Except String Offset := do + let spec : Format .any := datespec("xxx") + spec.parseBuilder some input + +end Offset +end TimeZone + namespace PlainDate /-- @@ -124,23 +167,23 @@ def format (date : PlainDate) (format : String) : String := | .error err => s!"error: {err}" | .ok res => let res := res.formatGeneric fun - | .YYYY | .YY => some date.year - | .MMMM | .MMM | .MM | .M => some date.month - | .DD | .D | .d => some date.day - | .EEEE | .EEE => some date.weekday + | .y _ => some date.year + | .MorL _ => some date.month + | .d _ => some date.day + | .E _ => some date.weekday | _ => none match res with | some res => res | none => "invalid time" /-- -Parses a date string in the American format (`MM/DD/YYYY`) and returns a `PlainDate`. +Parses a date string in the American format (`MM/DD/yyyy`) and returns a `PlainDate`. -/ def fromAmericanDateString (input : String) : Except String PlainDate := do Formats.americanDate.parseBuilder (fun m d y => PlainDate.ofYearMonthDay y m d) input /-- -Converts a Date in the American format (`MM/DD/YYYY`) into a `String`. +Converts a Date in the American format (`MM/DD/yyyy`) into a `String`. -/ def toAmericanDateString (input : PlainDate) : String := Formats.americanDate.formatBuilder input.month input.day input.year @@ -157,11 +200,17 @@ Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. def toSQLDateString (input : PlainDate) : String := Formats.sqlDate.formatBuilder input.year input.month input.day +/-- +Parses a date string into a `PlainDate` object. The input string must be in the SQL date format. +-/ +def fromLeanDateString (input : String) : Except String PlainDate := do + Formats.sqlDate.parseBuilder (PlainDate.ofYearMonthDay) input + /-- Converts a Date in the Lean format (`YYYY-MM-DD`) into a `String` with the format `date% YYY-MM-DD`. -/ def toLeanDateString (input : PlainDate) : String := - Formats.leanDate.formatBuilder input.year input.month input.day + Formats.sqlDate.formatBuilder input.year input.month input.day /-- Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `PlainDate`. @@ -171,10 +220,10 @@ def parse (input : String) : Except String PlainDate := <|> fromSQLDateString input instance : ToString PlainDate where - toString := toSQLDateString + toString := toLeanDateString instance : Repr PlainDate where - reprPrec data := Repr.addAppParen (toLeanDateString data) + reprPrec data := Repr.addAppParen ("date(\"" ++ toLeanDateString data ++ "\")") end PlainDate @@ -189,10 +238,10 @@ def format (time : PlainTime) (format : String) : String := | .error err => s!"error: {err}" | .ok res => let res := res.formatGeneric fun - | .HH | .H => some time.hour - | .mm | .m => some time.minute - | .sss => some (Internal.Bounded.LE.ofNat 0 (by decide)) - | .ss | .s => some time.second + | .H _ => some time.hour + | .m _ => some time.minute + | .n _ => some time.nano + | .s _ => some time.second | _ => none match res with | some res => res @@ -202,7 +251,7 @@ def format (time : PlainTime) (format : String) : String := Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. -/ def fromTime24Hour (input : String) : Except String PlainTime := - Formats.time24Hour.parseBuilder (fun h m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input + Formats.time24Hour.parseBuilder (fun h m s n => some (PlainTime.ofHourMinuteSecondsNano h m s.snd n)) input /-- Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). @@ -214,7 +263,8 @@ def toTime24Hour (input : PlainTime) : String := Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. -/ def fromLeanTime24Hour (input : String) : Except String PlainTime := - Formats.leanTime24Hour.parseBuilder (fun h m s n => PlainTime.ofHourMinuteSecondsNano? h.snd m s.snd n) input + Formats.leanTime24Hour.parseBuilder (fun h m s n => some (PlainTime.ofHourMinuteSecondsNano h m s.snd n)) input + <|> Formats.leanTime24HourNoNanos.parseBuilder (fun h m s => some (PlainTime.ofHourMinuteSecondsNano h m s.snd 0)) input /-- Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). @@ -227,8 +277,8 @@ Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `PlainT -/ def fromTime12Hour (input : String) : Except String PlainTime := do let builder h m s a : Option PlainTime := do - let value ← Internal.Bounded.ofInt? h.snd.val - PlainTime.ofHourMinuteSeconds? (leap₂ := false) (HourMarker.toAbsolute a value) m s.snd + let value ← Internal.Bounded.ofInt? h.val + some <| PlainTime.ofHourMinuteSeconds (HourMarker.toAbsolute a value) m s.snd Formats.time12Hour.parseBuilder builder input @@ -236,7 +286,7 @@ def fromTime12Hour (input : String) : Except String PlainTime := do Formats a `PlainTime` value into a 12-hour format string (`hh:mm:ss aa`). -/ def toTime12Hour (input : PlainTime) : String := - Formats.time12Hour.formatBuilder input.hour input.minute input.second (if input.hour.snd.val ≥ 12 then HourMarker.pm else HourMarker.am) + Formats.time12Hour.formatBuilder (input.hour.emod 12 (by decide) |>.add 1) input.minute input.second (if input.hour.val ≥ 12 then HourMarker.pm else HourMarker.am) /-- Parses a `String` in the `Time12Hour` or `Time24Hour` format and returns a `PlainTime`. @@ -249,13 +299,12 @@ instance : ToString PlainTime where toString := toLeanTime24Hour instance : Repr PlainTime where - reprPrec data := Repr.addAppParen (toString data) + reprPrec data := Repr.addAppParen ("time(\"" ++ toLeanTime24Hour data ++ "\")") end PlainTime namespace ZonedDateTime - /-- Formats a `ZonedDateTime` using a specific format. -/ @@ -318,6 +367,7 @@ Parses a `String` in the `DateTimeWithZone` format and returns a `DateTime` obje -/ def fromLeanDateTimeWithZoneString (input : String) : Except String ZonedDateTime := Formats.leanDateTimeWithZone.parse input + <|> Formats.leanDateTimeWithZoneNoNanos.parse input /-- Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notationg. @@ -337,7 +387,7 @@ instance : ToString ZonedDateTime where toString := toLeanDateTimeWithZoneString instance : Repr ZonedDateTime where - reprPrec data := Repr.addAppParen (toString data) + reprPrec data := Repr.addAppParen ("zoned(\"" ++ toLeanDateTimeWithZoneString data ++ "\")") end ZonedDateTime @@ -352,17 +402,15 @@ def format (date : PlainDateTime) (format : String) : String := | .error err => s!"error: {err}" | .ok res => let res := res.formatGeneric fun - | .YYYY | .YY => some date.year - | .MMMM | .MMM | .MM | .M => some date.month - | .DD | .D | .d => some date.day - | .EEEE | .EEE => some date.date.weekday - | .HH | .H => some date.time.hour - | .hh | .h => some date.time.hour - | .aa | .AA => some (if date.time.hour.snd.val > 12 then HourMarker.pm else HourMarker.am) - | .mm | .m => some date.time.minute - | .sssssssss => some date.time.nano - | .sss => some date.time.nano.toMillisecond - | .ss | .s => some date.time.second + | .y _ => some date.date.year + | .MorL _ => some date.date.month + | .d _ => some date.date.day + | .E _ => some date.date.weekday + | .H _ => some date.time.hour + | .m _ => some date.time.minute + | .n _ => some date.time.nano + | .S _ => some date.time.nano + | .s _ => some date.time.second | _ => none match res with | some res => res @@ -411,7 +459,7 @@ def toDateTimeString (pdt : PlainDateTime) : String := Parses a `String` in the `DateTime` format and returns a `PlainDateTime`. -/ def fromLeanDateTimeString (input : String) : Except String PlainDateTime := - Formats.leanDateTime24Hour.parse input + (Formats.leanDateTime24Hour.parse input <|> Formats.leanDateTime24HourNoNanos.parse input) |>.map DateTime.toPlainDateTime /-- @@ -432,7 +480,7 @@ instance : ToString PlainDateTime where toString := toLeanDateTimeString instance : Repr PlainDateTime where - reprPrec data := Repr.addAppParen (toString data) + reprPrec data := Repr.addAppParen ("datetime(\"" ++ toLeanDateTimeString data ++ "\")") end PlainDateTime @@ -508,10 +556,7 @@ def parse (date : String) : Except String (DateTime .GMT) := fromAscTimeString date <|> fromLongDateFormatString date -instance : ToString (DateTime tz) where - toString := toLeanDateTimeWithZoneString - instance : Repr (DateTime tz) where - reprPrec data := Repr.addAppParen (toString data) + reprPrec data := Repr.addAppParen (toLeanDateTimeWithZoneString data) end DateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 5061fbb0dc20..c580f3c339a4 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -10,6 +10,10 @@ import Std.Time.Time import Std.Time.Zoned import Std.Time.DateTime +/-! +This module defines the `Formatter` types. It is based on the Java's `DateTimeFormatter` format. +-/ + namespace Std namespace Time open Internal @@ -18,160 +22,505 @@ open Std.Internal.Parsec Lean PlainTime PlainDate TimeZone DateTime set_option linter.all true + +/-- +`Text` represents different text formatting styles. +-/ +inductive Text + /-- Short form (e.g., "Tue") -/ + | short + /-- Full form (e.g., "Tuesday") -/ + | full + /-- Narrow form (e.g., "T") -/ + | narrow + deriving Repr, Inhabited + +namespace Text + +/-- +`classify` classifies the number of pattern letters into a `Text` type. +-/ +def classify (num : Nat) : Option Text := + if num < 4 then + some (.short) + else if num = 4 then + some (.full) + else if num = 5 then + some (.narrow) + else + none + +end Text + +/-- +`Number` represents different number formatting styles. +-/ +structure Number where + /-- + The number of digits to pad, based on the count of pattern letters. + -/ + padding : Nat + deriving Repr, Inhabited + +/-- +`classifyNumberText` classifies the number of pattern letters into either a `Number` or `Text`. +-/ +def classifyNumberText : Nat → Option (Number ⊕ Text) + | n => if n < 3 then some (.inl ⟨n⟩) else .inr <$> (Text.classify n) + +/-- +`Fraction` represents the fraction of a second, which can either be full nanoseconds +or a truncated form with fewer digits. +-/ +inductive Fraction + /-- Nanosecond precision (up to 9 digits) -/ + | nano + /-- Fewer digits (truncated precision) -/ + | truncated (digits : Nat) + deriving Repr, Inhabited + +namespace Fraction + +/-- +`classify` classifies the number of pattern letters into either a `Fraction`. It's used for `nano`. +-/ +def classify (nat : Nat) : Option Fraction := + if nat < 9 then + some (.truncated nat) + else if nat = 9 then + some (.nano) + else + none + +end Fraction + +/-- +`Year` represents different year formatting styles based on the number of pattern letters. +-/ +inductive Year + /-- Two-digit year format (e.g., "23" for 2023) -/ + | twoDigit + /-- Four-digit year format (e.g., "2023") -/ + | fourDigit + /-- Extended year format for more than 4 digits (e.g., "002023") -/ + | extended (num : Nat) + deriving Repr, Inhabited + +namespace Year + +/-- +`classify` classifies the number of pattern letters into a `Year` format. +-/ +def classify (num : Nat) : Option Year := + if num = 2 then + some (.twoDigit) + else if num = 4 then + some (.fourDigit) + else if num > 4 ∨ num = 3 then + some (.extended num) + else + none + +end Year + +/-- +`ZoneId` represents different time zone ID formats based on the number of pattern letters. +-/ +inductive ZoneId + /-- Short form of time zone (e.g., "PST") -/ + | short + /-- Full form of time zone (e.g., "Pacific Standard Time") -/ + | full + deriving Repr, Inhabited + +namespace ZoneId + +/-- +`classify` classifies the number of pattern letters into a `ZoneId` format. +- If 2 letters, it returns the short form. +- If 4 letters, it returns the full form. +- Otherwise, it returns none. +-/ +def classify (num : Nat) : Option ZoneId := + if num = 2 then + some (.short) + else if num = 4 then + some (.full) + else + none + +end ZoneId + +/-- +`ZoneName` represents different zone name formats based on the number of pattern letters and +whether daylight saving time is considered. +-/ +inductive ZoneName + /-- Short form of zone name (e.g., "PST") -/ + | short + /-- Full form of zone name (e.g., "Pacific Standard Time") -/ + | full + deriving Repr, Inhabited + +namespace ZoneName + +/-- +`classify` classifies the number of pattern letters and the letter type ('z' or 'v') +into a `ZoneName` format. +- For 'z', if less than 4 letters, it returns the short form; if 4 letters, it returns the full form. +- For 'v', if 1 letter, it returns the short form; if 4 letters, it returns the full form. +- Otherwise, it returns none. +-/ +def classify (letter : Char) (num : Nat) : Option ZoneName := + if letter = 'z' then + if num < 4 then + some (.short) + else if num = 4 then + some (.full) + else + none + else if letter = 'v' then + if num = 1 then + some (.short) + else if num = 4 then + some (.full) + else + none + else + none + +end ZoneName + +/-- +`OffsetX` represents different offset formats based on the number of pattern letters. +The output will vary between the number of pattern letters, whether it's the hour, minute, second, +and whether colons are used. +-/ +inductive OffsetX + /-- Only the hour is output (e.g., "+01") -/ + | hour + /-- Hour and minute without colon (e.g., "+0130") -/ + | hourMinute + /-- Hour and minute with colon (e.g., "+01:30") -/ + | hourMinuteColon + /-- Hour, minute, and second without colon (e.g., "+013015") -/ + | hourMinuteSecond + /-- Hour, minute, and second with colon (e.g., "+01:30:15") -/ + | hourMinuteSecondColon + deriving Repr, Inhabited + +namespace OffsetX + +/-- +`classify` classifies the number of pattern letters into an `OffsetX` format. +-/ +def classify (num : Nat) : Option OffsetX := + if num = 1 then + some (.hour) + else if num = 2 then + some (.hourMinute) + else if num = 3 then + some (.hourMinuteColon) + else if num = 4 then + some (.hourMinuteSecond) + else if num = 5 then + some (.hourMinuteSecondColon) + else + none + +end OffsetX + +/-- +`OffsetO` represents localized offset text formats based on the number of pattern letters. +-/ +inductive OffsetO + /-- Short form of the localized offset (e.g., "GMT+8") -/ + | short + /-- Full form of the localized offset (e.g., "GMT+08:00") -/ + | full + deriving Repr, Inhabited + +namespace OffsetO + +/-- +`classify` classifies the number of pattern letters into an `OffsetO` format. +-/ +def classify (num : Nat) : Option OffsetO := + match num with + | 1 => some (.short) + | 4 => some (.full) + | _ => none + +end OffsetO + +/-- +`OffsetZ` represents different offset formats based on the number of pattern letters (capital 'Z'). +-/ +inductive OffsetZ + /-- Hour and minute without colon (e.g., "+0130") -/ + | hourMinute + /-- Localized offset text in full form (e.g., "GMT+08:00") -/ + | full + /-- Hour, minute, and second with colon (e.g., "+01:30:15") -/ + | hourMinuteSecondColon + deriving Repr, Inhabited + +namespace OffsetZ + /-- -The `Modifier` inductive type represents various formatting options for date and time components. -These modifiers are typically used in formatting functions to generate human-readable date and time strings. +`classify` classifies the number of pattern letters into an `OffsetZ` format. +-/ +def classify (num : Nat) : Option OffsetZ := + match num with + | 1 | 2 | 3 => some (.hourMinute) + | 4 => some (.full) + | 5 => some (.hourMinuteSecondColon) + | _ => none + +end OffsetZ + +/-- +The `Modifier` inductive type represents various formatting options for date and time components, +matching the format symbols used in date and time strings. +These modifiers can be applied in formatting functions to generate custom date and time outputs. -/ inductive Modifier /-- - `YYYY`: Four-digit year (e.g., 2024). + `G`: Era (e.g., AD, Anno Domini, A). -/ - | YYYY + | G (presentation : Text) /-- - `YY`: Two-digit year (e.g., 24 for 2024). + `y`: Year of era (e.g., 2004, 04). -/ - | YY + | y (presentation : Year) /-- - `MMMM`: Full month name (e.g., January, February). + `D`: Day of year (e.g., 189). -/ - | MMMM + | D (presentation : Number) /-- - `MMM`: Abbreviated month name (e.g., Jan, Feb). + `M`: Month of year as number or text (e.g., 7, 07, Jul, July, J). -/ - | MMM + | MorL (presentation : Number ⊕ Text) /-- - `MM`: Two-digit month (e.g., 01 for January). + `d`: Day of month (e.g., 10). -/ - | MM + | d (presentation : Number) /-- - `M`: One or two-digit month (e.g., 1 for January, 10 for October). + `Q`: Quarter of year as number or text (e.g., 3, 03, Q3, 3rd quarter). -/ - | M + | Qorq (presentation : Number ⊕ Text) /-- - `DD`: Two-digit day of the month (e.g., 01, 02). + `w`: Week of week-based year (e.g., 27). -/ - | DD + | w (presentation : Number) /-- - `D`: One or two-digit day of the month (e.g., 1, 2). + `W`: Week of month (e.g., 4). -/ - | D + | W (presentation : Number) /-- - `d`: One or two digit day of the month with space padding at the beggining (e.g. 1, 12). + `E`: Day of week as text (e.g., Tue, Tuesday, T). -/ - | d + | E (presentation : Text) /-- - `EEEE`: Full name of the day of the week (e.g., Monday, Tuesday). + `e`: Localized day of week as number or text (e.g., 2, 02, Tue, Tuesday, T). -/ - | EEEE + | eorc (presentation : Number ⊕ Text) /-- - `EEE`: Abbreviated day of the week (e.g., Mon, Tue). + `F`: Aligned week of month (e.g., 3). -/ - | EEE + | F (presentation : Number) /-- - `hh`: Two-digit hour in 24-hour format (e.g., 01, 02). + `a`: AM/PM of day (e.g., PM). -/ - | hh + | a (presentation : Text) /-- - `h`: One or two-digit hour in 24-hour format (e.g., 1, 2). + `B`: Period of day (e.g., in the morning). -/ - | h + | B (presentation : Text) /-- - `HH`: Two-digit hour in 12-hour format (e.g., 13, 14). + `h`: Clock hour of AM/PM (1-12) (e.g., 12). -/ - | HH + | h (presentation : Number) /-- - `H`: One or two-digit hour in 12-hour format (e.g., 1, 2). + `K`: Hour of AM/PM (0-11) (e.g., 0). -/ - | H + | K (presentation : Number) /-- - `AA`: Uppercase AM/PM indicator (e.g., AM, PM). + `k`: Clock hour of day (1-24) (e.g., 24). -/ - | AA + | k (presentation : Number) /-- - `aa`: Lowercase am/pm indicator (e.g., am, pm). + `H`: Hour of day (0-23) (e.g., 0). -/ - | aa + | H (presentation : Number) /-- - `mm`: Two-digit minute (e.g., 01, 02). + `m`: Minute of hour (e.g., 30). -/ - | mm + | m (presentation : Number) /-- - `m`: One or two-digit minute (e.g., 1, 2). + `s`: Second of minute (e.g., 55). -/ - | m + | s (presentation : Number) /-- - `sssssssss` Nine-digit nanosecond (e.g., 000000001). + `S`: Fraction of second (e.g., 978). -/ - | sssssssss + | S (presentation : Fraction) /-- - `sss`: Three-digit milliseconds (e.g., 001, 202). + `A`: Millisecond of day (e.g., 1234). -/ - | sss + | A (presentation : Number) /-- - `ss`: Two-digit second (e.g., 01, 02). + `n`: Nanosecond of second (e.g., 987654321). -/ - | ss + | n (presentation : Number) /-- - `s`: One or two-digit second (e.g., 1, 2). + `N`: Nanosecond of day (e.g., 1234000000). -/ - | s + | N (presentation : Number) /-- - `ZZZZZ`: Full timezone offset including hours and minutes (e.g., +03:00). + `V`: Time zone ID (e.g., America/Los_Angeles, Z, -08:30). -/ - | ZZZZZ + | V /-- - `ZZZZ`: Timezone offset including hours and minutes without the colon (e.g., +0300). + `z`: Time zone name (e.g., Pacific Standard Time, PST). -/ - | ZZZZ + | z (presentation : ZoneName) /-- - `ZZZ`: Like ZZZZ but with a special case "UTC" for UTC. + `O`: Localized zone offset (e.g., GMT+8, GMT+08:00, UTC-08:00). -/ - | ZZZ + | O (presentation : OffsetO) /-- - `Z`: Like ZZZZZ but with a special case "Z" for UTC. + `X`: Zone offset with 'Z' for zero (e.g., Z, -08, -0830, -08:30). -/ - | Z + | X (presentation : OffsetX) /-- - `z`: Name of the time-zone like (Brasilia Standard Time). + `x`: Zone offset without 'Z' (e.g., +0000, -08, -0830, -08:30). -/ - | z - deriving Repr + | x (presentation : OffsetX) + + /-- + `Z`: Zone offset with 'Z' for UTC (e.g., +0000, -0800, -08:00). + -/ + | Z (presentation : OffsetZ) + deriving Repr, Inhabited /-- -The part of a formatting string. a string is just a text and a modifier is in the format `%0T` where -0 is the quantity of left pad and `T` the `Modifier`. +`abstractParse` abstracts the parsing logic for any type that has a classify function. +It takes a constructor function to build the `Modifier` and a classify function that maps the pattern length to a specific type. +-/ +private def parseMod (constructor : α → Modifier) (classify : Nat → Option α) (p : String) : Parser Modifier := + let len := p.length + match classify len with + | some res => pure (constructor res) + | none => fail s!"invalid quantity of characters for '{p.get 0}'" + +private def parseText (constructor : Text → Modifier) (p : String) : Parser Modifier := + parseMod constructor Text.classify p + +private def parseFraction (constructor : Fraction → Modifier) (p : String) : Parser Modifier := + parseMod constructor Fraction.classify p + +private def parseNumber (constructor : Number → Modifier) (p : String) : Parser Modifier := + pure (constructor ⟨p.length⟩) + +private def parseYear (constructor : Year → Modifier) (p : String) : Parser Modifier := + parseMod constructor Year.classify p + +private def parseOffsetX (constructor : OffsetX → Modifier) (p : String) : Parser Modifier := + parseMod constructor OffsetX.classify p + +private def parseOffsetZ (constructor : OffsetZ → Modifier) (p : String) : Parser Modifier := + parseMod constructor OffsetZ.classify p + +private def parseOffsetO (constructor : OffsetO → Modifier) (p : String) : Parser Modifier := + parseMod constructor OffsetO.classify p + +private def parseZoneId (p : String) : Parser Modifier := + if p.length = 2 then pure .V else fail s!"invalid quantity of characters for '{p.get 0}'" + +private def parseNumberText (constructor : (Number ⊕ Text) → Modifier) (p : String) : Parser Modifier := + parseMod constructor classifyNumberText p + +private def parseZoneName (constructor : ZoneName → Modifier) (p : String) : Parser Modifier := + let len := p.length + match ZoneName.classify (p.get 0) len with + | some res => pure (constructor res) + | none => fail s!"invalid quantity of characters for '{p.get 0}'" + +private def parseModifier : Parser Modifier + := (parseText Modifier.G =<< many1Chars (pchar 'G')) + <|> parseYear Modifier.y =<< many1Chars (pchar 'y') + <|> parseNumber Modifier.D =<< many1Chars (pchar 'D') + <|> parseNumberText Modifier.MorL =<< many1Chars (pchar 'M') + <|> parseNumberText Modifier.MorL =<< many1Chars (pchar 'L') + <|> parseNumber Modifier.d =<< many1Chars (pchar 'd') + <|> parseNumberText Modifier.Qorq =<< many1Chars (pchar 'Q') + <|> parseNumberText Modifier.Qorq =<< many1Chars (pchar 'q') + <|> parseNumber Modifier.w =<< many1Chars (pchar 'w') + <|> parseNumber Modifier.W =<< many1Chars (pchar 'W') + <|> parseText Modifier.E =<< many1Chars (pchar 'E') + <|> parseNumberText Modifier.eorc =<< many1Chars (pchar 'e') + <|> parseNumberText Modifier.eorc =<< many1Chars (pchar 'c') + <|> parseNumber Modifier.F =<< many1Chars (pchar 'F') + <|> parseText Modifier.a =<< many1Chars (pchar 'a') + <|> parseText Modifier.B =<< many1Chars (pchar 'B') + <|> parseNumber Modifier.h =<< many1Chars (pchar 'h') + <|> parseNumber Modifier.K =<< many1Chars (pchar 'K') + <|> parseNumber Modifier.k =<< many1Chars (pchar 'k') + <|> parseNumber Modifier.H =<< many1Chars (pchar 'H') + <|> parseNumber Modifier.m =<< many1Chars (pchar 'm') + <|> parseNumber Modifier.s =<< many1Chars (pchar 's') + <|> parseFraction Modifier.S =<< many1Chars (pchar 'S') + <|> parseNumber Modifier.A =<< many1Chars (pchar 'A') + <|> parseNumber Modifier.n =<< many1Chars (pchar 'n') + <|> parseNumber Modifier.N =<< many1Chars (pchar 'N') + <|> parseZoneId =<< many1Chars (pchar 'V') + <|> parseZoneName Modifier.z =<< many1Chars (pchar 'z') + <|> parseOffsetO Modifier.O =<< many1Chars (pchar 'O') + <|> parseOffsetX Modifier.X =<< many1Chars (pchar 'X') + <|> parseOffsetX Modifier.x =<< many1Chars (pchar 'x') + <|> parseOffsetZ Modifier.Z =<< many1Chars (pchar 'Z') + +/-- +The part of a formatting string. A string is just a text and a modifier is in the format described in +the `Modifier` type. -/ inductive FormatPart - /-- A string literal. -/ + /-- + A string literal. + -/ | string (val : String) - /-- A modifier that renders some data into text. -/ + + /-- + A modifier that renders some data into text. + -/ | modifier (modifier : Modifier) deriving Repr @@ -229,50 +578,12 @@ structure Format (awareness : Awareness) where string : FormatString deriving Inhabited, Repr -private def isNonLetter : Char → Bool := not ∘ Char.isAlpha - -private def parseModifier : Parser Modifier - := pstring "YYYY" *> pure .YYYY - <|> pstring "YY" *> pure .YY - <|> pstring "MMMM" *> pure .MMMM - <|> pstring "MMM" *> pure .MMM - <|> pstring "MM" *> pure .MM - <|> pstring "M" *> pure .M - <|> pstring "DD" *> pure .DD - <|> pstring "D" *> pure .D - <|> pstring "d" *> pure .d - <|> pstring "EEEE" *> pure .EEEE - <|> pstring "EEE" *> pure .EEE - <|> pstring "hh" *> pure .hh - <|> pstring "h" *> pure .h - <|> pstring "HH" *> pure .HH - <|> pstring "H" *> pure .H - <|> pstring "AA" *> pure .AA - <|> pstring "aa" *> pure .aa - <|> pstring "mm" *> pure .mm - <|> pstring "m" *> pure .m - <|> pstring "sssssssss" *> pure .sssssssss - <|> pstring "sss" *> pure .sss - <|> pstring "ss" *> pure .ss - <|> pstring "s" *> pure .s - <|> pstring "ZZZZZ" *> pure .ZZZZZ - <|> pstring "ZZZZ" *> pure .ZZZZ - <|> pstring "ZZZ" *> pure .ZZZ - <|> pstring "Z" *> pure .Z - <|> pstring "z" *> pure .z - -private def isFormatStart : Char → Bool := Char.isAlpha - -private def pnumber : Parser Nat := do - let numbers ← manyChars digit - return numbers.foldl (fun acc char => acc * 10 + (char.toNat - 48)) 0 - private def parseFormatPart : Parser FormatPart := (.modifier <$> parseModifier) <|> (pchar '\\') *> any <&> (.string ∘ toString) <|> (pchar '\"' *> many1Chars (satisfy (· ≠ '\"')) <* pchar '\"') <&> .string <|> (pchar '\'' *> many1Chars (satisfy (· ≠ '\'')) <* pchar '\'') <&> .string - <|> many1Chars (satisfy (fun x => ¬isFormatStart x ∧ x ≠ '\'' ∧ x ≠ '\"')) <&> .string + <|> many1Chars (satisfy (fun x => ¬Char.isAlpha x ∧ x ≠ '\'' ∧ x ≠ '\"')) <&> .string private def specParser : Parser FormatString := (Array.toList <$> many parseFormatPart) <* eof @@ -282,48 +593,73 @@ private def specParse (s : String) : Except String FormatString := -- Pretty printer -private def unabbrevMonth (month : Month.Ordinal) : String := - match month.val, month.property with - | 1, _ => "January" - | 2, _ => "February" - | 3, _ => "March" - | 4, _ => "April" - | 5, _ => "May" - | 6, _ => "June" - | 7, _ => "July" - | 8, _ => "August" - | 9, _ => "September" - | 10, _ => "October" - | 11, _ => "November" - | 12, _ => "December" - -private def abbrevMonth (month : Month.Ordinal) : String := - match month.val, month.property with - | 1, _ => "Jan" - | 2, _ => "Feb" - | 3, _ => "Mar" - | 4, _ => "Apr" - | 5, _ => "May" - | 6, _ => "Jun" - | 7, _ => "Jul" - | 8, _ => "Aug" - | 9, _ => "Sep" - | 10, _ => "Oct" - | 11, _ => "Nov" - | 12, _ => "Dec" - -private def abbrevDayOfWeek (day : Weekday) : String := - match day with - | .sunday => "Sun" - | .monday => "Mon" - | .tuesday => "Tue" - | .wednesday => "Wed" - | .thursday => "Thu" - | .friday => "Fri" - | .saturday => "Sat" +private def leftPad (n : Nat) (a : Char) (s : String) : String := + "".pushn a (n - s.length) ++ s -private def dayOfWeek (day : Weekday) : String := - match day with +private def rightPad (n : Nat) (a : Char) (s : String) : String := + s ++ "".pushn a (n - s.length) + +private def truncate (size : Nat) (n : Int) : String := + let (sign, n) := if n < 0 then ("-", -n) else ("", n) + + let numStr := toString n + if numStr.length > size then + sign ++ numStr.drop (numStr.length - size) + else + sign ++ leftPad size '0' numStr + +private def rightTruncate (size : Nat) (n : Int) : String := + let (sign, n) := if n < 0 then ("-", -n) else ("", n) + + let numStr := toString n + if numStr.length > size then + sign ++ numStr.drop (numStr.length - size) + else + sign ++ rightPad size '0' numStr + +private def formatMonthLong : Month.Ordinal → String + | ⟨1, _⟩ => "January" + | ⟨2, _⟩ => "February" + | ⟨3, _⟩ => "March" + | ⟨4, _⟩ => "April" + | ⟨5, _⟩ => "May" + | ⟨6, _⟩ => "June" + | ⟨7, _⟩ => "July" + | ⟨8, _⟩ => "August" + | ⟨9, _⟩ => "September" + | ⟨10, _⟩ => "October" + | ⟨11, _⟩ => "November" + | ⟨12, _⟩ => "December" + +private def formatMonthShort : Month.Ordinal → String + | ⟨1, _⟩ => "Jan" + | ⟨2, _⟩ => "Feb" + | ⟨3, _⟩ => "Mar" + | ⟨4, _⟩ => "Apr" + | ⟨5, _⟩ => "May" + | ⟨6, _⟩ => "Jun" + | ⟨7, _⟩ => "Jul" + | ⟨8, _⟩ => "Aug" + | ⟨9, _⟩ => "Sep" + | ⟨10, _⟩ => "Oct" + | ⟨11, _⟩ => "Nov" + | ⟨12, _⟩ => "Dec" + +private def formatMonthNarrow : Month.Ordinal → String + | ⟨1, _⟩ => "J" + | ⟨2, _⟩ => "F" + | ⟨3, _⟩ => "M" + | ⟨4, _⟩ => "A" + | ⟨5, _⟩ => "M" + | ⟨6, _⟩ => "J" + | ⟨7, _⟩ => "J" + | ⟨8, _⟩ => "A" + | ⟨9, _⟩ => "S" + | ⟨10, _⟩ => "O" + | ⟨11, _⟩ => "N" + | ⟨12, _⟩ => "D" + +private def formatWeekdayLong : Weekday → String | .sunday => "Sunday" | .monday => "Monday" | .tuesday => "Tuesday" @@ -332,305 +668,723 @@ private def dayOfWeek (day : Weekday) : String := | .friday => "Friday" | .saturday => "Saturday" -private def leftPad (n : Nat) (a : Char) (s : String) : String := - "".pushn a (n - s.length) ++ s - -private def leftPadNum (n : Nat) (s : Int) : String := - let str := if s < 0 then toString (-s) else toString s - let start := if s < 0 then "-" else "" - start ++ "".pushn '0' (n - str.length) ++ str - -private def formatWithDate (date : DateTime tz) : Modifier → String - | .YYYY => s!"{leftPadNum 4 (date.year)}" - | .YY => s!"{leftPadNum 2 (date.year.toInt % 100)}" - | .MMMM => unabbrevMonth date.month - | .MMM => abbrevMonth date.month - | .MM => s!"{leftPadNum 2 (date.month.toInt)}" - | .M => s!"{date.month.toInt}" - | .DD => s!"{leftPadNum 2 (date.day.toInt)}" - | .D => s!"{date.day.toInt}" - | .d => s!"{leftPad 2 ' ' <| toString date.day.toInt}" - | .EEEE => dayOfWeek date.weekday - | .EEE => abbrevDayOfWeek date.weekday - | .hh => s!"{leftPadNum 2 (date.hour.toInt)}" - | .h => s!"{date.hour.toInt}" - | .HH => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{leftPadNum 2 <| hour}" - | .H => let hour := date.hour.val % 12; if hour == 0 then "12" else s!"{hour}" - | .AA => if date.hour.toInt < 12 then "AM" else "PM" - | .aa => if date.hour.toInt < 12 then "am" else "pm" - | .mm => s!"{leftPadNum 2 <| date.minute.toInt}" - | .m => s!"{date.minute.toInt}" - | .sssssssss => s!"{leftPadNum 9 <| date.nanoseconds.toInt}" - | .sss => s!"{leftPadNum 3 <| date.milliseconds.toInt}" - | .ss => s!"{leftPadNum 2 <| date.second.toInt}" - | .s => s!"{date.second.toInt}" - | .ZZZZZ => tz.offset.toIsoString true - | .ZZZZ => tz.offset.toIsoString false - | .ZZZ => if tz.offset.second.val = 0 then "UTC" else tz.offset.toIsoString false - | .Z => if tz.offset.second.val = 0 then "Z" else tz.offset.toIsoString true - | .z => tz.name - -private def formatPartWithDate (date : DateTime z) : FormatPart → String - | .string s => s - | .modifier t => formatWithDate date t - --- Parser +private def formatWeekdayShort : Weekday → String + | .sunday => "Sun" + | .monday => "Mon" + | .tuesday => "Tue" + | .wednesday => "Wed" + | .thursday => "Thu" + | .friday => "Fri" + | .saturday => "Sat" -@[simp] -private def SingleFormatType : Modifier → Type - | .YYYY | .YY => Year.Offset - | .MMMM | .MMM | .MM | .M => Month.Ordinal - | .DD | .D | .d => Day.Ordinal - | .EEEE | .EEE => Weekday - | .hh | .h | .HH | .H => Sigma Hour.Ordinal - | .AA | .aa => HourMarker - | .mm | .m => Minute.Ordinal - | .sssssssss => Nanosecond.Ordinal - | .sss => Millisecond.Ordinal - | .ss | .s => Sigma Second.Ordinal - | .ZZZZZ | .ZZZZ | .ZZZ | .Z => Offset - | .z => String - -private def formatPart (modifier : Modifier) (data : SingleFormatType modifier) : String := +private def formatWeekdayNarrow : Weekday → String + | .sunday => "S" + | .monday => "M" + | .tuesday => "T" + | .wednesday => "W" + | .thursday => "T" + | .friday => "F" + | .saturday => "S" + +private def formatEraShort : Year.Era → String + | .bce => "BCE" + | .ce => "CE" + +private def formatEraLong : Year.Era → String + | .bce => "Before Common Era" + | .ce => "Common Era" + +private def formatEraNarrow : Year.Era → String + | .bce => "B" + | .ce => "C" + +private def formatQuarterNumber : Bounded.LE 1 4 → String + |⟨1, _⟩ => "1" + |⟨2, _⟩ => "2" + |⟨3, _⟩ => "3" + |⟨4, _⟩ => "4" + +private def formatQuarterShort : Bounded.LE 1 4 → String + | ⟨1, _⟩ => "Q1" + | ⟨2, _⟩ => "Q2" + | ⟨3, _⟩ => "Q3" + | ⟨4, _⟩ => "Q4" + +private def formatQuarterLong : Bounded.LE 1 4 → String + | ⟨1, _⟩ => "1st quarter" + | ⟨2, _⟩ => "2nd quarter" + | ⟨3, _⟩ => "3rd quarter" + | ⟨4, _⟩ => "4th quarter" + +private def formatMarkerShort (marker : HourMarker) : String := + match marker with + | .am => "AM" + | .pm => "PM" + +private def formatMarkerLong (marker : HourMarker) : String := + match marker with + | .am => "Ante Meridiem" + | .pm => "Post Meridiem" + +private def formatMarkerNarrow (marker : HourMarker) : String := + match marker with + | .am => "A" + | .pm => "P" + +private def formatPeriodOfDayLong : Day.Ordinal.Period → String + | .morning => "in the morning" + | .afternoon => "in the afternoon" + | .evening => "in the evening" + | .night => "at night" + +private def formatPeriodOfDayShort : Day.Ordinal.Period → String + | .morning => "AM" + | .afternoon => "PM" + | .evening => "Eve" + | .night => "Night" + +private def formatPeriodOfDayNarrow : Day.Ordinal.Period → String + | .morning => "M" + | .afternoon => "A" + | .evening => "E" + | .night => "N" + +private def toSigned (data : Int) : String := + if data < 0 then toString data else "+" ++ toString data + +private def toIsoString (offset : Offset) (withMinutes : Bool) (withSeconds : Bool) (colon : Bool) : String := + let (sign, time) := if offset.second.val > 0 then ("+", offset.second) else ("-", -offset.second) + let time := PlainTime.ofSeconds time + let pad := leftPad 2 '0' ∘ toString + + let data := s!"{sign}{pad time.hour.val}" + let data := if withMinutes then s!"{data}{if colon then ":" else ""}{pad time.minute.val}" else data + let data := if withSeconds then s!"{data}{if colon then ":" else ""}{pad time.second.snd.val}" else data + + data + +private def TypeFormat : Modifier → Type + | .G _ => Year.Era + | .y _ => Year.Offset + | .D _ => Sigma Day.Ordinal.OfYear + | .MorL _ => Month.Ordinal + | .d _ => Day.Ordinal + | .Qorq _ => Month.Ordinal.Quarter + | .w _ => Week.Ordinal + | .W _ => Week.Ordinal.OfMonth + | .E _ => Weekday + | .eorc _ => Weekday + | .F _ => Week.Ordinal.OfMonth + | .a _ => HourMarker + | .B _ => Day.Ordinal.Period + | .h _ => Bounded.LE 1 12 + | .K _ => Bounded.LE 0 11 + | .k _ => Bounded.LE 1 24 + | .H _ => Hour.Ordinal + | .m _ => Minute.Ordinal + | .s _ => Sigma Second.Ordinal + | .S _ => Nanosecond.Ordinal + | .A _ => Millisecond.Offset + | .n _ => Nanosecond.Ordinal + | .N _ => Nanosecond.Offset + | .V => String + | .z _ => String + | .O _ => Offset + | .X _ => Offset + | .x _ => Offset + | .Z _ => Offset + +private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : String := match modifier with - | .YYYY => s!"{leftPadNum 4 (data.toInt)}" - | .YY => s!"{leftPadNum 2 (data.toInt % 100)}" - | .MMMM => unabbrevMonth data - | .MMM => abbrevMonth data - | .MM => s!"{leftPadNum 2 (data.toInt)}" - | .M => s!"{data.toInt}" - | .DD => s!"{leftPadNum 2 (data.toInt)}" - | .D => s!"{data.toInt}" - | .d => s!"{leftPad 2 ' ' <| toString data.toInt}" - | .EEEE => dayOfWeek data - | .EEE => abbrevDayOfWeek data - | .hh => s!"{leftPadNum 2 (data.snd.toInt)}" - | .h => s!"{data.snd.toInt}" - | .HH => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{leftPadNum 2 hour}" - | .H => let hour := data.snd.val % 12; if hour == 0 then "12" else s!"{hour}" - | .AA => match data with | .am => "AM" | .pm => "PM" - | .aa => match data with | .am => "am" | .pm => "pm" - | .mm => s!"{leftPadNum 2 data.toInt}" - | .m => s!"{data.toInt}" - | .sssssssss => s!"{leftPadNum 9 data.toInt}" - | .sss => s!"{leftPadNum 3 data.toInt}" - | .ss => s!"{leftPadNum 2 data.snd.toInt}" - | .s => s!"{data.snd.toInt}" - | .ZZZZZ => data.toIsoString true - | .ZZZZ => data.toIsoString false - | .ZZZ => if data.second.val = 0 then "UTC" else data.toIsoString false - | .Z => if data.second.val = 0 then "Z" else data.toIsoString true - | .z => data + | .G format => + match format with + | .short => formatEraShort data + | .full => formatEraLong data + | .narrow => formatEraNarrow data + | .y format => + match format with + | .twoDigit => truncate 2 (data.toInt % 100) + | .fourDigit => truncate 4 data.toInt + | .extended n => truncate n data.toInt + | .D format => + truncate format.padding data.snd.val + | .MorL format => + match format with + | .inl format => truncate format.padding data.val + | .inr .short => formatMonthShort data + | .inr .full => formatMonthLong data + | .inr .narrow => formatMonthNarrow data + | .d format => + truncate format.padding data.val + | .Qorq format => + match format with + | .inl format => truncate format.padding data.val + | .inr .short => formatQuarterShort data + | .inr .full => formatQuarterLong data + | .inr .narrow => formatQuarterNumber data + | .w format => + truncate format.padding data.val + | .W format => + truncate format.padding data.val + | .E format => + match format with + | .short => formatWeekdayShort data + | .full => formatWeekdayLong data + | .narrow => formatWeekdayNarrow data + | .eorc format => + match format with + | .inl format => truncate format.padding data.toOrdinal.val + | .inr .short => formatWeekdayShort data + | .inr .full => formatWeekdayLong data + | .inr .narrow => formatWeekdayNarrow data + | .F format => + truncate format.padding data.val + | .a format => + match format with + | .short => formatMarkerShort data + | .full => formatMarkerLong data + | .narrow => formatMarkerNarrow data + | .B format => + match format with + | .short => formatPeriodOfDayLong data + | .full => formatPeriodOfDayShort data + | .narrow => formatPeriodOfDayNarrow data + | .h format => truncate format.padding (data.val % 12) + | .K format => truncate format.padding (data.val % 12) + | .k format => truncate format.padding (data.val) + | .H format => truncate format.padding (data.val) + | .m format => truncate format.padding (data.val) + | .s format => truncate format.padding (data.snd.val) + | .S format => + match format with + | .nano => truncate 9 data.val + | .truncated n => rightTruncate n data.val + | .A format => + truncate format.padding data.val + | .n format => + truncate format.padding data.val + | .N format => + truncate format.padding data.val + | .V => data + | .z format => + match format with + | .short => data + | .full => data + | .O format => + match format with + | .short => s!"GMT{toSigned data.hour.val}" + | .full => s!"GMT{toIsoString data true false true}" + | .X format => + if data.second == 0 then + "Z" + else + match format with + | .hour => toIsoString data false false false + | .hourMinute => toIsoString data true false false + | .hourMinuteColon => toIsoString data true false true + | .hourMinuteSecond => toIsoString data true true false + | .hourMinuteSecondColon => toIsoString data true true true + | .x format => + match format with + | .hour => + toIsoString data (data.second.toMinutes.val % 60 ≠ 0) false false + | .hourMinute => + toIsoString data true false false + | .hourMinuteColon => + toIsoString data true (data.second.val % 60 ≠ 0) true + | .hourMinuteSecond => + toIsoString data true (data.second.val % 60 ≠ 0) false + | .hourMinuteSecondColon => + toIsoString data true true true + | .Z format => + match format with + | .hourMinute => + toIsoString data true false false + | .full => + if data.second.val = 0 + then "GMT" + else s!"GMT{toIsoString data true false true}" + | .hourMinuteSecondColon => + if data.second == 0 + then "Z" + else toIsoString data true (data.second.val % 60 ≠ 0) true + +private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := + match modifier with + | .G _ => date.era + | .y _ => date.year + | .D _ => Sigma.mk _ date.toOrdinal + | .MorL _ => date.month + | .d _ => date.day + | .Qorq _ => date.quarter + | .w _ => date.toWeekOfYear + | .W _ => date.toWeekOfMonth + | .E _ => date.weekday + | .eorc _ => date.weekday + | .F _ => date.toWeekOfMonth + | .a _ => HourMarker.ofOrdinal date.hour + | .B _ => date.getPeriod + | .h _ => (date.hour.sub 1).emod 12 (by decide) |>.add 1 + | .K _ => date.hour.emod 12 (by decide) + | .k _ => date.hour.add 1 + | .H _ => date.hour + | .m _ => date.minute + | .s _ => date.date.get.time.second + | .S _ => date.nanoseconds + | .A _ => date.date.get.time.toMilliseconds + | .n _ => date.nanoseconds + | .N _ => date.date.get.time.toNanoseconds + | .V => tz.name + | .z _ => tz.abbreviation + | .O _ => tz.offset + | .X _ => tz.offset + | .x _ => tz.offset + | .Z _ => tz.offset + +private def parseMonthLong : Parser Month.Ordinal + := pstring "January" *> pure ⟨1, by decide⟩ + <|> pstring "February" *> pure ⟨2, by decide⟩ + <|> pstring "March" *> pure ⟨3, by decide⟩ + <|> pstring "April" *> pure ⟨4, by decide⟩ + <|> pstring "May" *> pure ⟨5, by decide⟩ + <|> pstring "June" *> pure ⟨6, by decide⟩ + <|> pstring "July" *> pure ⟨7, by decide⟩ + <|> pstring "August" *> pure ⟨8, by decide⟩ + <|> pstring "September" *> pure ⟨9, by decide⟩ + <|> pstring "October" *> pure ⟨10, by decide⟩ + <|> pstring "November" *> pure ⟨11, by decide⟩ + <|> pstring "December" *> pure ⟨12, by decide⟩ + +private def parseMonthShort : Parser Month.Ordinal + := pstring "Jan" *> pure ⟨1, by decide⟩ + <|> pstring "Feb" *> pure ⟨2, by decide⟩ + <|> pstring "Mar" *> pure ⟨3, by decide⟩ + <|> pstring "Apr" *> pure ⟨4, by decide⟩ + <|> pstring "May" *> pure ⟨5, by decide⟩ + <|> pstring "Jun" *> pure ⟨6, by decide⟩ + <|> pstring "Jul" *> pure ⟨7, by decide⟩ + <|> pstring "Aug" *> pure ⟨8, by decide⟩ + <|> pstring "Sep" *> pure ⟨9, by decide⟩ + <|> pstring "Oct" *> pure ⟨10, by decide⟩ + <|> pstring "Nov" *> pure ⟨11, by decide⟩ + <|> pstring "Dec" *> pure ⟨12, by decide⟩ + +private def parseMonthNarrow : Parser Month.Ordinal + := pstring "J" *> pure ⟨1, by decide⟩ + <|> pstring "F" *> pure ⟨2, by decide⟩ + <|> pstring "M" *> pure ⟨3, by decide⟩ + <|> pstring "A" *> pure ⟨4, by decide⟩ + <|> pstring "M" *> pure ⟨5, by decide⟩ + <|> pstring "J" *> pure ⟨6, by decide⟩ + <|> pstring "J" *> pure ⟨7, by decide⟩ + <|> pstring "A" *> pure ⟨8, by decide⟩ + <|> pstring "S" *> pure ⟨9, by decide⟩ + <|> pstring "O" *> pure ⟨10, by decide⟩ + <|> pstring "N" *> pure ⟨11, by decide⟩ + <|> pstring "D" *> pure ⟨12, by decide⟩ + +private def parseWeekdayLong : Parser Weekday + := pstring "Sunday" *> pure Weekday.sunday + <|> pstring "Monday" *> pure Weekday.monday + <|> pstring "Tuesday" *> pure Weekday.tuesday + <|> pstring "Wednesday" *> pure Weekday.wednesday + <|> pstring "Thursday" *> pure Weekday.thursday + <|> pstring "Friday" *> pure Weekday.friday + <|> pstring "Saturday" *> pure Weekday.saturday + +private def parseWeekdayShort : Parser Weekday + := pstring "Sun" *> pure Weekday.sunday + <|> pstring "Mon" *> pure Weekday.monday + <|> pstring "Tue" *> pure Weekday.tuesday + <|> pstring "Wed" *> pure Weekday.wednesday + <|> pstring "Thu" *> pure Weekday.thursday + <|> pstring "Fri" *> pure Weekday.friday + <|> pstring "Sat" *> pure Weekday.saturday + +private def parseWeekdayNarrow : Parser Weekday + := pstring "S" *> pure Weekday.sunday + <|> pstring "M" *> pure Weekday.monday + <|> pstring "T" *> pure Weekday.tuesday + <|> pstring "W" *> pure Weekday.wednesday + <|> pstring "T" *> pure Weekday.thursday + <|> pstring "F" *> pure Weekday.friday + <|> pstring "S" *> pure Weekday.saturday + +private def parseEraShort : Parser Year.Era + := pstring "BCE" *> pure Year.Era.bce + <|> pstring "CE" *> pure Year.Era.ce + +private def parseEraLong : Parser Year.Era + := pstring "Before Common Era" *> pure Year.Era.bce + <|> pstring "Common Era" *> pure Year.Era.ce + +private def parseEraNarrow : Parser Year.Era + := pstring "B" *> pure Year.Era.bce + <|> pstring "C" *> pure Year.Era.ce + +private def parseQuarterNumber : Parser Month.Ordinal.Quarter + := pstring "1" *> pure ⟨1, by decide⟩ + <|> pstring "2" *> pure ⟨2, by decide⟩ + <|> pstring "3" *> pure ⟨3, by decide⟩ + <|> pstring "4" *> pure ⟨4, by decide⟩ + +private def parseQuarterLong : Parser Month.Ordinal.Quarter + := pstring "1st quarter" *> pure ⟨1, by decide⟩ + <|> pstring "2nd quarter" *> pure ⟨2, by decide⟩ + <|> pstring "3rd quarter" *> pure ⟨3, by decide⟩ + <|> pstring "4th quarter" *> pure ⟨4, by decide⟩ + +private def parseQuarterShort : Parser Month.Ordinal.Quarter + := pstring "Q1" *> pure ⟨1, by decide⟩ + <|> pstring "Q2" *> pure ⟨2, by decide⟩ + <|> pstring "Q3" *> pure ⟨3, by decide⟩ + <|> pstring "Q4" *> pure ⟨4, by decide⟩ + +private def parseMarkerShort : Parser HourMarker + := pstring "AM" *> pure HourMarker.am + <|> pstring "PM" *> pure HourMarker.pm + +private def parseMarkerLong : Parser HourMarker + := pstring "Ante Meridiem" *> pure HourMarker.am + <|> pstring "Post Meridiem" *> pure HourMarker.pm + +private def parseMarkerNarrow : Parser HourMarker + := pstring "A" *> pure HourMarker.am + <|> pstring "P" *> pure HourMarker.pm + +private def parsePeriodOfDayLong : Parser Day.Ordinal.Period + := pstring "in the morning" *> pure .morning + <|> pstring "in the afternoon" *> pure .afternoon + <|> pstring "in the evening" *> pure .evening + <|> pstring "at night" *> pure .night + +private def parsePeriodOfDayShort : Parser Day.Ordinal.Period + := pstring "AM" *> pure .morning + <|> pstring "PM" *> pure .afternoon + <|> pstring "Eve" *> pure .evening + <|> pstring "Night" *> pure .night + +private def parsePeriodOfDayNarrow : Parser Day.Ordinal.Period + := pstring "M" *> pure .morning + <|> pstring "A" *> pure .afternoon + <|> pstring "E" *> pure .evening + <|> pstring "N" *> pure .night + +private def exactly (parse : Parser α) (size : Nat) : Parser (Array α) := + let rec go (acc : Array α) (count : Nat) : Parser (Array α) := + if count ≥ size then + pure acc + else do + let res ← parse + go (acc.push res) count.succ + termination_by size - count + + go #[] 12 + +private def exactlyChars (parse : Parser Char) (size : Nat) : Parser String := + let rec go (acc : String) (count : Nat) : Parser String := + if count ≥ size then + pure acc + else do + let res ← parse + go (acc.push res) count.succ + termination_by size - count + + go "" 0 + +private def parseSigned (parser : Parser Nat) : Parser Int := do + let signed ← optional (pstring "-") + let res ← parser + if signed.isSome then + return -res + else + return res + +private def parseNum (size : Nat) : Parser Nat := + String.toNat! <$> exactlyChars (satisfy Char.isDigit) size + +private def parseFractionNum (size : Nat) (pad : Nat) : Parser Nat := + String.toNat! <$> rightPad pad '0' <$> exactlyChars (satisfy Char.isDigit) size + +private def parseIdentifier : Parser String := + many1Chars (satisfy (fun x => x.isAlpha ∨ x.isDigit ∨ x = '_' ∨ x = '-' ∨ x = '/')) + +private def parseNatToBounded { n m : Nat } (parser : Parser Nat) : Parser (Bounded.LE n m) := do + let res ← parser + if h : n ≤ res ∧ res ≤ m then + return Bounded.LE.ofNat' res h + else + fail s!"need a natural number in the interval of {n} to {m}" + +private inductive Reason + | yes + | no + | optional + +private def parseOffset (withMinutes : Reason) (withSeconds : Reason) (withColon : Bool) : Parser Offset := do + let sign : Second.Offset ← (pchar '+' *> pure 1) <|> (pchar '-' *> pure (-1)) + let hours : Hour.Offset ← UnitVal.mk <$> parseNum 2 + + let colon := if withColon then pchar ':' else pure ':' + + let parseUnit {n} (reason : Reason) : Parser (Option (UnitVal n)) := + match reason with + | .yes => some <$> (colon *> UnitVal.mk <$> parseNum 2) + | .no => pure none + | .optional => optional (colon *> UnitVal.mk <$> parseNum 2) + + let minutes : Option Minute.Offset ← parseUnit withMinutes + let seconds : Option Second.Offset ← parseUnit withSeconds + + let hours := hours.toSeconds + (minutes.getD 0).toSeconds + (seconds.getD 0) + + return Offset.ofSeconds (sign * hours) + +private def parseWith : (mod : Modifier) → Parser (TypeFormat mod) + | .G format => + match format with + | .short => parseEraShort + | .full => parseEraLong + | .narrow => parseEraNarrow + | .y format => + match format with + | .twoDigit => (2000 + ·) <$> (parseSigned <| parseNum 2) + | .fourDigit => parseSigned <| parseNum 4 + | .extended n => parseSigned <| parseNum n + | .D format => Sigma.mk true <$> parseNatToBounded (parseNum format.padding) + | .MorL format => + match format with + | .inl format => parseNatToBounded (parseNum format.padding) + | .inr .short => parseMonthShort + | .inr .full => parseMonthLong + | .inr .narrow => parseMonthNarrow + | .d format => parseNatToBounded (parseNum format.padding) + | .Qorq format => + match format with + | .inl format => parseNatToBounded (parseNum format.padding) + | .inr .short => parseQuarterShort + | .inr .full => parseQuarterLong + | .inr .narrow => parseQuarterNumber + | .w format => parseNatToBounded (parseNum format.padding) + | .W format => parseNatToBounded (parseNum format.padding) + | .E format => + match format with + | .short => parseWeekdayShort + | .full => parseWeekdayLong + | .narrow => parseWeekdayNarrow + | .eorc format => + match format with + | .inl format => Weekday.ofOrdinal <$> parseNatToBounded (parseNum format.padding) + | .inr .short => parseWeekdayShort + | .inr .full => parseWeekdayLong + | .inr .narrow => parseWeekdayNarrow + | .F format => parseNatToBounded (parseNum format.padding) + | .a format => + match format with + | .short => parseMarkerShort + | .full => parseMarkerLong + | .narrow => parseMarkerNarrow + | .B format => + match format with + | .short => parsePeriodOfDayLong + | .full => parsePeriodOfDayShort + | .narrow => parsePeriodOfDayNarrow + | .h format => parseNatToBounded (parseNum format.padding) + | .K format => parseNatToBounded (parseNum format.padding) + | .k format => parseNatToBounded (parseNum format.padding) + | .H format => parseNatToBounded (parseNum format.padding) + | .m format => parseNatToBounded (parseNum format.padding) + | .s format => Sigma.mk true <$> (parseNatToBounded (parseNum format.padding)) + | .S format => + match format with + | .nano => parseNatToBounded (parseNum 9) + | .truncated n => parseNatToBounded (parseFractionNum n 9) + | .A format => Millisecond.Offset.ofNat <$> (parseNum format.padding) + | .n format => parseNatToBounded (parseNum format.padding) + | .N format => Nanosecond.Offset.ofNat <$> (parseNum format.padding) + | .V => parseIdentifier + | .z format => + match format with + | .short => parseIdentifier + | .full => parseIdentifier + | .O format => + match format with + | .short => pstring "GMT" *> parseOffset .no .no false + | .full => pstring "GMT" *> parseOffset .yes .optional false + | .X format => + let p : Parser Offset := + match format with + | .hour => parseOffset .no .no false + | .hourMinute => parseOffset .yes .no false + | .hourMinuteColon => parseOffset .yes .no true + | .hourMinuteSecond => parseOffset .yes .yes false + | .hourMinuteSecondColon => parseOffset .yes .yes true + p <|> (pstring "Z" *> pure (Offset.ofSeconds 0)) + | .x format => + match format with + | .hour => + parseOffset .optional .no false + | .hourMinute => + parseOffset .yes .no false + | .hourMinuteColon => + parseOffset .yes .optional true + | .hourMinuteSecond => + parseOffset .yes .optional false + | .hourMinuteSecondColon => + parseOffset .yes .yes true + | .Z format => + match format with + | .hourMinute => + parseOffset .yes .no false + | .full => do + skipString "GMT" + let res ← optional (parseOffset .yes .no true) + return res.getD Offset.zero + | .hourMinuteSecondColon => + (skipString "Z" *> pure Offset.zero) + <|> (parseOffset .yes .optional true) + +private def formatPartWithDate (date : DateTime tz) (part : FormatPart) : String := + match part with + | .modifier mod => formatWith mod (dateFromModifier date) + | .string s => s @[simp] private def FormatType (result : Type) : FormatString → Type - | .modifier entry :: xs => (SingleFormatType entry) → (FormatType result xs) + | .modifier entry :: xs => (TypeFormat entry) → (FormatType result xs) | .string _ :: xs => (FormatType result xs) | [] => result -private def position : Parser Nat := fun s => (ParseResult.success s (s.pos.byteIdx)) - -private def size (data : Parser α) : Parser (α × Nat) := do - let st ← position - let res ← data - let en ← position - pure (res, en-st) - -private def transform (n: β → Option α) (p: Parser β) : Parser α := do - let res ← p - match n res with - | some n => pure n - | none => fail "cannot parse" - -private def parseMonth : Parser Month.Ordinal - := (pstring "Jan" *> pure 1) - <|> (pstring "Feb" *> pure 2) - <|> (pstring "Mar" *> pure 3) - <|> (pstring "Apr" *> pure 4) - <|> (pstring "May" *> pure 5) - <|> (pstring "Jun" *> pure 6) - <|> (pstring "Jul" *> pure 7) - <|> (pstring "Aug" *> pure 8) - <|> (pstring "Sep" *> pure 9) - <|> (pstring "Oct" *> pure 10) - <|> (pstring "Nov" *> pure 11) - <|> (pstring "Dec" *> pure 12) - -private def parseMonthUnabbrev : Parser Month.Ordinal - := (pstring "January" *> pure 1) - <|> (pstring "February" *> pure 2) - <|> (pstring "March" *> pure 3) - <|> (pstring "April" *> pure 4) - <|> (pstring "May" *> pure 5) - <|> (pstring "June" *> pure 6) - <|> (pstring "July" *> pure 7) - <|> (pstring "August" *> pure 8) - <|> (pstring "September" *> pure 9) - <|> (pstring "October" *> pure 10) - <|> (pstring "November" *> pure 11) - <|> (pstring "December" *> pure 12) - -private def parseWeekday : Parser Weekday - := (pstring "Mon" *> pure Weekday.monday) - <|> (pstring "Tue" *> pure Weekday.tuesday) - <|> (pstring "Wed" *> pure Weekday.wednesday) - <|> (pstring "Thu" *> pure Weekday.thursday) - <|> (pstring "Fri" *> pure Weekday.friday) - <|> (pstring "Sat" *> pure Weekday.saturday) - <|> (pstring "Sun" *> pure Weekday.sunday) - -private def parseWeekdayUnnabrev : Parser Weekday - := (pstring "Monday" *> pure Weekday.monday) - <|> (pstring "Tuesday" *> pure Weekday.tuesday) - <|> (pstring "Wednesday" *> pure Weekday.wednesday) - <|> (pstring "Thursday" *> pure Weekday.thursday) - <|> (pstring "Friday" *> pure Weekday.friday) - <|> (pstring "Saturday" *> pure Weekday.saturday) - <|> (pstring "Sunday" *> pure Weekday.sunday) - -private def parserUpperHourMarker : Parser HourMarker - := (pstring "AM" *> pure HourMarker.am) - <|> (pstring "PM" *> pure HourMarker.pm) - -private def parserLowerHourMarker : Parser HourMarker - := (pstring "am" *> pure HourMarker.am) - <|> (pstring "pm" *> pure HourMarker.pm) - -private def threeDigit : Parser Int := do - let digit1 ← digit - let digit2 ← digit - let digit3 ← digit - return String.toNat! s!"{digit1}{digit2}{digit3}" - -private def twoDigit : Parser Int := do - let digit1 ← digit - let digit2 ← digit - return String.toNat! s!"{digit1}{digit2}" - -private def parseYearTwo : Parser Int :=do - let year ← twoDigit - return if year < 70 then 2000 + year else 1900 + year - -private def timeOffset (colon: Bool) : Parser Offset := do - let sign : Int ← (pstring "-" *> pure (-1)) <|> (pstring "+" *> pure 1) - let hour ← twoDigit - if colon then discard <| pstring ":" - let minutes ← twoDigit - let res := (hour * 3600 + minutes * 60) * sign - pure (Offset.ofSeconds (UnitVal.ofInt res)) - -private def timeOrUTC (utcString: String) (colon: Bool) : Parser Offset := - (pstring utcString *> pure Offset.zero) <|> timeOffset colon - -private def number : Parser Nat := do - String.toNat! <$> many1Chars digit - -private def singleDigit : Parser Nat := do - let digit1 ← digit - return String.toNat! s!"{digit1}" - -private def fourDigit : Parser Int := do - let digit1 ← digit - let digit2 ← digit - let digit3 ← digit - let digit4 ← digit - return String.toNat! s!"{digit1}{digit2}{digit3}{digit4}" - -private def parserWithFormat : (typ: Modifier) → Parser (SingleFormatType typ) - | .YYYY => fourDigit - | .YY => parseYearTwo - | .MMMM => parseMonthUnabbrev - | .MMM => parseMonth - | .MM => transform Bounded.LE.ofInt twoDigit - | .M => transform Bounded.LE.ofInt number - | .DD => transform Bounded.LE.ofInt twoDigit - | .D => transform Bounded.LE.ofInt number - | .d => transform Bounded.LE.ofInt (orElse twoDigit (fun _ => pchar ' ' *> (singleDigit))) - | .EEEE => parseWeekdayUnnabrev - | .EEE => parseWeekday - | .hh => Sigma.mk true <$> transform Bounded.LE.ofInt twoDigit - | .h => Sigma.mk true <$> transform Bounded.LE.ofInt number - | .HH => do - let res : Bounded.LE 0 12 ← transform Bounded.LE.ofInt twoDigit - return Sigma.mk true (res.expandTop (by decide)) - | .H => do - let res : Bounded.LE 0 12 ← transform Bounded.LE.ofInt number - return Sigma.mk true (res.expandTop (by decide)) - | .AA => parserUpperHourMarker - | .aa => parserLowerHourMarker - | .mm => transform Bounded.LE.ofInt twoDigit - | .m => transform Bounded.LE.ofInt number - | .sssssssss => transform Bounded.LE.ofInt threeDigit - | .sss => transform Bounded.LE.ofInt threeDigit - | .ss => Sigma.mk true <$> transform Bounded.LE.ofInt twoDigit - | .s => Sigma.mk true <$> transform Bounded.LE.ofInt number - | .ZZZZZ => timeOffset true - | .ZZZZ => timeOffset false - | .ZZZ => timeOrUTC "UTC" false - | .Z => timeOrUTC "Z" true - | .z => many1Chars (satisfy (fun c => c == ' ' || c.isAlpha)) +namespace Format private structure DateBuilder where - tzName : String := "Greenwich Mean Time" - tz : Offset := Offset.zero - year : Year.Offset := 0 - month : Month.Ordinal := 1 - day : Day.Ordinal := 1 - hour : Sigma Hour.Ordinal := ⟨true, 0⟩ - minute : Minute.Ordinal := 0 - second : Sigma Second.Ordinal := ⟨true, 0⟩ - nanoseconds : Nanosecond.Ordinal := 0 - marker : Option HourMarker := none - -private def DateBuilder.build (builder : DateBuilder) (aw : Awareness) : Except String aw.type := do - - let hour ← - if let some marker := builder.marker then - match Bounded.ofInt? builder.hour.snd.val with - | some res => pure ⟨true, marker.toAbsolute res⟩ - | none => .error "The 24-hour is out of the range and cannot be transformed into a 12-hour with a marker." + G : Option Year.Era := none + y : Option Year.Offset := none + D : Option (Sigma Day.Ordinal.OfYear) := none + MorL : Option Month.Ordinal := none + d : Option Day.Ordinal := none + Qorq : Option Month.Ordinal.Quarter := none + w : Option Week.Ordinal := none + W : Option Week.Ordinal.OfMonth := none + E : Option Weekday := none + eorc : Option Weekday := none + F : Option Week.Ordinal.OfMonth := none + a : Option HourMarker := none + B : Option Day.Ordinal.Period := none + h : Option (Bounded.LE 1 12) := none + K : Option (Bounded.LE 0 11) := none + k : Option (Bounded.LE 1 24) := none + H : Option Hour.Ordinal := none + m : Option Minute.Ordinal := none + s : Option (Sigma Second.Ordinal) := none + S : Option Nanosecond.Ordinal := none + A : Option Millisecond.Offset := none + n : Option Nanosecond.Ordinal := none + N : Option Nanosecond.Offset := none + V : Option String := none + z : Option String := none + zabbrev : Option String := none + O : Option Offset := none + X : Option Offset := none + x : Option Offset := none + Z : Option Offset := none + +namespace DateBuilder + +private def insert (date : DateBuilder) (modifier : Modifier) (data : TypeFormat modifier) : DateBuilder := + match modifier with + | .G _ => { date with G := some data } + | .y _ => { date with y := some data } + | .D _ => { date with D := some data } + | .MorL _ => { date with MorL := some data } + | .d _ => { date with d := some data } + | .Qorq _ => { date with Qorq := some data } + | .w _ => { date with w := some data } + | .W _ => { date with W := some data } + | .E _ => { date with E := some data } + | .eorc _ => { date with eorc := some data } + | .F _ => { date with F := some data } + | .a _ => { date with a := some data } + | .B _ => { date with B := some data } + | .h _ => { date with h := some data } + | .K _ => { date with K := some data } + | .k _ => { date with k := some data } + | .H _ => { date with H := some data } + | .m _ => { date with m := some data } + | .s _ => { date with s := some data } + | .S _ => { date with S := some data } + | .A _ => { date with A := some data } + | .n _ => { date with n := some data } + | .N _ => { date with N := some data } + | .V => { date with V := some data } + | .z .full => { date with z := some data } + | .z .short => { date with zabbrev := some data } + | .O _ => { date with O := some data } + | .X _ => { date with X := some data } + | .x _ => { date with x := some data } + | .Z _ => { date with Z := some data } + +private def build (builder : DateBuilder) (aw : Awareness) : Option aw.type := + let tz : TimeZone := { + offset := builder.O <|> builder.X <|> builder.x <|> builder.Z |>.getD Offset.zero, + name := builder.V <|> builder.z |>.getD "Unknown", + abbreviation := builder.zabbrev |>.getD "Unknown", + isDST := false, + } + + let year := builder.y |>.getD 0 + let month := builder.MorL |>.getD 0 + let day := builder.d |>.getD 0 + + let era := builder.G |>.getD .ce + + let year := + match era with + | .ce => year + | .bce => -(year + 1) + + let hour : Option (Bounded.LE 0 23) := + if let some marker := builder.a then + marker.toAbsolute <$> builder.h + <|> marker.toAbsolute <$> ((Bounded.LE.add · 1) <$> builder.K) else - pure builder.hour + none + + let hour := + hour <|> ( + let one : Option (Bounded.LE 0 23) := builder.H + let other : Option (Bounded.LE 0 23) := (Bounded.LE.sub · 1) <$> builder.k + (one <|> other)) + |>.getD ⟨0, by decide⟩ + + let minute := builder.m |>.getD 0 + let second := builder.s |>.getD ⟨false, 0⟩ + let nano := (builder.n <|> builder.S) |>.getD 0 + + let time : PlainTime + := PlainTime.ofNanoseconds <$> builder.N + <|> PlainTime.ofMilliseconds <$> builder.A + |>.getD (PlainTime.mk hour minute second nano) + + let datetime : Option PlainDateTime := + if valid : year.Valid month day then + let date : PlainDate := { year, month, day, valid } + some { date, time } + else + none - if let .isTrue p := inferInstanceAs (Decidable (ValidTime hour.snd builder.minute builder.second.snd)) then - let build := DateTime.ofLocalDateTime { - date := PlainDate.clip builder.year builder.month builder.day - time := PlainTime.mk hour builder.minute builder.second builder.nanoseconds p - } + match aw with + | .only newTz => (DateTime.ofLocalDateTime · newTz) <$> datetime + | .any => (ZonedDateTime.ofLocalDateTime · tz) <$> datetime - match aw with - | .only tz => .ok (build tz) - | .any => - let tz₁ := TimeZone.mk builder.tz builder.tzName false - .ok ⟨tz₁, build tz₁⟩ - else - .error "invalid leap seconds {} {} {}" - -private def addDataInDateTime (data : DateBuilder) (typ : Modifier) (value : SingleFormatType typ) : Except String DateBuilder := - match typ with - | .AA | .aa => do pure { data with marker := some value } - | .YYYY | .YY => .ok { data with year := value } - | .MMMM | .MMM | .MM | .M => .ok { data with month := value } - | .DD | .D | .d => .ok { data with day := value } - | .EEEE | .EEE => .ok data - | .hh | .h | .HH | .H => .ok { data with hour := value } - | .mm | .m => .ok { data with minute := value } - | .sssssssss => .ok { data with nanoseconds := value } - | .sss => .ok { data with nanoseconds := Nanosecond.Ordinal.ofMillisecond value } - | .ss | .s => .ok { data with second := value } - | .ZZZZZ | .ZZZZ | .ZZZ - | .Z => .ok { data with tz := value } - | .z => .ok { data with tzName := value } - -private def formatParser (date : DateBuilder) : FormatPart → Parser DateBuilder - | .modifier mod => do - let res ← addDataInDateTime date mod <$> parserWithFormat mod - match res with - | .error err => fail err - | .ok res => pure res - | .string s => skipString s *> pure date +end DateBuilder -namespace Format +private def parseWithDate (date : DateBuilder) (mod : FormatPart) : Parser DateBuilder := do + match mod with + | .modifier s => do + let res ← parseWith s + return date.insert s res + | .string s => pstring s *> pure date /-- Constructs a new `Format` specification for a date-time string. Modifiers can be combined to create @@ -652,45 +1406,22 @@ def spec! (input : String) : Format tz := Formats the date using the format into a String. -/ def format (format : Format aw) (date : DateTime tz) : String := - let func := + let mapper (part : FormatPart) := match aw with - | .any => formatPartWithDate date - | .only tz => formatPartWithDate (date.convertTimeZone tz) + | .any => formatPartWithDate date part + | .only tz => formatPartWithDate (date.convertTimeZone tz) part - format.string.map func + format.string.map mapper |> String.join -/-- -Formats the date using the format into a String. --/ -def formatBuilder (format : Format aw) : FormatType String format.string := - let rec go (data : String) : (format : FormatString) → FormatType String format - | .modifier x :: xs => fun res => go (data ++ formatPart x res) xs - | .string x :: xs => go (data ++ x) xs - | [] => data - go "" format.string - -/-- -Formats the date using the format into a String. --/ -def formatGeneric (format : Format aw) (getInfo : (typ : Modifier) → Option (SingleFormatType typ)) : Option String := - let rec go (data : String) : (format : FormatString) → Option String - | .modifier x :: xs => do go (data ++ formatPart x (← getInfo x)) xs - | .string x :: xs => go (data ++ x) xs - | [] => data - go "" format.string - -/-- -Parser for a ZonedDateTime. --/ -def parser (format : FormatString) (aw : Awareness) : Parser (aw.type) := - let rec go (date : DateBuilder) (x : FormatString) : Parser aw.type := +private def parser (format : FormatString) (aw : Awareness) : Parser (aw.type) := + let rec go (builder : DateBuilder) (x : FormatString) : Parser aw.type := match x with - | x :: xs => formatParser date x >>= (go · xs) + | x :: xs => parseWithDate builder x >>= (go · xs) | [] => - match date.build aw with - | .ok res => pure res - | .error err => fail err + match builder.build aw with + | some res => pure res + | none => fail "could not parse the date" go {} format /-- @@ -700,12 +1431,12 @@ def builderParser (format: FormatString) (func: FormatType (Option α) format) : let rec go (format : FormatString) (func: FormatType (Option α) format) : Parser α := match format with | .modifier x :: xs => do - let res ← parserWithFormat x + let res ← parseWith x go xs (func res) | .string s :: xs => skipString s *> (go xs func) | [] => match func with - | some res => pure res + | some res => eof *> pure res | none => fail "invalid date." go format func @@ -729,6 +1460,7 @@ Parses and instead of using a builder to build a date, it uses a builder functio def parseBuilder (format : Format aw) (builder : FormatType (Option α) format.string) (input : String) : Except String α := (builderParser format.string builder).run input + /-- Parses and instead of using a builder to build a date, it uses a builder function instead. -/ @@ -737,6 +1469,23 @@ def parseBuilder! [Inhabited α] (format : Format aw) (builder : FormatType (Op | .ok res => res | .error err => panic! err -end Format -end Time -end Std + +/-- +Formats the date using the format into a String. +-/ +def formatGeneric (format : Format aw) (getInfo : (typ : Modifier) → Option (TypeFormat typ)) : Option String := + let rec go (data : String) : (format : FormatString) → Option String + | .modifier x :: xs => do go (data ++ formatWith x (← getInfo x)) xs + | .string x :: xs => go (data ++ x) xs + | [] => data + go "" format.string + +/-- +Formats the date using the format into a String. +-/ +def formatBuilder (format : Format aw) : FormatType String format.string := + let rec go (data : String) : (format : FormatString) → FormatType String format + | .modifier x :: xs => fun res => go (data ++ formatWith x res) xs + | .string x :: xs => go (data ++ x) xs + | [] => data + go "" format.string diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index 61f80855988b..8f4400308dde 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -170,7 +170,7 @@ def toNat' (n : Bounded.LE lo hi) (h : lo ≥ 0) : Nat := let h₁ := (Int.le_trans h n.property.left) match n.val, h₁ with | .ofNat n, _ => n - | .negSucc _, h => nomatch h + | .negSucc _, h => by contradiction /-- Convert a `Bounded.LE` to an Int. @@ -182,7 +182,7 @@ def toInt (n : Bounded.LE lo hi) : Int := /-- Convert a `Bounded.LE` to a `Fin`. -/ -@[inline] +@[inline, simp] def toFin (n : Bounded.LE lo hi) (h₀ : 0 ≤ lo) (h₁ : lo < hi) : Fin (hi + 1).toNat := by let h := n.property.right let h₁ := Int.le_trans h₀ n.property.left @@ -252,7 +252,7 @@ def truncate (bounded : Bounded.LE n m) : Bounded.LE 0 (m - n) := by Adjust the bounds of a `Bounded` by changing the higher bound if another value `j` satisfies the same constraint. -/ -@[inline] +@[inline, simp] def truncateTop (bounded : Bounded.LE n m) (h : bounded.val ≤ j) : Bounded.LE n j := by refine ⟨bounded.val, And.intro ?_ ?_⟩ · exact bounded.property.left @@ -280,13 +280,22 @@ def neg (bounded : Bounded.LE n m) : Bounded.LE (-m) (-n) := by /-- Adjust the bounds of a `Bounded` by adding a constant value to both the lower and upper bounds. -/ -@[inline] +@[inline, simp] def add (bounded : Bounded.LE n m) (num : Int) : Bounded.LE (n + num) (m + num) := by refine ⟨bounded.val + num, And.intro ?_ ?_⟩ all_goals apply (Int.add_le_add · (Int.le_refl num)) · exact bounded.property.left · exact bounded.property.right +/-- +Adjust the bounds of a `Bounded` by adding a constant value to both the lower and upper bounds. +-/ +@[inline] +def addProven (bounded : Bounded.LE n m) (h₀ : bounded.val + num ≤ m) (h₁ : num ≥ 0) : Bounded.LE n m := by + refine ⟨bounded.val + num, And.intro ?_ ?_⟩ + · exact Int.le_trans bounded.property.left (Int.le_add_of_nonneg_right h₁) + · exact h₀ + /-- Adjust the bounds of a `Bounded` by adding a constant value to the upper bounds. -/ @@ -321,7 +330,7 @@ def addBounds (bounded : Bounded.LE n m) (bounded₂ : Bounded.LE i j) : Bounded /-- Adjust the bounds of a `Bounded` by subtracting a constant value to both the lower and upper bounds. -/ -@[inline] +@[inline, simp] def sub (bounded : Bounded.LE n m) (num : Int) : Bounded.LE (n - num) (m - num) := add bounded (-num) @@ -407,6 +416,18 @@ def succ (bounded : Bounded.LE lo hi) (h : bounded.val < hi) : Bounded.LE lo hi let left := bounded.property.left ⟨bounded.val + 1, And.intro (by omega) (by omega)⟩ +/-- +Returns the absolute value of the bounded number `bo` with bounds `-(i - 1)` to `i - 1`. The result +will be a new bounded number with bounds `0` to `i - 1`. +-/ +def abs (bo : Bounded.LE (- (i - 1)) (i - 1)) : Bounded.LE 0 (i - 1) := + if h : bo.val ≥ 0 then + bo.truncateBottom h + else by + let r := bo.truncateTop (Int.le_of_lt (Int.not_le.mp h)) |>.neg + rw [Int.neg_neg] at r + exact r + end LE end Bounded end Internal diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 370ad5a7a289..2d19629f206e 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -63,6 +63,13 @@ Divides the `UnitVal` by an `Int`, resulting in a new `UnitVal` with an adjusted def ediv (unit : UnitVal a) (divisor : Int) : UnitVal (a * divisor) := ⟨unit.val.ediv divisor⟩ +/-- +Divides the `UnitVal` by an `Int`, resulting in a new `UnitVal` with an adjusted ratio. +-/ +@[inline] +def div (unit : UnitVal a) (divisor : Int) : UnitVal (a * divisor) := + ⟨unit.val.div divisor⟩ + /-- Adds two `UnitVal` values of the same ratio. -/ diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index cf36393ce394..820e9d2e2b00 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -8,260 +8,172 @@ import Std.Time.Date import Std.Time.Time import Std.Time.Zoned import Std.Time.DateTime -import Std.Time.Format.Basic +import Std.Time.Format import Lean.Parser namespace Std namespace Time open Lean Parser Command Std Lean.Parser -/-- -Category of units that are valid inside a date. --/ -declare_syntax_cat date_component +private def convertText : Text → MacroM (TSyntax `term) + | .short => `(Std.Time.Text.short) + | .full => `(Std.Time.Text.full) + | .narrow => `(Std.Time.Text.narrow) -/-- -Raw numbers --/ -syntax num : date_component +private def convertNumber : Number → MacroM (TSyntax `term) + | ⟨padding⟩ => `(Std.Time.Number.mk $(quote padding)) -/-- -Strings for dates like `"1970"-"1"-"1"` --/ -syntax str : date_component +private def convertFraction : Fraction → MacroM (TSyntax `term) + | .nano => `(Std.Time.Fraction.nano) + | .truncated digits => `(Std.Time.Fraction.truncated $(quote digits)) -/-- -Variables inside of the date. --/ -syntax ident : date_component - -private def parseComponent (lower : Option Int := none) (upper : Option Int := none) : TSyntax `date_component -> MacroM (TSyntax `term) - | `(date_component| $num:num) => do - let res := Int.ofNat num.getNat - if let some lower := lower then - if res < lower then - Macro.throwErrorAt num s!"the number should be bigger than {lower}" - - if let some upper := upper then - if res > upper then - Macro.throwErrorAt num s!"the number should be lower than {upper}" - - `($num) - | `(date_component| $str:str) => `($(Syntax.mkNumLit str.getString)) - | `(date_component| $ident:ident) => `($ident:ident) - | syn => Macro.throwErrorAt syn "unsupported syntax" - -/-- -Category of dates like `DD-MM-YYYY`. --/ -declare_syntax_cat date +private def convertYear : Year → MacroM (TSyntax `term) + | .twoDigit => `(Std.Time.Year.twoDigit) + | .fourDigit => `(Std.Time.Year.fourDigit) + | .extended n => `(Std.Time.Year.extended $(quote n)) -/-- -Date in `DD-MM-YYYY` format. --/ -syntax date_component noWs "-" noWs date_component noWs "-" noWs date_component : date +private def convertZoneName : ZoneName → MacroM (TSyntax `term) + | .short => `(Std.Time.ZoneName.short) + | .full => `(Std.Time.ZoneName.full) -private def parseDate : TSyntax `date -> MacroM (TSyntax `term) - | `(date|$year:date_component-$month:date_component-$day:date_component) => do - `(Std.Time.PlainDate.mk $(← parseComponent none none year) $(← parseComponent (some 1) (some 12) month) $(← parseComponent (some 1) (some 31) day) (by decide)) - | syn => Macro.throwErrorAt syn "unsupported type" +private def convertOffsetX : OffsetX → MacroM (TSyntax `term) + | .hour => `(Std.Time.OffsetX.hour) + | .hourMinute => `(Std.Time.OffsetX.hourMinute) + | .hourMinuteColon => `(Std.Time.OffsetX.hourMinuteColon) + | .hourMinuteSecond => `(Std.Time.OffsetX.hourMinuteSecond) + | .hourMinuteSecondColon => `(Std.Time.OffsetX.hourMinuteSecondColon) -/-- -Category of times like `HH:mm:ss`. --/ -declare_syntax_cat time - -/-- -Date in `HH-mm-ss` format. --/ -syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component : time - -/-- -Date in `HH-mm-ss,sssssssss` format. --/ -syntax date_component noWs ":" noWs date_component noWs ":" noWs date_component noWs "," noWs date_component : time +private def convertOffsetO : OffsetO → MacroM (TSyntax `term) + | .short => `(Std.Time.OffsetO.short) + | .full => `(Std.Time.OffsetO.full) -private def parseTime : TSyntax `time -> MacroM (TSyntax `term) - | `(time| $hour:date_component:$minute:date_component:$second:date_component,$nanos:date_component) => do - `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ $(← parseComponent (some 0) (some 999999999) nanos) (by decide)) - | `(time| $hour:date_component:$minute:date_component:$second:date_component) => do - `(Std.Time.PlainTime.mk ⟨true, $(← parseComponent (some 0) (some 24) hour)⟩ $(← parseComponent (some 0) (some 59) minute) ⟨true, $(← parseComponent (some 0) (some 60) second)⟩ 0 (by decide)) - | syn => Macro.throwErrorAt syn "unsupported syntax" +private def convertOffsetZ : OffsetZ → MacroM (TSyntax `term) + | .hourMinute => `(Std.Time.OffsetZ.hourMinute) + | .full => `(Std.Time.OffsetZ.full) + | .hourMinuteSecondColon => `(Std.Time.OffsetZ.hourMinuteSecondColon) -/-- -Category of date and time together. --/ -declare_syntax_cat datetime - -/-- -Date and time in `YYYY-MM-DDTHH:mm:ss,nano` format. --/ -syntax date &"T" time : datetime - -/-- -Date but using timestamp. --/ -syntax date_component : datetime - -private def parseDateTime : TSyntax `datetime -> MacroM (TSyntax `term) - | `(datetime| $date:date T $time:time) => do - `(Std.Time.PlainDateTime.mk $(← parseDate date) $(← parseTime time)) - | `(datetime|$tm:date_component) => do - `(Std.Time.PlainDateTime.ofUTCTimestamp $(← parseComponent none none tm)) - | syn => Macro.throwErrorAt syn "unsupported syntax" - -/-- -Category of offsets like `+00:00`. --/ -declare_syntax_cat offset +private def convertModifier : Modifier → MacroM (TSyntax `term) + | .G p => do `(Std.Time.Modifier.G $(← convertText p)) + | .y p => do `(Std.Time.Modifier.y $(← convertYear p)) + | .D p => do `(Std.Time.Modifier.D $(← convertNumber p)) + | .MorL p => + match p with + | .inl num => do `(Std.Time.Modifier.MorL (.inl $(← convertNumber num))) + | .inr txt => do `(Std.Time.Modifier.MorL (.inr $(← convertText txt))) + | .d p => do `(Std.Time.Modifier.d $(← convertNumber p)) + | .Qorq p => + match p with + | .inl num => do `(Std.Time.Modifier.Qorq (.inl $(← convertNumber num))) + | .inr txt => do `(Std.Time.Modifier.Qorq (.inr $(← convertText txt))) + | .w p => do `(Std.Time.Modifier.w $(← convertNumber p)) + | .W p => do `(Std.Time.Modifier.W $(← convertNumber p)) + | .E p => do `(Std.Time.Modifier.E $(← convertText p)) + | .eorc p => + match p with + | .inl num => do `(Std.Time.Modifier.eorc (.inl $(← convertNumber num))) + | .inr txt => do `(Std.Time.Modifier.eorc (.inr $(← convertText txt))) + | .F p => do `(Std.Time.Modifier.F $(← convertNumber p)) + | .a p => do `(Std.Time.Modifier.a $(← convertText p)) + | .B p => do `(Std.Time.Modifier.B $(← convertText p)) + | .h p => do `(Std.Time.Modifier.h $(← convertNumber p)) + | .K p => do `(Std.Time.Modifier.K $(← convertNumber p)) + | .k p => do `(Std.Time.Modifier.k $(← convertNumber p)) + | .H p => do `(Std.Time.Modifier.H $(← convertNumber p)) + | .m p => do `(Std.Time.Modifier.m $(← convertNumber p)) + | .s p => do `(Std.Time.Modifier.s $(← convertNumber p)) + | .S p => do `(Std.Time.Modifier.S $(← convertFraction p)) + | .A p => do `(Std.Time.Modifier.A $(← convertNumber p)) + | .n p => do `(Std.Time.Modifier.n $(← convertNumber p)) + | .N p => do `(Std.Time.Modifier.N $(← convertNumber p)) + | .V => `(Std.Time.Modifier.V) + | .z p => do `(Std.Time.Modifier.z $(← convertZoneName p)) + | .O p => do `(Std.Time.Modifier.O $(← convertOffsetO p)) + | .X p => do `(Std.Time.Modifier.X $(← convertOffsetX p)) + | .x p => do `(Std.Time.Modifier.x $(← convertOffsetX p)) + | .Z p => do `(Std.Time.Modifier.Z $(← convertOffsetZ p)) -/-- -Timezone in `+HH:mm`, or `-HH:mm` format. --/ -syntax ("+" <|> "-") date_component ":" date_component : offset +private def convertFormatPart : FormatPart → MacroM (TSyntax `term) + | .string s => `(.string $(Syntax.mkStrLit s)) + | .modifier mod => do `(.modifier $(← convertModifier mod)) -private def parseOffset : TSyntax `offset -> MacroM (TSyntax `term) - | `(offset| +$hour:date_component:$minutes:date_component) => do - `(Std.Time.TimeZone.Offset.ofHoursAndMinutes $(← parseComponent (some 0) (some 24) hour) $(← parseComponent (some 0) (some 59) minutes)) - | `(offset| -$hour:date_component:$minutes:date_component) => do - `(Std.Time.TimeZone.Offset.ofHoursAndMinutes (- $(← parseComponent (some 0) (some 24) hour)) (-$(← parseComponent (some 0) (some 59) minutes))) - | syn => Macro.throwErrorAt syn "unsupported syntax" +def syntaxNat (n : Nat) : MacroM (TSyntax `term) := do + let info ← MonadRef.mkInfoFromRefPos + pure { raw := Syntax.node1 info `num (Lean.Syntax.atom info (toString n)) } -/-- -Category of timezones like `Z`, `UTC`, `GMT` and `+HH:mm`. --/ -declare_syntax_cat zone +def syntaxString (n : String) : MacroM (TSyntax `term) := do + let info ← MonadRef.mkInfoFromRefPos + pure { raw := Syntax.node1 info `str (Lean.Syntax.atom info (toString n)) } -/-- -Timezone in `+HHmm`, or `-HHmm` format. --/ -syntax (offset <|> &"UTC" <|> &"GMT" <|> &"Z" <|> ("(" term ")")) : zone - -private def parseZone : TSyntax `zone -> MacroM (TSyntax `term) - | `(zone| ($term)) => do `($term) - | `(zone| UTC) => do `(Std.Time.TimeZone.UTC) - | `(zone| GMT) => do `(Std.Time.TimeZone.GMT) - | `(zone| Z) => do `(Std.Time.TimeZone.UTC) - | `(zone| $offset:offset) => do `(Std.Time.TimeZone.mk $(← parseOffset offset) "Unknown" false) - | syn => Macro.throwErrorAt syn "unsupported syntax" - -/-- -Datetimes with zone. --/ -declare_syntax_cat zoned +def syntaxInt (n : Int) : MacroM (TSyntax `term) := do + match n with + | .ofNat n => `(Int.ofNat $(Syntax.mkNumLit <| toString n)) + | .negSucc n => `(Int.negSucc $(Syntax.mkNumLit <| toString n)) -/-- -Datetime with zone in `YYYY-MM-DDTHH:mm:ssZ` format. --/ -syntax datetime zone : zoned +def syntaxBounded (n : Int) : MacroM (TSyntax `term) := do + `(Std.Time.Internal.Bounded.LE.ofNatWrapping $(← syntaxInt n) (by decide)) -private def parseZoned : TSyntax `zoned -> MacroM (TSyntax `term) - | `(zoned| $timestamp:num $zone) => do - `(Std.Time.DateTime.ofUTCTimestamp $timestamp $(← parseZone zone)) - | `(zoned| $datetime:datetime $zone) => do - `(Std.Time.DateTime.ofLocalDateTime $(← parseDateTime datetime) $(← parseZone zone)) - | syn => Macro.throwErrorAt syn "unsupported syntax" +def syntaxVal (n : Int) : MacroM (TSyntax `term) := do + `(Std.Time.Internal.UnitVal.ofInt $(← syntaxInt n)) -/-- -Syntax for zones. --/ -syntax "zone%" zone : term +def convertOffset (offset : Std.Time.TimeZone.Offset) : MacroM (TSyntax `term) := do + `(Std.Time.TimeZone.Offset.mk $(← syntaxVal offset.hour.val) $(← syntaxVal offset.second.val)) -/-- -Zoned date times. --/ -syntax "date%" zoned : term +def convertTimezone (tz : Std.Time.TimeZone) : MacroM (TSyntax `term) := do + `(Std.Time.TimeZone.mk $(← convertOffset tz.offset) $(Syntax.mkStrLit tz.name) $(Syntax.mkStrLit tz.abbreviation) false) -/-- -Normal datetimes. --/ -syntax "date%" datetime : term +def convertPlainDate (d : Std.Time.PlainDate) : MacroM (TSyntax `term) := do + `(Std.Time.PlainDate.clip $(← syntaxInt d.year) $(← syntaxBounded d.month.val) $(← syntaxBounded d.day.val)) -/-- -Macro for defining dates. --/ -syntax (name := dateDate) "date%" date : term +def convertPlainTime (d : Std.Time.PlainTime) : MacroM (TSyntax `term) := do + `(Std.Time.PlainTime.mk $(← syntaxBounded d.hour.val) $(← syntaxBounded d.minute.val) ⟨true, $(← syntaxBounded d.second.snd.val)⟩ $(← syntaxBounded d.nano.val)) -/-- -Macro for defining times. --/ -syntax "time%" time : term +def convertPlainDateTime (d : Std.Time.PlainDateTime) : MacroM (TSyntax `term) := do + `(Std.Time.PlainDateTime.mk $(← convertPlainDate d.date) $(← convertPlainTime d.time)) -/-- -Macro for creating offsets --/ -syntax "offset%" offset : term +def convertZonedDateTime (d : Std.Time.ZonedDateTime) : MacroM (TSyntax `term) := do + `(Std.Time.ZonedDateTime.mk $(← convertTimezone d.timezone) (DateTime.ofLocalDateTime $(← convertPlainDateTime d.snd.date.get) $(← convertTimezone d.timezone))) -/-- -Macro for creating timezones --/ -syntax "timezone%" ident noWs "/" noWs ident offset : term +syntax "zoned(" str ")" : term -/-- -Macro for creating timezones with string --/ -syntax "timezone%" str offset : term +syntax "datetime(" str ")" : term -macro_rules - | `(date% $date:date) => parseDate date - | `(date% $datetime:datetime) => parseDateTime datetime - | `(date% $zoned:zoned) => parseZoned zoned - | `(time% $time:time) => parseTime time - | `(offset% $offset:offset) => parseOffset offset - | `(timezone% $str $offset:offset) => do - do `(Std.Time.TimeZone.mk $(← parseOffset offset) $str false) - | `(timezone% $name/$sub $offset:offset) => do - let name := s!"{name.getId.toString}/{sub.getId.toString}" - let syn := Syntax.mkStrLit name - do `(Std.Time.TimeZone.mk $(← parseOffset offset) $syn false) +syntax "date(" str ")" : term -private def convertModifier : Modifier → MacroM (TSyntax `term) - | .YYYY => `(Std.Time.Modifier.YYYY) - | .YY => `(Std.Time.Modifier.YY) - | .MMMM => `(Std.Time.Modifier.MMMM) - | .MMM => `(Std.Time.Modifier.MMM) - | .MM => `(Std.Time.Modifier.MM) - | .M => `(Std.Time.Modifier.M) - | .DD => `(Std.Time.Modifier.DD) - | .D => `(Std.Time.Modifier.D) - | .d => `(Std.Time.Modifier.d) - | .EEEE => `(Std.Time.Modifier.EEEE) - | .EEE => `(Std.Time.Modifier.EEE) - | .hh => `(Std.Time.Modifier.hh) - | .h => `(Std.Time.Modifier.h) - | .HH => `(Std.Time.Modifier.HH) - | .H => `(Std.Time.Modifier.H) - | .AA => `(Std.Time.Modifier.AA) - | .aa => `(Std.Time.Modifier.aa) - | .mm => `(Std.Time.Modifier.mm) - | .m => `(Std.Time.Modifier.m) - | .sssssssss => `(Std.Time.Modifier.sssssssss) - | .sss => `(Std.Time.Modifier.sss) - | .ss => `(Std.Time.Modifier.ss) - | .s => `(Std.Time.Modifier.s) - | .ZZZZZ => `(Std.Time.Modifier.ZZZZZ) - | .ZZZZ => `(Std.Time.Modifier.ZZZZ) - | .ZZZ => `(Std.Time.Modifier.ZZZ) - | .Z => `(Std.Time.Modifier.Z) - | .z => `(Std.Time.Modifier.z) +syntax "time(" str ")" : term -private def convertFormatPart : FormatPart → MacroM (TSyntax `term) - | .string s => `(.string $(Syntax.mkStrLit s)) - | .modifier mod => do `(.modifier $(← convertModifier mod)) +syntax "offset(" str ")" : term -/-- -Syntax for defining a date spec at compile time. --/ -syntax "date-spec%" str : term +syntax "timezone(" str ")" : term macro_rules - | `(date-spec% $format_string) => do - let input := format_string.getString - let format : Except String (Format .any) := Format.spec input - match format with - | .ok res => - let alts ← res.string.mapM convertFormatPart - let alts := alts.foldl Syntax.TSepArray.push (Syntax.TSepArray.mk #[] (sep := ",")) - `(⟨[$alts,*]⟩) - | .error err => - Macro.throwErrorAt format_string s!"cannot compile spec: {err}" + | `(zoned( $date:str )) => do + match ZonedDateTime.fromLeanDateTimeWithZoneString date.getString with + | .ok res => do + return ← convertZonedDateTime res + | .error res => Macro.throwErrorAt date s!"error: {res}" + + | `(datetime( $date:str )) => do + match PlainDateTime.fromLeanDateTimeString date.getString with + | .ok res => do + return ← convertPlainDateTime res + | .error res => Macro.throwErrorAt date s!"error: {res}" + + | `(date( $date:str )) => do + match PlainDate.fromSQLDateString date.getString with + | .ok res => return ← convertPlainDate res + | .error res => Macro.throwErrorAt date s!"error: {res}" + + | `(time( $time:str )) => do + match PlainTime.fromLeanTime24Hour time.getString with + | .ok res => return ← convertPlainTime res + | .error res => Macro.throwErrorAt time s!"error: {res}" + + | `(offset( $offset:str )) => do + match TimeZone.Offset.fromOffset offset.getString with + | .ok res => return ← convertOffset res + | .error res => Macro.throwErrorAt offset s!"error: {res}" + + | `(timezone( $tz:str )) => do + match TimeZone.fromTimeZone tz.getString with + | .ok res => return ← convertTimezone res + | .error res => Macro.throwErrorAt tz s!"error: {res}" diff --git a/src/Std/Time/Notation/Spec.lean b/src/Std/Time/Notation/Spec.lean new file mode 100644 index 000000000000..f6aefc82ff82 --- /dev/null +++ b/src/Std/Time/Notation/Spec.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.Date +import Std.Time.Time +import Std.Time.Zoned +import Std.Time.DateTime +import Std.Time.Format.Basic +import Lean.Parser + +namespace Std +namespace Time +open Lean Parser Command Std Lean.Parser + +private def convertText : Text → MacroM (TSyntax `term) + | .short => `(Std.Time.Text.short) + | .full => `(Std.Time.Text.full) + | .narrow => `(Std.Time.Text.narrow) + +private def convertNumber : Number → MacroM (TSyntax `term) + | ⟨padding⟩ => `(Std.Time.Number.mk $(quote padding)) + +private def convertFraction : Fraction → MacroM (TSyntax `term) + | .nano => `(Std.Time.Fraction.nano) + | .truncated digits => `(Std.Time.Fraction.truncated $(quote digits)) + +private def convertYear : Year → MacroM (TSyntax `term) + | .twoDigit => `(Std.Time.Year.twoDigit) + | .fourDigit => `(Std.Time.Year.fourDigit) + | .extended n => `(Std.Time.Year.extended $(quote n)) + +private def convertZoneName : ZoneName → MacroM (TSyntax `term) + | .short => `(Std.Time.ZoneName.short) + | .full => `(Std.Time.ZoneName.full) + +private def convertOffsetX : OffsetX → MacroM (TSyntax `term) + | .hour => `(Std.Time.OffsetX.hour) + | .hourMinute => `(Std.Time.OffsetX.hourMinute) + | .hourMinuteColon => `(Std.Time.OffsetX.hourMinuteColon) + | .hourMinuteSecond => `(Std.Time.OffsetX.hourMinuteSecond) + | .hourMinuteSecondColon => `(Std.Time.OffsetX.hourMinuteSecondColon) + +private def convertOffsetO : OffsetO → MacroM (TSyntax `term) + | .short => `(Std.Time.OffsetO.short) + | .full => `(Std.Time.OffsetO.full) + +private def convertOffsetZ : OffsetZ → MacroM (TSyntax `term) + | .hourMinute => `(Std.Time.OffsetZ.hourMinute) + | .full => `(Std.Time.OffsetZ.full) + | .hourMinuteSecondColon => `(Std.Time.OffsetZ.hourMinuteSecondColon) + +private def convertModifier : Modifier → MacroM (TSyntax `term) + | .G p => do `(Std.Time.Modifier.G $(← convertText p)) + | .y p => do `(Std.Time.Modifier.y $(← convertYear p)) + | .D p => do `(Std.Time.Modifier.D $(← convertNumber p)) + | .MorL p => + match p with + | .inl num => do `(Std.Time.Modifier.MorL (.inl $(← convertNumber num))) + | .inr txt => do `(Std.Time.Modifier.MorL (.inr $(← convertText txt))) + | .d p => do `(Std.Time.Modifier.d $(← convertNumber p)) + | .Qorq p => + match p with + | .inl num => do `(Std.Time.Modifier.Qorq (.inl $(← convertNumber num))) + | .inr txt => do `(Std.Time.Modifier.Qorq (.inr $(← convertText txt))) + | .w p => do `(Std.Time.Modifier.w $(← convertNumber p)) + | .W p => do `(Std.Time.Modifier.W $(← convertNumber p)) + | .E p => do `(Std.Time.Modifier.E $(← convertText p)) + | .eorc p => + match p with + | .inl num => do `(Std.Time.Modifier.eorc (.inl $(← convertNumber num))) + | .inr txt => do `(Std.Time.Modifier.eorc (.inr $(← convertText txt))) + | .F p => do `(Std.Time.Modifier.F $(← convertNumber p)) + | .a p => do `(Std.Time.Modifier.a $(← convertText p)) + | .B p => do `(Std.Time.Modifier.B $(← convertText p)) + | .h p => do `(Std.Time.Modifier.h $(← convertNumber p)) + | .K p => do `(Std.Time.Modifier.K $(← convertNumber p)) + | .k p => do `(Std.Time.Modifier.k $(← convertNumber p)) + | .H p => do `(Std.Time.Modifier.H $(← convertNumber p)) + | .m p => do `(Std.Time.Modifier.m $(← convertNumber p)) + | .s p => do `(Std.Time.Modifier.s $(← convertNumber p)) + | .S p => do `(Std.Time.Modifier.S $(← convertFraction p)) + | .A p => do `(Std.Time.Modifier.A $(← convertNumber p)) + | .n p => do `(Std.Time.Modifier.n $(← convertNumber p)) + | .N p => do `(Std.Time.Modifier.N $(← convertNumber p)) + | .V => `(Std.Time.Modifier.V) + | .z p => do `(Std.Time.Modifier.z $(← convertZoneName p)) + | .O p => do `(Std.Time.Modifier.O $(← convertOffsetO p)) + | .X p => do `(Std.Time.Modifier.X $(← convertOffsetX p)) + | .x p => do `(Std.Time.Modifier.x $(← convertOffsetX p)) + | .Z p => do `(Std.Time.Modifier.Z $(← convertOffsetZ p)) + +private def convertFormatPart : FormatPart → MacroM (TSyntax `term) + | .string s => `(.string $(Syntax.mkStrLit s)) + | .modifier mod => do `(.modifier $(← convertModifier mod)) + +/-- +Syntax for defining a date spec at compile time. +-/ +syntax "datespec(" str ")" : term + +macro_rules + | `(datespec( $format_string:str )) => do + let input := format_string.getString + let format : Except String (Format .any) := Format.spec input + match format with + | .ok res => + let alts ← res.string.mapM convertFormatPart + let alts := alts.foldl Syntax.TSepArray.push (Syntax.TSepArray.mk #[] (sep := ",")) + `(⟨[$alts,*]⟩) + | .error err => + Macro.throwErrorAt format_string s!"cannot compile spec: {err}" diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 449ae283a5ca..9529d0bf9444 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -23,16 +23,25 @@ inductive HourMarker | pm deriving Repr, BEq +namespace HourMarker + +/-- +`ofOrdinal` converts an `Hour.Ordinal` value to an `HourMarker`, indicating whether it is AM or PM. +-/ +def ofOrdinal (time : Hour.Ordinal) : HourMarker := + if time.val ≥ 12 then + .pm + else + .am + /-- Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker`. -/ -def HourMarker.toAbsolute (marker : HourMarker) (time : Bounded.LE 0 12) : Hour.Ordinal l := +def toAbsolute (marker : HourMarker) (time : Bounded.LE 1 12) : Hour.Ordinal := match marker with - | .am => by - refine time.expandTop ?_ - split <;> decide - | .pm => by - have time := time.add 12 |>.emod 24 (by decide) - cases l - · exact time - · exact time.expandTop (by decide) + | .am => time.sub 1 |>.expandTop (by decide) + | .pm => time.add 11 |>.emod 24 (by decide) + +end HourMarker +end Time +end Std diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index b8663cdc712b..388beb495ebb 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -13,14 +13,14 @@ open Internal set_option linter.all true /-- -Represents a specific point in time, including hours, minutes, seconds, and nanoseconds. +Represents a specific point in a day, including hours, minutes, seconds, and nanoseconds. -/ structure PlainTime where /-- `Hour` component of the `PlainTime` -/ - hour : Sigma Hour.Ordinal + hour : Hour.Ordinal /-- `Minute` component of the `PlainTime` @@ -37,87 +37,52 @@ structure PlainTime where -/ nano : Nanosecond.Ordinal - /-- - The prove that if it includes a leap second than it needs to be exactly 23 : 59 : 60 - -/ - proof : (second.snd.val = 60 → hour.snd.val = 23 ∧ minute.val = 59) - ∧ (hour.snd.val = 24 → second.snd.val = 0 ∧ minute.val = 0) - instance : Inhabited PlainTime where - default := ⟨Sigma.mk false 0, 0, Sigma.mk false 0, 0, by simp; decide⟩ + default := ⟨0, 0, Sigma.mk false 0, 0, by decide⟩ instance : BEq PlainTime where - beq x y := x.hour.snd.val == y.hour.snd.val && x.minute == y.minute + beq x y := x.hour.val == y.hour.val && x.minute == y.minute && x.second.snd.val == y.second.snd.val && x.nano == y.nano namespace PlainTime -/-- -Checks if the hour is valid if it has a leap second or leap hour. --/ -@[simp] -abbrev ValidTime (hour : Hour.Ordinal l) (minute : Minute.Ordinal) (second : Second.Ordinal l₁) : Prop := - (second.val = 60 → hour.val = 23 ∧ minute.val = 59) - ∧ (hour.val = 24 → second.val = 0 ∧ minute.val = 0) - -/-- -Creates a `PlainTime` value from hours, minutes, and seconds. --/ -def ofHourMinuteSecondsNano (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (nano : Nanosecond.Ordinal) (proof : ValidTime hour minute second) : PlainTime := - ⟨Sigma.mk leap₂ hour, minute, Sigma.mk leap second, nano, proof⟩ - -/-- -Creates a `PlainTime` value from hours, minutes, and seconds. Return `none` if its invalid. --/ -def ofHourMinuteSeconds? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : Option PlainTime := - if h : ValidTime hour minute second - then some <| ofHourMinuteSecondsNano hour minute second 0 h - else none - /-- Creates a `PlainTime` value from hours, minutes, seconds and nanoseconds. -/ -def ofValidHourMinuteSecondsNano (hour : Hour.Ordinal false) (minute : Minute.Ordinal) (second : Second.Ordinal false) (nano : Nanosecond.Ordinal) : PlainTime := by - refine ⟨Sigma.mk false hour, minute, Sigma.mk false second, nano, ?_⟩ - constructor - exact fun x => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr second.property.right)) x) - exact fun x => nomatch (Int.ne_iff_lt_or_gt.mpr (Or.inl (Int.lt_add_one_iff.mpr hour.property.right)) x) +@[inline] +def ofHourMinuteSecondsNano (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (nano : Nanosecond.Ordinal) : PlainTime := + ⟨hour, minute, Sigma.mk leap second, nano⟩ /-- -Creates a `PlainTime` value from hours, minutes, and seconds. Return `none` if its invalid. +Creates a `PlainTime` value from hours, minutes, and seconds. -/ -def ofHourMinuteSecondsNano? (hour : Hour.Ordinal leap₂) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (nano : Nanosecond.Ordinal) : Option PlainTime := - if h : ValidTime hour minute second - then some <| ofHourMinuteSecondsNano hour minute second nano h - else none +@[inline] +def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : PlainTime := + ofHourMinuteSecondsNano hour minute second 0 /-- Converts a `PlainTime` value to the total number of milliseconds. -/ def toMilliseconds (time : PlainTime) : Millisecond.Offset := - let secs := - time.hour.snd.toOffset.toSeconds + - time.minute.toOffset.toSeconds + - time.second.snd.toOffset - let millis := secs.mul 1000 - Millisecond.Offset.ofInt (millis.val + time.nano.val) + time.hour.toOffset.toMilliseconds + + time.minute.toOffset.toMilliseconds + + time.second.snd.toOffset.toMilliseconds + + time.nano.toOffset.toMilliseconds /-- Converts a `PlainTime` value to the total number of nanoseconds. -/ def toNanoseconds (time : PlainTime) : Nanosecond.Offset := - let secs := - time.hour.snd.toOffset.toSeconds + - time.minute.toOffset.toSeconds + - time.second.snd.toOffset - let nanos := secs.mul 1000000000 - Nanosecond.Offset.ofInt (nanos.val + time.nano.val) + time.hour.toOffset.toNanoseconds + + time.minute.toOffset.toNanoseconds + + time.second.snd.toOffset.toNanoseconds + + time.nano.toOffset /-- Converts a `PlainTime` value to the total number of seconds. -/ def toSeconds (time : PlainTime) : Second.Offset := - time.hour.snd.toOffset.toSeconds + + time.hour.toOffset.toSeconds + time.minute.toOffset.toSeconds + time.second.snd.toOffset @@ -125,7 +90,7 @@ def toSeconds (time : PlainTime) : Second.Offset := Converts a `PlainTime` value to the total number of minutes. -/ def toMinutes (time : PlainTime) : Minute.Offset := - time.hour.snd.toOffset.toMinutes + + time.hour.toOffset.toMinutes + time.minute.toOffset + time.second.snd.toOffset.toMinutes @@ -133,8 +98,7 @@ def toMinutes (time : PlainTime) : Minute.Offset := Converts a `PlainTime` value to the total number of hours. -/ def toHours (time : PlainTime) : Hour.Offset := - let hour : Hour.Offset := time.minute.toOffset.ediv 60 - time.hour.snd.toOffset + hour + time.second.snd.toOffset.toHours + time.hour.toOffset /-- Creates a `PlainTime` value from a total number of nanoseconds. @@ -146,17 +110,21 @@ def ofNanoseconds (nanos : Nanosecond.Offset) : PlainTime := have minutes := (Bounded.LE.byEmod totalSeconds.val 3600 (by decide)).ediv 60 (by decide) have seconds := Bounded.LE.byEmod totalSeconds.val 60 (by decide) let nanos := Bounded.LE.byEmod nanos.val 1000000000 (by decide) - ofValidHourMinuteSecondsNano hours minutes seconds nanos + PlainTime.mk hours minutes (Sigma.mk false seconds) nanos + +/-- +Creates a `PlainTime` value from a total number of millisecond. +-/ +@[inline] +def ofMilliseconds (millis : Millisecond.Offset) : PlainTime := + ofNanoseconds millis.toNanoseconds /-- Creates a `PlainTime` value from a total number of seconds. -/ @[inline] def ofSeconds (secs : Second.Offset) : PlainTime := - have hours := Bounded.LE.byEmod (secs.val / 3600) 24 (by decide) - have minutes := (Bounded.LE.byEmod secs.val 3600 (by decide)).ediv 60 (by decide) - have seconds := Bounded.LE.byEmod secs.val 60 (by decide) - ofValidHourMinuteSecondsNano hours minutes seconds 0 + ofNanoseconds secs.toNanoseconds /-- Adds seconds to a `PlainTime`. @@ -191,6 +159,7 @@ def subMinutes (time : PlainTime) (minutesToSub : Minute.Offset) : PlainTime := /-- Adds hours to a `PlainTime`. -/ +@[inline] def addHours (time : PlainTime) (hoursToAdd : Hour.Offset) : PlainTime := let total := time.toSeconds + hoursToAdd.toSeconds ofSeconds total @@ -215,12 +184,31 @@ Subtracts nanoseconds from a `PlainTime`. def subNanoseconds (time : PlainTime) (nanosToSub : Nanosecond.Offset) : PlainTime := addNanoseconds time (-nanosToSub) +/-- +Adds milliseconds to a `PlainTime`. +-/ +def addMilliseconds (time : PlainTime) (millisToAdd : Millisecond.Offset) : PlainTime := + let total := time.toMilliseconds + millisToAdd + ofMilliseconds total + +/-- +Subtracts milliseconds from a `PlainTime`. +-/ +def subMilliseconds (time : PlainTime) (millisToSub : Millisecond.Offset) : PlainTime := + addMilliseconds time (-millisToSub) + instance : HAdd PlainTime Nanosecond.Offset PlainTime where hAdd := addNanoseconds instance : HSub PlainTime Nanosecond.Offset PlainTime where hSub := subNanoseconds +instance : HAdd PlainTime Millisecond.Offset PlainTime where + hAdd := addMilliseconds + +instance : HSub PlainTime Millisecond.Offset PlainTime where + hSub := subMilliseconds + instance : HAdd PlainTime Second.Offset PlainTime where hAdd := addSeconds diff --git a/src/Std/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean index e699a3670a96..ab5f9d282d0c 100644 --- a/src/Std/Time/Time/Unit/Basic.lean +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -32,7 +32,7 @@ Converts a `Nanosecond.Offset` to a `Millisecond.Offset`. -/ @[inline] def toMilliseconds (offset : Nanosecond.Offset) : Millisecond.Offset := - offset.ediv 1000000 + offset.div 1000000 /-- Converts a `Millisecond.Offset` to a `Nanosecond.Offset`. @@ -46,7 +46,7 @@ Converts a `Nanosecond.Offset` to a `Second.Offset`. -/ @[inline] def toSeconds (offset : Nanosecond.Offset) : Second.Offset := - offset.ediv 1000000000 + offset.div 1000000000 /-- Converts a `Second.Offset` to a `Nanosecond.Offset`. @@ -60,7 +60,7 @@ Converts a `Nanosecond.Offset` to a `Minute.Offset`. -/ @[inline] def toMinutes (offset : Nanosecond.Offset) : Minute.Offset := - offset.ediv 60000000000 + offset.div 60000000000 /-- Converts a `Minute.Offset` to a `Nanosecond.Offset`. @@ -74,7 +74,7 @@ Converts a `Nanosecond.Offset` to an `Hour.Offset`. -/ @[inline] def toHours (offset : Nanosecond.Offset) : Hour.Offset := - offset.ediv 3600000000000 + offset.div 3600000000000 /-- Converts an `Hour.Offset` to a `Nanosecond.Offset`. @@ -99,14 +99,14 @@ Converts a `Nanosecond.Offset` to a `Millisecond.Offset`. -/ @[inline] def ofNanoseconds (offset : Nanosecond.Offset) : Millisecond.Offset := - offset.ediv 1000000 + offset.div 1000000 /-- Converts a `Millisecond.Offset` to a `Second.Offset`. -/ @[inline] def toSeconds (offset : Millisecond.Offset) : Second.Offset := - offset.ediv 1000 + offset.div 1000 /-- Converts a `Second.Offset` to a `Millisecond.Offset`. @@ -120,7 +120,7 @@ Converts a `Millisecond.Offset` to a `Minute.Offset`. -/ @[inline] def toMinutes (offset : Millisecond.Offset) : Minute.Offset := - offset.ediv 60000 + offset.div 60000 /-- Converts a `Minute.Offset` to a `Millisecond.Offset`. @@ -134,7 +134,7 @@ Converts a `Millisecond.Offset` to an `Hour.Offset`. -/ @[inline] def toHours (offset : Millisecond.Offset) : Hour.Offset := - offset.ediv 3600000 + offset.div 3600000 /-- Converts an `Hour.Offset` to a `Millisecond.Offset`. @@ -159,7 +159,7 @@ Converts a `Nanosecond.Offset` to a `Second.Offset`. -/ @[inline] def ofNanoseconds (offset : Nanosecond.Offset) : Second.Offset := - offset.ediv 1000000000 + offset.div 1000000000 /-- Converts a `Second.Offset` to a `Millisecond.Offset`. @@ -173,14 +173,14 @@ Converts a `Millisecond.Offset` to a `Second.Offset`. -/ @[inline] def ofMilliseconds (offset : Millisecond.Offset) : Second.Offset := - offset.ediv 1000 + offset.div 1000 /-- Converts a `Second.Offset` to a `Minute.Offset`. -/ @[inline] def toMinutes (offset : Second.Offset) : Minute.Offset := - offset.ediv 60 + offset.div 60 /-- Converts a `Minute.Offset` to a `Second.Offset`. @@ -194,7 +194,7 @@ Converts a `Second.Offset` to an `Hour.Offset`. -/ @[inline] def toHours (offset : Second.Offset) : Hour.Offset := - offset.ediv 3600 + offset.div 3600 /-- Converts an `Hour.Offset` to a `Second.Offset`. @@ -219,7 +219,7 @@ Converts a `Nanosecond.Offset` to a `Minute.Offset`. -/ @[inline] def ofNanoseconds (offset : Nanosecond.Offset) : Minute.Offset := - offset.ediv 60000000000 + offset.div 60000000000 /-- Converts a `Minute.Offset` to a `Millisecond.Offset`. @@ -233,7 +233,7 @@ Converts a `Millisecond.Offset` to a `Minute.Offset`. -/ @[inline] def ofMilliseconds (offset : Millisecond.Offset) : Minute.Offset := - offset.ediv 60000 + offset.div 60000 /-- Converts a `Minute.Offset` to a `Second.Offset`. @@ -247,14 +247,14 @@ Converts a `Second.Offset` to a `Minute.Offset`. -/ @[inline] def ofSeconds (offset : Second.Offset) : Minute.Offset := - offset.ediv 60 + offset.div 60 /-- Converts a `Minute.Offset` to an `Hour.Offset`. -/ @[inline] def toHours (offset : Minute.Offset) : Hour.Offset := - offset.ediv 60 + offset.div 60 /-- Converts an `Hour.Offset` to a `Minute.Offset`. @@ -279,7 +279,7 @@ Converts a `Nanosecond.Offset` to an `Hour.Offset`. -/ @[inline] def ofNanoseconds (offset : Nanosecond.Offset) : Hour.Offset := - offset.ediv 3600000000000 + offset.div 3600000000000 /-- Converts an `Hour.Offset` to a `Millisecond.Offset`. @@ -293,7 +293,7 @@ Converts a `Millisecond.Offset` to an `Hour.Offset`. -/ @[inline] def ofMilliseconds (offset : Millisecond.Offset) : Hour.Offset := - offset.ediv 3600000 + offset.div 3600000 /-- Converts an `Hour.Offset` to a `Second.Offset`. @@ -307,7 +307,7 @@ Converts a `Second.Offset` to an `Hour.Offset`. -/ @[inline] def ofSeconds (offset : Second.Offset) : Hour.Offset := - offset.ediv 3600 + offset.div 3600 /-- Converts an `Hour.Offset` to a `Minute.Offset`. @@ -321,7 +321,7 @@ Converts a `Minute.Offset` to an `Hour.Offset`. -/ @[inline] def ofMinutes (offset : Minute.Offset) : Hour.Offset := - offset.ediv 60 + offset.div 60 end Hour.Offset diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 4454e4e663ae..a09c5c4d94cf 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -18,22 +18,25 @@ open Internal set_option linter.all true /-- -`Ordinal` represents a bounded value for hours, ranging from 0 to 24 or 23. The upper bound is 24 to -account for valid timestamps like 24:00:00 with leap seconds. +`Ordinal` represents a bounded value for hours, ranging from 0 to 23. -/ -def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 24 else 23)) +def Ordinal := Bounded.LE 0 23 + deriving Repr, BEq, LE, LT -instance : ToString (Ordinal leap) where +instance : ToString Ordinal where toString x := toString x.val -instance : Repr (Ordinal l) where - reprPrec r l := reprPrec r.val l +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 0 (0 + (23 : Nat))) n) -instance : OfNat (Ordinal leap) n := by - have inst := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (23 : Nat))) n) - cases leap - · exact inst - · exact ⟨inst.ofNat.expandTop (by decide)⟩ +instance : Inhabited Ordinal where + default := 0 + +instance {x y : Ordinal} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance {x y : Ordinal} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) /-- `Offset` represents an offset in hours, defined as an `Int`. This can be used to express durations @@ -47,28 +50,31 @@ instance : OfNat Offset n := namespace Ordinal +/-- +Converts an `Ordinal` into a relative month in the range of 1 to 12. +-/ +def toRelative (ordinal : Ordinal) : Bounded.LE 1 12 := + (ordinal.add 11).emod 12 (by decide) |>.add 1 + /-- Creates an `Ordinal` from a natural number, ensuring the value is within the valid bounds for hours. -/ @[inline] -def ofNat (data : Nat) (h : data ≤ (if leap then 24 else 23)) : Ordinal leap := +def ofNat (data : Nat) (h : data ≤ 23) : Ordinal := Bounded.LE.ofNat data h /-- -Creates an `Ordinal` from a `Fin` value, ensuring the value is within bounds depending on whether -leap seconds are considered. +Creates an `Ordinal` from a `Fin` value. -/ @[inline] -def ofFin (data : Fin (if leap then 25 else 24)) : Ordinal leap := - match leap with - | true => Bounded.LE.ofFin data - | false => Bounded.LE.ofFin data +def ofFin (data : Fin 24) : Ordinal := + Bounded.LE.ofFin data /-- Converts an `Ordinal` to an `Offset`, which represents the duration in hours as an integer value. -/ @[inline] -def toOffset (ordinal : Ordinal leap) : Offset := +def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val end Ordinal diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 389473cc8e19..4ffe993743c5 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -6,6 +6,7 @@ Authors: Sofia Rodrigues prelude import Std.Internal.Rat import Std.Time.Internal +import Std.Time.Time.Unit.Nanosecond namespace Std namespace Time @@ -30,6 +31,12 @@ instance : OfNat Ordinal n := instance : Inhabited Ordinal where default := 0 +instance {x y : Ordinal} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance {x y : Ordinal} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + /-- `Offset` represents a duration offset in milliseconds. It is defined as an `Int` value, where each unit corresponds to one millisecond. @@ -57,6 +64,37 @@ def ofInt (data : Int) : Offset := UnitVal.mk data end Offset +namespace Ordinal + +/-- +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 0 ≤ data ∧ data ≤ 999) : Ordinal := + Bounded.LE.mk data h + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≤ 999) : Ordinal := + Bounded.LE.ofNat data h + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +-/ +@[inline] +def ofFin (data : Fin 1000) : Ordinal := + Bounded.LE.ofFin data + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +@[inline] +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal end Millisecond end Time end Std diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index 346e5eff5988..59c8d31be401 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -20,7 +20,7 @@ set_option linter.all true `Ordinal` represents a bounded value for minutes, ranging from 0 to 59. This is useful for representing the minute component of a time. -/ def Ordinal := Bounded.LE 0 59 - deriving Repr, BEq, LE + deriving Repr, BEq, LE, LT instance : ToString Ordinal where toString x := toString x.val @@ -31,6 +31,12 @@ instance : OfNat Ordinal n := instance : Inhabited Ordinal where default := 0 +instance {x y : Ordinal} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance {x y : Ordinal} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + /-- `Offset` represents a duration offset in minutes. It is defined as an `Int` value. -/ @@ -43,22 +49,28 @@ instance : OfNat Offset n := namespace Ordinal /-- -Creates an `Ordinal` from a natural number, ensuring the value is within the valid range (0 to 59). +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 0 ≤ data ∧ data ≤ 59) : Ordinal := + Bounded.LE.mk data h + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. -/ @[inline] -def ofNat (data : Nat) (h : data ≤ 59 := by decide) : Ordinal := +def ofNat (data : Nat) (h : data ≤ 59) : Ordinal := Bounded.LE.ofNat data h /-- -Creates an `Ordinal` from a `Fin` value, ensuring it is within the valid range (0 to 59). +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. -/ @[inline] def ofFin (data : Fin 60) : Ordinal := Bounded.LE.ofFin data /-- -Converts an `Ordinal` to an `Offset`. The `Ordinal` value is converted to an `Offset` by interpreting -it as the number of minutes. +Converts an `Ordinal` to an `Offset`. -/ @[inline] def toOffset (ordinal : Ordinal) : Offset := diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index 2cfe052c313d..21fb6ceaea0e 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -5,7 +5,7 @@ Authors: Sofia Rodrigues -/ prelude import Std.Internal.Rat -import Std.Time.Time.Unit.Millisecond +import Std.Time.Internal namespace Std namespace Time @@ -30,28 +30,11 @@ instance : OfNat Ordinal n where instance : Inhabited Ordinal where default := 0 -namespace Ordinal - -/-- -`Ordinal` represents a bounded value for nanoseconds in a day, which ranges between 0 and 86400000000000. --/ -def OfDay := Bounded.LE 0 86400000000000 - deriving Repr, BEq, LE, LT - -/-- -Converts a `Nanosecond.Ordinal` value to `Millisecond.Ordinal`. --/ -def toMillisecond (nano : Ordinal) : Millisecond.Ordinal := - nano.ediv 1000000 (by decide) +instance {x y : Ordinal} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) -/-- -Converts a `Millisecond.Ordinal` value to `Nanosecond.Ordinal`. --/ -def ofMillisecond (nano : Millisecond.Ordinal) : Nanosecond.Ordinal := - nano.mul_pos 1000000 (by decide) - |>.expandTop (by decide) - -end Ordinal +instance {x y : Ordinal} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) /-- `Offset` represents a time offset in nanoseconds and is defined as an `Int`. @@ -59,7 +42,8 @@ end Ordinal def Offset : Type := UnitVal (1 / 1000000000) deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString -instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +instance : OfNat Offset n := + ⟨UnitVal.ofNat n⟩ namespace Offset @@ -97,6 +81,44 @@ def toOffset (span : Span) : Offset := UnitVal.mk span.val end Span + +namespace Ordinal + +/-- +`Ordinal` represents a bounded value for nanoseconds in a day, which ranges between 0 and 86400000000000. +-/ +def OfDay := Bounded.LE 0 86400000000000 + deriving Repr, BEq, LE, LT + +/-- +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 0 ≤ data ∧ data ≤ 999999999) : Ordinal := + Bounded.LE.mk data h + +/-- +Creates an `Ordinal` from a natural number, ensuring the value is within bounds. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≤ 999999999) : Ordinal := + Bounded.LE.ofNat data h + +/-- +Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. +-/ +@[inline] +def ofFin (data : Fin 1000000000) : Ordinal := + Bounded.LE.ofFin data + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +@[inline] +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal end Nanosecond end Time end Std diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index b8b5d0093183..cf86e6f76c99 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -21,6 +21,15 @@ for potential leap second. -/ def Ordinal (leap : Bool) := Bounded.LE 0 (.ofNat (if leap then 60 else 59)) +instance : BEq (Ordinal leap) where + beq x y := BEq.beq x.val y.val + +instance : LE (Ordinal leap) where + le x y := LE.le x.val y.val + +instance : LT (Ordinal leap) where + lt x y := LT.lt x.val y.val + instance : ToString (Ordinal leap) where toString x := toString x.val @@ -33,8 +42,11 @@ instance : OfNat (Ordinal leap) n := by · exact inst · exact ⟨inst.ofNat.expandTop (by decide)⟩ -instance : OfNat (Ordinal true) 60 where - ofNat := Bounded.LE.mk (Int.ofNat 60) (by decide) +instance {x y : Ordinal l} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance {x y : Ordinal l} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) /-- `Offset` represents an offset in seconds. It is defined as an `Int`. @@ -42,7 +54,8 @@ instance : OfNat (Ordinal true) 60 where def Offset : Type := UnitVal 1 deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString -instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ +instance : OfNat Offset n := + ⟨UnitVal.ofNat n⟩ namespace Offset @@ -64,6 +77,13 @@ end Offset namespace Ordinal +/-- +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 0 ≤ data ∧ data ≤ Int.ofNat (if leap then 60 else 59)) : Ordinal leap := + Bounded.LE.mk data h + /-- Creates an `Ordinal` from a natural number, ensuring the value is within bounds. -/ diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 06a110acc27c..1c5e0ccf3858 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -41,14 +41,14 @@ Converts a `PlainDate` to a `DateTime` -/ @[inline] def ofPlainDate (pd : PlainDate) (tz : TimeZone) : DateTime tz := - DateTime.ofTimestamp (Timestamp.ofPlainDate pd) tz + DateTime.ofTimestamp (Timestamp.ofPlainDateAssumingUTC pd) tz /-- Converts a `DateTime` to a `PlainDate` -/ @[inline] def toPlainDate (dt : DateTime tz) : PlainDate := - Timestamp.toPlainDate dt.toTimestamp + Timestamp.toPlainDateAssumingUTC dt.toTimestamp /-- Converts a `PlainTime` to a `DateTime` @@ -71,8 +71,8 @@ Calculates the duration between a given `DateTime` and a specified date. ```lean example : Duration := - let startDate := date% 2023-1-1:05:10:20UTC - let endDate := date% 2023-3-15:05:10:20UTC + let startDate := date("2023-1-1:05:10:20UTC") + let endDate := date("2023-3-15:05:10:20UTC") endDate.since startDate ``` -/ @@ -99,7 +99,7 @@ Converts a `PlainDate` to a `ZonedDateTime` -/ @[inline] def ofPlainDate (pd : PlainDate) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainDate pd) tz⟩ + ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainDateAssumingUTC pd) tz⟩ /-- Converts a `ZonedDateTime` to a `PlainDate` @@ -115,6 +115,13 @@ Converts a `PlainTime` to a `ZonedDateTime` def ofPlainTime (pt : PlainTime) (tz : TimeZone) : ZonedDateTime := ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainTime pt) tz⟩ +/-- +Converts a `PlainDateTime` to a `ZonedDateTime` assuming the Plain Date is Local. +-/ +@[inline] +def ofLocalDateTime (pd : PlainDateTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofLocalDateTime pd tz⟩ + /-- Converts a `ZonedDateTime` to a `PlainTime` -/ diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index f23f11342a09..179a82765b85 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -6,9 +6,11 @@ Authors: Sofia Rodrigues prelude import Std.Time.DateTime import Std.Time.Zoned.TimeZone +import Std.Internal namespace Std namespace Time +open Internal set_option linter.all true @@ -42,7 +44,7 @@ Creates a new `DateTime` out of a `Timestamp` that is in a `TimeZone`. -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := - DateTime.mk tm (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds).toPlainDateTime) + DateTime.mk tm (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds).toPlainDateTimeAssumingUTC) /-- Creates a new zone aware `Timestamp` out of a `DateTime`. @@ -64,7 +66,7 @@ to UTC. If you're using hte PlainDateTime -/ @[inline] def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : DateTime tz := - let tm := Timestamp.ofPlainDateTime date + let tm := Timestamp.ofPlainDateTimeAssumingUTC date DateTime.mk tm (Thunk.mk fun _ => date.addSeconds tz.toSeconds) /-- @@ -73,7 +75,7 @@ to the -/ @[inline] def ofLocalDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := - let tm := Timestamp.ofPlainDateTime date + let tm := Timestamp.ofPlainDateTimeAssumingUTC date DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk fun _ => date) /-- @@ -250,7 +252,7 @@ def day (dt : DateTime tz) : Day.Ordinal := Getter for the `Hour` inside of a `DateTime` -/ @[inline] -def hour (dt : DateTime tz) : Hour.Ordinal dt.date.get.time.hour.fst := +def hour (dt : DateTime tz) : Hour.Ordinal := dt.date.get.hour /-- @@ -272,7 +274,7 @@ Getter for the `Milliseconds` inside of a `DateTime` -/ @[inline] def milliseconds (dt : DateTime tz) : Millisecond.Ordinal := - dt.date.get.time.nano.toMillisecond + dt.date.get.time.nano.emod 1000 (by decide) /-- Getter for the `Nanosecond` inside of a `DateTime` @@ -288,14 +290,26 @@ Gets the `Weekday` of a DateTime. def weekday (dt : DateTime tz) : Weekday := dt.date.get.date.weekday +/-- +Gets the `Weekday` of a DateTime. +-/ +@[inline] +def getPeriod (dt : DateTime tz) : Day.Ordinal.Period := + if dt.hour ≥ 21 ∨ dt.hour ≤ 4 then + .night + else if dt.hour ≥ 17 then + .evening + else if dt.hour ≥ 12 then + .afternoon + else + .morning + + /-- Determines the era of the given `DateTime` based on its year. -/ def era (date : DateTime tz) : Year.Era := - if date.year.toInt ≥ 0 then - .ce - else - .bce + date.year.era /-- Checks if the `DateTime` is in a leap year. @@ -303,6 +317,36 @@ Checks if the `DateTime` is in a leap year. def inLeapYear (date : DateTime tz) : Bool := date.year.isLeap +/-- +Determines the ordinal day of the year for the given `DateTime`. +-/ +def toOrdinal (date : DateTime tz) : Day.Ordinal.OfYear date.year.isLeap := + Month.Ordinal.toOrdinal ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ + +/-- +Determines the week of the year for the given `DateTime`. +-/ +def toWeekOfYear (date : DateTime tz) : Week.Ordinal := + let res := Month.Ordinal.toOrdinal ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ |>.ediv 7 (by decide) |>.add 1 + match date.date.get.date.year.isLeap, res with + | true, res => res + | false, res => res + +/-- +Determines the week of the month for the given `DateTime`. The week of the month is calculated based +on the day of the month and the weekday. Each week starts on Sunday because the entire library is +based on the Gregorian Calendar. +-/ +def toWeekOfMonth (date : DateTime tz) : Bounded.LE 1 6 := + let weekday := date.weekday.toOrdinal + date.day.addBounds (weekday.sub 1) |>.ediv 7 (by decide) |>.add 1 + +/-- +Determines the quarter of the year for the given `DateTime`. +-/ +def quarter (date : DateTime tz) : Bounded.LE 1 4 := + date.month.sub 1 |>.ediv 3 (by decide) |>.add 1 + instance : ToTimestamp (DateTime tz) where toTimestamp dt := dt.toTimestamp diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 5ca3dcbbf50c..5b53d6730bdc 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -17,7 +17,6 @@ set_option linter.all true Represents a timezone offset with an hour and second component. -/ structure Offset where - private mk :: /-- The timezone offset in Hours. diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index f46a0fe2045f..5b9d62ab6eac 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -26,6 +26,11 @@ structure TimeZone where -/ name : String + /-- + The abbreviation of the time zone/ + -/ + abbreviation : String + /-- Day light saving flag. -/ @@ -44,25 +49,25 @@ opaque getSystemTimezone : IO TimeZone A zeroed `Timezone` representing UTC (no offset). -/ def UTC : TimeZone := - TimeZone.mk (Offset.zero) "Coordinated Universal Time" false + TimeZone.mk (Offset.zero) "Coordinated Universal Time" "UTC" false /-- A zeroed `Timezone` representing GMT (no offset). -/ def GMT : TimeZone := - TimeZone.mk (Offset.zero) "Greenwich Mean Time" false + TimeZone.mk (Offset.zero) "Greenwich Mean Time" "GMT" false /-- Creates a `Timestamp` from a given number of hour. -/ -def ofHours (name : String) (n : Hour.Offset) (isDST : Bool := false) : TimeZone := - TimeZone.mk (Offset.ofHours n) name isDST +def ofHours (name : String) (abbreviation : String) (n : Hour.Offset) (isDST : Bool := false) : TimeZone := + TimeZone.mk (Offset.ofHours n) name abbreviation isDST /-- Creates a `Timestamp` from a given number of second. -/ -def ofSeconds (name : String) (n : Second.Offset) (isDST : Bool := false) : TimeZone := - TimeZone.mk (Offset.ofSeconds n) name isDST +def ofSeconds (name : String) (abbreviation : String) (n : Second.Offset) (isDST : Bool := false) : TimeZone := + TimeZone.mk (Offset.ofSeconds n) name abbreviation isDST /-- Gets the number of seconds in a timezone offset. diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 6962f85c62b0..fff16b721a9a 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -86,7 +86,7 @@ namespace LocalTimeType Gets the `TimeZone` offset from a `LocalTimeType`. -/ def getTimeZone (time : LocalTimeType) : TimeZone := - ⟨time.gmtOffset, time.abbreviation, time.isDst⟩ + ⟨time.gmtOffset, time.identifier, time.abbreviation, time.isDst⟩ end LocalTimeType @@ -149,9 +149,10 @@ namespace Transition Create a TimeZone from a Transition. -/ def createTimeZoneFromTransition (transition : Transition) : TimeZone := + let name := transition.localTimeType.identifier let offset := transition.localTimeType.gmtOffset - let name := transition.localTimeType.abbreviation - TimeZone.mk offset name transition.localTimeType.isDst + let abbreviation := transition.localTimeType.abbreviation + TimeZone.mk offset name abbreviation transition.localTimeType.isDst /-- Applies the transition to a Timestamp. diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index bf9ae9857716..1244028d0e73 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -30,6 +30,13 @@ instance : Inhabited ZonedDateTime where namespace ZonedDateTime open DateTime +/-- +Creates a new `ZonedDateTime` out of a `DateTime` and `TimeZone` +-/ +@[inline] +def mk (tz : TimeZone) (datetime : DateTime tz) : ZonedDateTime := + ⟨tz, datetime⟩ + /-- Creates a new `ZonedDateTime` out of a `Timestamp` -/ @@ -99,8 +106,8 @@ def day (zdt : ZonedDateTime) : Day.Ordinal := Getter for the `Hour` inside of a `ZonedDateTime` -/ @[inline] -def hour (zdt : ZonedDateTime) : Hour.Ordinal zdt.snd.date.get.time.hour.fst := - zdt.snd.date.get.time.hour.snd +def hour (zdt : ZonedDateTime) : Hour.Ordinal := + zdt.snd.date.get.time.hour /-- Getter for the `Minute` inside of a `ZonedDateTime` @@ -116,6 +123,13 @@ Getter for the `Second` inside of a `ZonedDateTime` def second (zdt : ZonedDateTime) : Second.Ordinal zdt.snd.date.get.time.second.fst := zdt.snd.date.get.time.second.snd +/-- +Getter for the `Nanosecond` inside of a `ZonedDateTime` +-/ +@[inline] +def nano (zdt : ZonedDateTime) : Nanosecond.Ordinal := + zdt.snd.date.get.time.nano + /-- Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` -/ @@ -206,6 +220,20 @@ Add `Year.Offset` to a `ZonedDateTime`, it clips the day to the last valid day o def addYearsClip (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := Sigma.mk dt.fst (dt.snd.addYearsClip years) +/-- +Subtract `Year.Offset` from a `ZonedDateTime`, this function clips the day to the last valid day of that month. +-/ +@[inline] +def subYearsClip (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subYearsClip years) + +/-- +Subtract `Year.Offset` from a `ZonedDateTime`, this function rolls over any excess days into the previous month. +-/ +@[inline] +def subYearsRollOver (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := + Sigma.mk dt.fst (dt.snd.subYearsRollOver years) + /-- Add `Hour.Offset` to a `ZonedDateTime`, adjusting the date if necessary. -/ diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index 80d10ac8d69a..2c92b96813d4 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -1,160 +1,160 @@ import Std.Time open Std.Time -def date := date% 1970-01-20 +def date := date("1970-01-20") /-- -info: date% 1970-02-01 +info: date("1970-02-01") -/ #guard_msgs in #eval date + (12 : Day.Offset) /-- -info: date% 1970-01-08 +info: date("1970-01-08") -/ #guard_msgs in #eval date - (12 : Day.Offset) -def datetime := date% 2000-01-20 T 03:02:01 +def datetime := datetime("2000-01-20T03:02:01") /-- -info: date% 2000-01-20T04:02:01,000000000 +info: datetime("2000-01-20T04:02:01.000000000") -/ #guard_msgs in #eval datetime + (1 : Hour.Offset) /-- -info: date% 2000-01-20T02:02:01,000000000 +info: datetime("2000-01-20T02:02:01.000000000") -/ #guard_msgs in #eval datetime - (1 : Hour.Offset) /-- -info: date% 2000-01-20T03:12:01,000000000 +info: datetime("2000-01-20T03:12:01.000000000") -/ #guard_msgs in #eval datetime + (10 : Minute.Offset) /-- -info: date% 2000-01-20T02:52:01,000000000 +info: datetime("2000-01-20T02:52:01.000000000") -/ #guard_msgs in #eval datetime - (10 : Minute.Offset) /-- -info: date% 2000-01-20T03:03:01,000000000 +info: datetime("2000-01-20T03:03:01.000000000") -/ #guard_msgs in #eval datetime + (60 : Second.Offset) /-- -info: date% 2000-01-20T03:01:01,000000000 +info: datetime("2000-01-20T03:01:01.000000000") -/ #guard_msgs in #eval datetime - (60 : Second.Offset) /-- -info: date% 2000-01-21T03:02:01,000000000 +info: datetime("2000-01-21T03:02:01.000000000") -/ #guard_msgs in #eval datetime + (1 : Day.Offset) /-- -info: date% 2000-01-19T03:02:01,000000000 +info: datetime("2000-01-19T03:02:01.000000000") -/ #guard_msgs in #eval datetime - (1 : Day.Offset) -def time := time% 13:02:01 +def time := time("13:02:01") /-- -info: time% 14:02:01,000000000 +info: time("14:02:01.000000000") -/ #guard_msgs in #eval time + (1 : Hour.Offset) /-- -info: time% 12:02:01,000000000 +info: time("12:02:01.000000000") -/ #guard_msgs in #eval time - (1 : Hour.Offset) /-- -info: time% 13:12:01,000000000 +info: time("13:12:01.000000000") -/ #guard_msgs in #eval time + (10 : Minute.Offset) /-- -info: time% 12:52:01,000000000 +info: time("12:52:01.000000000") -/ #guard_msgs in #eval time - (10 : Minute.Offset) /-- -info: time% 13:03:01,000000000 +info: time("13:03:01.000000000") -/ #guard_msgs in #eval time + (60 : Second.Offset) /-- -info: time% 13:01:01,000000000 +info: time("13:01:01.000000000") -/ #guard_msgs in #eval time - (60 : Second.Offset) -def datetimetz := date% 2000-01-20 T 06:02:01-03:00 +def datetimetz := zoned("2000-01-20T06:02:01-03:00") /-- -info: date% 2000-01-20T06:02:01,000000000-03:00 +info: zoned("2000-01-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz /-- -info: date% 2000-01-22T06:02:01,000000000-03:00 +info: zoned("2000-01-22T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz + (2 : Day.Offset) /-- -info: date% 2000-01-19T06:02:01,000000000-03:00 +info: zoned("2000-01-19T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz - (1 : Day.Offset) /-- -info: date% 2000-01-20T07:02:01,000000000-03:00 +info: zoned("2000-01-20T07:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz + (1 : Hour.Offset) /-- -info: date% 2000-01-20T05:02:01,000000000-03:00 +info: zoned("2000-01-20T05:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz - (1 : Hour.Offset) /-- -info: date% 2000-01-20T06:12:01,000000000-03:00 +info: zoned("2000-01-20T06:12:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz + (10 : Minute.Offset) /-- -info: date% 2000-01-20T05:52:01,000000000-03:00 +info: zoned("2000-01-20T05:52:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz - (10 : Minute.Offset) /-- -info: date% 2000-01-20T06:03:01,000000000-03:00 +info: zoned("2000-01-20T06:03:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz + (60 : Second.Offset) /-- -info: date% 2000-01-20T06:01:01,000000000-03:00 +info: zoned("2000-01-20T06:01:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz - (60 : Second.Offset) @@ -163,23 +163,23 @@ info: date% 2000-01-20T06:01:01,000000000-03:00 info: "3600s" -/ #guard_msgs in -#eval (date% 2000-12-20 T 00:0:00-03:00) - (date% 2000-12-20 T 00:00:00-02:00) +#eval zoned("2000-12-20T00:00:00-03:00") - zoned("2000-12-20T00:00:00-02:00") def now := PlainDateTime.ofUTCTimestamp ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ def after := PlainDateTime.ofUTCTimestamp ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ def duration := after.since now /-- -info: "03:40:38.907328169" +info: "15:40:38.907328169" -/ #guard_msgs in -#eval now.format "HH:mm:ss.sssssssss" +#eval now.format "HH:mm:ss.SSSSSSSSS" /-- -info: "03.40:39.907641224" +info: "15.40:39.907641224" -/ #guard_msgs in -#eval after.format "HH.mm:ss.sssssssss" +#eval after.format "HH.mm:ss.SSSSSSSSS" /-- info: "1.000313055s" diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 7f8db5c2abbd..8e72a2873f43 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -1,33 +1,34 @@ import Std.Time open Std.Time -def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" -def RFC1123 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" -def ShortDate : Format .any := date-spec% "MM/DD/YYYY" -def LongDate : Format .any := date-spec% "MMMM D, YYYY" -def ShortDateTime : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss" -def LongDateTime : Format .any := date-spec% "MMMM D, YYYY h:mm aa" -def Time24Hour : Format .any := date-spec% "hh:mm:ss" -def Time12Hour : Format .any := date-spec% "HH:mm:ss aa" -def FullDayTimeZone : Format .any := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss ZZZZ" -def CustomDayTime : Format .any := date-spec% "EEE D MMM YYYY hh:mm" +def ISO8601UTC : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZ") +def RFC1123 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def ShortDate : Format .any := datespec("MM/dd/yyyy") +def LongDate : Format .any := datespec("MMMM D, yyyy") +def ShortDateTime : Format .any := datespec("MM/dd/yyyy HH:mm:ss") +def LongDateTime : Format .any := datespec("MMMM D, yyyy h:mm aa") +def Time24Hour : Format .any := datespec("HH:mm:ss") +def Time12Hour : Format .any := datespec("hh:mm:ss aa") +def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") +def CustomDayTime : Format .any := datespec("EEE dd MMM yyyy HH:mm") -- Dates -def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 -def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 +def brTZ : TimeZone := timezone("America/Sao_Paulo -03:00") +def jpTZ : TimeZone := timezone("Asia/Tokyo +09:00") -def date₁ := date% 2014-06-16 T 03:03:03(brTZ) -def time₁ := time% 14:11:01 -def time₂ := time% 03:11:01 +def date₁ := zoned("2014-06-16T03:03:03-03:00") + +def time₁ := time("14:11:01") +def time₂ := time("03:11:01") /-- info: "Monday, June 16, 2014 03:03:03 -0300" -/ #guard_msgs in -#eval FullDayTimeZone.format date₁ +#eval FullDayTimeZone.format date₁.snd -def tm := date₁.timestamp +def tm := date₁.snd.timestamp def date₂ := DateTime.ofTimestamp tm brTZ /-- @@ -178,17 +179,18 @@ info: "14:11:01" #guard_msgs in #eval Time24Hour.formatBuilder time₁.hour time₁.minute time₁.second +def l := Time12Hour.formatBuilder time₁.hour.toRelative time₁.minute time₁.second (if time₁.hour.val > 12 then HourMarker.pm else HourMarker.am) + /-- -info: "02:11:01 pm" +info: "02:11:01 PM" -/ #guard_msgs in -#eval Time12Hour.formatBuilder time₁.hour time₁.minute time₁.second (if time₁.hour.snd.val > 12 then HourMarker.pm else HourMarker.am) - +#eval l /-- -info: "03:11:01 am" +info: "03:11:01 AM" -/ #guard_msgs in -#eval Time12Hour.formatBuilder time₂.hour time₂.minute time₂.second (if time₂.hour.snd.val > 12 then HourMarker.pm else HourMarker.am) +#eval Time12Hour.formatBuilder time₂.hour.toRelative time₂.minute time₂.second (if time₂.hour.val > 12 then HourMarker.pm else HourMarker.am) /-- info: "06/16/2014" @@ -221,49 +223,49 @@ info: "-0221-09-04" #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) /-- -info: date% -0221-09-04 +info: date("-0221-09-04") -/ #guard_msgs in -#eval (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) +#eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ /-- -info: "-0221-09-04" +info: date("-0221-09-04") -/ #guard_msgs in -#eval toString (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) +#eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ /-- -info: date% 2002-07-14 +info: date("2002-07-14") -/ #guard_msgs in -#eval date% 2002-07-14 +#eval date("2002-07-14") /-- -info: time% 14:13:12,000000000 +info: time("14:13:12.000000000") -/ #guard_msgs in -#eval time% 14:13:12 +#eval time("14:13:12") /-- -info: date% 2002-07-14T14:13:12,000000000 +info: zoned("2002-07-14T14:13:12.000000000Z") -/ #guard_msgs in -#eval date% 2002-07-14 T 14:13:12 +#eval zoned("2002-07-14T14:13:12Z") /-- -info: date% 2002-07-14T14:13:12,000000000+09:00 +info: zoned("2002-07-14T14:13:12.000000000+09:00") -/ #guard_msgs in -#eval date% 2002-07-14 T 14:13:12+09:00 +#eval zoned("2002-07-14T14:13:12+09:00") /-- info: "2002-07-14" -/ #guard_msgs in -#eval (date% 2002-07-14 T 14:13:12+09:00).format "YYYY-MM-DD" +#eval zoned("2002-07-14T14:13:12+09:00").format "yyyy-MM-dd" /-- info: "14-13-12" -/ #guard_msgs in -#eval (date% 2002-07-14 T 14:13:12+09:00).format "hh-mm-ss" +#eval zoned("2002-07-14T14:13:12+09:00").format "HH-mm-ss" diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index aee24cc90e49..e79d6e202617 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -2,15 +2,16 @@ import Std.Time import Init open Std.Time -def ShortDateTime : Format .any := date-spec% "DD/MM/YYYY hh:mm:ss" -def ShortDate : Format .any := date-spec% "DD/MM/YYYY" +def ShortDateTime : Format .any := datespec("dd/MM/yyyy HH:mm:ss") +def ShortDate : Format .any := datespec("dd/MM/yyyy") + def format (PlainDate : PlainDateTime) : String := ShortDateTime.formatBuilder PlainDate.day PlainDate.month PlainDate.year PlainDate.time.hour PlainDate.minute PlainDate.time.second def format₂ (PlainDate : PlainDate) : String := ShortDate.formatBuilder PlainDate.day PlainDate.month PlainDate.year -def date₁ := date% 1993-11-19 T 09:08:07 -def date₂ := date% 1993-05-09 T 12:59:59 -def date₃ := date% 2024-08-16 -def date₄ := date% 1500-08-16 +def date₁ := datetime("1993-11-19T09:08:07") +def date₂ := datetime("1993-05-09T12:59:59") +def date₃ := date("2024-08-16") +def date₄ := date("1500-08-16") def tm₁ := 753700087 def tm₂ := 736952399 diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index 8a003a652dc6..9f4eba8a858c 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -1,275 +1,275 @@ import Std.Time open Std.Time -def date := date% 1970-01-20 +def date := date("1970-01-20") /-- -info: date% 1970-02-01 +info: date("1970-02-01") -/ #guard_msgs in #eval date.addDays 12 /-- -info: date% 1970-02-20 +info: date("1970-02-20") -/ #guard_msgs in #eval date.addMonthsClip 1 /-- -info: date% 1971-01-20 +info: date("1971-01-20") -/ #guard_msgs in #eval date.addYearsRollOver 1 /-- -info: date% 1969-01-20 +info: date("1969-01-20") -/ #guard_msgs in #eval date.subYearsClip 1 /-- -info: date% 1969-12-20 +info: date("1969-12-20") -/ #guard_msgs in #eval date.subMonthsClip 1 -def datetime := date% 2000-01-20 T 03:02:01 +def datetime := datetime("2000-01-20T03:02:01") /-- -info: date% 2000-01-20T04:02:01,000000000 +info: datetime("2000-01-20T04:02:01.000000000") -/ #guard_msgs in #eval datetime.addHours 1 /-- -info: date% 2000-01-20T02:02:01,000000000 +info: datetime("2000-01-20T02:02:01.000000000") -/ #guard_msgs in #eval datetime.subHours 1 /-- -info: date% 2000-01-20T03:12:01,000000000 +info: datetime("2000-01-20T03:12:01.000000000") -/ #guard_msgs in #eval datetime.addMinutes 10 /-- -info: date% 2000-01-20T02:52:01,000000000 +info: datetime("2000-01-20T02:52:01.000000000") -/ #guard_msgs in #eval datetime.subMinutes 10 /-- -info: date% 2000-01-20T03:03:01,000000000 +info: datetime("2000-01-20T03:03:01.000000000") -/ #guard_msgs in #eval datetime.addSeconds 60 /-- -info: date% 2000-01-20T03:01:01,000000000 +info: datetime("2000-01-20T03:01:01.000000000") -/ #guard_msgs in #eval datetime.subSeconds 60 /-- -info: date% 2000-01-21T03:02:01,000000000 +info: datetime("2000-01-21T03:02:01.000000000") -/ #guard_msgs in #eval datetime.addDays 1 /-- -info: date% 2000-01-19T03:02:01,000000000 +info: datetime("2000-01-19T03:02:01.000000000") -/ #guard_msgs in #eval datetime.subDays 1 /-- -info: date% 2000-02-20T03:02:01,000000000 +info: datetime("2000-02-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.addMonthsClip 1 /-- -info: date% 1999-12-20T03:02:01,000000000 +info: datetime("1999-12-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.subMonthsClip 1 /-- -info: date% 2000-02-20T03:02:01,000000000 +info: datetime("2000-02-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.addMonthsRollOver 1 /-- -info: date% 1999-12-20T03:02:01,000000000 +info: datetime("1999-12-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.subMonthsRollOver 1 /-- -info: date% 2001-01-20T03:02:01,000000000 +info: datetime("2001-01-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.addYearsClip 1 /-- -info: date% 1999-01-20T03:02:01,000000000 +info: datetime("1999-01-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.subYearsClip 1 /-- -info: date% 2001-01-20T03:02:01,000000000 +info: datetime("2001-01-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.addYearsRollOver 1 /-- -info: date% 1999-01-20T03:02:01,000000000 +info: datetime("1999-01-20T03:02:01.000000000") -/ #guard_msgs in #eval datetime.subYearsRollOver 1 -def time := time% 13:02:01 +def time := time("13:02:01") /-- -info: time% 14:02:01,000000000 +info: time("14:02:01.000000000") -/ #guard_msgs in #eval time.addHours 1 /-- -info: time% 12:02:01,000000000 +info: time("12:02:01.000000000") -/ #guard_msgs in #eval time.subHours 1 /-- -info: time% 13:12:01,000000000 +info: time("13:12:01.000000000") -/ #guard_msgs in #eval time.addMinutes 10 /-- -info: time% 12:52:01,000000000 +info: time("12:52:01.000000000") -/ #guard_msgs in #eval time.subMinutes 10 /-- -info: time% 13:03:01,000000000 +info: time("13:03:01.000000000") -/ #guard_msgs in #eval time.addSeconds 60 /-- -info: time% 13:01:01,000000000 +info: time("13:01:01.000000000") -/ #guard_msgs in #eval time.subSeconds 60 -def datetimetz := date% 2000-01-20 T 06:02:01-03:00 +def datetimetz := zoned("2000-01-20T06:02:01-03:00") /-- -info: date% 2000-01-20T06:02:01,000000000-03:00 +info: zoned("2000-01-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz /-- -info: date% 2000-01-22T06:02:01,000000000-03:00 +info: zoned("2000-01-22T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addDays 2 /-- -info: date% 2000-01-19T06:02:01,000000000-03:00 +info: zoned("2000-01-19T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.subDays 1 /-- -info: date% 2000-02-20T06:02:01,000000000-03:00 +info: zoned("2000-02-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addMonthsClip 1 /-- -info: date% 1999-12-20T06:02:01,000000000-03:00 +info: zoned("1999-12-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.subMonthsClip 1 /-- -info: date% 2000-02-20T06:02:01,000000000-03:00 +info: zoned("2000-02-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addMonthsRollOver 1 /-- -info: date% 1999-12-20T06:02:01,000000000-03:00 +info: zoned("1999-12-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.subMonthsRollOver 1 /-- -info: date% 2001-01-20T06:02:01,000000000-03:00 +info: zoned("2001-01-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: date% 2001-01-20T06:02:01,000000000-03:00 +info: zoned("2001-01-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addYearsClip 1 /-- -info: date% 2001-01-20T06:02:01,000000000-03:00 +info: zoned("2001-01-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addYearsRollOver 1 /-- -info: date% 1999-01-20T06:02:01,000000000-03:00 +info: zoned("1999-01-20T06:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.subYearsRollOver 1 /-- -info: date% 2000-01-20T07:02:01,000000000-03:00 +info: zoned("2000-01-20T07:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addHours 1 /-- -info: date% 2000-01-20T05:02:01,000000000-03:00 +info: zoned("2000-01-20T05:02:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.subHours 1 /-- -info: date% 2000-01-20T06:12:01,000000000-03:00 +info: zoned("2000-01-20T06:12:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addMinutes 10 /-- -info: date% 2000-01-20T05:52:01,000000000-03:00 +info: zoned("2000-01-20T05:52:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.subMinutes 10 /-- -info: date% 2000-01-20T06:03:01,000000000-03:00 +info: zoned("2000-01-20T06:03:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.addSeconds 60 /-- -info: date% 2000-01-20T06:01:01,000000000-03:00 +info: zoned("2000-01-20T06:01:01.000000000-03:00") -/ #guard_msgs in #eval datetimetz.subSeconds 60 @@ -278,31 +278,30 @@ info: date% 2000-01-20T06:01:01,000000000-03:00 info: "86400s" -/ #guard_msgs in -#eval (date% 2000-1-20).since (date% 2000-1-19) +#eval (date("2000-01-20").since (date("2000-01-19"))) /-- info: "86399s" -/ #guard_msgs in -#eval (date% 2000-1-20 T 0:0:0).since (date% 2000-1-19 T 0:0:1) +#eval (datetime("2000-01-20T00:00:00").since (datetime("2000-01-19T00:00:01"))) /-- info: "86399s" -/ #guard_msgs in -#eval (date% 2000-1-20 T 0:0:0UTC).since (date% 2000-1-19 T 0:0:1UTC) - -def now := date% 2024-08-29 T 10:56:43,276801081+02:00 +#eval (zoned("2000-01-20T00:00:00Z").since (zoned("2000-01-19T00:00:01Z"))) +def now := zoned("2024-08-29T10:56:43.276801081+02:00") /-- -info: date% 2024-08-29T10:56:43,276801081+02:00 +info: zoned("2024-08-29T10:56:43.276801081+02:00") -/ #guard_msgs in #eval now /-- -info: date% 2024-08-30T10:56:43,276801081+02:00 +info: zoned("2024-08-30T10:56:43.276801081+02:00") -/ #guard_msgs in #eval now.addDays 1 diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 33961dc215e1..06e9028f7369 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -1,41 +1,43 @@ import Std.Time open Std.Time -def ISO8601UTC : Format .any := date-spec% "YYYY-MM-DD'T'hh:mm:ss.sssZ" -def RFC1123 : Format .any := date-spec% "EEE, DD MMM YYYY hh:mm:ss ZZZ" -def ShortDate : Format .any := date-spec% "MM/DD/YYYY" -def LongDate : Format .any := date-spec% "MMMM D, YYYY" -def ShortDateTime : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss" -def LongDateTime : Format .any := date-spec% "MMMM D, YYYY h:mm AA" -def Time24Hour : Format .any := date-spec% "hh:mm:ss" -def Time12Hour : Format .any := date-spec% "HH:mm:ss aa" -def FullDayTimeZone : Format .any := date-spec% "EEEE, MMMM D, YYYY hh:mm:ss ZZZZ" -def CustomDayTime : Format .any := date-spec% "EEE D MMM YYYY hh:mm" - -def Full12HourWrong : Format .any := date-spec% "MM/DD/YYYY hh:mm:ss aa Z" +def ISO8601UTC : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") +def RFC1123 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def ShortDate : Format .any := datespec("MM/dd/yyyy") +def LongDate : Format .any := datespec("MMMM D, yyyy") +def ShortDateTime : Format .any := datespec("MM/dd/yyyy HH:mm:ss") +def LongDateTime : Format .any := datespec("MMMM dd, yyyy hh:mm aa") +def Time24Hour : Format .any := datespec("HH:mm:ss") +def Time12Hour : Format .any := datespec("hh:mm:ss aa") +def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss XXX") +def CustomDayTime : Format .any := datespec("EEE dd MMM yyyy HH:mm") + +def Full12HourWrong : Format .any := datespec("MM/dd/yyyy HH:mm:ss aa XXX") -- Dates -def brTZ : TimeZone := timezone% "America/Sao_Paulo" -03:00 -def jpTZ : TimeZone := timezone% "Asia/Tokyo" +09:00 -def date₁ := date% 2014-06-16 T 03:03:03(brTZ) -def time₁ := time% 14:11:01 -def time₂ := time% 03:11:01 +def brTZ : TimeZone := timezone("America/Sao_Paulo -03:00") +def jpTZ : TimeZone := timezone("Asia/Tokyo +09:00") + +def date₁ := zoned("2014-06-16T03:03:03-03:00") + +def time₁ := time("14:11:01") +def time₂ := time("03:11:01") /-- -info: "2014-06-16T03:03:03.000-03:00" +info: "2014-06-16T03:03:03.000000100-03:00" -/ #guard_msgs in #eval - let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03.000-03:00" + let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03.000000100-03:00" ISO8601UTC.format t.snd -def tm := date₁.timestamp +def tm := date₁.snd.timestamp def date₂ := DateTime.ofTimestamp tm brTZ /-- -info: "2014-06-16T03:03:03.000-03:00" +info: "2014-06-16T03:03:03.000000000-03:00" -/ #guard_msgs in #eval @@ -46,7 +48,7 @@ def tm₃ := date₁.toTimestamp def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- -info: "2014-06-16T00:00:00.000Z" +info: "2014-06-16T00:00:00.000000000Z" -/ #guard_msgs in #eval @@ -67,7 +69,7 @@ def dateJP := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) jpT def dateUTC := DateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch tm₄) .UTC /-- -info: "2024-08-15T13:28:12.000-03:00" +info: "2024-08-15T13:28:12.000000000-03:00" -/ #guard_msgs in #eval @@ -75,11 +77,11 @@ info: "2024-08-15T13:28:12.000-03:00" ISO8601UTC.format t.snd /-- -info: "2024-08-16T01:28:00.000Z" +info: "2024-08-16T01:28:00.000000000Z" -/ #guard_msgs in #eval - let t : ZonedDateTime := LongDateTime.parse! "August 16, 2024 1:28 AM" + let t : ZonedDateTime := LongDateTime.parse! "August 16, 2024 01:28 AM" ISO8601UTC.format t.snd /-- @@ -147,7 +149,7 @@ def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ def dateUTC₁ := DateTime.ofLocalDateTime PlainDate .UTC /-- -info: "2024-08-15T14:03:47.000-03:00" +info: "2024-08-15T14:03:47.000000000-03:00" -/ #guard_msgs in #eval @@ -155,7 +157,7 @@ info: "2024-08-15T14:03:47.000-03:00" ISO8601UTC.format t.snd /-- -info: "2024-08-15T14:03:47.000+09:00" +info: "2024-08-15T14:03:47.000000000+09:00" -/ #guard_msgs in #eval @@ -163,7 +165,7 @@ info: "2024-08-15T14:03:47.000+09:00" ISO8601UTC.format t1.snd /-- -info: "2014-06-16T03:03:03.000-03:00" +info: "2014-06-16T03:03:03.000000000-03:00" -/ #guard_msgs in #eval @@ -171,11 +173,11 @@ info: "2014-06-16T03:03:03.000-03:00" ISO8601UTC.format t2.snd /-- -info: Except.ok "1993-05-10T10:30:23.000+03:00" +info: Except.ok "1993-05-10T10:30:23.000000000+03:00" -/ #guard_msgs in #eval - let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 am +03:00" + let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 AM +03:00" (ISO8601UTC.format ·.snd) <$> t2 /-- @@ -183,7 +185,7 @@ info: Except.ok "1993-05-10T22:30:23.000+03:00" -/ #guard_msgs in #eval - let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 pm +03:00" + let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 PM +03:00" (ISO8601UTC.format ·.snd) <$> t2 /-- @@ -191,7 +193,7 @@ info: Except.error "offset 29: The 24-hour is out of the range and cannot be tra -/ #guard_msgs in #eval - let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 am +03:00" + let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 AM +03:00" (ISO8601UTC.format ·.snd) <$> t2 /-- @@ -199,5 +201,5 @@ info: Except.error "offset 29: The 24-hour is out of the range and cannot be tra -/ #guard_msgs in #eval - let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 pm +03:00" + let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 PM +03:00" (ISO8601UTC.format ·.snd) <$> t2 From d152573b341d21fffdfd591bc5b212254b415c8f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 20 Sep 2024 13:51:55 -0300 Subject: [PATCH 065/118] fix: set functions and hour marker --- src/Std/Time/DateTime/PlainDateTime.lean | 84 +++++++++++++++++++-- src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Internal/Bounded.lean | 7 ++ src/Std/Time/Time/HourMarker.lean | 22 +++++- src/Std/Time/Time/PlainTime.lean | 28 +++++++ src/Std/Time/Zoned/DateTime.lean | 76 +++++++++++++++++++ src/Std/Time/Zoned/ZonedDateTime.lean | 76 +++++++++++++++++++ tests/lean/run/timeParse.lean | 24 +++--- tests/lean/run/timeSet.lean | 93 ++++++++++++++++++++++++ 9 files changed, 389 insertions(+), 23 deletions(-) create mode 100644 tests/lean/run/timeSet.lean diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 71a659b6f1b7..179f816f49fb 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -126,6 +126,82 @@ def ofUTCTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do time := PlainTime.ofHourMinuteSecondsNano (leap := false) (hour.expandTop (by decide)) minute second nano } +/-- +Creates a new `PlainDateTime` by adjusting the day of the month to the given `days` value, with any +out-of-range days clipped to the nearest valid date. +-/ +@[inline] +def withDaysClip (dt : PlainDateTime) (days : Day.Ordinal) : PlainDateTime := + { dt with date := PlainDate.clip dt.date.year dt.date.month days } + +/-- +Creates a new `PlainDateTime` by adjusting the day of the month to the given `days` value, with any +out-of-range days rolled over to the next month or year as needed. +-/ +@[inline] +def withDaysRollOver (dt : PlainDateTime) (days : Day.Ordinal) : PlainDateTime := + { dt with date := PlainDate.rollOver dt.date.year dt.date.month days } + +/-- +Creates a new `PlainDateTime` by adjusting the month to the given `month` value. +The day remains unchanged, and any invalid days for the new month will be handled according to the `clip` behavior. +-/ +@[inline] +def withMonthClip (dt : PlainDateTime) (month : Month.Ordinal) : PlainDateTime := + { dt with date := PlainDate.clip dt.date.year month dt.date.day } + +/-- +Creates a new `PlainDateTime` by adjusting the month to the given `month` value. +The day is rolled over to the next valid month if necessary. +-/ +@[inline] +def withMonthRollOver (dt : PlainDateTime) (month : Month.Ordinal) : PlainDateTime := + { dt with date := PlainDate.rollOver dt.date.year month dt.date.day } + +/-- +Creates a new `PlainDateTime` by adjusting the year to the given `year` value. The month and day remain unchanged, +and any invalid days for the new year will be handled according to the `clip` behavior. +-/ +@[inline] +def withYearClip (dt : PlainDateTime) (year : Year.Offset) : PlainDateTime := + { dt with date := PlainDate.clip year dt.date.month dt.date.day } + +/-- +Creates a new `PlainDateTime` by adjusting the year to the given `year` value. The month and day are rolled +over to the next valid month and day if necessary. +-/ +@[inline] +def withYearRollOver (dt : PlainDateTime) (year : Year.Offset) : PlainDateTime := + { dt with date := PlainDate.rollOver year dt.date.month dt.date.day } + +/-- +Creates a new `PlainDateTime` by adjusting the `hour` component of its `time` to the given value. +-/ +@[inline] +def withHour (dt : PlainDateTime) (hour : Hour.Ordinal) : PlainDateTime := + { dt with time := { dt.time with hour := hour } } + +/-- +Creates a new `PlainDateTime` by adjusting the `minute` component of its `time` to the given value. +-/ +@[inline] +def withMinute (dt : PlainDateTime) (minute : Minute.Ordinal) : PlainDateTime := + { dt with time := { dt.time with minute := minute } } + +/-- +Creates a new `PlainDateTime` by adjusting the `second` component of its `time` to the given value. +-/ +@[inline] +def withSecond (dt : PlainDateTime) (second : Sigma Second.Ordinal) : PlainDateTime := + { dt with time := { dt.time with second := second } } + +/-- +Creates a new `PlainDateTime` by adjusting the `nano` component of its `time` to the given value. +-/ +@[inline] +def withNano (dt : PlainDateTime) (nano : Nanosecond.Ordinal) : PlainDateTime := + { dt with time := { dt.time with nano := nano } } + /-- Adds a `Day.Offset` to a `PlainDateTime`. -/ @@ -140,14 +216,6 @@ Subtracts a `Day.Offset` from a `PlainDateTime`. def subDays (dt : PlainDateTime) (days : Day.Offset) : PlainDateTime := { dt with date := dt.date.subDays days } -/-- -Sets the `Day.Offset` from a `PlainDateTime`. --/ -@[inline] -def withDaysClip (dt : PlainDateTime) (days : Day.Ordinal) : PlainDateTime := - { dt with date := dt.date.withDaysClip days } - - /-- Adds a `Week.Offset` to a `PlainDateTime`. -/ diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index c580f3c339a4..0e0079bc62c9 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -924,7 +924,7 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .F _ => date.toWeekOfMonth | .a _ => HourMarker.ofOrdinal date.hour | .B _ => date.getPeriod - | .h _ => (date.hour.sub 1).emod 12 (by decide) |>.add 1 + | .h _ => HourMarker.toRelative date.hour |>.fst | .K _ => date.hour.emod 12 (by decide) | .k _ => date.hour.add 1 | .H _ => date.hour diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index 8f4400308dde..a2ad9a8fdc1f 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -394,6 +394,13 @@ def ediv (bounded : Bounded.LE n m) (num : Int) (h : num > 0) : Bounded.LE (n / def eq {n : Int} : Bounded.LE n n := ⟨n, And.intro (Int.le_refl n) (Int.le_refl n)⟩ +/-- +Expand the top and bottom of the bounded. +-/ +@[inline] +def expand (bounded : Bounded.LE lo hi) (h : hi ≤ nhi) (h₁ : nlo ≤ lo) : Bounded.LE nlo nhi := + ⟨bounded.val, And.intro (Int.le_trans h₁ bounded.property.left) (Int.le_trans bounded.property.right h)⟩ + /-- Expand the bottom of the bounded to a number `nhi` is `hi` is less or equal to the previous higher bound. -/ diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 9529d0bf9444..eaaa2f6577d1 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -39,8 +39,26 @@ Converts a 12-hour clock time to a 24-hour clock time based on the `HourMarker`. -/ def toAbsolute (marker : HourMarker) (time : Bounded.LE 1 12) : Hour.Ordinal := match marker with - | .am => time.sub 1 |>.expandTop (by decide) - | .pm => time.add 11 |>.emod 24 (by decide) + | .am => if time.val = 12 then 0 else time.expand (by decide) (by decide) + | .pm => if time.val = 12 then 12 else time.add 12 |>.emod 24 (by decide) + +/-- +Converts a 24-hour clock time to a 12-hour clock time with an `HourMarker`. +-/ +def toRelative (hour : Hour.Ordinal) : Bounded.LE 1 12 × HourMarker := + if h₀ : hour.val = 0 then + (⟨12, by decide⟩, .am) + else if h₁ : hour.val ≤ 12 then + if hour.val = 12 then + (⟨12, by decide⟩, .pm) + else + Int.ne_iff_lt_or_gt.mp h₀ |>.by_cases + (λh => nomatch (Int.not_le.mpr h) hour.property.left) + (λh => (⟨hour.val, And.intro h h₁⟩, .am)) + else + let h := Int.not_le.mp h₁ + let t := hour |>.truncateBottom h |>.sub 12 + (t.expandTop (by decide), .pm) end HourMarker end Time diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 388beb495ebb..bfb1cac1a00d 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -197,6 +197,34 @@ Subtracts milliseconds from a `PlainTime`. def subMilliseconds (time : PlainTime) (millisToSub : Millisecond.Offset) : PlainTime := addMilliseconds time (-millisToSub) +/-- +Creates a new `PlainTime` by adjusting the `second` component to the given value. +-/ +@[inline] +def withSecond (pt : PlainTime) (second : Sigma Second.Ordinal) : PlainTime := + { pt with second := second } + +/-- +Creates a new `PlainTime` by adjusting the `minute` component to the given value. +-/ +@[inline] +def withMinute (pt : PlainTime) (minute : Minute.Ordinal) : PlainTime := + { pt with minute := minute } + +/-- +Creates a new `PlainTime` by adjusting the `nano` component to the given value. +-/ +@[inline] +def withNano (pt : PlainTime) (nano : Nanosecond.Ordinal) : PlainTime := + { pt with nano := nano } + +/-- +Creates a new `PlainTime` by adjusting the `hour` component to the given value. +-/ +@[inline] +def withHour (pt : PlainTime) (hour : Hour.Ordinal) : PlainTime := + { pt with hour := hour } + instance : HAdd PlainTime Nanosecond.Offset PlainTime where hAdd := addNanoseconds diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 179a82765b85..f10c3a786b96 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -220,6 +220,82 @@ Subtract `Year.Offset` from to a `DateTime`, it clips the day to the last valid def subYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := ofLocalDateTime (dt.date.get.subYearsClip years) tz +/-- +Creates a new `DateTime tz` by adjusting the day of the month to the given `days` value, with any +out-of-range days clipped to the nearest valid date. +-/ +@[inline] +def withDaysClip (dt : DateTime tz) (days : Day.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withDaysClip days) tz + +/-- +Creates a new `DateTime tz` by adjusting the day of the month to the given `days` value, with any +out-of-range days rolled over to the next month or year as needed. +-/ +@[inline] +def withDaysRollOver (dt : DateTime tz) (days : Day.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withDaysRollOver days) tz + +/-- +Creates a new `DateTime tz` by adjusting the month to the given `month` value. +The day remains unchanged, and any invalid days for the new month will be handled according to the `clip` behavior. +-/ +@[inline] +def withMonthClip (dt : DateTime tz) (month : Month.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withMonthClip month) tz + +/-- +Creates a new `DateTime tz` by adjusting the month to the given `month` value. +The day is rolled over to the next valid month if necessary. +-/ +@[inline] +def withMonthRollOver (dt : DateTime tz) (month : Month.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withMonthRollOver month) tz + +/-- +Creates a new `DateTime tz` by adjusting the year to the given `year` value. The month and day remain unchanged, +and any invalid days for the new year will be handled according to the `clip` behavior. +-/ +@[inline] +def withYearClip (dt : DateTime tz) (year : Year.Offset) : DateTime tz := + .ofLocalDateTime (dt.date.get.withYearClip year) tz + +/-- +Creates a new `DateTime tz` by adjusting the year to the given `year` value. The month and day are rolled +over to the next valid month and day if necessary. +-/ +@[inline] +def withYearRollOver (dt : DateTime tz) (year : Year.Offset) : DateTime tz := + .ofLocalDateTime (dt.date.get.withYearRollOver year) tz + +/-- +Creates a new `DateTime tz` by adjusting the `hour` component. +-/ +@[inline] +def withHour (dt : DateTime tz) (hour : Hour.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withHour hour) tz + +/-- +Creates a new `DateTime tz` by adjusting the `minute` component. +-/ +@[inline] +def withMinute (dt : DateTime tz) (minute : Minute.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withMinute minute) tz + +/-- +Creates a new `DateTime tz` by adjusting the `second` component. +-/ +@[inline] +def withSecond (dt : DateTime tz) (second : Sigma Second.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withSecond second) tz + +/-- +Creates a new `DateTime tz` by adjusting the `nano` component. +-/ +@[inline] +def withNano (dt : DateTime tz) (nano : Nanosecond.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withNano nano) tz + /-- Converts a `Timestamp` to a `PlainDateTime` -/ diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 1244028d0e73..d6480f0e08d1 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -296,6 +296,82 @@ def era (date : ZonedDateTime) : Year.Era := else .bce +/-- +Creates a new `ZonedDateTime` by adjusting the day of the month to the given `days` value, with any +out-of-range days clipped to the nearest valid date. +-/ +@[inline] +def withDaysClip (dt : ZonedDateTime) (days : Day.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withDaysClip days⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the day of the month to the given `days` value, with any +out-of-range days rolled over to the next month or year as needed. +-/ +@[inline] +def withDaysRollOver (dt : ZonedDateTime) (days : Day.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withDaysRollOver days⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the month to the given `month` value. +The day remains unchanged, and any invalid days for the new month will be handled according to the `clip` behavior. +-/ +@[inline] +def withMonthClip (dt : ZonedDateTime) (month : Month.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withMonthClip month⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the month to the given `month` value. +The day is rolled over to the next valid month if necessary. +-/ +@[inline] +def withMonthRollOver (dt : ZonedDateTime) (month : Month.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withMonthRollOver month⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the year to the given `year` value. The month and day remain unchanged, +and any invalid days for the new year will be handled according to the `clip` behavior. +-/ +@[inline] +def withYearClip (dt : ZonedDateTime) (year : Year.Offset) : ZonedDateTime := + ⟨dt.fst, dt.snd.withYearClip year⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the year to the given `year` value. The month and day are rolled +over to the next valid month and day if necessary. +-/ +@[inline] +def withYearRollOver (dt : ZonedDateTime) (year : Year.Offset) : ZonedDateTime := + ⟨dt.fst, dt.snd.withYearRollOver year⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the `hour` component. +-/ +@[inline] +def withHour (dt : ZonedDateTime) (hour : Hour.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withHour hour⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the `minute` component. +-/ +@[inline] +def withMinute (dt : ZonedDateTime) (minute : Minute.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withMinute minute⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the `second` component. +-/ +@[inline] +def withSecond (dt : ZonedDateTime) (second : Sigma Second.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withSecond second⟩ + +/-- +Creates a new `ZonedDateTime` by adjusting the `nano` component. +-/ +@[inline] +def withNano (dt : ZonedDateTime) (nano : Nanosecond.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withNano nano⟩ + /-- Checks if the `DateTime` is in a leap year. -/ diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 06e9028f7369..f843a5a818c1 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -9,10 +9,10 @@ def ShortDateTime : Format .any := datespec("MM/dd/yyyy HH:mm:ss") def LongDateTime : Format .any := datespec("MMMM dd, yyyy hh:mm aa") def Time24Hour : Format .any := datespec("HH:mm:ss") def Time12Hour : Format .any := datespec("hh:mm:ss aa") -def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss XXX") +def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") def CustomDayTime : Format .any := datespec("EEE dd MMM yyyy HH:mm") -def Full12HourWrong : Format .any := datespec("MM/dd/yyyy HH:mm:ss aa XXX") +def Full12HourWrong : Format .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") -- Dates @@ -85,7 +85,7 @@ info: "2024-08-16T01:28:00.000000000Z" ISO8601UTC.format t.snd /-- -info: "-0001-12-31T22:28:12.000+09:00" +info: "0000-12-30T22:28:12.000000000+09:00" -/ #guard_msgs in #eval @@ -93,11 +93,11 @@ info: "-0001-12-31T22:28:12.000+09:00" ISO8601UTC.format (t.snd.convertTimeZone jpTZ) /-- -info: "-0001-12-31T09:28:12.000-03:00" +info: "0000-12-29T21:28:12.000000000-03:00" -/ #guard_msgs in #eval - let t1 : ZonedDateTime := Time12Hour.parse! "12:28:12 am" + let t1 : ZonedDateTime := Time12Hour.parse! "12:28:12 AM" ISO8601UTC.format (t1.snd.convertTimeZone brTZ) /-- @@ -109,7 +109,7 @@ info: "Thu 15 Aug 2024 16:28" CustomDayTime.format t2.snd /-- -info: "2024-08-16T13:28:00.000Z" +info: "2024-08-16T13:28:00.000000000Z" -/ #guard_msgs in #eval @@ -117,7 +117,7 @@ info: "2024-08-16T13:28:00.000Z" ISO8601UTC.format t5.snd /-- -info: "2024-08-16T01:28:12.000+09:00" +info: "2024-08-16T01:28:12.000000000+09:00" -/ #guard_msgs in #eval @@ -125,7 +125,7 @@ info: "2024-08-16T01:28:12.000+09:00" ISO8601UTC.format (t6.snd.convertTimeZone jpTZ) /-- -info: "2024-08-16T01:28:12.000+09:00" +info: "2024-08-16T01:28:12.000000000+09:00" -/ #guard_msgs in #eval @@ -142,7 +142,7 @@ def localTm : Second.Offset := 1723730627 /-- This PlainDate is relative to the local time. -/ -def PlainDate : PlainDateTime := Timestamp.toPlainDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) +def PlainDate : PlainDateTime := Timestamp.toPlainDateTimeAssumingUTC (Timestamp.ofSecondsSinceUnixEpoch localTm) def dateBR₁ := DateTime.ofLocalDateTime PlainDate brTZ def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ @@ -181,7 +181,7 @@ info: Except.ok "1993-05-10T10:30:23.000000000+03:00" (ISO8601UTC.format ·.snd) <$> t2 /-- -info: Except.ok "1993-05-10T22:30:23.000+03:00" +info: Except.ok "1993-05-10T22:30:23.000000000+03:00" -/ #guard_msgs in #eval @@ -189,7 +189,7 @@ info: Except.ok "1993-05-10T22:30:23.000+03:00" (ISO8601UTC.format ·.snd) <$> t2 /-- -info: Except.error "offset 29: The 24-hour is out of the range and cannot be transformed into a 12-hour with a marker." +info: Except.error "offset 13: need a natural number in the interval of 1 to 12" -/ #guard_msgs in #eval @@ -197,7 +197,7 @@ info: Except.error "offset 29: The 24-hour is out of the range and cannot be tra (ISO8601UTC.format ·.snd) <$> t2 /-- -info: Except.error "offset 29: The 24-hour is out of the range and cannot be transformed into a 12-hour with a marker." +info: Except.error "offset 13: need a natural number in the interval of 1 to 12" -/ #guard_msgs in #eval diff --git a/tests/lean/run/timeSet.lean b/tests/lean/run/timeSet.lean new file mode 100644 index 000000000000..13ffd31ba8c2 --- /dev/null +++ b/tests/lean/run/timeSet.lean @@ -0,0 +1,93 @@ +import Std.Time +open Std.Time + +def ISO8601UTC : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") +def RFC1123 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def ShortDate : Format .any := datespec("MM/dd/yyyy") +def LongDate : Format .any := datespec("MMMM D, yyyy") +def ShortDateTime : Format .any := datespec("MM/dd/yyyy HH:mm:ss") +def LongDateTime : Format .any := datespec("MMMM dd, yyyy hh:mm aa") +def Time24Hour : Format .any := datespec("HH:mm:ss") +def Time12Hour : Format .any := datespec("hh:mm:ss aa") +def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") +def CustomDayTime : Format .any := datespec("EEE dd MMM yyyy HH:mm") + +def Full12HourWrong : Format .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") + +-- Dates + +def brTZ : TimeZone := timezone("America/Sao_Paulo -03:00") +def jpTZ : TimeZone := timezone("Asia/Tokyo +09:00") + +def date₁ := zoned("2014-06-16T03:03:03-03:00") + +def time₁ := time("14:11:01") +def time₂ := time("03:11:01") + +/-- +info: "2014-06-16T03:03:03.000000100-03:00" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03.000000100-03:00" + ISO8601UTC.format t.snd + +/-- +info: zoned("2014-06-30T03:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withDaysClip 31 + +/-- +info: zoned("2014-07-01T03:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withDaysRollOver 31 + +/-- +info: zoned("2014-05-16T03:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withMonthClip 5 + +/-- +info: zoned("2014-05-16T03:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withMonthRollOver 5 + +/-- +info: zoned("2016-06-16T03:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withYearClip 2016 + +/-- +info: zoned("2016-06-16T03:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withYearRollOver 2016 + +/-- +info: zoned("2014-06-16T19:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withHour 19 + +/-- +info: zoned("2014-06-16T03:45:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withMinute 45 + +/-- +info: zoned("2014-06-16T03:03:59.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withSecond ⟨true, 59⟩ + +/-- +info: zoned("2014-06-16T03:03:03.000000002-03:00") +-/ +#guard_msgs in +#eval date₁.withNano 2 From 594aa87ee04a1da9dd9df664032e0ad9d5c15eb8 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 20 Sep 2024 14:10:24 -0300 Subject: [PATCH 066/118] fix: deprecations --- src/Std/Time/Date/PlainDate.lean | 16 ++++++++-------- src/Std/Time/Date/Unit/Year.lean | 2 +- src/Std/Time/Duration.lean | 8 ++++---- src/Std/Time/Internal/Bounded.lean | 8 ++++---- src/Std/Time/Notation/Spec.lean | 3 +-- src/Std/Time/Zoned/Offset.lean | 2 +- 6 files changed, 19 insertions(+), 20 deletions(-) diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 102c3cd6aa63..4c47ae880d56 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -78,13 +78,13 @@ Creates a `PlainDate` from the number of days since the UNIX epoch (January 1st, -/ def ofDaysSinceUNIXEpoch (day : Day.Offset) : PlainDate := let z := day.toInt + 719468 - let era := (if z ≥ 0 then z else z - 146096).div 146097 + let era := (if z ≥ 0 then z else z - 146096).tdiv 146097 let doe := z - era * 146097 - let yoe := (doe - doe.div 1460 + doe.div 36524 - doe.div 146096).div 365 + let yoe := (doe - doe.tdiv 1460 + doe.tdiv 36524 - doe.tdiv 146096).tdiv 365 let y := yoe + era * 400 - let doy := doe - (365 * yoe + yoe.div 4 - yoe.div 100) - let mp := (5 * doy + 2).div 153 - let d := doy - (153 * mp + 2).div 5 + 1 + let doy := doe - (365 * yoe + yoe.tdiv 4 - yoe.tdiv 100) + let mp := (5 * doy + 2).tdiv 153 + let d := doy - (153 * mp + 2).tdiv 5 + 1 let m := mp + (if mp < 10 then 3 else -9) let y := y + (if m <= 2 then 1 else 0) .clip y (.clip m (by decide)) (.clip d (by decide)) @@ -131,12 +131,12 @@ Converts a `PlainDate` to the number of days since the UNIX epoch. -/ def toDaysSinceUNIXEpoch (date : PlainDate) : Day.Offset := let y : Int := if date.month.toInt > 2 then date.year else date.year.toInt - 1 - let era : Int := (if y ≥ 0 then y else y - 399).div 400 + let era : Int := (if y ≥ 0 then y else y - 399).tdiv 400 let yoe : Int := y - era * 400 let m : Int := date.month.toInt let d : Int := date.day.toInt - let doy := (153 * (m + (if m > 2 then -3 else 9)) + 2).div 5 + d - 1 - let doe := yoe * 365 + yoe.div 4 - yoe.div 100 + doy + let doy := (153 * (m + (if m > 2 then -3 else 9)) + 2).tdiv 5 + d - 1 + let doe := yoe * 365 + yoe.tdiv 4 - yoe.tdiv 100 + doy .ofInt (era * 146097 + doe - 719468) diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index 4cf634a8281e..33d8234ac12c 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -74,7 +74,7 @@ Determines if a year is a leap year in the proleptic Gregorian calendar. -/ @[inline] def isLeap (y : Offset) : Bool := - y.toInt.mod 4 = 0 ∧ (y.toInt.mod 100 ≠ 0 ∨ y.toInt.mod 400 = 0) + y.toInt.tmod 4 = 0 ∧ (y.toInt.tmod 100 ≠ 0 ∨ y.toInt.tmod 400 = 0) /-- Returns the `Era` of the `Year`. diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 16f1d1770d18..f54bf09e80c3 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -87,11 +87,11 @@ def ofNanoseconds (s : Nanosecond.Offset) : Duration := by refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ cases Int.le_total s.val 0 next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) - next n => exact Or.inl (And.intro (Int.ediv_nonneg n (by decide)) (Int.mod_nonneg 1000000000 n)) + next n => exact Or.inl (And.intro (Int.ediv_nonneg n (by decide)) (Int.tmod_nonneg 1000000000 n)) where - mod_nonpos : ∀ {a : Int} (b : Int), (a ≤ 0) → (b ≥ 0) → 0 ≥ a.mod b - | .negSucc m, .ofNat n, _, _ => Int.neg_le_neg (Int.mod_nonneg (↑n) (Int.ofNat_le.mpr (Nat.zero_le (m + 1)))) - | 0, n, _, _ => Int.eq_iff_le_and_ge.mp (Int.zero_mod n) |>.left + mod_nonpos : ∀ {a : Int} (b : Int), (a ≤ 0) → (b ≥ 0) → 0 ≥ a.tmod b + | .negSucc m, .ofNat n, _, _ => Int.neg_le_neg (Int.tmod_nonneg (↑n) (Int.ofNat_le.mpr (Nat.zero_le (m + 1)))) + | 0, n, _, _ => Int.eq_iff_le_and_ge.mp (Int.zero_tmod n) |>.left /-- Checks if the duration is zero seconds and zero nanoseconds. diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index a2ad9a8fdc1f..dc7223c24f55 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -224,8 +224,8 @@ Creates a new `Bounded.LE` using a the Truncating modulus of a number. -/ @[inline] def byMod (b : Int) (i : Int) (hi : 0 < i) : Bounded.LE (- (i - 1)) (i - 1) := by - refine ⟨b.mod i, And.intro ?_ ?_⟩ - · simp [Int.mod] + refine ⟨b.tmod i, And.intro ?_ ?_⟩ + · simp [Int.tmod] split <;> try contradiction next m n => let h := Int.emod_nonneg (a := m) (b := n) (Int.ne_of_gt hi) @@ -235,9 +235,9 @@ def byMod (b : Int) (i : Int) (hi : 0 < i) : Bounded.LE (- (i - 1)) (i - 1) := b exact (Int.le_sub_one_of_lt hi) next m n => apply Int.neg_le_neg - have h := Int.mod_lt_of_pos (m + 1) hi + have h := Int.tmod_lt_of_pos (m + 1) hi exact Int.le_sub_one_of_lt h - · exact Int.le_sub_one_of_lt (Int.mod_lt_of_pos b hi) + · exact Int.le_sub_one_of_lt (Int.tmod_lt_of_pos b hi) /-- Adjust the bounds of a `Bounded` by setting the lower bound to zero and the maximum value to (m - n). diff --git a/src/Std/Time/Notation/Spec.lean b/src/Std/Time/Notation/Spec.lean index f6aefc82ff82..c00f7105cb65 100644 --- a/src/Std/Time/Notation/Spec.lean +++ b/src/Std/Time/Notation/Spec.lean @@ -9,11 +9,10 @@ import Std.Time.Time import Std.Time.Zoned import Std.Time.DateTime import Std.Time.Format.Basic -import Lean.Parser namespace Std namespace Time -open Lean Parser Command Std Lean.Parser +open Lean Parser Command Std private def convertText : Text → MacroM (TSyntax `term) | .short => `(Std.Time.Text.short) diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 5b53d6730bdc..70e60ff5e234 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -38,7 +38,7 @@ and minute components are separated by a colon (e.g., "+01:00" or "+0100"). def toIsoString (offset : Offset) (colon : Bool) : String := let (sign, time) := if offset.second.val > 0 then ("+", offset.second) else ("-", -offset.second) let hour : Hour.Offset := time.ediv 3600 - let minute := Int.ediv (Int.mod time.val 3600) 60 + let minute := Int.ediv (Int.tmod time.val 3600) 60 let hourStr := if hour.val < 10 then s!"0{hour.val}" else toString hour.val let minuteStr := if minute < 10 then s!"0{minute}" else toString minute if colon then s!"{sign}{hourStr}:{minuteStr}" From 5367936e2f8c99105c43bb1546d12ee119fbc1b7 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 20 Sep 2024 14:12:54 -0300 Subject: [PATCH 067/118] feat: remove lean parser from notation --- src/Std/Time/Notation.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 820e9d2e2b00..60ee6df8b127 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -9,11 +9,10 @@ import Std.Time.Time import Std.Time.Zoned import Std.Time.DateTime import Std.Time.Format -import Lean.Parser namespace Std namespace Time -open Lean Parser Command Std Lean.Parser +open Lean Parser Command Std private def convertText : Text → MacroM (TSyntax `term) | .short => `(Std.Time.Text.short) From 18feda10e7ecf0a3e4487e7a4f61f0e6160ef10c Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 20 Sep 2024 14:29:02 -0300 Subject: [PATCH 068/118] fix: rename div to tdiv --- src/Std/Time/Internal/UnitVal.lean | 2 +- src/Std/Time/Zoned/DateTime.lean | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 2d19629f206e..b16ba0bc9b2d 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -68,7 +68,7 @@ Divides the `UnitVal` by an `Int`, resulting in a new `UnitVal` with an adjusted -/ @[inline] def div (unit : UnitVal a) (divisor : Int) : UnitVal (a * divisor) := - ⟨unit.val.div divisor⟩ + ⟨unit.val.tdiv divisor⟩ /-- Adds two `UnitVal` values of the same ratio. diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index f10c3a786b96..e7949a8e6f4e 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -70,8 +70,8 @@ def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : DateTime DateTime.mk tm (Thunk.mk fun _ => date.addSeconds tz.toSeconds) /-- -Creates a new `DateTime` out of a `PlainDateTime`. It assumes that the `PlainDateTime` it's relative -to the +Creates a new `DateTime` from a `PlainDateTime`, assuming that the `PlainDateTime` +is relative to the given `TimeZone`. -/ @[inline] def ofLocalDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := From b213fc7e65abed01402b92d70f4095becd8b7535 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 20 Sep 2024 14:44:08 -0300 Subject: [PATCH 069/118] fix: tests --- tests/lean/run/timeFormats.lean | 2 +- tests/lean/run/timeTzifParse.lean | 6 +++--- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 8e72a2873f43..9e638ab4e352 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -131,7 +131,7 @@ def localTm : Second.Offset := 1723730627 /-- This PlainDate is relative to the local time. -/ -def PlainDate : PlainDateTime := Timestamp.toPlainDateTime (Timestamp.ofSecondsSinceUnixEpoch localTm) +def PlainDate : PlainDateTime := Timestamp.toPlainDateTimeAssumingUTC (Timestamp.ofSecondsSinceUnixEpoch localTm) def dateBR₁ := DateTime.ofLocalDateTime PlainDate brTZ def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index efd25062834d..10ac2ca911d6 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -44,7 +44,7 @@ info: 91 info: 3 -/ #guard_msgs in -#eval code.v1.PlainTimeTypes.size +#eval code.v1.localTimeTypes.size /-- info: 0 @@ -59,13 +59,13 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: some (date% 1969-12-30T21:00:00,000000000-03:00) +info: some (zoned("1969-12-30T21:00:00.000000000-03:00")) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules /-- -info: some (date% 2012-12-10T00:35:47,000000000-02:00) +info: some (zoned("2012-12-10T00:35:47.000000000-02:00")) -/ #guard_msgs in #eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From 9e74aac3aeb6b79b2ca0e3abc39d64f42559fbe2 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 25 Sep 2024 23:23:55 -0300 Subject: [PATCH 070/118] fix: a bunch of comments and inlines --- src/Std/Time/Date/PlainDate.lean | 22 ++++----- src/Std/Time/Date/Unit/Day.lean | 58 +++++++++++++----------- src/Std/Time/Date/Unit/Month.lean | 18 ++++---- src/Std/Time/Date/Unit/Week.lean | 6 ++- src/Std/Time/Date/Unit/Weekday.lean | 13 ++++-- src/Std/Time/Date/Unit/Year.lean | 12 ++--- src/Std/Time/DateTime.lean | 14 +++--- src/Std/Time/DateTime/PlainDateTime.lean | 20 ++++---- src/Std/Time/DateTime/Timestamp.lean | 3 ++ src/Std/Time/Duration.lean | 26 ++++++++++- src/Std/Time/Format/Basic.lean | 30 ++++++------ src/Std/Time/Internal/Bounded.lean | 18 +++++--- src/Std/Time/Internal/UnitVal.lean | 10 ++++ src/Std/Time/Notation/Spec.lean | 4 +- src/Std/Time/Time/PlainTime.lean | 6 +++ src/Std/Time/Zoned/DateTime.lean | 19 +++----- 16 files changed, 163 insertions(+), 116 deletions(-) diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 4c47ae880d56..f0b191405d4f 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -22,26 +22,19 @@ and day components, with validation to ensure the date is valid. -/ structure PlainDate where - /-- - The year component of the date. It is represented as an `Offset` type from `Year`. - -/ + /-- The year component of the date. It is represented as an `Offset` type from `Year`. -/ year : Year.Offset - /-- - The month component of the date. It is represented as an `Ordinal` type from `Month`. - -/ + /-- The month component of the date. It is represented as an `Ordinal` type from `Month`. -/ month : Month.Ordinal - /-- - The day component of the date. It is represented as an `Ordinal` type from `Day`. - -/ + /-- The day component of the date. It is represented as an `Ordinal` type from `Day`. -/ day : Day.Ordinal - /-- - Validates the date by ensuring that the year, month, and day form a correct and valid date. - -/ + /-- Validates the date by ensuring that the year, month, and day form a correct and valid date. -/ valid : year.Valid month day + instance : BEq PlainDate where beq x y := x.day == y.day && x.month == y.month && x.year == y.year @@ -51,6 +44,7 @@ namespace PlainDate Creates a `PlainDate` by clipping the day to ensure validity. This function forces the date to be valid by adjusting the day to fit within the valid range to fit the given month and year. -/ +@[inline] def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := let day := month.clipDay year.isLeap day PlainDate.mk year month day Month.Ordinal.clipDay_valid @@ -61,6 +55,7 @@ instance : Inhabited PlainDate where /-- Creates a new `PlainDate` from year, month, and day components. -/ +@[inline] def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option PlainDate := if valid : year.Valid month day then some (PlainDate.mk year month day valid) @@ -69,6 +64,7 @@ def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordin /-- Creates a `PlainDate` from a year and a day ordinal within that year. -/ +@[inline] def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : PlainDate := let ⟨⟨month, day⟩, proof⟩ := Month.Ordinal.ofOrdinal ordinal ⟨year, month, day, proof⟩ @@ -123,6 +119,7 @@ def era (date : PlainDate) : Year.Era := /-- Checks if the `PlainDate` is in a leap year. -/ +@[inline] def inLeapYear (date : PlainDate) : Bool := date.year.isLeap @@ -143,6 +140,7 @@ def toDaysSinceUNIXEpoch (date : PlainDate) : Day.Offset := /-- Calculates the difference in years between a `PlainDate` and a given year. -/ +@[inline] def yearsSince (date : PlainDate) (year : Year.Offset) : Year.Offset := date.year - year diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 79f6b4af3572..751ff3c05676 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -34,26 +34,17 @@ instance {x y : Ordinal} : Decidable (x < y) := instance : Inhabited Ordinal where default := 1 /-- -`Offset` represents an offset in days. It is defined as an `Int` with a base unit of 86400 (the number of seconds in a day). -This type supports arithmetic operations like addition, subtraction, multiplication, and division, and also comparisons like less than or equal. +`Offset` represents an offset in days. It is defined as an `Int` with a base unit of 86400 +(the number of seconds in a day). -/ def Offset : Type := UnitVal 86400 deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString -/-- -Provides an instance for creating an `Offset` from a natural number (`OfNat`), converting the input to the base unit (days). --/ instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ -/-- -Provides a decidable instance to check if one `Offset` is less than or equal to another. --/ instance {x y : Offset} : Decidable (x ≤ y) := inferInstanceAs (Decidable (x.val ≤ y.val)) -/-- -Provides a decidable instance to check if one `Offset` is strictly less than another. --/ instance {x y : Offset} : Decidable (x < y) := inferInstanceAs (Decidable (x.val < y.val)) @@ -78,33 +69,50 @@ def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ (if leap then 366 else 365) end OfYear /-- -`Period` is an enumeration representing different times of the day: morning, afternoon, evening, and night. +`Period` is an enumeration representing different times of the day : morning, afternoon, evening, and night. -/ inductive Period - /-- Represents the morning period. -/ + /--Represents the morning period. -/ | morning - /-- Represents the afternoon period. -/ + /--Represents the afternoon period. -/ | afternoon - /-- Represents the evening period. -/ + /--Represents the evening period. -/ | evening - /-- Represents the night period. -/ + /--Represents the night period. -/ | night + deriving Repr, BEq, Inhabited + +namespace Period /-- -Instance to allow creation of an `Ordinal.OfYear` from a natural number, ensuring the value is -within the bounds of the year, which depends on whether it's a leap year or not. +Determines the `Period` of the day based on the given hour + +- If the hour is between 20 and 4, it returns `night`. +- If the hour is between 17 and 20, it returns `evening`. +- If the hour is between 12 and 17, it returns `afternoon`. +- If the hour is between 5 and 12, it reutrns `morning`. -/ +@[inline] +def fromHour (hour : Hour.Ordinal) : Day.Ordinal.Period := + if hour ≥ 20 ∨ hour ≤ 4 then + .night + else if hour ≥ 17 then + .evening + else if hour ≥ 12 then + .afternoon + else + .morning + +end Period + instance : OfNat (Ordinal.OfYear leap) n := match leap with | true => inferInstanceAs (OfNat (Bounded.LE 1 (1 + (365 : Nat))) n) | false => inferInstanceAs (OfNat (Bounded.LE 1 (1 + (364 : Nat))) n) -/-- -Provides a default value for `Ordinal.OfYear`, defaulting to day 1. --/ instance : Inhabited (Ordinal.OfYear leap) where default := by refine ⟨1, And.intro (by decide) ?_⟩ @@ -138,8 +146,8 @@ namespace OfYear /-- Converts an `OfYear` ordinal to a `Offset`. -/ -def toOffset (of: OfYear leap) : Offset := - UnitVal.mk of.val +def toOffset (ofYear : OfYear leap) : Offset := + UnitVal.mk ofYear.val end OfYear end Ordinal @@ -153,10 +161,6 @@ Converts an `Ordinal` to an `Offset`. def toOrdinal (off : Offset) (h : off.val ≥ 1 ∧ off.val ≤ 31) : Ordinal := Bounded.LE.mk off.val h -theorem toOffset_toOrdinal {d : Ordinal} : ∃h, d.toOffset.toOrdinal h = d := by - simp [Ordinal.toOffset, toOrdinal, Bounded.LE.mk, UnitVal.ofInt] - exists d.property - /-- Creates an `Offset` from a natural number. -/ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index be9ac02cec6a..6a60a5488d65 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -39,15 +39,16 @@ def Offset : Type := Int instance : OfNat Offset n := ⟨Int.ofNat n⟩ -namespace Ordinal /-- -`OfYear` represents the number of days in a year, accounting for leap years. It ensures that -the day is within the correct bounds—either 1 to 365 for regular years or 1 to 366 for leap years. +`Quarter` represents a value between 1 and 4, inclusive, corresponding to the four quarters of a year. -/ def Quarter := Bounded.LE 1 4 -end Ordinal +namespace Quarter + +end Quarter + namespace Offset /-- @@ -205,14 +206,14 @@ def toDays (leap : Bool) (month : Ordinal) : Day.Offset := Size in days of each month if the year is not a leap year. -/ @[inline] -def monthSizesNonLeap : { val : Array Day.Ordinal // val.size = 12 } := +private def monthSizesNonLeap : { val : Array Day.Ordinal // val.size = 12 } := ⟨#[31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31], by decide⟩ /-- -Size in days of each month if the year is not a leap year. +Returns the cumulative size in days of each month for a non-leap year. -/ @[inline] -def cumulativeSizes : { val : Array Day.Offset // val.size = 12 } := +private def cumulativeSizes : { val : Array Day.Offset // val.size = 12 } := ⟨#[0, 31, 59, 90, 120, 151, 181, 212, 243, 273, 304, 334], by decide⟩ /-- @@ -238,8 +239,7 @@ theorem days_gt_27 (leap : Bool) (i : Month.Ordinal) : days leap i > 27 := by decide /-- -Gets the number of days in a month without requiring a proof of the validity of the ordinal in a -month and year. +Returns the number of days until the `month`. -/ def cumulativeDays (leap : Bool) (month : Ordinal) : Day.Offset := by let ⟨months, p⟩ := cumulativeSizes diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 5e3bb35294da..84e0e1475757 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -31,7 +31,7 @@ instance : Inhabited Ordinal where default := 1 /-- -`Offset` represents an offset in weeks. It is defined as an `Int`. +`Offset` represents an offset in weeks. -/ def Offset : Type := UnitVal (86400 * 7) deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString @@ -46,22 +46,24 @@ correct bounds—either 1 to 6, representing the possible weeks in a month. -/ def OfMonth := Bounded.LE 1 6 - /-- Creates an `Ordinal` from a natural number, ensuring the value is within bounds. -/ +@[inline] def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ 53 := by decide) : Ordinal := Bounded.LE.ofNat' data h /-- Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds. -/ +@[inline] def ofFin (data : Fin 54) : Ordinal := Bounded.LE.ofFin' data (by decide) /-- Converts an `Ordinal` to an `Offset`. -/ +@[inline] def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index 976fd02dfeed..2a6437859572 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -15,22 +15,27 @@ open Internal set_option linter.all true /-- -Defines the enumeration for days of the week. Each variant corresponds to a day of the week, from Monday to Sunday. -It starts on the sunday because the library is based on the Gregorian Calendar. +Defines the enumeration for days of the week. Each variant corresponds to a day of the week. -/ inductive Weekday /-- Sunday. -/ | sunday + /-- Monday. -/ | monday + /-- Tuesday. -/ | tuesday + /-- Wednesday. -/ | wednesday + /-- Thursday. -/ | thursday + /-- Friday. -/ | friday + /-- Saturday. -/ | saturday deriving Repr, Inhabited, BEq @@ -59,7 +64,7 @@ def ofOrdinal : Ordinal → Weekday | 7 => .saturday /-- -Converts a `Weekday` to a `Nat`. +Converts a `Weekday` to a `Ordinal`. -/ def toOrdinal : Weekday → Ordinal | .sunday => 1 @@ -102,7 +107,7 @@ Converts a `Nat` to a `Weekday`. Panics if the value provided is invalid. def ofNat! (n : Nat) : Weekday := match ofNat? n with | some res => res - | none => panic! "invalid weekday" + | none => panic! "invalid weekday" /-- Gets the next `Weekday`. diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index 33d8234ac12c..b594ba752761 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -21,14 +21,10 @@ set_option linter.all true Defines the different eras. -/ inductive Era - /-- - The era before the Common Era (BCE), always represents a date before year 0. - -/ + /-- The era before the Common Era (BCE), always represents a date before year 0. -/ | bce - /-- - The Common Era (CE), represents dates from year 0 onwards. - -/ + /-- The Common Era (CE), represents dates from year 0 onwards. -/ | ce /-- @@ -63,7 +59,7 @@ def toInt (offset : Offset) : Int := offset /-- -Converts the `Year` offset to a `Month` offset (i.e., multiplies by 12). +Converts the `Year` offset to a `Month` offset. -/ @[inline] def toMonths (val : Offset) : Month.Offset := @@ -88,7 +84,7 @@ def era (year : Offset) : Era := Checks if the given date is valid for the specified year, month, and day. -/ @[inline] -abbrev Valid (year : Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := +abbrev Valid (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := day ≤ month.days year.isLeap instance : Decidable (Valid year month day) := diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 1da4b00f1482..93b5a1a4d7ab 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -58,7 +58,7 @@ def ofPlainTime (pt : PlainTime) : Timestamp := Converts a `Timestamp` to a `PlainTime` -/ @[inline] -def toPlainTime (timestamp : Timestamp) : PlainTime := +def getTime (timestamp : Timestamp) : PlainTime := let nanos := timestamp.toNanosecondsSinceUnixEpoch PlainTime.ofNanoseconds nanos @@ -100,29 +100,29 @@ namespace PlainDateTime Converts a `PlainDate` to a `Timestamp` -/ @[inline] -def ofPlainDate (pd : PlainDate) : PlainDateTime := - PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainDateAssumingUTC pd) +def ofPlainDate (date : PlainDate) : PlainDateTime := + { date, time := PlainTime.midnight } /-- Converts a `PlainDateTime` to a `PlainDate` -/ @[inline] def toPlainDate (pdt : PlainDateTime) : PlainDate := - Timestamp.toPlainDateAssumingUTC pdt.toTimestampAssumingUTC + pdt.date /-- Converts a `PlainTime` to a `PlainDateTime` -/ @[inline] -def ofPlainTime (pt : PlainTime) : PlainDateTime := - PlainDateTime.ofUTCTimestamp (Timestamp.ofPlainTime pt) +def ofPlainTime (time : PlainTime) : PlainDateTime := + { date := ⟨1, 1, 1, by decide⟩, time } /-- Converts a `PlainDateTime` to a `PlainTime` -/ @[inline] def toPlainTime (pdt : PlainDateTime) : PlainTime := - Timestamp.toPlainTime pdt.toTimestampAssumingUTC + pdt.time instance : ToTimestamp PlainDateTime where toTimestamp := Timestamp.ofPlainDateTimeAssumingUTC diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 179f816f49fb..2a1eb3d4f019 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -39,7 +39,7 @@ Converts a `PlainDateTime` to a `Timestamp` -/ def toTimestampAssumingUTC (dt : PlainDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch - let nanos : Nanosecond.Offset := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 + let nanos := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 let nanos := nanos.val + dt.time.nano.val Timestamp.ofNanosecondsSinceUnixEpoch (Nanosecond.Offset.ofInt nanos) @@ -402,6 +402,13 @@ Getter for the `Second` inside of a `PlainDateTime`. def second (dt : PlainDateTime) : Second.Ordinal dt.time.second.fst := dt.time.second.snd +/-- +Getter for the `Nanosecond.Ordinal` inside of a `PlainDateTime`. +-/ +@[inline] +def nanosecond (dt : PlainDateTime) : Nanosecond.Ordinal := + dt.time.nano + /-- Determines the era of the given `PlainDateTime` based on its year. -/ @@ -414,16 +421,10 @@ def era (date : PlainDateTime) : Year.Era := /-- Checks if the `PlainDateTime` is in a leap year. -/ +@[inline] def inLeapYear (date : PlainDateTime) : Bool := date.year.isLeap -/-- -Getter for the `Second` inside of a `PlainDateTime`. --/ -@[inline] -def nanosecond (dt : PlainDateTime) : Nanosecond.Ordinal := - dt.time.nano - instance : HAdd PlainDateTime Day.Offset PlainDateTime where hAdd := addDays @@ -460,6 +461,9 @@ instance : HAdd PlainDateTime Nanosecond.Offset PlainDateTime where instance : HSub PlainDateTime Nanosecond.Offset PlainDateTime where hSub := subNanoseconds +instance : HAdd PlainDateTime Duration PlainDateTime where + hAdd x y := addNanoseconds x y.toNanoseconds + end PlainDateTime end Time end Std diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 43d39522b264..36bf380c1e1d 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -196,6 +196,9 @@ Subtracts a `Week.Offset` from a `Timestamp` def subWeeks (t : Timestamp) (d : Week.Offset) : Timestamp := ⟨t.val - d⟩ +instance : HAdd Timestamp Duration Timestamp where + hAdd x y := ⟨x.val + y⟩ + instance : HAdd Timestamp Day.Offset Timestamp where hAdd := addDays diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index f54bf09e80c3..586439f4eb49 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -82,7 +82,6 @@ def ofSeconds (s : Second.Offset) : Duration := by /-- Creates a new `Duration` out of `Nanosecond.Offset`. -/ -@[inline] def ofNanoseconds (s : Nanosecond.Offset) : Duration := by refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ cases Int.le_total s.val 0 @@ -285,6 +284,31 @@ instance : HSub Duration Duration Duration where instance : HAdd Duration Duration Duration where hAdd := add +instance : Coe Nanosecond.Offset Duration where + coe := ofNanoseconds + +instance : Coe Second.Offset Duration where + coe := ofSeconds + +instance : Coe Minute.Offset Duration where + coe := ofSeconds ∘ Minute.Offset.toSeconds + +instance : Coe Hour.Offset Duration where + coe := ofSeconds ∘ Hour.Offset.toSeconds + +instance : Coe Week.Offset Duration where + coe := ofSeconds ∘ Day.Offset.toSeconds ∘ Week.Offset.toDays + +instance : Coe Day.Offset Duration where + coe := ofSeconds ∘ Day.Offset.toSeconds + +instance : HMul Int Duration Duration where + hMul x y := Duration.ofNanoseconds <| Nanosecond.Offset.ofInt (y.toNanoseconds.val * x) + +instance : HMul Duration Int Duration where + hMul y x := Duration.ofNanoseconds <| Nanosecond.Offset.ofInt (y.toNanoseconds.val * x) + + end Duration end Time end Std diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 0e0079bc62c9..c157933e9ba8 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -698,19 +698,19 @@ private def formatEraNarrow : Year.Era → String | .bce => "B" | .ce => "C" -private def formatQuarterNumber : Bounded.LE 1 4 → String +private def formatQuarterNumber : Month.Quarter → String |⟨1, _⟩ => "1" |⟨2, _⟩ => "2" |⟨3, _⟩ => "3" |⟨4, _⟩ => "4" -private def formatQuarterShort : Bounded.LE 1 4 → String +private def formatQuarterShort : Month.Quarter → String | ⟨1, _⟩ => "Q1" | ⟨2, _⟩ => "Q2" | ⟨3, _⟩ => "Q3" | ⟨4, _⟩ => "Q4" -private def formatQuarterLong : Bounded.LE 1 4 → String +private def formatQuarterLong : Month.Quarter → String | ⟨1, _⟩ => "1st quarter" | ⟨2, _⟩ => "2nd quarter" | ⟨3, _⟩ => "3rd quarter" @@ -769,7 +769,7 @@ private def TypeFormat : Modifier → Type | .D _ => Sigma Day.Ordinal.OfYear | .MorL _ => Month.Ordinal | .d _ => Day.Ordinal - | .Qorq _ => Month.Ordinal.Quarter + | .Qorq _ => Month.Quarter | .w _ => Week.Ordinal | .W _ => Week.Ordinal.OfMonth | .E _ => Weekday @@ -917,13 +917,13 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .MorL _ => date.month | .d _ => date.day | .Qorq _ => date.quarter - | .w _ => date.toWeekOfYear - | .W _ => date.toWeekOfMonth + | .w _ => date.weekOfYear + | .W _ => date.weekOfMonth | .E _ => date.weekday | .eorc _ => date.weekday - | .F _ => date.toWeekOfMonth + | .F _ => date.weekOfMonth | .a _ => HourMarker.ofOrdinal date.hour - | .B _ => date.getPeriod + | .B _ => date.period | .h _ => HourMarker.toRelative date.hour |>.fst | .K _ => date.hour.emod 12 (by decide) | .k _ => date.hour.add 1 @@ -1022,19 +1022,19 @@ private def parseEraNarrow : Parser Year.Era := pstring "B" *> pure Year.Era.bce <|> pstring "C" *> pure Year.Era.ce -private def parseQuarterNumber : Parser Month.Ordinal.Quarter +private def parseQuarterNumber : Parser Month.Quarter := pstring "1" *> pure ⟨1, by decide⟩ <|> pstring "2" *> pure ⟨2, by decide⟩ <|> pstring "3" *> pure ⟨3, by decide⟩ <|> pstring "4" *> pure ⟨4, by decide⟩ -private def parseQuarterLong : Parser Month.Ordinal.Quarter +private def parseQuarterLong : Parser Month.Quarter := pstring "1st quarter" *> pure ⟨1, by decide⟩ <|> pstring "2nd quarter" *> pure ⟨2, by decide⟩ <|> pstring "3rd quarter" *> pure ⟨3, by decide⟩ <|> pstring "4th quarter" *> pure ⟨4, by decide⟩ -private def parseQuarterShort : Parser Month.Ordinal.Quarter +private def parseQuarterShort : Parser Month.Quarter := pstring "Q1" *> pure ⟨1, by decide⟩ <|> pstring "Q2" *> pure ⟨2, by decide⟩ <|> pstring "Q3" *> pure ⟨3, by decide⟩ @@ -1095,10 +1095,7 @@ private def exactlyChars (parse : Parser Char) (size : Nat) : Parser String := private def parseSigned (parser : Parser Nat) : Parser Int := do let signed ← optional (pstring "-") let res ← parser - if signed.isSome then - return -res - else - return res + return if signed.isSome then -res else res private def parseNum (size : Nat) : Parser Nat := String.toNat! <$> exactlyChars (satisfy Char.isDigit) size @@ -1263,7 +1260,7 @@ private structure DateBuilder where D : Option (Sigma Day.Ordinal.OfYear) := none MorL : Option Month.Ordinal := none d : Option Day.Ordinal := none - Qorq : Option Month.Ordinal.Quarter := none + Qorq : Option Month.Quarter := none w : Option Week.Ordinal := none W : Option Week.Ordinal.OfMonth := none E : Option Weekday := none @@ -1294,6 +1291,7 @@ namespace DateBuilder private def insert (date : DateBuilder) (modifier : Modifier) (data : TypeFormat modifier) : DateBuilder := match modifier with | .G _ => { date with G := some data } + | .y .twoDigit => { date with y := some (Year.Offset.ofInt (data.toInt + 2000)) } | .y _ => { date with y := some data } | .D _ => { date with D := some data } | .MorL _ => { date with MorL := some data } diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index dc7223c24f55..b1e9df28aaae 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -46,8 +46,11 @@ integers that `lo ≤ val ≤ hi`. -/ abbrev LE := @Bounded LE.le -def cast {rel : Int → Int → Prop} {lo₁ lo₂ hi₁ hi₂ : Int} (h₁ : lo₁ = lo₂) (h₂ : hi₁ = hi₂) - (b : Bounded rel lo₁ hi₁) : Bounded rel lo₂ hi₂ := +/-- +Casts the boundaries of the `Bounded` using equivalences. +-/ +@[inline] +def cast {rel : Int → Int → Prop} {lo₁ lo₂ hi₁ hi₂ : Int} (h₁ : lo₁ = lo₂) (h₂ : hi₁ = hi₂) (b : Bounded rel lo₁ hi₁) : Bounded rel lo₂ hi₂ := .mk b.val ⟨h₁ ▸ b.property.1, h₂ ▸ b.property.2⟩ /-- @@ -79,7 +82,7 @@ namespace LE Convert a `Nat` to a `Bounded.LE` by wrapping it. -/ @[inline] -def ofNatWrapping { lo hi : Int } (val : Int) (h : lo ≤ hi): Bounded.LE lo hi := by +def ofNatWrapping { lo hi : Int } (val : Int) (h : lo ≤ hi) : Bounded.LE lo hi := by let range := hi - lo + 1 have range_pos := Int.add_pos_of_nonneg_of_pos (b := 1) (Int.sub_nonneg_of_le h) (by decide) have not_zero := Int.ne_iff_lt_or_gt.mpr (Or.inl range_pos) @@ -395,7 +398,7 @@ def eq {n : Int} : Bounded.LE n n := ⟨n, And.intro (Int.le_refl n) (Int.le_refl n)⟩ /-- -Expand the top and bottom of the bounded. +Expand the range of a bounded value. -/ @[inline] def expand (bounded : Bounded.LE lo hi) (h : hi ≤ nhi) (h₁ : nlo ≤ lo) : Bounded.LE nlo nhi := @@ -406,14 +409,14 @@ Expand the bottom of the bounded to a number `nhi` is `hi` is less or equal to t -/ @[inline] def expandTop (bounded : Bounded.LE lo hi) (h : hi ≤ nhi) : Bounded.LE lo nhi := - ⟨bounded.val, And.intro bounded.property.left (Int.le_trans bounded.property.right h)⟩ + expand bounded h (Int.le_refl lo) /-- Expand the bottom of the bounded to a number `nlo` if `lo` is greater or equal to the previous lower bound. -/ @[inline] def expandBottom (bounded : Bounded.LE lo hi) (h : nlo ≤ lo) : Bounded.LE nlo hi := - ⟨bounded.val, And.intro (Int.le_trans h bounded.property.left) bounded.property.right⟩ + expand bounded (Int.le_refl hi) h /-- Adds one to the value of the bounded if the value is less than the higher bound of the bounded number. @@ -427,7 +430,8 @@ def succ (bounded : Bounded.LE lo hi) (h : bounded.val < hi) : Bounded.LE lo hi Returns the absolute value of the bounded number `bo` with bounds `-(i - 1)` to `i - 1`. The result will be a new bounded number with bounds `0` to `i - 1`. -/ -def abs (bo : Bounded.LE (- (i - 1)) (i - 1)) : Bounded.LE 0 (i - 1) := +@[inline] +def abs (bo : Bounded.LE (-i) i) : Bounded.LE 0 i := if h : bo.val ≥ 0 then bo.truncateBottom h else by diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index b16ba0bc9b2d..897ffea8351b 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -93,16 +93,26 @@ def convert (val : UnitVal a) : UnitVal b := ofInt <| val.toInt * ratio.num / ratio.den instance : OfNat (UnitVal α) n where ofNat := ⟨Int.ofNat n⟩ + instance : Repr (UnitVal α) where reprPrec x p := reprPrec x.val p + instance : LE (UnitVal α) where le x y := x.val ≤ y.val + instance : LT (UnitVal α) where lt x y := x.val < y.val + instance : Add (UnitVal α) where add x y := ⟨x.val + y.val⟩ + instance : Sub (UnitVal α) where sub x y := ⟨x.val - y.val⟩ + instance : Mul (UnitVal α) where mul x y := ⟨x.val * y.val⟩ + instance : Div (UnitVal α) where div x y := ⟨x.val / y.val⟩ + instance : Neg (UnitVal α) where neg x := ⟨-x.val⟩ + instance : ToString (UnitVal n) where toString n := toString n.val + end UnitVal end Internal end Time diff --git a/src/Std/Time/Notation/Spec.lean b/src/Std/Time/Notation/Spec.lean index c00f7105cb65..72d44dca9748 100644 --- a/src/Std/Time/Notation/Spec.lean +++ b/src/Std/Time/Notation/Spec.lean @@ -15,8 +15,8 @@ namespace Time open Lean Parser Command Std private def convertText : Text → MacroM (TSyntax `term) - | .short => `(Std.Time.Text.short) - | .full => `(Std.Time.Text.full) + | .short => `(Std.Time.Text.short) + | .full => `(Std.Time.Text.full) | .narrow => `(Std.Time.Text.narrow) private def convertNumber : Number → MacroM (TSyntax `term) diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index bfb1cac1a00d..01ffb26d5c18 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -46,6 +46,12 @@ instance : BEq PlainTime where namespace PlainTime +/-- +Creates a `PlainTime` value representing midnight (00:00:00.000). +-/ +def midnight : PlainTime := + ⟨0, 0, ⟨true, 0⟩, 0⟩ + /-- Creates a `PlainTime` value from hours, minutes, seconds and nanoseconds. -/ diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index e7949a8e6f4e..f14392410500 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -367,19 +367,12 @@ def weekday (dt : DateTime tz) : Weekday := dt.date.get.date.weekday /-- -Gets the `Weekday` of a DateTime. +Gets the `Period` of a `DateTime`, corresponding to the part of the day (e.g. night, morning, +afternoon, evening) based on the hour. -/ @[inline] -def getPeriod (dt : DateTime tz) : Day.Ordinal.Period := - if dt.hour ≥ 21 ∨ dt.hour ≤ 4 then - .night - else if dt.hour ≥ 17 then - .evening - else if dt.hour ≥ 12 then - .afternoon - else - .morning - +def period (dt : DateTime tz) : Day.Ordinal.Period := + Day.Ordinal.Period.fromHour dt.hour /-- Determines the era of the given `DateTime` based on its year. @@ -402,7 +395,7 @@ def toOrdinal (date : DateTime tz) : Day.Ordinal.OfYear date.year.isLeap := /-- Determines the week of the year for the given `DateTime`. -/ -def toWeekOfYear (date : DateTime tz) : Week.Ordinal := +def weekOfYear (date : DateTime tz) : Week.Ordinal := let res := Month.Ordinal.toOrdinal ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ |>.ediv 7 (by decide) |>.add 1 match date.date.get.date.year.isLeap, res with | true, res => res @@ -413,7 +406,7 @@ Determines the week of the month for the given `DateTime`. The week of the month on the day of the month and the weekday. Each week starts on Sunday because the entire library is based on the Gregorian Calendar. -/ -def toWeekOfMonth (date : DateTime tz) : Bounded.LE 1 6 := +def weekOfMonth (date : DateTime tz) : Week.Ordinal.OfMonth := let weekday := date.weekday.toOrdinal date.day.addBounds (weekday.sub 1) |>.ediv 7 (by decide) |>.add 1 From fb8f037757c342182f5df5e661b3469c2d6bb4d5 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 30 Sep 2024 08:28:53 -0300 Subject: [PATCH 071/118] fix: grammar correction --- src/Std/Time/Zoned/DateTime.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index f14392410500..e082cac78471 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -62,7 +62,7 @@ def convertTimeZone (date : DateTime tz) (tz₁ : TimeZone) : DateTime tz₁ := /-- Creates a new `DateTime` out of a `PlainDateTime`. It assumes that the `PlainDateTime` is relative -to UTC. If you're using hte PlainDateTime +to UTC. -/ @[inline] def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : DateTime tz := From d353fa6a3ee7150688a54a4e8a0d26a745a7d6c3 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 10 Oct 2024 21:06:13 -0300 Subject: [PATCH 072/118] feat: add leap second parser --- src/Std/Time.lean | 3 +- src/Std/Time/DateTime/Timestamp.lean | 12 ++++ src/Std/Time/Duration.lean | 13 +++- src/Std/Time/Format.lean | 8 +-- src/Std/Time/Format/Basic.lean | 5 +- src/Std/Time/Internal/UnitVal.lean | 6 ++ src/Std/Time/Time/PlainTime.lean | 1 + src/Std/Time/Time/Unit/Nanosecond.lean | 3 + src/Std/Time/Zoned.lean | 3 +- src/Std/Time/Zoned/Database.lean | 8 +-- src/Std/Time/Zoned/Database/Basic.lean | 23 +++--- src/Std/Time/Zoned/Database/Leap.lean | 97 ++++++++++++++++++++++++++ src/Std/Time/Zoned/Database/TZdb.lean | 11 +++ src/Std/Time/Zoned/ZoneRules.lean | 47 +------------ src/Std/Time/Zoned/ZonedDateTime.lean | 10 ++- 15 files changed, 178 insertions(+), 72 deletions(-) create mode 100644 src/Std/Time/Zoned/Database/Leap.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index ff922c5c1381..62eeefc49da8 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -10,6 +10,7 @@ import Std.Time.Format import Std.Time.DateTime import Std.Time.Notation import Std.Time.Duration +import Std.Time.Zoned.Database namespace Std namespace Time @@ -17,7 +18,7 @@ namespace Time /-! # Time -The Lean4 API for date, time, and duration functionalities, following the ISO8601 standards. +The Lean4 API for date, time, and duration functionalities. # Overview diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 36bf380c1e1d..1cd9434a678d 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -27,6 +27,12 @@ structure Timestamp where val : Duration deriving Repr, BEq, Inhabited +instance : LE Timestamp where + le x y := x.val ≤ y.val + +instance { x y : Timestamp } : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + instance : OfNat Timestamp n where ofNat := ⟨OfNat.ofNat n⟩ @@ -238,4 +244,10 @@ instance : HSub Timestamp Nanosecond.Offset Timestamp where instance : HSub Timestamp Timestamp Duration where hSub x y := x.val - y.val +instance : HAdd Timestamp Duration Timestamp where + hAdd x y := x.addNanoseconds y.toNanoseconds + +instance : HSub Timestamp Duration Timestamp where + hSub x y := x.subNanoseconds y.toNanoseconds + end Timestamp diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 586439f4eb49..1b867bfbde91 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -34,6 +34,7 @@ structure Duration where proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) deriving Repr + instance : ToString Duration where toString s := let (sign, secs, nanos) := @@ -115,6 +116,12 @@ def toNanoseconds (duration : Duration) : Nanosecond.Offset := let nanos := nanos + (.ofInt duration.nano.val) nanos +instance : LE Duration where + le d1 d2 := d1.toNanoseconds ≤ d2.toNanoseconds + +instance {x y : Duration} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.toNanoseconds ≤ y.toNanoseconds)) + /-- Converts a `Duration` to a `Minute.Offset` -/ @@ -303,11 +310,13 @@ instance : Coe Day.Offset Duration where coe := ofSeconds ∘ Day.Offset.toSeconds instance : HMul Int Duration Duration where - hMul x y := Duration.ofNanoseconds <| Nanosecond.Offset.ofInt (y.toNanoseconds.val * x) + hMul i d := Duration.ofNanoseconds <| Nanosecond.Offset.ofInt (d.toNanoseconds.val * i) instance : HMul Duration Int Duration where - hMul y x := Duration.ofNanoseconds <| Nanosecond.Offset.ofInt (y.toNanoseconds.val * x) + hMul d i := Duration.ofNanoseconds <| Nanosecond.Offset.ofInt (d.toNanoseconds.val * i) +instance : HAdd PlainTime Duration PlainTime where + hAdd pt d := PlainTime.ofNanoseconds (d.toNanoseconds + pt.toNanoseconds) end Duration end Time diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index c7f5b4c95c0c..ef6210d83456 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -372,8 +372,8 @@ def fromLeanDateTimeWithZoneString (input : String) : Except String ZonedDateTim /-- Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notationg. -/ -def toLeanDateTimeWithZoneString (pdt : ZonedDateTime) : String := - Formats.leanDateTimeWithZone.format pdt.snd +def toLeanDateTimeWithZoneString (zdt : ZonedDateTime) : String := + Formats.leanDateTimeWithZone.formatBuilder zdt.year zdt.month zdt.day zdt.hour zdt.minute zdt.snd.date.get.time.second zdt.nanosecond zdt.fst.offset /-- Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. @@ -453,7 +453,7 @@ def fromDateTimeString (input : String) : Except String PlainDateTime := Formats a `PlainDateTime` value into a `DateTime` format string. -/ def toDateTimeString (pdt : PlainDateTime) : String := - Formats.dateTime24Hour.format (DateTime.ofPlainDateTimeAssumingUTC pdt .UTC) + Formats.dateTime24Hour.formatBuilder pdt.year pdt.month pdt.day pdt.hour pdt.minute pdt.time.second pdt.nanosecond /-- Parses a `String` in the `DateTime` format and returns a `PlainDateTime`. @@ -466,7 +466,7 @@ def fromLeanDateTimeString (input : String) : Except String PlainDateTime := Formats a `PlainDateTime` value into a `DateTime` format string. -/ def toLeanDateTimeString (pdt : PlainDateTime) : String := - Formats.leanDateTime24Hour.format (DateTime.ofPlainDateTimeAssumingUTC pdt .UTC) + Formats.leanDateTime24Hour.formatBuilder pdt.year pdt.month pdt.day pdt.hour pdt.minute pdt.time.second pdt.nanosecond /-- Parses a `String` in the `AscTime` or `LongDate` format and returns a `PlainDateTime`. diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index c157933e9ba8..1c41bd8a9b83 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -955,7 +955,10 @@ private def parseMonthLong : Parser Month.Ordinal <|> pstring "November" *> pure ⟨11, by decide⟩ <|> pstring "December" *> pure ⟨12, by decide⟩ -private def parseMonthShort : Parser Month.Ordinal +/-- +Parses a short value of a `Month.Ordinal` +-/ +def parseMonthShort : Parser Month.Ordinal := pstring "Jan" *> pure ⟨1, by decide⟩ <|> pstring "Feb" *> pure ⟨2, by decide⟩ <|> pstring "Mar" *> pure ⟨3, by decide⟩ diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 897ffea8351b..97c9bf45a8f6 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -26,6 +26,12 @@ structure UnitVal (α : Rat) where val : Int deriving Inhabited, BEq +instance : LE (UnitVal x) where + le x y := x.val ≤ y.val + +instance { x y : UnitVal z }: Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + namespace UnitVal /-- diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 01ffb26d5c18..4658c7d6227a 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -36,6 +36,7 @@ structure PlainTime where `Nanoseconds` component of the `PlainTime` -/ nano : Nanosecond.Ordinal + deriving Repr instance : Inhabited PlainTime where default := ⟨0, 0, Sigma.mk false 0, 0, by decide⟩ diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index 21fb6ceaea0e..8d635c5a8f50 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -42,6 +42,9 @@ instance {x y : Ordinal} : Decidable (x < y) := def Offset : Type := UnitVal (1 / 1000000000) deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString +instance { x y : Offset } : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 1c5e0ccf3858..73a0eddc68c3 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -5,11 +5,9 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Zoned.DateTime -import Std.Time.Zoned.Database import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.ZonedDateTime -import Std.Time.Zoned.Database.Basic namespace Std namespace Time @@ -148,3 +146,4 @@ def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch end ZonedDateTime + diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index cb27b13377e7..68400253e900 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -15,11 +15,11 @@ open TimeZone.ZoneRules set_option linter.all true -/-- +/- Apply leap seconds rules and transition rules on a UTC Timestamp to make it aware of the timezone. -/ -def applyLeapSecondsOnUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO Timestamp := do - (applyLeapSeconds tm ·) <$> Database.localRules db +-- def applyLeapSecondsOnUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO Timestamp := do +-- (applyLeapSeconds tm ·) <$> Database.localRules db /-- Gets the TimeZone at the local timezone. @@ -39,5 +39,5 @@ Get the local ZonedDataTime given a UTC `Timestamp`. def ofUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime := do let rules ← Database.localRules db let tz ← IO.ofExcept <| timezoneAt rules tm - let tm := applyLeapSeconds tm rules + --let tm := applyLeapSeconds tm rules return ZonedDateTime.ofTimestamp tm tz diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index e3e4ee8a0229..a1a843ee98e6 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -5,6 +5,7 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Zoned.ZoneRules +import Std.Time.Zoned.Database.Leap import Std.Time.Zoned.Database.TzIf namespace Std @@ -23,10 +24,15 @@ class Database (α : Type) where load : α → String → IO TimeZone.ZoneRules /-- - Loads a `ZoneRules` that is set by the local computer. + Loads a `ZoneRules` that is set by the local machine. -/ localRules : α → IO TimeZone.ZoneRules + /-- + Loads a `LeapSecond` array from the local machine. + -/ + leapSeconds : α → IO (Array TimeZone.LeapSecond) + namespace TimeZone /-- @@ -44,15 +50,9 @@ def convertUt : Bool → UTLocal | false => .local /-- -Converts a `TZif.LeapSecond` structure to a `LeapSecond` structure. --/ -def convertLeapSecond (tz : TZif.LeapSecond) : LeapSecond := - { transitionTime := .ofInt tz.transitionTime, correction := Second.Offset.ofInt tz.correction } - -/-- -Converts a PlainTime. +Converts a given time index into a `LocalTimeType` by using a time zone (`tz`) and its identifier. -/ -def convertPlainTime (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do +def convertLocalTimeType (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do let localType ← tz.localTimeTypes.get? index let abbreviation ← tz.abbreviations.getD index "Unknown" let wallflag := convertWall (tz.stdWallIndicators.getD index true) @@ -83,19 +83,18 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := let mut times : Array LocalTimeType := #[] for i in [0:tz.header.typecnt.toNat] do - if let some result := convertPlainTime i tz id + if let some result := convertLocalTimeType i tz id then times := times.push result else .error s!"cannot convert local time {i} of the file" let mut transitions := #[] - let leapSeconds := tz.leapSeconds.map convertLeapSecond for i in [0:tz.transitionTimes.size] do if let some result := convertTransition times i tz then transitions := transitions.push result else .error s!"cannot convert transtiion {i} of the file" - .ok { leapSeconds, transitions, localTimes := times } + .ok { transitions, localTimes := times } /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. diff --git a/src/Std/Time/Zoned/Database/Leap.lean b/src/Std/Time/Zoned/Database/Leap.lean new file mode 100644 index 000000000000..07671733e109 --- /dev/null +++ b/src/Std/Time/Zoned/Database/Leap.lean @@ -0,0 +1,97 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +import Init.Data.Range +import Std.Internal.Parsec +import Std.Internal.Parsec.ByteArray + +import Std.Time.Internal.Bounded +import Std.Time.Format.Basic +import Std.Time.Time +import Std.Time.Date + +namespace Std +namespace Time +namespace TimeZone +open Std.Internal.Parsec Std.Internal.Parsec.String + +set_option linter.all true + +/-- +Represents a leap second entry with details such as the year, month, day, time, whether it's a positive leap second, +and a stationary string (presumably to capture additional information). +-/ +structure LeapSecond where + + /-- + The timestamp when the leap second occurs. + -/ + timestamp : Timestamp + + /-- + Indicates whether the leap second is positive (`true` for positive, `false` for negative). + -/ + positive : Bool + + /-- + A string representing the stationary field, which could be used for additional information + or metadata about the leap second. + -/ + stationary : String + +private def skipComment : Parser Unit := do + skipChar '#' + discard <| many (satisfy (· ≠ '\n')) + discard ∘ optional <| skipChar '\n' + return () + +private def failWith (opt : Option α) : Parser α := + match opt with + | none => fail "invalid number" + | some res => pure res + +private def parseLeapSecond : Parser LeapSecond := do + skipString "Leap" + ws + let year ← digits + ws + let month ← Std.Time.parseMonthShort + ws + let day ← digits + ws + + let hour : Hour.Ordinal ← failWith =<< Internal.Bounded.ofInt? <$> digits + skipChar ':' + let minute : Minute.Ordinal ← failWith =<< Internal.Bounded.ofInt? <$> digits + skipChar ':' + let second : Second.Ordinal true ← failWith =<< Internal.Bounded.ofInt? <$> digits + + ws + let positive ← (pchar '+' *> pure true) + ws + let stationary ← manyChars (satisfy Char.isAlpha) + skipChar '\n' + + dbg_trace s!"{year}--{month}--{day}" + + let day ← failWith <| Internal.Bounded.ofInt? day + let time : PlainTime ← failWith <| PlainTime.ofHourMinuteSeconds hour minute second + let date : PlainDate ← failWith <| PlainDate.ofYearMonthDay (Year.Offset.ofNat year) month day + let pdt := PlainDateTime.mk date time + + return { timestamp := pdt.toTimestampAssumingUTC, positive, stationary } + +private def parseComments : Parser (Array Unit) := + many1 (ws *> skipComment) + +/-- +Parses a sequence of leap second entries. +-/ +def parseLeapSeconds : Parser (Array LeapSecond) := do + discard <| many (ws *> skipComment) + let res ← many parseLeapSecond + discard <| many (ws *> skipComment) + eof + return res diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index e1c19ce999ff..2a85dbcf76e5 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -83,6 +83,17 @@ Reads timezone rules from disk based on the provided file path and timezone ID. def readRulesFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := do parseTZIfFromDisk (System.FilePath.join path id) id +/-- +Reads leap seconds form from disk based on the provided file path and timezone ID. +-/ +def readLeapsFromDisk (path : System.FilePath) : IO (Array LeapSecond) := do + let res ← parseLeapSeconds.run <$> IO.FS.readFile path + + if let .ok res := res + then return res + else throw (IO.userError "cannot read the id of the path.") + instance : Database TZdb where load db id := readRulesFromDisk db.zonesPath id localRules db := localRules db.localPath + leapSeconds db := readLeapsFromDisk (db.zonesPath / "leapseconds") diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index fff16b721a9a..048cbef73480 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -90,22 +90,6 @@ def getTimeZone (time : LocalTimeType) : TimeZone := end LocalTimeType -/-- -Represents a leap second event, including the time of the transition and the correction applied. --/ -structure LeapSecond where - - /-- - The time when the leap second transition occurs. - -/ - transitionTime : Second.Offset - - /-- - The correction applied during the leap second event. - -/ - correction : Second.Offset - deriving Repr, Inhabited - /-- Represents a time zone transition, mapping a time to a local time type. -/ @@ -123,7 +107,8 @@ structure Transition where deriving Repr, Inhabited /-- -Represents the rules for a time zone, abstracting away binary data and focusing on key transitions and types. +Represents the rules for a time zone, abstracting away binary data and focusing on key transitions +and types. -/ structure ZoneRules where @@ -137,10 +122,6 @@ structure ZoneRules where -/ transitions : Array Transition - /-- - The array of leap seconds affecting the time zone. - -/ - leapSeconds : Array LeapSecond deriving Repr, Inhabited namespace Transition @@ -159,8 +140,7 @@ Applies the transition to a Timestamp. -/ def apply (timestamp : Timestamp) (transition : Transition) : Timestamp := let offsetInSeconds := transition.localTimeType.gmtOffset.hour.mul 3600 |>.add transition.localTimeType.gmtOffset.second - let PlainTimestamp := timestamp.addSeconds offsetInSeconds - PlainTimestamp + timestamp.addSeconds offsetInSeconds end Transition @@ -183,24 +163,3 @@ def timezoneAt (rules : ZoneRules) (tm : Timestamp) : Except String TimeZone := if let some transition := rules.findTransitionForTimestamp tm then .ok transition.createTimeZoneFromTransition else .error "cannot find local timezone." - -/-- -Applies a `LeapSecond` sequence in a `Timestamp`. --/ -def applyLeapSeconds (tm : Timestamp) (leapSeconds : ZoneRules) : Timestamp := Id.run do - let mut currentTime := tm - let leapSeconds := leapSeconds.leapSeconds - for i in [:leapSeconds.size] do - let leapSec := leapSeconds.get! i - if currentTime.toSecondsSinceUnixEpoch.val >= leapSec.transitionTime.val then - currentTime := tm.addSeconds (.ofInt leapSec.correction.val) - return currentTime - -/-- -Adjust a UTC timestamp according to `ZoneRules`. --/ -protected def applyToTimestamp (rules : ZoneRules) (tm : Timestamp) : Timestamp := - let tm := applyLeapSeconds tm rules - if let some transition := findTransitionForTimestamp rules tm - then transition.apply tm - else tm diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index d6480f0e08d1..b1beb3b485c9 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -57,9 +57,15 @@ Creates a new `ZonedDateTime` out of a `Timestamp` @[inline] def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do let transition ← rules.findTransitionForTimestamp tm - let tm := rules.applyLeapSeconds tm return ofTimestamp tm transition.localTimeType.getTimeZone +/-- +Converts a `PlainDateTime` to a `ZonedDateTime` using the given `ZoneRules`. +-/ +@[inline] +def ofLocalDateTimeWithZoneRules (pdt : PlainDateTime) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do + ofZoneRules (pdt.toTimestampAssumingUTC) rules + /-- Changes the `TimeZone` to a new one. -/ @@ -127,7 +133,7 @@ def second (zdt : ZonedDateTime) : Second.Ordinal zdt.snd.date.get.time.second.f Getter for the `Nanosecond` inside of a `ZonedDateTime` -/ @[inline] -def nano (zdt : ZonedDateTime) : Nanosecond.Ordinal := +def nanosecond (zdt : ZonedDateTime) : Nanosecond.Ordinal := zdt.snd.date.get.time.nano /-- From cc59eebebc1cc33ba64012b7dda7924c78300d1c Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 14 Oct 2024 09:24:42 -0300 Subject: [PATCH 073/118] fix: a lot of comment fixes and format change --- src/Std/Time/Date/PlainDate.lean | 92 ++++++++++----------- src/Std/Time/Date/Unit/Basic.lean | 18 ---- src/Std/Time/Date/Unit/Day.lean | 80 +++++++++--------- src/Std/Time/Date/Unit/Month.lean | 16 +++- src/Std/Time/Date/Unit/Week.lean | 12 ++- src/Std/Time/Date/Unit/Weekday.lean | 2 +- src/Std/Time/Date/Unit/Year.lean | 10 ++- src/Std/Time/DateTime.lean | 8 +- src/Std/Time/DateTime/PlainDateTime.lean | 22 +++-- src/Std/Time/DateTime/Timestamp.lean | 72 +++++++++------- src/Std/Time/Duration.lean | 22 ++++- src/Std/Time/Format.lean | 100 +++++++++++------------ src/Std/Time/Format/Basic.lean | 80 +++++++++++------- src/Std/Time/Notation.lean | 68 ++++++++++++--- src/Std/Time/Notation/Spec.lean | 2 +- src/Std/Time/Time/HourMarker.lean | 6 +- src/Std/Time/Time/PlainTime.lean | 28 +++++-- src/Std/Time/Time/Unit/Hour.lean | 2 +- src/Std/Time/Time/Unit/Millisecond.lean | 3 +- src/Std/Time/Time/Unit/Minute.lean | 2 +- src/Std/Time/Time/Unit/Nanosecond.lean | 2 +- src/Std/Time/Time/Unit/Second.lean | 12 +-- src/Std/Time/Zoned.lean | 15 ++-- src/Std/Time/Zoned/Database.lean | 15 +--- src/Std/Time/Zoned/Database/TZdb.lean | 4 +- src/Std/Time/Zoned/Database/TzIf.lean | 18 ++-- src/Std/Time/Zoned/DateTime.lean | 28 +++---- src/Std/Time/Zoned/Offset.lean | 21 +++-- src/Std/Time/Zoned/TimeZone.lean | 7 +- src/Std/Time/Zoned/ZoneRules.lean | 3 +- src/Std/Time/Zoned/ZonedDateTime.lean | 45 +++++----- src/runtime/io.cpp | 5 +- tests/lean/run/timeFormats.lean | 22 ++--- tests/lean/run/timeLocalDateTime.lean | 4 +- tests/lean/run/timeParse.lean | 24 +++--- tests/lean/run/timeSet.lean | 32 ++++---- 36 files changed, 511 insertions(+), 391 deletions(-) diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index f0b191405d4f..36b913f3e244 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -34,7 +34,6 @@ structure PlainDate where /-- Validates the date by ensuring that the year, month, and day form a correct and valid date. -/ valid : year.Valid month day - instance : BEq PlainDate where beq x y := x.day == y.day && x.month == y.month && x.year == y.year @@ -110,11 +109,9 @@ def weekday (date : PlainDate) : Weekday := /-- Determines the era of the given `PlainDate` based on its year. -/ +@[inline] def era (date : PlainDate) : Year.Era := - if date.year.toInt ≥ 0 then - .ce - else - .bce + date.year.era /-- Checks if the `PlainDate` is in a leap year. @@ -162,7 +159,6 @@ def subDays (date : PlainDate) (days : Day.Offset) : PlainDate := /-- Adds a given number of weeks to a `PlainDate`. -/ - @[inline] def addWeeks (date : PlainDate) (weeks : Week.Offset) : PlainDate := let dateDays := date.toDaysSinceUNIXEpoch @@ -175,6 +171,7 @@ Subtracts a given number of weeks from a `PlainDate`. @[inline] def subWeeks (date : PlainDate) (weeks : Week.Offset) : PlainDate := addWeeks date (-weeks) + /-- Adds a given number of months to a `PlainDate`, clipping the day to the last valid day of the month. -/ @@ -198,46 +195,6 @@ Creates a `PlainDate` by rolling over the extra days to the next month. def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := clip year month 1 |>.addDays (day.toOffset - 1) -/-- -Creates a new `PlainDate` by adjusting the day of the month to the given `days` value, with any -out-of-range days clipped to the nearest valid date. --/ -@[inline] -def withDaysClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := - clip dt.year dt.month days - -/-- -Creates a new `PlainDate` by adjusting the day of the month to the given `days` value, with any -out-of-range days rolled over to the next month or year as needed. --/ -@[inline] -def withDaysRollOver (dt : PlainDate) (days : Day.Ordinal) : PlainDate := - rollOver dt.year dt.month days - -/-- -Creates a new `PlainDate` by adjusting the month to the given `month` value. -The day remains unchanged, and any invalid days for the new month will be handled according to the `clip` behavior. --/ -@[inline] -def withMonthClip (dt : PlainDate) (month : Month.Ordinal) : PlainDate := - clip dt.year month dt.day - -/-- -Creates a new `PlainDate` by adjusting the month to the given `month` value. -The day is rolled over to the next valid month if necessary. --/ -@[inline] -def withMonthRollOver (dt : PlainDate) (month : Month.Ordinal) : PlainDate := - rollOver dt.year month dt.day - -/-- -Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day remain unchanged, -and any invalid days for the new year will be handled according to the `clip` behavior. --/ -@[inline] -def withYearClip (dt : PlainDate) (year : Year.Offset) : PlainDate := - clip year dt.month dt.day - /-- Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day are rolled over to the next valid month and day if necessary. @@ -250,7 +207,8 @@ def withYearRollOver (dt : PlainDate) (year : Year.Offset) : PlainDate := Adds a given number of months to a `PlainDate`, rolling over any excess days into the following month. -/ def addMonthsRollOver (date : PlainDate) (months : Month.Offset) : PlainDate := - addMonthsClip (clip date.year date.month 1) months |>.addDays (date.day.toOffset - 1) + addMonthsClip (clip date.year date.month 1) months + |>.addDays (date.day.toOffset - 1) /-- Subtracts `Month.Offset` from a `PlainDate`, rolling over excess days as needed. @@ -287,6 +245,46 @@ Subtracts `Year.Offset` from a `PlainDate`, clipping the day to the last valid d def subYearsClip (date : PlainDate) (years : Year.Offset) : PlainDate := addMonthsClip date (- years.mul 12) +/-- +Creates a new `PlainDate` by adjusting the day of the month to the given `days` value, with any +out-of-range days clipped to the nearest valid date. +-/ +@[inline] +def withDaysClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := + clip dt.year dt.month days + +/-- +Creates a new `PlainDate` by adjusting the day of the month to the given `days` value, with any +out-of-range days rolled over to the next month or year as needed. +-/ +@[inline] +def withDaysRollOver (dt : PlainDate) (days : Day.Ordinal) : PlainDate := + rollOver dt.year dt.month days + +/-- +Creates a new `PlainDate` by adjusting the month to the given `month` value. +The day remains unchanged, and any invalid days for the new month will be handled according to the `clip` behavior. +-/ +@[inline] +def withMonthClip (dt : PlainDate) (month : Month.Ordinal) : PlainDate := + clip dt.year month dt.day + +/-- +Creates a new `PlainDate` by adjusting the month to the given `month` value. +The day is rolled over to the next valid month if necessary. +-/ +@[inline] +def withMonthRollOver (dt : PlainDate) (month : Month.Ordinal) : PlainDate := + rollOver dt.year month dt.day + +/-- +Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day remain unchanged, +and any invalid days for the new year will be handled according to the `clip` behavior. +-/ +@[inline] +def withYearClip (dt : PlainDate) (year : Year.Offset) : PlainDate := + clip year dt.month dt.day + instance : HAdd PlainDate Day.Offset PlainDate where hAdd := addDays diff --git a/src/Std/Time/Date/Unit/Basic.lean b/src/Std/Time/Date/Unit/Basic.lean index b16bb3219082..92b33bf07d62 100644 --- a/src/Std/Time/Date/Unit/Basic.lean +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -39,21 +39,3 @@ def toWeeks (day : Day.Offset) : Week.Offset := day.ediv 7 end Day.Offset - -namespace Week.Offset - -/-- -Convert `Week.Offset` into `Day.Offset`. --/ -@[inline] -def toDays (week : Week.Offset) : Day.Offset := - week.mul 7 - -/-- -Convert `Day.Offset` into `Week.Offset`. --/ -@[inline] -def ofDays (day : Day.Offset) : Week.Offset := - day.ediv 7 - -end Week.Offset diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 751ff3c05676..97e23aae649e 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -48,47 +48,27 @@ instance {x y : Offset} : Decidable (x ≤ y) := instance {x y : Offset} : Decidable (x < y) := inferInstanceAs (Decidable (x.val < y.val)) -namespace Ordinal - -/-- -`OfYear` represents the day ordinal within a year, which can be bounded between 1 and 365 or 366, -depending on whether it's a leap year. --/ -def OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) - -namespace OfYear - -/-- -Creates an ordinal for a specific day within the year, ensuring that the provided day (`data`) -is within the valid range for the year, which can be 1 to 365 or 366 for leap years. --/ -@[inline] -def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ (if leap then 366 else 365) := by decide) : OfYear leap := - Bounded.LE.ofNat' data h - -end OfYear - /-- `Period` is an enumeration representing different times of the day : morning, afternoon, evening, and night. -/ inductive Period - /--Represents the morning period. -/ + /-- Represents the morning period. -/ | morning - /--Represents the afternoon period. -/ + /-- Represents the afternoon period. -/ | afternoon - /--Represents the evening period. -/ + /-- Represents the evening period. -/ | evening - /--Represents the night period. -/ + /-- Represents the night period. -/ | night deriving Repr, BEq, Inhabited namespace Period /-- -Determines the `Period` of the day based on the given hour +Determines the `Period` of the day based on the given hour. - If the hour is between 20 and 4, it returns `night`. - If the hour is between 17 and 20, it returns `evening`. @@ -96,7 +76,7 @@ Determines the `Period` of the day based on the given hour - If the hour is between 5 and 12, it reutrns `morning`. -/ @[inline] -def fromHour (hour : Hour.Ordinal) : Day.Ordinal.Period := +def fromHour (hour : Hour.Ordinal) : Day.Period := if hour ≥ 20 ∨ hour ≤ 4 then .night else if hour ≥ 17 then @@ -108,6 +88,26 @@ def fromHour (hour : Hour.Ordinal) : Day.Ordinal.Period := end Period +namespace Ordinal + +/-- +`OfYear` represents the day ordinal within a year, which can be bounded between 1 and 365 or 366, +depending on whether it's a leap year. +-/ +def OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) + +namespace OfYear + +/-- +Creates an ordinal for a specific day within the year, ensuring that the provided day (`data`) +is within the valid range for the year, which can be 1 to 365 or 366 for leap years. +-/ +@[inline] +def ofNat (data : Nat) (h : data ≥ 1 ∧ data ≤ (if leap then 366 else 365) := by decide) : OfYear leap := + Bounded.LE.ofNat' data h + +end OfYear + instance : OfNat (Ordinal.OfYear leap) n := match leap with | true => inferInstanceAs (OfNat (Bounded.LE 1 (1 + (365 : Nat))) n) @@ -155,94 +155,94 @@ end Ordinal namespace Offset /-- -Converts an `Ordinal` to an `Offset`. +Converts an `Offset` to an `Ordinal`. -/ @[inline] -def toOrdinal (off : Offset) (h : off.val ≥ 1 ∧ off.val ≤ 31) : Ordinal := +def toOrdinal (off : Day.Offset) (h : off.val ≥ 1 ∧ off.val ≤ 31) : Ordinal := Bounded.LE.mk off.val h /-- Creates an `Offset` from a natural number. -/ @[inline] -def ofNat (data : Nat) : Offset := +def ofNat (data : Nat) : Day.Offset := UnitVal.mk data /-- Creates an `Offset` from an integer. -/ @[inline] -def ofInt (data : Int) : Offset := +def ofInt (data : Int) : Day.Offset := UnitVal.mk data /-- Convert `Day.Offset` into `Nanosecond.Offset`. -/ @[inline] -def toNanoseconds (days : Offset) : Nanosecond.Offset := +def toNanoseconds (days : Day.Offset) : Nanosecond.Offset := days.mul 86400000000000 /-- Convert `Nanosecond.Offset` into `Day.Offset`. -/ @[inline] -def ofNanoseconds (ns : Nanosecond.Offset) : Offset := +def ofNanoseconds (ns : Nanosecond.Offset) : Day.Offset := ns.ediv 86400000000000 /-- Convert `Day.Offset` into `Millisecond.Offset`. -/ @[inline] -def toMilliseconds (days : Offset) : Millisecond.Offset := +def toMilliseconds (days : Day.Offset) : Millisecond.Offset := days.mul 86400000 /-- Convert `Millisecond.Offset` into `Day.Offset`. -/ @[inline] -def ofMilliseconds (ms : Millisecond.Offset) : Offset := +def ofMilliseconds (ms : Millisecond.Offset) : Day.Offset := ms.ediv 86400000 /-- Convert `Day.Offset` into `Second.Offset`. -/ @[inline] -def toSeconds (days : Offset) : Second.Offset := +def toSeconds (days : Day.Offset) : Second.Offset := days.mul 86400 /-- Convert `Second.Offset` into `Day.Offset`. -/ @[inline] -def ofSeconds (secs : Second.Offset) : Offset := +def ofSeconds (secs : Second.Offset) : Day.Offset := secs.ediv 86400 /-- Convert `Day.Offset` into `Minute.Offset`. -/ @[inline] -def toMinutes (days : Offset) : Minute.Offset := +def toMinutes (days : Day.Offset) : Minute.Offset := days.mul 1440 /-- Convert `Minute.Offset` into `Day.Offset`. -/ @[inline] -def ofMinutes (minutes : Minute.Offset) : Offset := +def ofMinutes (minutes : Minute.Offset) : Day.Offset := minutes.ediv 1440 /-- Convert `Day.Offset` into `Hour.Offset`. -/ @[inline] -def toHours (days : Offset) : Hour.Offset := +def toHours (days : Day.Offset) : Hour.Offset := days.mul 24 /-- Convert `Hour.Offset` into `Day.Offset`. -/ @[inline] -def ofHours (hours : Hour.Offset) : Offset := +def ofHours (hours : Hour.Offset) : Day.Offset := hours.ediv 24 end Offset diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 6a60a5488d65..c2fa99d985a7 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -30,6 +30,12 @@ instance : OfNat Ordinal n := instance : Inhabited Ordinal where default := 1 +instance {x y : Ordinal} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance {x y : Ordinal} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + /-- `Offset` represents an offset in months. It is defined as an `Int`. -/ @@ -39,7 +45,6 @@ def Offset : Type := Int instance : OfNat Offset n := ⟨Int.ofNat n⟩ - /-- `Quarter` represents a value between 1 and 4, inclusive, corresponding to the four quarters of a year. -/ @@ -47,6 +52,15 @@ def Quarter := Bounded.LE 1 4 namespace Quarter +/-- +Determine the `Quarter` by the month. +-/ +def ofMonth (month : Month.Ordinal) : Quarter := + month + |>.sub 1 + |>.ediv 3 (by decide) + |>.add 1 + end Quarter namespace Offset diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 84e0e1475757..4127392bbf11 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -27,6 +27,12 @@ instance : ToString Ordinal where instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (52 : Nat))) n) +instance {x y : Ordinal} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +instance {x y : Ordinal} : Decidable (x < y) := + inferInstanceAs (Decidable (x.val < y.val)) + instance : Inhabited Ordinal where default := 1 @@ -67,6 +73,9 @@ Converts an `Ordinal` to an `Offset`. def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val +end Ordinal +namespace Offset + /-- Convert `Week.Offset` into `Millisecond.Offset`. -/ @@ -151,8 +160,7 @@ Convert `Day.Offset` into `Week.Offset`. def ofDays (hours : Day.Offset) : Week.Offset := hours.ediv 7 -end Ordinal - +end Offset end Week end Time end Std diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index 2a6437859572..18e94e4d02ad 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -43,7 +43,7 @@ inductive Weekday namespace Weekday /-- -`Ordinal` represents a bounded value for weekdays, which ranges between 1 and 7.. +`Ordinal` represents a bounded value for weekdays, which ranges between 1 and 7. -/ def Ordinal := Bounded.LE 1 7 diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index b594ba752761..fb64d6d996f0 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -31,7 +31,15 @@ inductive Era `Offset` represents a year offset, defined as an `Int`. -/ def Offset : Type := Int - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString, LT, LE + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + +instance {x y : Offset} : Decidable (x ≤ y) := + let x : Int := x + inferInstanceAs (Decidable (x ≤ y)) + +instance {x y : Offset} : Decidable (x < y) := + let x : Int := x + inferInstanceAs (Decidable (x < y)) instance : OfNat Offset n := ⟨Int.ofNat n⟩ diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 93b5a1a4d7ab..6625be2dca6e 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -79,8 +79,8 @@ Calculates the duration between a given `PlainDate` and a specified date. ```lean def example : Duration := - let startDate := date% 2023-1-1 - let endDate := date% 2023-3-15 + let startDate := date("2023-1-1") + let endDate := date("2023-3-15") endDate.since startDate ``` -/ @@ -137,8 +137,8 @@ Calculates the duration between a given `PlainDateTime` and a specified date. ```lean example : Duration := - let startDate := date% 2023-1-1:05:10:20 - let endDate := date% 2023-3-15:05:10:20 + let startDate := datetime("2023-1-1:05:10:20") + let endDate := datetime("2023-3-15:05:10:20") endDate.since startDate ``` -/ diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 2a1eb3d4f019..7640f1a5d82a 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -143,8 +143,8 @@ def withDaysRollOver (dt : PlainDateTime) (days : Day.Ordinal) : PlainDateTime : { dt with date := PlainDate.rollOver dt.date.year dt.date.month days } /-- -Creates a new `PlainDateTime` by adjusting the month to the given `month` value. -The day remains unchanged, and any invalid days for the new month will be handled according to the `clip` behavior. +Creates a new `PlainDateTime` by adjusting the month to the given `month` value, with any +out-of-range days clipped to the nearest valid date. -/ @[inline] def withMonthClip (dt : PlainDateTime) (month : Month.Ordinal) : PlainDateTime := @@ -159,8 +159,8 @@ def withMonthRollOver (dt : PlainDateTime) (month : Month.Ordinal) : PlainDateTi { dt with date := PlainDate.rollOver dt.date.year month dt.date.day } /-- -Creates a new `PlainDateTime` by adjusting the year to the given `year` value. The month and day remain unchanged, -and any invalid days for the new year will be handled according to the `clip` behavior. +Creates a new `PlainDateTime` by adjusting the year to the given `year` value. The month and day +remain unchanged, with any out-of-range days clipped to the nearest valid date. -/ @[inline] def withYearClip (dt : PlainDateTime) (year : Year.Offset) : PlainDateTime := @@ -178,28 +178,28 @@ def withYearRollOver (dt : PlainDateTime) (year : Year.Offset) : PlainDateTime : Creates a new `PlainDateTime` by adjusting the `hour` component of its `time` to the given value. -/ @[inline] -def withHour (dt : PlainDateTime) (hour : Hour.Ordinal) : PlainDateTime := +def withHours (dt : PlainDateTime) (hour : Hour.Ordinal) : PlainDateTime := { dt with time := { dt.time with hour := hour } } /-- Creates a new `PlainDateTime` by adjusting the `minute` component of its `time` to the given value. -/ @[inline] -def withMinute (dt : PlainDateTime) (minute : Minute.Ordinal) : PlainDateTime := +def withMinutes (dt : PlainDateTime) (minute : Minute.Ordinal) : PlainDateTime := { dt with time := { dt.time with minute := minute } } /-- Creates a new `PlainDateTime` by adjusting the `second` component of its `time` to the given value. -/ @[inline] -def withSecond (dt : PlainDateTime) (second : Sigma Second.Ordinal) : PlainDateTime := +def withSeconds (dt : PlainDateTime) (second : Sigma Second.Ordinal) : PlainDateTime := { dt with time := { dt.time with second := second } } /-- Creates a new `PlainDateTime` by adjusting the `nano` component of its `time` to the given value. -/ @[inline] -def withNano (dt : PlainDateTime) (nano : Nanosecond.Ordinal) : PlainDateTime := +def withNanoseconds (dt : PlainDateTime) (nano : Nanosecond.Ordinal) : PlainDateTime := { dt with time := { dt.time with nano := nano } } /-- @@ -412,11 +412,9 @@ def nanosecond (dt : PlainDateTime) : Nanosecond.Ordinal := /-- Determines the era of the given `PlainDateTime` based on its year. -/ +@[inline] def era (date : PlainDateTime) : Year.Era := - if date.year.toInt ≥ 0 then - .ce - else - .bce + date.date.era /-- Checks if the `PlainDateTime` is in a leap year. diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 1cd9434a678d..7095f29877a9 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -37,7 +37,7 @@ instance : OfNat Timestamp n where ofNat := ⟨OfNat.ofNat n⟩ instance : ToString Timestamp where - toString s := toString s.val + toString s := toString s.val.toMilliseconds instance : Repr Timestamp where reprPrec s := reprPrec (toString s) @@ -47,6 +47,7 @@ Type class to transform to `Timestamp`. It's used internally to generate timesta some absolute date types. -/ class ToTimestamp (α : Type) where + /-- Transforms into a `Timestamp`. -/ @@ -64,38 +65,42 @@ Fetches the current duration from the system. opaque now : IO Timestamp /-- -Converts a `Timestamp` to a `Minute.Offset` +Converts a `Timestamp` to minutes as `Minute.Offset`. -/ +@[inline] def toMinutes (tm : Timestamp) : Minute.Offset := tm.val.second.ediv 60 /-- -Converts a `Timestamp` to a `Day.Offset` +Converts a `Timestamp` to days as `Day.Offset`. -/ +@[inline] def toDays (tm : Timestamp) : Day.Offset := tm.val.second.ediv 86400 /-- -Creates a new `Timestamp` out of `Second.Offset`. +Creates a `Timestamp` from a `Second.Offset` since the Unix epoch. -/ +@[inline] def ofSecondsSinceUnixEpoch (secs : Second.Offset) : Timestamp := ⟨Duration.ofSeconds secs⟩ /-- -Creates a new `Timestamp` out of `Nanosecond.Offset`. +Creates a `Timestamp` from a `Nanosecond.Offset` since the Unix epoch. -/ +@[inline] def ofNanosecondsSinceUnixEpoch (secs : Nanosecond.Offset) : Timestamp := ⟨Duration.ofNanoseconds secs⟩ /-- -Converts a `Timestamp` into its equivalent `Second.Offset`. +Converts a `Timestamp` to seconds as `Second.Offset`. -/ @[inline] def toSecondsSinceUnixEpoch (t : Timestamp) : Second.Offset := t.val.second /-- -Converts a `Timestamp` from a `Nanosecond.Offset` +Converts a `Timestamp` to nanoseconds as `Nanosecond.Offset`. -/ @[inline] def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := @@ -104,7 +109,7 @@ def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := nanos /-- -Calculates a `Timestamp` out of two `Timestamp`s. +Calculates the duration from the given `Timestamp` to the current time. -/ @[inline] def since (f : Timestamp) : IO Duration := do @@ -112,98 +117,115 @@ def since (f : Timestamp) : IO Duration := do return Std.Time.Duration.sub cur.val f.val /-- -Creates a duration out of the `Timestamp` since the unix epoch. +Returns the `Duration` represented by the `Timestamp` since the Unix epoch. -/ @[inline] def toDurationSinceUnixEpoch (tm : Timestamp) : Duration := tm.val /-- -Adds a `Nanosecond.Offset` to a `Timestamp` +Adds a `Nanosecond.Offset` to the given `Timestamp`. -/ @[inline] def addNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := ⟨t.val + s⟩ /-- -Adds a `Nanosecond.Offset` to a `Timestamp` +Subtracts a `Nanosecond.Offset` from the given `Timestamp`. -/ @[inline] def subNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := ⟨t.val - s⟩ /-- -Adds a `Second.Offset` to a `Timestamp` +Adds a `Second.Offset` to the given `Timestamp`. -/ @[inline] def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := ⟨t.val + s⟩ /-- -Subtracts a `Second.Offset` from a `Timestamp` +Subtracts a `Second.Offset` from the given `Timestamp`. -/ @[inline] def subSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := ⟨t.val - s⟩ /-- -Adds a `Minute.Offset` to a `Timestamp` +Adds a `Minute.Offset` to the given `Timestamp`. -/ @[inline] def addMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := ⟨t.val + m⟩ /-- -Subtracts a `Minute.Offset` from a `Timestamp` +Subtracts a `Minute.Offset` from the given `Timestamp`. -/ @[inline] def subMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := ⟨t.val - m⟩ /-- -Adds an `Hour.Offset` to a `Timestamp` +Adds an `Hour.Offset` to the given `Timestamp`. -/ @[inline] def addHours (t : Timestamp) (h : Hour.Offset) : Timestamp := ⟨t.val + h⟩ /-- -Subtracts an `Hour.Offset` from a `Timestamp` +Subtracts an `Hour.Offset` from the given `Timestamp`. -/ @[inline] def subHours (t : Timestamp) (h : Hour.Offset) : Timestamp := ⟨t.val - h⟩ /-- -Adds a `Day.Offset` to a `Timestamp` +Adds a `Day.Offset` to the given `Timestamp`. -/ @[inline] def addDays (t : Timestamp) (d : Day.Offset) : Timestamp := ⟨t.val + d⟩ /-- -Subtracts a `Day.Offset` from a `Timestamp` +Subtracts a `Day.Offset` from the given `Timestamp`. -/ @[inline] def subDays (t : Timestamp) (d : Day.Offset) : Timestamp := ⟨t.val - d⟩ /-- -Adds a `Week.Offset` to a `Timestamp` +Adds a `Week.Offset` to the given `Timestamp`. -/ @[inline] def addWeeks (t : Timestamp) (d : Week.Offset) : Timestamp := ⟨t.val + d⟩ /-- -Subtracts a `Week.Offset` from a `Timestamp` +Subtracts a `Week.Offset` from the given `Timestamp`. -/ @[inline] def subWeeks (t : Timestamp) (d : Week.Offset) : Timestamp := ⟨t.val - d⟩ +/-- +Adds a `Duration` to the given `Timestamp`. +-/ +@[inline] +def addDuration (t : Timestamp) (d : Duration) : Timestamp := + ⟨t.val + d⟩ + +/-- +Subtracts a `Duration` from the given `Timestamp`. +-/ +@[inline] +def subDuration (t : Timestamp) (d : Duration) : Timestamp := + ⟨t.val - d⟩ + instance : HAdd Timestamp Duration Timestamp where - hAdd x y := ⟨x.val + y⟩ + hAdd := addDuration + +instance : HSub Timestamp Duration Timestamp where + hSub := subDuration instance : HAdd Timestamp Day.Offset Timestamp where hAdd := addDays @@ -244,10 +266,4 @@ instance : HSub Timestamp Nanosecond.Offset Timestamp where instance : HSub Timestamp Timestamp Duration where hSub x y := x.val - y.val -instance : HAdd Timestamp Duration Timestamp where - hAdd x y := x.addNanoseconds y.toNanoseconds - -instance : HSub Timestamp Duration Timestamp where - hSub x y := x.subNanoseconds y.toNanoseconds - end Timestamp diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 1b867bfbde91..b2d2d945df82 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -34,7 +34,6 @@ structure Duration where proof : (second.val ≥ 0 ∧ nano.val ≥ 0) ∨ (second.val ≤ 0 ∧ nano.val ≤ 0) deriving Repr - instance : ToString Duration where toString s := let (sign, secs, nanos) := @@ -93,6 +92,13 @@ def ofNanoseconds (s : Nanosecond.Offset) : Duration := by | .negSucc m, .ofNat n, _, _ => Int.neg_le_neg (Int.tmod_nonneg (↑n) (Int.ofNat_le.mpr (Nat.zero_le (m + 1)))) | 0, n, _, _ => Int.eq_iff_le_and_ge.mp (Int.zero_tmod n) |>.left +/-- +Creates a new `Duration` out of `Millisecond.Offset`. +-/ +@[inline] +def ofMillisecond (s : Millisecond.Offset) : Duration := + ofNanoseconds (s.mul 1000000) + /-- Checks if the duration is zero seconds and zero nanoseconds. -/ @@ -101,14 +107,24 @@ def isZero (d : Duration) : Bool := d.second.val = 0 ∧ d.nano.val = 0 /-- -Converts a `Duration` from a `Nanosecond.Offset` +Converts a `Duration` to a `Second.Offset` -/ @[inline] def toSeconds (duration : Duration) : Second.Offset := duration.second /-- -Converts a `Duration` from a `Nanosecond.Offset` +Converts a `Duration` to a `Millisecond.Offset` +-/ +@[inline] +def toMilliseconds (duration : Duration) : Millisecond.Offset := + let secMillis := duration.second.mul 1000 + let nanosMillis := duration.nano.ediv 1000000 (by decide) + let millis := secMillis + (.ofInt nanosMillis.val) + millis + +/-- +Converts a `Duration` to a `Nanosecond.Offset` -/ @[inline] def toNanoseconds (duration : Duration) : Nanosecond.Offset := diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index ef6210d83456..7fb9acd57331 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -17,121 +17,121 @@ set_option linter.all true /-- The ISO8601 format, which is always 24 or 27 characters long, used for representing date and time in -a standardized format. The format follows the pattern `yyyy-MM-dd:HH::mm:ss.SSSZ`. +a standardized format. The format follows the pattern `yyyy-MM-dd'T'HH:mm:ssZ`. -/ -def iso8601 : Format .any := datespec("yyyy-MM-dd:HH::mm:ss.SSSZ") +def iso8601 : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm.ssZ") /-- The americanDate format, which follows the pattern `MM-dd-yyyy`. -/ -def americanDate : Format .any := datespec("MM-dd-yyyy") +def americanDate : GenericFormat .any := datespec("MM-dd-yyyy") /-- The europeanDate format, which follows the pattern `dd-MM-yyyy`. -/ -def europeanDate : Format .any := datespec("dd-MM-yyyy") +def europeanDate : GenericFormat .any := datespec("dd-MM-yyyy") /-- The time12Hour format, which follows the pattern `hh:mm:ss aa` for representing time in a 12-hour clock format with an upper case AM/PM marker. -/ -def time12Hour : Format .any := datespec("hh:mm:ss aa") +def time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") /-- -The Time24Hour format, which follows the pattern `HH:mm:ss:SSSSSSSSS` for representing time +The Time24Hour format, which follows the pattern `HH:mm:ss` for representing time in a 24-hour clock format. -/ -def time24Hour : Format .any := datespec("HH:mm:ss.SSSSSSSSS") +def time24Hour : GenericFormat .any := datespec("HH:mm:ss") /-- The DateTimeZone24Hour format, which follows the pattern `yyyy-MM-dd:HH:mm:ss.SSSSSSSSS` for representing date, time, and time zone. -/ -def dateTime24Hour : Format (.only .GMT) := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSS") +def dateTime24Hour : GenericFormat (.only .GMT) := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSS") /-- The DateTimeWithZone format, which follows the pattern `yyyy-MM-dd:HH:mm:ss.SSSSSSSSSZZZ` for representing date, time, and time zone. -/ -def dateTimeWithZone : Format .any := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSSZZZ") +def dateTimeWithZone : GenericFormat .any := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSSZZZ") /-- -The Time24Hour format, which follows the pattern `HH:mm:ss` for representing time +The leanTime24Hour format, which follows the pattern `HH:mm:ss.SSSSSSSSS` for representing time in a 24-hour clock format. It uses the default value that can be parsed with the notation of dates. -/ -def leanTime24Hour : Format .any := datespec("HH:mm:ss.SSSSSSSSS") +def leanTime24Hour : GenericFormat .any := datespec("HH:mm:ss.SSSSSSSSS") /-- -The Time24Hour format, which follows the pattern `HH:mm:ss` for representing time +The leanTime24HourNoNanos format, which follows the pattern `HH:mm:ss` for representing time in a 24-hour clock format. It uses the default value that can be parsed with the notation of dates. -/ -def leanTime24HourNoNanos : Format .any := datespec("HH:mm:ss") +def leanTime24HourNoNanos : GenericFormat .any := datespec("HH:mm:ss") /-- -The DateTimeZone24Hour format, which follows the pattern `YYYY-MM-DD HH:mm:ss:sssssssss` for +The leanDateTime24Hour format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSS` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTime24Hour : Format (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSS") +def leanDateTime24Hour : GenericFormat (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSS") /-- -The DateTime24HourNoNanos format, which follows the pattern `YYYY-MM-DD HH:mm:ss` for +The leanDateTime24HourNoNanos format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ss` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTime24HourNoNanos : Format (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss") +def leanDateTime24HourNoNanos : GenericFormat (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss") /-- -The DateTimeWithZone format, which follows the pattern `YYYY-MM-DD HH:mm:ss.SSSSSSSSS` +The DateTimeWithZone format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithZone : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ") +def leanDateTimeWithZone : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ") /-- -The DateTimeWithZoneNoNanos format, which follows the pattern `YYYY-MM-DD HH:mm:ssZZZZZ` +The DateTimeWithZoneNoNanos format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ssZZZZZ` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithZoneNoNanos : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ssZZZZZ") +def leanDateTimeWithZoneNoNanos : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ssZZZZZ") /-- -The Lean Date format, which follows the pattern `yyyy-MM-DD`. It uses the default value that can be parsed with the +The Lean Date format, which follows the pattern `yyyy-MM-dd`. It uses the default value that can be parsed with the notation of dates. -/ -def leanDate : Format .any := datespec("yyyy-MM-dd") +def leanDate : GenericFormat .any := datespec("yyyy-MM-dd") /-- -The SQLDate format, which follows the pattern `yyyy-MM-DD` and is commonly used +The SQLDate format, which follows the pattern `yyyy-MM-dd` and is commonly used in SQL databases to represent dates. -/ -def sqlDate : Format .any := datespec("yyyy-MM-dd") +def sqlDate : GenericFormat .any := datespec("yyyy-MM-dd") /-- The LongDateFormat, which follows the pattern `EEEE, MMMM D, yyyy HH:mm:ss` for representing a full date and time with the day of the week and month name. -/ -def longDateFormat : Format (.only .GMT) := datespec("EEEE, MMMM D, yyyy HH:mm:ss") +def longDateFormat : GenericFormat (.only .GMT) := datespec("EEEE, MMMM D, yyyy HH:mm:ss") /-- The AscTime format, which follows the pattern `EEE MMM d HH:mm:ss yyyy`. This format is often used in older systems for logging and time-stamping events. -/ -def ascTime : Format (.only .GMT) := datespec("EEE MMM d HH:mm:ss yyyy") +def ascTime : GenericFormat (.only .GMT) := datespec("EEE MMM d HH:mm:ss yyyy") /-- The RFC822 format, which follows the pattern `eee, dd MMM yyyy HH:mm:ss ZZZ`. This format is used in email headers and HTTP headers. -/ -def rfc822 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def rfc822 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") /-- The RFC850 format, which follows the pattern `eee, dd-MMM-YY HH:mm:ss ZZZ`. This format is an older standard for representing date and time in headers. -/ -def rfc850 : Format .any := datespec("eee, dd-MM-yyyy HH:mm:ss ZZZ") +def rfc850 : GenericFormat .any := datespec("eee, dd-MM-yyyy HH:mm:ss ZZZ") end Formats @@ -141,7 +141,7 @@ namespace TimeZone Parses a string into a `TimeZone` object. The input string must be in the format `"VV ZZZZZ"`. -/ def fromTimeZone (input : String) : Except String TimeZone := do - let spec : Format .any := datespec("VV ZZZZZ") + let spec : GenericFormat .any := datespec("VV ZZZZZ") spec.parseBuilder (λid off => some (TimeZone.mk off id "Unknown" false)) input namespace Offset @@ -150,7 +150,7 @@ namespace Offset Parses a string representing an offset into an `Offset` object. The input string must follow the `"xxx"` format. -/ def fromOffset (input : String) : Except String Offset := do - let spec : Format .any := datespec("xxx") + let spec : GenericFormat .any := datespec("xxx") spec.parseBuilder some input end Offset @@ -162,7 +162,7 @@ namespace PlainDate Formats a `PlainDate` using a specific format. -/ def format (date : PlainDate) (format : String) : String := - let format : Except String (Format .any) := Format.spec format + let format : Except String (GenericFormat .any) := GenericFormat.spec format match format with | .error err => s!"error: {err}" | .ok res => @@ -177,40 +177,40 @@ def format (date : PlainDate) (format : String) : String := | none => "invalid time" /-- -Parses a date string in the American format (`MM/DD/yyyy`) and returns a `PlainDate`. +Parses a date string in the American format (`MM-dd-yyyy`) and returns a `PlainDate`. -/ def fromAmericanDateString (input : String) : Except String PlainDate := do Formats.americanDate.parseBuilder (fun m d y => PlainDate.ofYearMonthDay y m d) input /-- -Converts a Date in the American format (`MM/DD/yyyy`) into a `String`. +Converts a date in the American format (`MM-dd-yyyy`) into a `String`. -/ def toAmericanDateString (input : PlainDate) : String := Formats.americanDate.formatBuilder input.month input.day input.year /-- -Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. +Parses a date string in the SQL format (`yyyy-MM-dd`) and returns a `PlainDate`. -/ def fromSQLDateString (input : String) : Except String PlainDate := do Formats.sqlDate.parseBuilder (PlainDate.ofYearMonthDay) input /-- -Converts a Date in the SQL format (`YYYY-MM-DD`) into a `String`. +Converts a date in the SQL format (`yyyy-MM-dd`) into a `String`. -/ def toSQLDateString (input : PlainDate) : String := Formats.sqlDate.formatBuilder input.year input.month input.day /-- -Parses a date string into a `PlainDate` object. The input string must be in the SQL date format. +Parses a date string in the Lean format (`yyyy-MM-dd`) and returns a `PlainDate`. -/ def fromLeanDateString (input : String) : Except String PlainDate := do - Formats.sqlDate.parseBuilder (PlainDate.ofYearMonthDay) input + Formats.leanDate.parseBuilder (PlainDate.ofYearMonthDay) input /-- -Converts a Date in the Lean format (`YYYY-MM-DD`) into a `String` with the format `date% YYY-MM-DD`. +Converts a date in the Lean format (`yyyy-MM-dd`) into a `String`. -/ def toLeanDateString (input : PlainDate) : String := - Formats.sqlDate.formatBuilder input.year input.month input.day + Formats.leanDate.formatBuilder input.year input.month input.day /-- Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `PlainDate`. @@ -233,7 +233,7 @@ namespace PlainTime Formats a `PlainTime` using a specific format. -/ def format (time : PlainTime) (format : String) : String := - let format : Except String (Format .any) := Format.spec format + let format : Except String (GenericFormat .any) := GenericFormat.spec format match format with | .error err => s!"error: {err}" | .ok res => @@ -248,26 +248,26 @@ def format (time : PlainTime) (format : String) : String := | none => "invalid time" /-- -Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. +Parses a time string in the 24-hour format (`HH:mm:ss`) and returns a `PlainTime`. -/ def fromTime24Hour (input : String) : Except String PlainTime := - Formats.time24Hour.parseBuilder (fun h m s n => some (PlainTime.ofHourMinuteSecondsNano h m s.snd n)) input + Formats.time24Hour.parseBuilder (fun h m s => some (PlainTime.ofHourMinuteSeconds h m s.snd)) input /-- -Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). +Formats a `PlainTime` value into a 24-hour format string (`HH:mm:ss`). -/ def toTime24Hour (input : PlainTime) : String := - Formats.time24Hour.formatBuilder input.hour input.minute input.second input.nano + Formats.time24Hour.formatBuilder input.hour input.minute input.second /-- -Parses a time string in the 24-hour format (`hh:mm:ss.sssssssss`) and returns a `PlainTime`. +Parses a time string in the lean 24-hour format (`HH:mm:ss.SSSSSSSSS` or `HH:mm:ss`) and returns a `PlainTime`. -/ def fromLeanTime24Hour (input : String) : Except String PlainTime := Formats.leanTime24Hour.parseBuilder (fun h m s n => some (PlainTime.ofHourMinuteSecondsNano h m s.snd n)) input <|> Formats.leanTime24HourNoNanos.parseBuilder (fun h m s => some (PlainTime.ofHourMinuteSecondsNano h m s.snd 0)) input /-- -Formats a `PlainTime` value into a 24-hour format string (`hh:mm:ss.sssssssss`). +Formats a `PlainTime` value into a 24-hour format string (`HH:mm:ss.SSSSSSSSS`). -/ def toLeanTime24Hour (input : PlainTime) : String := Formats.leanTime24Hour.formatBuilder input.hour input.minute input.second input.nano @@ -309,7 +309,7 @@ namespace ZonedDateTime Formats a `ZonedDateTime` using a specific format. -/ def format (data: ZonedDateTime) (format : String) : String := - let format : Except String (Format .any) := Format.spec format + let format : Except String (GenericFormat .any) := GenericFormat.spec format match format with | .error err => s!"error: {err}" | .ok res => res.format data.snd @@ -397,7 +397,7 @@ namespace PlainDateTime Formats a `PlainDateTime` using a specific format. -/ def format (date : PlainDateTime) (format : String) : String := - let format : Except String (Format .any) := Format.spec format + let format : Except String (GenericFormat .any) := GenericFormat.spec format match format with | .error err => s!"error: {err}" | .ok res => @@ -490,7 +490,7 @@ namespace DateTime Formats a `DateTime` using a specific format. -/ def format (data: DateTime tz) (format : String) : String := - let format : Except String (Format .any) := Format.spec format + let format : Except String (GenericFormat .any) := GenericFormat.spec format match format with | .error err => s!"error: {err}" | .ok res => res.format data diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 1c41bd8a9b83..46447af2ec3c 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -570,7 +570,7 @@ end Awareness /-- A specification on how to format a data or parse some string. -/ -structure Format (awareness : Awareness) where +structure GenericFormat (awareness : Awareness) where /-- The format that is not aware of the timezone. @@ -731,19 +731,19 @@ private def formatMarkerNarrow (marker : HourMarker) : String := | .am => "A" | .pm => "P" -private def formatPeriodOfDayLong : Day.Ordinal.Period → String +private def formatPeriodOfDayLong : Day.Period → String | .morning => "in the morning" | .afternoon => "in the afternoon" | .evening => "in the evening" | .night => "at night" -private def formatPeriodOfDayShort : Day.Ordinal.Period → String +private def formatPeriodOfDayShort : Day.Period → String | .morning => "AM" | .afternoon => "PM" | .evening => "Eve" | .night => "Night" -private def formatPeriodOfDayNarrow : Day.Ordinal.Period → String +private def formatPeriodOfDayNarrow : Day.Period → String | .morning => "M" | .afternoon => "A" | .evening => "E" @@ -776,7 +776,7 @@ private def TypeFormat : Modifier → Type | .eorc _ => Weekday | .F _ => Week.Ordinal.OfMonth | .a _ => HourMarker - | .B _ => Day.Ordinal.Period + | .B _ => Day.Period | .h _ => Bounded.LE 1 12 | .K _ => Bounded.LE 0 11 | .k _ => Bounded.LE 1 24 @@ -1055,19 +1055,19 @@ private def parseMarkerNarrow : Parser HourMarker := pstring "A" *> pure HourMarker.am <|> pstring "P" *> pure HourMarker.pm -private def parsePeriodOfDayLong : Parser Day.Ordinal.Period +private def parsePeriodOfDayLong : Parser Day.Period := pstring "in the morning" *> pure .morning <|> pstring "in the afternoon" *> pure .afternoon <|> pstring "in the evening" *> pure .evening <|> pstring "at night" *> pure .night -private def parsePeriodOfDayShort : Parser Day.Ordinal.Period +private def parsePeriodOfDayShort : Parser Day.Period := pstring "AM" *> pure .morning <|> pstring "PM" *> pure .afternoon <|> pstring "Eve" *> pure .evening <|> pstring "Night" *> pure .night -private def parsePeriodOfDayNarrow : Parser Day.Ordinal.Period +private def parsePeriodOfDayNarrow : Parser Day.Period := pstring "M" *> pure .morning <|> pstring "A" *> pure .afternoon <|> pstring "E" *> pure .evening @@ -1255,7 +1255,7 @@ private def FormatType (result : Type) : FormatString → Type | .string _ :: xs => (FormatType result xs) | [] => result -namespace Format +namespace GenericFormat private structure DateBuilder where G : Option Year.Era := none @@ -1270,7 +1270,7 @@ private structure DateBuilder where eorc : Option Weekday := none F : Option Week.Ordinal.OfMonth := none a : Option HourMarker := none - B : Option Day.Ordinal.Period := none + B : Option Day.Period := none h : Option (Bounded.LE 1 12) := none K : Option (Bounded.LE 0 11) := none k : Option (Bounded.LE 1 24) := none @@ -1388,25 +1388,25 @@ private def parseWithDate (date : DateBuilder) (mod : FormatPart) : Parser DateB | .string s => pstring s *> pure date /-- -Constructs a new `Format` specification for a date-time string. Modifiers can be combined to create +Constructs a new `GenericFormat` specification for a date-time string. Modifiers can be combined to create custom formats, such as "YYYY, MMMM, D". -/ -def spec (input : String) : Except String (Format tz) := do +def spec (input : String) : Except String (GenericFormat tz) := do let string ← specParser.run input return ⟨string⟩ /-- -Builds a new `Format` specification for a Date-time, panics if the input string results in an error. +Builds a `GenericFormat` from the input string. If parsing fails, it will panic -/ -def spec! (input : String) : Format tz := +def spec! (input : String) : GenericFormat tz := match specParser.run input with | .ok res => ⟨res⟩ | .error res => panic! res /-- -Formats the date using the format into a String. +Formats a `DateTime` value into a string using the given `GenericFormat`. -/ -def format (format : Format aw) (date : DateTime tz) : String := +def format (format : GenericFormat aw) (date : DateTime tz) : String := let mapper (part : FormatPart) := match aw with | .any => formatPartWithDate date part @@ -1442,39 +1442,37 @@ def builderParser (format: FormatString) (func: FormatType (Option α) format) : go format func /-- -Parses the input string into a `ZoneDateTime` +Parses the input string into a `ZoneDateTime`. -/ -def parse (format : Format aw) (input : String) : Except String aw.type := +def parse (format : GenericFormat aw) (input : String) : Except String aw.type := (parser format.string aw <* eof).run input /-- -Parses the input string into a `ZoneDateTime`, panics if its wrong +Parses the input string into a `ZoneDateTime` and panics if its wrong. -/ -def parse! (format : Format aw) (input : String) : aw.type := +def parse! (format : GenericFormat aw) (input : String) : aw.type := match parse format input with | .ok res => res | .error err => panic! err /-- -Parses and instead of using a builder to build a date, it uses a builder function instead. +Parses an input string using a builder function to produce a value. -/ -def parseBuilder (format : Format aw) (builder : FormatType (Option α) format.string) (input : String) : Except String α := +def parseBuilder (format : GenericFormat aw) (builder : FormatType (Option α) format.string) (input : String) : Except String α := (builderParser format.string builder).run input - /-- -Parses and instead of using a builder to build a date, it uses a builder function instead. +Parses an input string using a builder function, panicking on errors. -/ -def parseBuilder! [Inhabited α] (format : Format aw) (builder : FormatType (Option α) format.string) (input : String) : α := +def parseBuilder! [Inhabited α] (format : GenericFormat aw) (builder : FormatType (Option α) format.string) (input : String) : α := match parseBuilder format builder input with | .ok res => res | .error err => panic! err - /-- -Formats the date using the format into a String. +Formats the date using the format into a String, using a `getInfo` function to get the information needed to build the `String`. -/ -def formatGeneric (format : Format aw) (getInfo : (typ : Modifier) → Option (TypeFormat typ)) : Option String := +def formatGeneric (format : GenericFormat aw) (getInfo : (typ : Modifier) → Option (TypeFormat typ)) : Option String := let rec go (data : String) : (format : FormatString) → Option String | .modifier x :: xs => do go (data ++ formatWith x (← getInfo x)) xs | .string x :: xs => go (data ++ x) xs @@ -1482,11 +1480,33 @@ def formatGeneric (format : Format aw) (getInfo : (typ : Modifier) → Option (T go "" format.string /-- -Formats the date using the format into a String. +Constructs a `FormatType` function to format a date into a string using a `GenericFormat`. -/ -def formatBuilder (format : Format aw) : FormatType String format.string := +def formatBuilder (format : GenericFormat aw) : FormatType String format.string := let rec go (data : String) : (format : FormatString) → FormatType String format | .modifier x :: xs => fun res => go (data ++ formatWith x res) xs | .string x :: xs => go (data ++ x) xs | [] => data go "" format.string + +end GenericFormat + +/-- +Typeclass for formatting and parsing values with the given format type. +-/ +class Format (f : Type) (typ : Type → f → Type) where + /-- + Converts a format `f` into a string. + -/ + format : (fmt : f) → typ String fmt + + /-- + Parses a string into a format using the provided format type `f`. + -/ + parse : (fmt : f) → typ (Option α) fmt → String → Except String α + +instance : Format (GenericFormat aw) (FormatType · ·.string) where + format := GenericFormat.formatBuilder + parse := GenericFormat.parseBuilder + +end Time diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 60ee6df8b127..33ece36c104f 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -14,6 +14,8 @@ namespace Std namespace Time open Lean Parser Command Std +set_option linter.all true + private def convertText : Text → MacroM (TSyntax `term) | .short => `(Std.Time.Text.short) | .full => `(Std.Time.Text.full) @@ -95,55 +97,97 @@ private def convertFormatPart : FormatPart → MacroM (TSyntax `term) | .string s => `(.string $(Syntax.mkStrLit s)) | .modifier mod => do `(.modifier $(← convertModifier mod)) -def syntaxNat (n : Nat) : MacroM (TSyntax `term) := do +private def syntaxNat (n : Nat) : MacroM (TSyntax `term) := do let info ← MonadRef.mkInfoFromRefPos pure { raw := Syntax.node1 info `num (Lean.Syntax.atom info (toString n)) } -def syntaxString (n : String) : MacroM (TSyntax `term) := do +private def syntaxString (n : String) : MacroM (TSyntax `term) := do let info ← MonadRef.mkInfoFromRefPos pure { raw := Syntax.node1 info `str (Lean.Syntax.atom info (toString n)) } -def syntaxInt (n : Int) : MacroM (TSyntax `term) := do +private def syntaxInt (n : Int) : MacroM (TSyntax `term) := do match n with | .ofNat n => `(Int.ofNat $(Syntax.mkNumLit <| toString n)) | .negSucc n => `(Int.negSucc $(Syntax.mkNumLit <| toString n)) -def syntaxBounded (n : Int) : MacroM (TSyntax `term) := do +private def syntaxBounded (n : Int) : MacroM (TSyntax `term) := do `(Std.Time.Internal.Bounded.LE.ofNatWrapping $(← syntaxInt n) (by decide)) -def syntaxVal (n : Int) : MacroM (TSyntax `term) := do +private def syntaxVal (n : Int) : MacroM (TSyntax `term) := do `(Std.Time.Internal.UnitVal.ofInt $(← syntaxInt n)) -def convertOffset (offset : Std.Time.TimeZone.Offset) : MacroM (TSyntax `term) := do - `(Std.Time.TimeZone.Offset.mk $(← syntaxVal offset.hour.val) $(← syntaxVal offset.second.val)) +private def convertOffset (offset : Std.Time.TimeZone.Offset) : MacroM (TSyntax `term) := do + `(Std.Time.TimeZone.Offset.ofSeconds $(← syntaxVal offset.second.val)) -def convertTimezone (tz : Std.Time.TimeZone) : MacroM (TSyntax `term) := do +private def convertTimezone (tz : Std.Time.TimeZone) : MacroM (TSyntax `term) := do `(Std.Time.TimeZone.mk $(← convertOffset tz.offset) $(Syntax.mkStrLit tz.name) $(Syntax.mkStrLit tz.abbreviation) false) -def convertPlainDate (d : Std.Time.PlainDate) : MacroM (TSyntax `term) := do +private def convertPlainDate (d : Std.Time.PlainDate) : MacroM (TSyntax `term) := do `(Std.Time.PlainDate.clip $(← syntaxInt d.year) $(← syntaxBounded d.month.val) $(← syntaxBounded d.day.val)) -def convertPlainTime (d : Std.Time.PlainTime) : MacroM (TSyntax `term) := do +private def convertPlainTime (d : Std.Time.PlainTime) : MacroM (TSyntax `term) := do `(Std.Time.PlainTime.mk $(← syntaxBounded d.hour.val) $(← syntaxBounded d.minute.val) ⟨true, $(← syntaxBounded d.second.snd.val)⟩ $(← syntaxBounded d.nano.val)) -def convertPlainDateTime (d : Std.Time.PlainDateTime) : MacroM (TSyntax `term) := do +private def convertPlainDateTime (d : Std.Time.PlainDateTime) : MacroM (TSyntax `term) := do `(Std.Time.PlainDateTime.mk $(← convertPlainDate d.date) $(← convertPlainTime d.time)) -def convertZonedDateTime (d : Std.Time.ZonedDateTime) : MacroM (TSyntax `term) := do +private def convertZonedDateTime (d : Std.Time.ZonedDateTime) : MacroM (TSyntax `term) := do `(Std.Time.ZonedDateTime.mk $(← convertTimezone d.timezone) (DateTime.ofLocalDateTime $(← convertPlainDateTime d.snd.date.get) $(← convertTimezone d.timezone))) +/-- +Defines a syntax for zoned datetime values. It expects a string representing a datetime with +timezone information. + +Example: +`zoned("2024-10-13T15:00:00-03:00")` +-/ syntax "zoned(" str ")" : term +/-- +Defines a syntax for datetime values without timezone. The input should be a string in an +ISO8601-like format. + +Example: +`datetime("2024-10-13T15:00:00")` +-/ syntax "datetime(" str ")" : term +/-- +Defines a syntax for date-only values. The input string represents a date in formats like "YYYY-MM-DD". + +Example: +`date("2024-10-13")` +-/ syntax "date(" str ")" : term +/-- +Defines a syntax for time-only values. The string should represent a time, either in 24-hour or +12-hour format. + +Example: +`time("15:00:00")` or `time("03:00:00 PM")` +-/ syntax "time(" str ")" : term +/-- +Defines a syntax for UTC offset values. The string should indicate the time difference from UTC +(e.g., "-03:00"). + +Example: +`offset("-03:00")` +-/ syntax "offset(" str ")" : term +/-- +Defines a syntax for timezone identifiers. The input string should be a valid timezone name or +abbreviation. + +Example: +`timezone("America/Sao_Paulo")` +-/ syntax "timezone(" str ")" : term + macro_rules | `(zoned( $date:str )) => do match ZonedDateTime.fromLeanDateTimeWithZoneString date.getString with diff --git a/src/Std/Time/Notation/Spec.lean b/src/Std/Time/Notation/Spec.lean index 72d44dca9748..4b49af7bbb2e 100644 --- a/src/Std/Time/Notation/Spec.lean +++ b/src/Std/Time/Notation/Spec.lean @@ -103,7 +103,7 @@ syntax "datespec(" str ")" : term macro_rules | `(datespec( $format_string:str )) => do let input := format_string.getString - let format : Except String (Format .any) := Format.spec input + let format : Except String (GenericFormat .any) := GenericFormat.spec input match format with | .ok res => let alts ← res.string.mapM convertFormatPart diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index eaaa2f6577d1..7b0246eeaf58 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -43,7 +43,7 @@ def toAbsolute (marker : HourMarker) (time : Bounded.LE 1 12) : Hour.Ordinal := | .pm => if time.val = 12 then 12 else time.add 12 |>.emod 24 (by decide) /-- -Converts a 24-hour clock time to a 12-hour clock time with an `HourMarker`. +Converts a 24-hour clock time to a 12-hour clock time with a `HourMarker`. -/ def toRelative (hour : Hour.Ordinal) : Bounded.LE 1 12 × HourMarker := if h₀ : hour.val = 0 then @@ -53,8 +53,8 @@ def toRelative (hour : Hour.Ordinal) : Bounded.LE 1 12 × HourMarker := (⟨12, by decide⟩, .pm) else Int.ne_iff_lt_or_gt.mp h₀ |>.by_cases - (λh => nomatch (Int.not_le.mpr h) hour.property.left) - (λh => (⟨hour.val, And.intro h h₁⟩, .am)) + (nomatch Int.not_le.mpr · <| hour.property.left) + (⟨hour.val, And.intro · h₁⟩, .am) else let h := Int.not_le.mp h₁ let t := hour |>.truncateBottom h |>.sub 12 diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 4658c7d6227a..232651963c81 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -48,20 +48,20 @@ instance : BEq PlainTime where namespace PlainTime /-- -Creates a `PlainTime` value representing midnight (00:00:00.000). +Creates a `PlainTime` value representing midnight (00:00:00.000000000). -/ def midnight : PlainTime := ⟨0, 0, ⟨true, 0⟩, 0⟩ /-- -Creates a `PlainTime` value from hours, minutes, seconds and nanoseconds. +Creates a `PlainTime` value from the provided hours, minutes, seconds and nanoseconds components. -/ @[inline] def ofHourMinuteSecondsNano (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal leap) (nano : Nanosecond.Ordinal) : PlainTime := ⟨hour, minute, Sigma.mk leap second, nano⟩ /-- -Creates a `PlainTime` value from hours, minutes, and seconds. +Creates a `PlainTime` value from the provided hours, minutes, and seconds. -/ @[inline] def ofHourMinuteSeconds (hour : Hour.Ordinal) (minute : Minute.Ordinal) (second : Second.Ordinal leap) : PlainTime := @@ -133,6 +133,20 @@ Creates a `PlainTime` value from a total number of seconds. def ofSeconds (secs : Second.Offset) : PlainTime := ofNanoseconds secs.toNanoseconds +/-- +Creates a `PlainTime` value from a total number of minutes. +-/ +@[inline] +def ofMinutes (secs : Minute.Offset) : PlainTime := + ofNanoseconds secs.toNanoseconds + +/-- +Creates a `PlainTime` value from a total number of hours. +-/ +@[inline] +def ofHours (hour : Hour.Offset) : PlainTime := + ofNanoseconds hour.toNanoseconds + /-- Adds seconds to a `PlainTime`. -/ @@ -208,28 +222,28 @@ def subMilliseconds (time : PlainTime) (millisToSub : Millisecond.Offset) : Plai Creates a new `PlainTime` by adjusting the `second` component to the given value. -/ @[inline] -def withSecond (pt : PlainTime) (second : Sigma Second.Ordinal) : PlainTime := +def withSeconds (pt : PlainTime) (second : Sigma Second.Ordinal) : PlainTime := { pt with second := second } /-- Creates a new `PlainTime` by adjusting the `minute` component to the given value. -/ @[inline] -def withMinute (pt : PlainTime) (minute : Minute.Ordinal) : PlainTime := +def withMinutes (pt : PlainTime) (minute : Minute.Ordinal) : PlainTime := { pt with minute := minute } /-- Creates a new `PlainTime` by adjusting the `nano` component to the given value. -/ @[inline] -def withNano (pt : PlainTime) (nano : Nanosecond.Ordinal) : PlainTime := +def withNanoseconds (pt : PlainTime) (nano : Nanosecond.Ordinal) : PlainTime := { pt with nano := nano } /-- Creates a new `PlainTime` by adjusting the `hour` component to the given value. -/ @[inline] -def withHour (pt : PlainTime) (hour : Hour.Ordinal) : PlainTime := +def withHours (pt : PlainTime) (hour : Hour.Ordinal) : PlainTime := { pt with hour := hour } instance : HAdd PlainTime Nanosecond.Offset PlainTime where diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index a09c5c4d94cf..e5a6eac664ec 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -51,7 +51,7 @@ instance : OfNat Offset n := namespace Ordinal /-- -Converts an `Ordinal` into a relative month in the range of 1 to 12. +Converts an `Ordinal` into a relative hour in the range of 1 to 12. -/ def toRelative (ordinal : Ordinal) : Bounded.LE 1 12 := (ordinal.add 11).emod 12 (by decide) |>.add 1 diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 4ffe993743c5..e900df7ab78d 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -38,8 +38,7 @@ instance {x y : Ordinal} : Decidable (x < y) := inferInstanceAs (Decidable (x.val < y.val)) /-- -`Offset` represents a duration offset in milliseconds. It is defined as an `Int` value, -where each unit corresponds to one millisecond. +`Offset` represents a duration offset in milliseconds. -/ def Offset : Type := UnitVal (1 / 1000) deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index 59c8d31be401..8e0447b05a26 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -38,7 +38,7 @@ instance {x y : Ordinal} : Decidable (x < y) := inferInstanceAs (Decidable (x.val < y.val)) /-- -`Offset` represents a duration offset in minutes. It is defined as an `Int` value. +`Offset` represents a duration offset in minutes. -/ def Offset : Type := UnitVal 60 deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index 8d635c5a8f50..c56bcda46f6c 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -37,7 +37,7 @@ instance {x y : Ordinal} : Decidable (x < y) := inferInstanceAs (Decidable (x.val < y.val)) /-- -`Offset` represents a time offset in nanoseconds and is defined as an `Int`. +`Offset` represents a time offset in nanoseconds. -/ def Offset : Type := UnitVal (1 / 1000000000) deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index cf86e6f76c99..13ddc12d6915 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -60,17 +60,17 @@ instance : OfNat Offset n := namespace Offset /-- -Creates an `Offset` from a natural number. +Creates an `Second.Offset` from a natural number. -/ @[inline] -def ofNat (data : Nat) : Offset := +def ofNat (data : Nat) : Second.Offset := UnitVal.mk data /-- -Creates an `Offset` from an integer. +Creates an `Second.Offset` from an integer. -/ @[inline] -def ofInt (data : Int) : Offset := +def ofInt (data : Int) : Second.Offset := UnitVal.mk data end Offset @@ -101,10 +101,10 @@ def ofFin (data : Fin (if leap then 61 else 60)) : Ordinal leap := | false => Bounded.LE.ofFin data /-- -Converts an `Ordinal` to an `Offset`. +Converts an `Ordinal` to an `Second.Offset`. -/ @[inline] -def toOffset (ordinal : Ordinal leap) : Offset := +def toOffset (ordinal : Ordinal leap) : Second.Offset := UnitVal.ofInt ordinal.val end Ordinal diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 73a0eddc68c3..55b2b7850a24 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -14,6 +14,8 @@ namespace Time namespace PlainDateTime +set_option linter.all true + /-- Creaates a new `PlainDateTime` out of a `Timestamp` and a `TimeZone`. -/ @@ -22,7 +24,7 @@ def ofTimestamp (stamp : Timestamp) (tz : TimeZone) : PlainDateTime := PlainDateTime.ofUTCTimestamp stamp /-- -Get the current monotonic time. +Get the current time. -/ @[inline] def now : IO PlainDateTime := do @@ -35,7 +37,7 @@ end PlainDateTime namespace DateTime /-- -Converts a `PlainDate` to a `DateTime` +Converts a `PlainDate` with a `TimeZone` to a `DateTime` -/ @[inline] def ofPlainDate (pd : PlainDate) (tz : TimeZone) : DateTime tz := @@ -46,10 +48,10 @@ Converts a `DateTime` to a `PlainDate` -/ @[inline] def toPlainDate (dt : DateTime tz) : PlainDate := - Timestamp.toPlainDateAssumingUTC dt.toTimestamp + Timestamp.toPlainDateAssumingUTC dt.toUTCTimestamp /-- -Converts a `PlainTime` to a `DateTime` +Converts a `PlainTime` and a `TimeZone` to a `DateTime` -/ @[inline] def ofPlainTime (pt : PlainTime) (tz : TimeZone) : DateTime tz := @@ -76,7 +78,7 @@ example : Duration := -/ @[inline] def since [ToTimestamp α] (date : DateTime tz) (since : α) : Duration := - let date := date.toTimestamp + let date := date.toUTCTimestamp let since := ToTimestamp.toTimestamp since Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch @@ -141,9 +143,8 @@ def example : Duration := -/ @[inline] def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := - let date := date.toTimestamp + let date := date.toUTCTimestamp let since := ToTimestamp.toTimestamp since Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch end ZonedDateTime - diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index 68400253e900..8577a47cdbcc 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -15,29 +15,22 @@ open TimeZone.ZoneRules set_option linter.all true -/- -Apply leap seconds rules and transition rules on a UTC Timestamp to make it aware of the timezone. --/ --- def applyLeapSecondsOnUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO Timestamp := do --- (applyLeapSeconds tm ·) <$> Database.localRules db - /-- -Gets the TimeZone at the local timezone. +Gets the `TimeZone` at the local timezone. -/ -def getPlainTimeZoneAt [Database α] (db : α) (tm : Timestamp) : IO TimeZone := do +def getLocalTimeZoneAt [Database α] (db : α) (tm : Timestamp) : IO TimeZone := do (IO.ofExcept <| timezoneAt · tm) =<< Database.localRules db /-- -Gets the TimeZone at a timezone. +Gets the TimeZone at a timezone using a `Database`. -/ def getTimeZoneAt [Database α] (db : α) (id : String) (tm : Timestamp) : IO TimeZone := do (IO.ofExcept <| timezoneAt · tm) =<< Database.load db id /-- -Get the local ZonedDataTime given a UTC `Timestamp`. +Get the local `ZonedDataTime` given a UTC `Timestamp`. -/ def ofUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime := do let rules ← Database.localRules db let tz ← IO.ofExcept <| timezoneAt rules tm - --let tm := applyLeapSeconds tm rules return ZonedDateTime.ofTimestamp tm tz diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 2a85dbcf76e5..38ec57cff5b6 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -71,9 +71,9 @@ def idFromPath (path : System.FilePath) : Option String := do Retrieves the timezone rules from the local timezone data file. -/ def localRules (path : System.FilePath) : IO ZoneRules := do - let PlainTimePath ← IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } + let localTimePath ← IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } - if let some id := idFromPath PlainTimePath + if let some id := idFromPath localTimePath then parseTZIfFromDisk path id else throw (IO.userError "cannot read the id of the path.") diff --git a/src/Std/Time/Zoned/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean index 0b36e1f1e22c..3822d582f84c 100644 --- a/src/Std/Time/Zoned/Database/TzIf.lean +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -66,7 +66,7 @@ structure Header where /-- Represents the local time type information, including offset and daylight saving details. -/ -structure PlainTimeType where +structure LocalTimeType where /-- The GMT offset in seconds for this local time type. @@ -123,7 +123,7 @@ structure TZifV1 where /-- The array of local time type structures. -/ - localTimeTypes : Array PlainTimeType + localTimeTypes : Array LocalTimeType /-- The array of abbreviation strings used by local time types. @@ -216,8 +216,8 @@ private def parseHeader : Parser Header := <*> pu32 <*> pu32 -private def parsePlainTimeType : Parser PlainTimeType := - PlainTimeType.mk +private def parseLocalTimeType : Parser LocalTimeType := + LocalTimeType.mk <$> pi32 <*> pbool <*> pu8 @@ -233,10 +233,10 @@ private def parseTransitionTimes (size : Parser Int32) (n : UInt32) : Parser (Ar private def parseTransitionIndices (n : UInt32) : Parser (Array UInt8) := manyN (n.toNat) pu8 -private def parseLocalTimeType (n : UInt32) : Parser (Array PlainTimeType) := - manyN (n.toNat) parsePlainTimeType +private def parseLocalTimeTypes (n : UInt32) : Parser (Array LocalTimeType) := + manyN (n.toNat) parseLocalTimeType -private def parseAbbreviations (times : Array PlainTimeType) (n : UInt32) : Parser (Array String) := do +private def parseAbbreviations (times : Array LocalTimeType) (n : UInt32) : Parser (Array String) := do let mut strings := #[] let mut current := "" let mut chars ← manyN n.toNat pu8 @@ -264,7 +264,7 @@ private def parseTZifV1 : Parser TZifV1 := do let transitionTimes ← parseTransitionTimes pi32 header.timecnt let transitionIndices ← parseTransitionIndices header.timecnt - let localTimeTypes ← parseLocalTimeType header.typecnt + let localTimeTypes ← parseLocalTimeTypes header.typecnt let abbreviations ← parseAbbreviations localTimeTypes header.charcnt let leapSeconds ← parseLeapSeconds pi32 header.leapcnt let stdWallIndicators ← parseIndicators header.isstdcnt @@ -299,7 +299,7 @@ private def parseTZifV2 : Parser (Option TZifV2) := optional do let transitionTimes ← parseTransitionTimes pi64 header.timecnt let transitionIndices ← parseTransitionIndices header.timecnt - let localTimeTypes ← parseLocalTimeType header.typecnt + let localTimeTypes ← parseLocalTimeTypes header.typecnt let abbreviations ← parseAbbreviations localTimeTypes header.charcnt let leapSeconds ← parseLeapSeconds pi64 header.leapcnt let stdWallIndicators ← parseIndicators header.isstdcnt diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index e082cac78471..3f07f2e2f65f 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -47,10 +47,10 @@ def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := DateTime.mk tm (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds).toPlainDateTimeAssumingUTC) /-- -Creates a new zone aware `Timestamp` out of a `DateTime`. +Creates a UTC `Timestamp` out of a `DateTime`. -/ @[inline] -def toTimestamp (date : DateTime tz) : Timestamp := +def toUTCTimestamp (date : DateTime tz) : Timestamp := date.timestamp /-- @@ -272,29 +272,29 @@ def withYearRollOver (dt : DateTime tz) (year : Year.Offset) : DateTime tz := Creates a new `DateTime tz` by adjusting the `hour` component. -/ @[inline] -def withHour (dt : DateTime tz) (hour : Hour.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withHour hour) tz +def withHours (dt : DateTime tz) (hour : Hour.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withHours hour) tz /-- Creates a new `DateTime tz` by adjusting the `minute` component. -/ @[inline] -def withMinute (dt : DateTime tz) (minute : Minute.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withMinute minute) tz +def withMinutes (dt : DateTime tz) (minute : Minute.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withMinutes minute) tz /-- Creates a new `DateTime tz` by adjusting the `second` component. -/ @[inline] -def withSecond (dt : DateTime tz) (second : Sigma Second.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withSecond second) tz +def withSeconds (dt : DateTime tz) (second : Sigma Second.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withSeconds second) tz /-- Creates a new `DateTime tz` by adjusting the `nano` component. -/ @[inline] -def withNano (dt : DateTime tz) (nano : Nanosecond.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withNano nano) tz +def withNanoseconds (dt : DateTime tz) (nano : Nanosecond.Ordinal) : DateTime tz := + .ofLocalDateTime (dt.date.get.withNanoseconds nano) tz /-- Converts a `Timestamp` to a `PlainDateTime` @@ -371,8 +371,8 @@ Gets the `Period` of a `DateTime`, corresponding to the part of the day (e.g. ni afternoon, evening) based on the hour. -/ @[inline] -def period (dt : DateTime tz) : Day.Ordinal.Period := - Day.Ordinal.Period.fromHour dt.hour +def period (dt : DateTime tz) : Day.Period := + Day.Period.fromHour dt.hour /-- Determines the era of the given `DateTime` based on its year. @@ -417,7 +417,7 @@ def quarter (date : DateTime tz) : Bounded.LE 1 4 := date.month.sub 1 |>.ediv 3 (by decide) |>.add 1 instance : ToTimestamp (DateTime tz) where - toTimestamp dt := dt.toTimestamp + toTimestamp dt := dt.toUTCTimestamp instance : HAdd (DateTime tz) (Day.Offset) (DateTime tz) where hAdd := addDays @@ -456,7 +456,7 @@ instance : HSub (DateTime tz) (Nanosecond.Offset) (DateTime tz) where hSub := subNanoseconds instance : HSub (DateTime tz) (DateTime tz₁) Duration where - hSub x y := x.toTimestamp - y.toTimestamp + hSub x y := x.toUTCTimestamp - y.toUTCTimestamp end DateTime end Time diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 70e60ff5e234..70a8774c9fa5 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -27,7 +27,18 @@ structure Offset where The same timezone offset in seconds. -/ second : Second.Offset - deriving Repr, Inhabited, BEq + + /-- + The proof that both are equal + -/ + proof : second.toHours = hour + deriving Repr + +instance : Inhabited Offset where + default := ⟨0, 0, rfl⟩ + +instance : BEq Offset where + beq x y := BEq.beq x.second y.second namespace Offset @@ -48,26 +59,26 @@ def toIsoString (offset : Offset) (colon : Bool) : String := A zero `Offset` representing UTC (no offset). -/ def zero : Offset := - { hour := 0, second := 0 } + { hour := 0, second := 0, proof := rfl } /-- Creates an `Offset` from a given number of hour. -/ def ofHours (n : Hour.Offset) : Offset := - mk n n.toSeconds + mk n n.toSeconds (by simp [Hour.Offset.toSeconds, Second.Offset.toHours, UnitVal.mul, UnitVal.div]; rfl) /-- Creates an `Offset` from a given number of hour and minuets. -/ def ofHoursAndMinutes (n : Hour.Offset) (m : Minute.Offset) : Offset := let secs := n.toSeconds + m.toSeconds - mk secs.toHours secs + mk secs.toHours secs rfl /-- Creates an `Offset` from a given number of second. -/ def ofSeconds (n : Second.Offset) : Offset := - mk n.toHours n + mk n.toHours n rfl end Offset end TimeZone diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index 5b9d62ab6eac..f88dd2a77f1a 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -12,7 +12,8 @@ namespace Time set_option linter.all true /-- -An enumeration representing different time zones. +A TimeZone structure that stores the timezone offset, the name, abbreviation and if it's in daylight +saving time. -/ structure TimeZone where @@ -27,7 +28,7 @@ structure TimeZone where name : String /-- - The abbreviation of the time zone/ + The abbreviation of the time zone. -/ abbreviation : String @@ -40,7 +41,7 @@ structure TimeZone where namespace TimeZone /-- -Fetches the current timestamp from the system. +Fetches the current timzeon from the system. -/ @[extern "lean_get_timezone_offset"] opaque getSystemTimezone : IO TimeZone diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 048cbef73480..05762da2c24a 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -107,8 +107,7 @@ structure Transition where deriving Repr, Inhabited /-- -Represents the rules for a time zone, abstracting away binary data and focusing on key transitions -and types. +Represents the rules for a time zone. -/ structure ZoneRules where diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index b1beb3b485c9..9e7cfa1a2230 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -14,7 +14,7 @@ set_option linter.all true /-- An existential version of `DateTime` that encapsulates a `DateTime` value without explicitly storing -the time zone. +the time zone in the type. -/ def ZonedDateTime := Sigma DateTime @@ -31,28 +31,28 @@ namespace ZonedDateTime open DateTime /-- -Creates a new `ZonedDateTime` out of a `DateTime` and `TimeZone` +Creates a new `ZonedDateTime` out of a `DateTime` and `TimeZone`. -/ @[inline] def mk (tz : TimeZone) (datetime : DateTime tz) : ZonedDateTime := ⟨tz, datetime⟩ /-- -Creates a new `ZonedDateTime` out of a `Timestamp` +Creates a new `ZonedDateTime` out of a `Timestamp` and a `TimeZone`. -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := ⟨tz, DateTime.ofTimestamp tm tz⟩ /-- -Creates a new `Timestamp` out of a `ZonedDateTime` +Creates a new UTC `Timestamp` out of a `ZonedDateTime`. -/ @[inline] -def toTimestamp (date : ZonedDateTime) : Timestamp := - date.snd.toTimestamp +def toUTCTimestamp (date : ZonedDateTime) : Timestamp := + date.snd.toUTCTimestamp /-- -Creates a new `ZonedDateTime` out of a `Timestamp` +Creates a new `ZonedDateTime` out of a `Timestamp` and `ZoneRules`. -/ @[inline] def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do @@ -71,7 +71,7 @@ Changes the `TimeZone` to a new one. -/ @[inline] def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := - ofTimestamp date.toTimestamp tz₁ + ofTimestamp date.toUTCTimestamp tz₁ /-- Creates a new `ZonedDateTime` out of a `PlainDateTime` @@ -293,14 +293,11 @@ def subNanoseconds (dt : ZonedDateTime) (nanoseconds : Nanosecond.Offset) : Zone Sigma.mk dt.fst (dt.snd.subNanoseconds nanoseconds) /-- -Determines the era of the given `PlainDate` based on its year. +Determines the era of the given `ZonedDateTime` based on its year. -/ @[inline] def era (date : ZonedDateTime) : Year.Era := - if date.year.toInt ≥ 0 then - .ce - else - .bce + date.snd.era /-- Creates a new `ZonedDateTime` by adjusting the day of the month to the given `days` value, with any @@ -354,38 +351,38 @@ def withYearRollOver (dt : ZonedDateTime) (year : Year.Offset) : ZonedDateTime : Creates a new `ZonedDateTime` by adjusting the `hour` component. -/ @[inline] -def withHour (dt : ZonedDateTime) (hour : Hour.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withHour hour⟩ +def withHours (dt : ZonedDateTime) (hour : Hour.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withHours hour⟩ /-- Creates a new `ZonedDateTime` by adjusting the `minute` component. -/ @[inline] -def withMinute (dt : ZonedDateTime) (minute : Minute.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withMinute minute⟩ +def withMinutes (dt : ZonedDateTime) (minute : Minute.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withMinutes minute⟩ /-- Creates a new `ZonedDateTime` by adjusting the `second` component. -/ @[inline] -def withSecond (dt : ZonedDateTime) (second : Sigma Second.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withSecond second⟩ +def withSeconds (dt : ZonedDateTime) (second : Sigma Second.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withSeconds second⟩ /-- Creates a new `ZonedDateTime` by adjusting the `nano` component. -/ @[inline] -def withNano (dt : ZonedDateTime) (nano : Nanosecond.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withNano nano⟩ +def withNanoseconds (dt : ZonedDateTime) (nano : Nanosecond.Ordinal) : ZonedDateTime := + ⟨dt.fst, dt.snd.withNanoseconds nano⟩ /-- -Checks if the `DateTime` is in a leap year. +Checks if the `ZonedDateTime` is in a leap year. -/ def inLeapYear (date : ZonedDateTime) : Bool := date.year.isLeap instance : ToTimestamp ZonedDateTime where - toTimestamp dt := dt.toTimestamp + toTimestamp dt := dt.toUTCTimestamp instance : HAdd ZonedDateTime Day.Offset ZonedDateTime where hAdd := addDays @@ -424,7 +421,7 @@ instance : HSub ZonedDateTime Nanosecond.Offset ZonedDateTime where hSub := subNanoseconds instance : HSub ZonedDateTime ZonedDateTime Duration where - hSub x y := x.toTimestamp - y.toTimestamp + hSub x y := x.toUTCTimestamp - y.toUTCTimestamp end ZonedDateTime end Time diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index f2742eff00b9..819b156b66d5 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -567,10 +567,11 @@ extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { lean_ctor_set(lean_offset, 0, lean_int_to_int(offset_hour)); lean_ctor_set(lean_offset, 1, lean_int_to_int(offset_seconds)); - lean_object *lean_tz = lean_alloc_ctor(0, 2, 1); + lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); lean_ctor_set(lean_tz, 0, lean_offset); lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked("Unknown")); - lean_ctor_set_uint8(lean_tz, sizeof(void*)*2, tm_info.tm_isdst); + lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked("Unknown")); + lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, tm_info.tm_isdst); return lean_io_result_mk_ok(lean_tz); } diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 9e638ab4e352..4a61533921fb 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -1,16 +1,16 @@ import Std.Time open Std.Time -def ISO8601UTC : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZ") -def RFC1123 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") -def ShortDate : Format .any := datespec("MM/dd/yyyy") -def LongDate : Format .any := datespec("MMMM D, yyyy") -def ShortDateTime : Format .any := datespec("MM/dd/yyyy HH:mm:ss") -def LongDateTime : Format .any := datespec("MMMM D, yyyy h:mm aa") -def Time24Hour : Format .any := datespec("HH:mm:ss") -def Time12Hour : Format .any := datespec("hh:mm:ss aa") -def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") -def CustomDayTime : Format .any := datespec("EEE dd MMM yyyy HH:mm") +def ISO8601UTC : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZ") +def RFC1123 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def ShortDate : GenericFormat .any := datespec("MM/dd/yyyy") +def LongDate : GenericFormat .any := datespec("MMMM D, yyyy") +def ShortDateTime : GenericFormat .any := datespec("MM/dd/yyyy HH:mm:ss") +def LongDateTime : GenericFormat .any := datespec("MMMM D, yyyy h:mm aa") +def Time24Hour : GenericFormat .any := datespec("HH:mm:ss") +def Time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") +def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") +def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM yyyy HH:mm") -- Dates @@ -37,7 +37,7 @@ info: "Monday, June 16, 2014 03:03:03 -0300" #guard_msgs in #eval FullDayTimeZone.format date₂ -def tm₃ := date₁.toTimestamp +def tm₃ := date₁.toUTCTimestamp def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index e79d6e202617..60bf46e59313 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -2,8 +2,8 @@ import Std.Time import Init open Std.Time -def ShortDateTime : Format .any := datespec("dd/MM/yyyy HH:mm:ss") -def ShortDate : Format .any := datespec("dd/MM/yyyy") +def ShortDateTime : GenericFormat .any := datespec("dd/MM/yyyy HH:mm:ss") +def ShortDate : GenericFormat .any := datespec("dd/MM/yyyy") def format (PlainDate : PlainDateTime) : String := ShortDateTime.formatBuilder PlainDate.day PlainDate.month PlainDate.year PlainDate.time.hour PlainDate.minute PlainDate.time.second def format₂ (PlainDate : PlainDate) : String := ShortDate.formatBuilder PlainDate.day PlainDate.month PlainDate.year diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index f843a5a818c1..4b8c4586a1a1 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -1,18 +1,18 @@ import Std.Time open Std.Time -def ISO8601UTC : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") -def RFC1123 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") -def ShortDate : Format .any := datespec("MM/dd/yyyy") -def LongDate : Format .any := datespec("MMMM D, yyyy") -def ShortDateTime : Format .any := datespec("MM/dd/yyyy HH:mm:ss") -def LongDateTime : Format .any := datespec("MMMM dd, yyyy hh:mm aa") -def Time24Hour : Format .any := datespec("HH:mm:ss") -def Time12Hour : Format .any := datespec("hh:mm:ss aa") -def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") -def CustomDayTime : Format .any := datespec("EEE dd MMM yyyy HH:mm") - -def Full12HourWrong : Format .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") +def ISO8601UTC : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") +def RFC1123 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def ShortDate : GenericFormat .any := datespec("MM/dd/yyyy") +def LongDate : GenericFormat .any := datespec("MMMM D, yyyy") +def ShortDateTime : GenericFormat .any := datespec("MM/dd/yyyy HH:mm:ss") +def LongDateTime : GenericFormat .any := datespec("MMMM dd, yyyy hh:mm aa") +def Time24Hour : GenericFormat .any := datespec("HH:mm:ss") +def Time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") +def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") +def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM yyyy HH:mm") + +def Full12HourWrong : GenericFormat .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") -- Dates diff --git a/tests/lean/run/timeSet.lean b/tests/lean/run/timeSet.lean index 13ffd31ba8c2..36020cf2d137 100644 --- a/tests/lean/run/timeSet.lean +++ b/tests/lean/run/timeSet.lean @@ -1,18 +1,18 @@ import Std.Time open Std.Time -def ISO8601UTC : Format .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") -def RFC1123 : Format .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") -def ShortDate : Format .any := datespec("MM/dd/yyyy") -def LongDate : Format .any := datespec("MMMM D, yyyy") -def ShortDateTime : Format .any := datespec("MM/dd/yyyy HH:mm:ss") -def LongDateTime : Format .any := datespec("MMMM dd, yyyy hh:mm aa") -def Time24Hour : Format .any := datespec("HH:mm:ss") -def Time12Hour : Format .any := datespec("hh:mm:ss aa") -def FullDayTimeZone : Format .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") -def CustomDayTime : Format .any := datespec("EEE dd MMM yyyy HH:mm") - -def Full12HourWrong : Format .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") +def ISO8601UTC : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") +def RFC1123 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def ShortDate : GenericFormat .any := datespec("MM/dd/yyyy") +def LongDate : GenericFormat .any := datespec("MMMM D, yyyy") +def ShortDateTime : GenericFormat .any := datespec("MM/dd/yyyy HH:mm:ss") +def LongDateTime : GenericFormat .any := datespec("MMMM dd, yyyy hh:mm aa") +def Time24Hour : GenericFormat .any := datespec("HH:mm:ss") +def Time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") +def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") +def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM yyyy HH:mm") + +def Full12HourWrong : GenericFormat .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") -- Dates @@ -72,22 +72,22 @@ info: zoned("2016-06-16T03:03:03.000000000-03:00") info: zoned("2014-06-16T19:03:03.000000000-03:00") -/ #guard_msgs in -#eval date₁.withHour 19 +#eval date₁.withHours 19 /-- info: zoned("2014-06-16T03:45:03.000000000-03:00") -/ #guard_msgs in -#eval date₁.withMinute 45 +#eval date₁.withMinutes 45 /-- info: zoned("2014-06-16T03:03:59.000000000-03:00") -/ #guard_msgs in -#eval date₁.withSecond ⟨true, 59⟩ +#eval date₁.withSeconds ⟨true, 59⟩ /-- info: zoned("2014-06-16T03:03:03.000000002-03:00") -/ #guard_msgs in -#eval date₁.withNano 2 +#eval date₁.withNanoseconds 2 From e8c48f49400c2dcac84a630260441bec5db8fe7a Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 15 Oct 2024 00:38:56 -0300 Subject: [PATCH 074/118] feat: icu.h for timezone support on windows --- script/prepare-llvm-mingw.sh | 4 +- src/CMakeLists.txt | 4 + src/Std/Time/Zoned.lean | 7 +- src/Std/Time/Zoned/Database.lean | 29 ++-- src/Std/Time/Zoned/Database/Basic.lean | 9 +- src/Std/Time/Zoned/Database/Leap.lean | 18 ++- src/Std/Time/Zoned/Database/TZdb.lean | 17 ++- src/Std/Time/Zoned/Database/Windows.lean | 46 +++++++ src/Std/Time/Zoned/TimeZone.lean | 6 - src/Std/Time/Zoned/ZoneRules.lean | 20 ++- src/Std/Time/Zoned/ZonedDateTime.lean | 2 +- src/runtime/io.cpp | 166 +++++++++++++++++++++-- 12 files changed, 266 insertions(+), 62 deletions(-) create mode 100644 src/Std/Time/Zoned/Database/Windows.lean diff --git a/script/prepare-llvm-mingw.sh b/script/prepare-llvm-mingw.sh index 749fc9d9fac4..a0c839f22419 100644 --- a/script/prepare-llvm-mingw.sh +++ b/script/prepare-llvm-mingw.sh @@ -35,9 +35,9 @@ cp /clang64/lib/lib{m,bcrypt,mingw32,moldname,mingwex,msvcrt,pthread,advapi32,sh echo -n " -DLEAN_STANDALONE=ON" echo -n " -DCMAKE_C_COMPILER=$PWD/stage1/bin/clang.exe -DCMAKE_C_COMPILER_WORKS=1 -DCMAKE_CXX_COMPILER=$PWD/llvm/bin/clang++.exe -DCMAKE_CXX_COMPILER_WORKS=1 -DLEAN_CXX_STDLIB='-lc++ -lc++abi'" echo -n " -DSTAGE0_CMAKE_C_COMPILER=clang -DSTAGE0_CMAKE_CXX_COMPILER=clang++" -echo -n " -DLEAN_EXTRA_CXX_FLAGS='--sysroot $PWD/llvm -idirafter /clang64/include/'" +echo -n " -DLEAN_EXTRA_CXX_FLAGS='-idirafter \"${WINDOWS_SDK_DIR}/Include/10.0.22621.0/um\" --sysroot $PWD/llvm -idirafter /clang64/include/'" echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang.exe" -echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -static-libgcc -Wl,-Bstatic -lgmp -luv -lunwind -Wl,-Bdynamic -fuse-ld=lld'" +echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -static-libgcc -Wl,-Bstatic -lgmp -luv -licu -lunwind -Wl,-Bdynamic -fuse-ld=lld'" # when not using the above flags, link GMP dynamically/as usual echo -n " -DLEAN_EXTRA_LINKER_FLAGS='-lgmp -luv -lucrtbase'" # do not set `LEAN_CC` for tests diff --git a/src/CMakeLists.txt b/src/CMakeLists.txt index 6ccb39580d49..f18efbebe05b 100644 --- a/src/CMakeLists.txt +++ b/src/CMakeLists.txt @@ -252,6 +252,10 @@ if(NOT LEAN_STANDALONE) string(APPEND LEAN_EXTRA_LINKER_FLAGS " ${LIBUV_LIBRARIES}") endif() +set(WINDOWS_SDK_DIR "/c/Program Files (x86)/Windows Kits/10") +string(APPEND LEAN_EXTRA_LINKER_FLAGS " -licu") +string(APPEND CMAKE_CXX_FLAGS " -idirafter \"${WINDOWS_SDK_DIR}/Include/10.0.22621.0/um\"") + # ccache if(CCACHE AND NOT CMAKE_CXX_COMPILER_LAUNCHER AND NOT CMAKE_C_COMPILER_LAUNCHER) find_program(CCACHE_PATH ccache) diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 55b2b7850a24..2c5becfa172b 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -8,6 +8,7 @@ import Std.Time.Zoned.DateTime import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.ZonedDateTime +import Std.Time.Zoned.Database namespace Std namespace Time @@ -29,7 +30,7 @@ Get the current time. @[inline] def now : IO PlainDateTime := do let tm ← Timestamp.now - let tz ← TimeZone.getSystemTimezone + let tz ← Database.defaultGetCurrentTimeZone return PlainDateTime.ofTimestamp tm tz end PlainDateTime @@ -86,12 +87,12 @@ end DateTime namespace ZonedDateTime /-- -Gets the current `DateTime`. +Gets the current `ZonedDateTime`. -/ @[inline] def now : IO ZonedDateTime := do let date ← Timestamp.now - let tz ← TimeZone.getSystemTimezone + let tz ← Database.defaultGetCurrentTimeZone return ofTimestamp date tz /-- diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index 8577a47cdbcc..72af7352a13f 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -7,6 +7,8 @@ prelude import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.Database.Basic import Std.Time.Zoned.Database.TZdb +import Std.Time.Zoned.Database.Windows +import Init.System.Platform namespace Std namespace Time @@ -16,21 +18,26 @@ open TimeZone.ZoneRules set_option linter.all true /-- -Gets the `TimeZone` at the local timezone. +Gets the time zone for a given location and timestamp, handling Windows and non-Windows platforms. -/ -def getLocalTimeZoneAt [Database α] (db : α) (tm : Timestamp) : IO TimeZone := do - (IO.ofExcept <| timezoneAt · tm) =<< Database.localRules db +def defaultGetTimeZoneAt : String → Timestamp → IO TimeZone := + if System.Platform.isWindows + then Database.getTimeZoneAt WindowsDb.default + else Database.getTimeZoneAt TZdb.default /-- -Gets the TimeZone at a timezone using a `Database`. +Gets the local time zone for a specific timestamp, accounting for platform differences. -/ -def getTimeZoneAt [Database α] (db : α) (id : String) (tm : Timestamp) : IO TimeZone := do - (IO.ofExcept <| timezoneAt · tm) =<< Database.load db id +def defaultGetLocalTimeZoneAt : Timestamp → IO TimeZone := + if System.Platform.isWindows + then Database.getLocalTimeZoneAt WindowsDb.default + else Database.getLocalTimeZoneAt TZdb.default /-- -Get the local `ZonedDataTime` given a UTC `Timestamp`. +Retrieves the current local time zone based on the system platform and the current timestamp. -/ -def ofUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO ZonedDateTime := do - let rules ← Database.localRules db - let tz ← IO.ofExcept <| timezoneAt rules tm - return ZonedDateTime.ofTimestamp tm tz +def defaultGetCurrentTimeZone : IO TimeZone := do + let now <- Timestamp.now + if System.Platform.isWindows + then Database.getLocalTimeZoneAt WindowsDb.default now + else Database.getLocalTimeZoneAt TZdb.default now diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index a1a843ee98e6..9bf6208b7418 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -21,17 +21,12 @@ class Database (α : Type) where /-- Loads a `ZoneRules` by its id. -/ - load : α → String → IO TimeZone.ZoneRules + getTimeZoneAt : α → String → Timestamp → IO TimeZone /-- Loads a `ZoneRules` that is set by the local machine. -/ - localRules : α → IO TimeZone.ZoneRules - - /-- - Loads a `LeapSecond` array from the local machine. - -/ - leapSeconds : α → IO (Array TimeZone.LeapSecond) + getLocalTimeZoneAt : α → Timestamp → IO TimeZone namespace TimeZone diff --git a/src/Std/Time/Zoned/Database/Leap.lean b/src/Std/Time/Zoned/Database/Leap.lean index 07671733e109..d6d30683cd69 100644 --- a/src/Std/Time/Zoned/Database/Leap.lean +++ b/src/Std/Time/Zoned/Database/Leap.lean @@ -8,9 +8,9 @@ import Std.Internal.Parsec import Std.Internal.Parsec.ByteArray import Std.Time.Internal.Bounded -import Std.Time.Format.Basic import Std.Time.Time import Std.Time.Date +import Std.Time.DateTime namespace Std namespace Time @@ -52,12 +52,26 @@ private def failWith (opt : Option α) : Parser α := | none => fail "invalid number" | some res => pure res +def parseMonthShort : Parser Month.Ordinal + := pstring "Jan" *> pure ⟨1, by decide⟩ + <|> pstring "Feb" *> pure ⟨2, by decide⟩ + <|> pstring "Mar" *> pure ⟨3, by decide⟩ + <|> pstring "Apr" *> pure ⟨4, by decide⟩ + <|> pstring "May" *> pure ⟨5, by decide⟩ + <|> pstring "Jun" *> pure ⟨6, by decide⟩ + <|> pstring "Jul" *> pure ⟨7, by decide⟩ + <|> pstring "Aug" *> pure ⟨8, by decide⟩ + <|> pstring "Sep" *> pure ⟨9, by decide⟩ + <|> pstring "Oct" *> pure ⟨10, by decide⟩ + <|> pstring "Nov" *> pure ⟨11, by decide⟩ + <|> pstring "Dec" *> pure ⟨12, by decide⟩ + private def parseLeapSecond : Parser LeapSecond := do skipString "Leap" ws let year ← digits ws - let month ← Std.Time.parseMonthShort + let month ← parseMonthShort ws let day ← digits ws diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 38ec57cff5b6..6c0d3a1a7d4b 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -93,7 +93,18 @@ def readLeapsFromDisk (path : System.FilePath) : IO (Array LeapSecond) := do then return res else throw (IO.userError "cannot read the id of the path.") +/-- +Gets the `TimeZone` at the local timezone. +-/ +def getLocalTimeZoneAt (db : TZdb) (tm : Timestamp) : IO TimeZone := do + (IO.ofExcept <| Transition.timezoneAt ·.transitions tm) =<< localRules db.localPath + +/-- +Gets the TimeZone at a timezone using a `Database`. +-/ +def getTimeZoneAt (db : TZdb) (id : String) (tm : Timestamp) : IO TimeZone := do + (IO.ofExcept <| Transition.timezoneAt ·.transitions tm) =<< readRulesFromDisk db.zonesPath id + instance : Database TZdb where - load db id := readRulesFromDisk db.zonesPath id - localRules db := localRules db.localPath - leapSeconds db := readLeapsFromDisk (db.zonesPath / "leapseconds") + getLocalTimeZoneAt := TZdb.getLocalTimeZoneAt + getTimeZoneAt := TZdb.getTimeZoneAt diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean new file mode 100644 index 000000000000..c5790cdb947a --- /dev/null +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -0,0 +1,46 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Time.DateTime +import Std.Time.Zoned.TimeZone +import Std.Time.Zoned.ZoneRules +import Std.Time.Zoned.Database.Basic + +namespace Std +namespace Time +namespace Database + +set_option linter.all true + +/-- +Fetches the timezone information from the current windows machine. +-/ +@[extern "lean_get_windows_timezone_at"] +opaque Windows.getTimeZoneAt : String -> UInt64 -> IO TimeZone + +/-- +Fetches the timezone at a timestamp. +-/ +@[extern "lean_get_timezone_offset_at"] +opaque Windows.getLocalTimezoneAt : UInt64 -> IO TimeZone + +/-- +Represents a Time Zone Database that we get from ICU available on Windows SDK. +-/ +structure WindowsDb where + +namespace WindowsDb +open TimeZone + +/-- +Returns a default `WindowsDb` instance. +-/ +@[inline] +def default : WindowsDb := {} + +instance : Database WindowsDb where + getTimeZoneAt _ id tm := Windows.getTimeZoneAt id (tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64) + getLocalTimeZoneAt _ tm := Windows.getLocalTimezoneAt (tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64) diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index f88dd2a77f1a..c4b63456d038 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -40,12 +40,6 @@ structure TimeZone where namespace TimeZone -/-- -Fetches the current timzeon from the system. --/ -@[extern "lean_get_timezone_offset"] -opaque getSystemTimezone : IO TimeZone - /-- A zeroed `Timezone` representing UTC (no offset). -/ diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 05762da2c24a..733a9f604c4b 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -141,24 +141,20 @@ def apply (timestamp : Timestamp) (transition : Transition) : Timestamp := let offsetInSeconds := transition.localTimeType.gmtOffset.hour.mul 3600 |>.add transition.localTimeType.gmtOffset.second timestamp.addSeconds offsetInSeconds -end Transition - -namespace ZoneRules - /-- -Finds the transition corresponding to a given timestamp in `ZoneRules`. +Finds the transition corresponding to a given timestamp in `Array Transition`. If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. -/ -def findTransitionForTimestamp (zoneRules : ZoneRules) (timestamp : Timestamp) : Option Transition := +def findTransitionForTimestamp (transitions : Array Transition) (timestamp : Timestamp) : Option Transition := let value := timestamp.toSecondsSinceUnixEpoch - if let some idx := zoneRules.transitions.findIdx? (fun t => t.time.val > value.val) - then zoneRules.transitions.get? (idx - 1) - else zoneRules.transitions.back? + if let some idx := transitions.findIdx? (fun t => t.time.val > value.val) + then transitions.get? (idx - 1) + else transitions.back? /-- -Find the current `TimeZone` out of a `Transition` in a `ZoneRules` +Find the current `TimeZone` out of a `Transition` in a `Array Transition` -/ -def timezoneAt (rules : ZoneRules) (tm : Timestamp) : Except String TimeZone := - if let some transition := rules.findTransitionForTimestamp tm +def timezoneAt(transitions : Array Transition) (tm : Timestamp) : Except String TimeZone := + if let some transition := findTransitionForTimestamp transitions tm then .ok transition.createTimeZoneFromTransition else .error "cannot find local timezone." diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 9e7cfa1a2230..ce8d437d56a1 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -56,7 +56,7 @@ Creates a new `ZonedDateTime` out of a `Timestamp` and `ZoneRules`. -/ @[inline] def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do - let transition ← rules.findTransitionForTimestamp tm + let transition ← TimeZone.Transition.findTransitionForTimestamp rules.transitions tm return ofTimestamp tm transition.localTimeType.getTimeZone /-- diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 819b156b66d5..bf087a70d222 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -5,6 +5,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Sebastian Ullrich */ #if defined(LEAN_WINDOWS) +#include +#include +#include #include #include #define NOMINMAX // prevent ntdef.h from defining min/max macros @@ -538,40 +541,173 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } +/* Std.Time.TimeZone.getWindowsTimeZoneAt : String -> UInt64 -> IO Timezone */ +extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str, uint64_t timestamp_secs, obj_arg /* w */) { +#if defined(LEAN_WINDOWS) + UErrorCode status = U_ZERO_ERROR; + const char* dst_name = lean_string_cstr(timezone_str); + + UChar tzID[256]; + u_strFromUTF8(tzID, sizeof(tzID) / sizeof(tzID[0]), NULL, dst_name, strlen(dst_name), &status); + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to read identifier"))); + } + + UCalendar *cal = ucal_open(tzID, -1, NULL, UCAL_GREGORIAN, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); + } + + ucal_setMillis(cal, timestamp_secs * 1000, &status); + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); + } + + UChar display_name[32]; + int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to read abbreaviation"))); + } + + char display_name_str[256]; + u_strToUTF8(display_name_str, sizeof(display_name_str), NULL, display_name, display_name_len, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get abbreviation to cstr"))); + } + + int32_t zone_offset = ucal_get(cal, UCAL_ZONE_OFFSET, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); + } + + int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); + + zone_offset += dst_offset; + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); + } + + int offset_hour = zone_offset / 3600000; + int offset_seconds = zone_offset / 1000; + int is_dst = dst_offset != 0; + + lean_object *lean_offset = lean_alloc_ctor(0, 2, 0); + lean_ctor_set(lean_offset, 0, lean_int_to_int(offset_hour)); + lean_ctor_set(lean_offset, 1, lean_int_to_int(offset_seconds)); + + lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); + lean_ctor_set(lean_tz, 0, lean_offset); + lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked(dst_name)); + lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked(display_name_str)); + lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); + + return lean_io_result_mk_ok(lean_tz); +#else + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get timezone, its windows only."))); +#endif +} + /* Std.Time.TimeZone.getCurrentTimezone : IO Timezone */ -extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset(obj_arg /* w */) { +extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset_at(uint64_t timestamp_secs, obj_arg /* w */) { using namespace std::chrono; +#if defined(LEAN_WINDOWS) + UErrorCode status = U_ZERO_ERROR; + UCalendar *cal = ucal_open(NULL, -1, NULL, UCAL_GREGORIAN, &status); - auto now = system_clock::now(); - auto now_time_t = system_clock::to_time_t(now); + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); + } - std::tm tm_info; -#if defined(LEAN_WINDOWS) - errno_t err = localtime_s(&tm_info, &now_time_t); + ucal_setMillis(cal, timestamp_secs * 1000, &status); + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); + } + + UChar tzId[256]; + int32_t tzIdLength = ucal_getTimeZoneID(cal, tzId, sizeof(tzId)/sizeof(tzId[0]), &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); + } + + char dst_name[256]; + u_strToUTF8(dst_name, sizeof(dst_name), NULL, tzId, tzIdLength, &status); + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); + } + + UChar display_name[32]; + int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); - if (err != 0) { - return lean_io_result_mk_error(lean_decode_io_error(err, mk_string(""))); + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to read abbreaviation"))); } + + char display_name_str[32]; + u_strToUTF8(display_name_str, sizeof(display_name_str), NULL, display_name, display_name_len, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get abbreviation to cstr"))); + } + + int32_t zone_offset = ucal_get(cal, UCAL_ZONE_OFFSET, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); + } + + int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); + + zone_offset += dst_offset; + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); + } + + int offset_hour = zone_offset / 3600000; + int offset_seconds = zone_offset / 1000; + int is_dst = dst_offset != 0; #else - struct tm *tm_ptr = localtime_r(&now_time_t, &tm_info); + std::time_t input_time = static_cast(timestamp_secs); + std::tm tm_info; + struct tm *tm_ptr = localtime_r(&input_time, &tm_info); if (tm_ptr == NULL) { return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string(""))); } -#endif - + int offset_hour = tm_info.tm_gmtoff / 3600; int offset_seconds = tm_info.tm_gmtoff; - + int is_dst = tm_info.tm_isdst; + char dst_name[] = "Unknown"; + char display_name_str[32] = "Unknown"; +#endif lean_object *lean_offset = lean_alloc_ctor(0, 2, 0); lean_ctor_set(lean_offset, 0, lean_int_to_int(offset_hour)); lean_ctor_set(lean_offset, 1, lean_int_to_int(offset_seconds)); lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); lean_ctor_set(lean_tz, 0, lean_offset); - lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked("Unknown")); - lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked("Unknown")); - lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, tm_info.tm_isdst); + lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked(dst_name)); + lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked(display_name_str)); + lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); return lean_io_result_mk_ok(lean_tz); } From facbda15db8264a7b19633e4329f81ada3697c38 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 17 Oct 2024 19:20:27 -0300 Subject: [PATCH 075/118] fix: a bunch of fixes on comments and io --- src/CMakeLists.txt | 8 +- src/Std/Time.lean | 52 +-- src/Std/Time/Date/PlainDate.lean | 46 ++- src/Std/Time/Date/Unit/Day.lean | 43 -- src/Std/Time/Date/Unit/Month.lean | 20 +- src/Std/Time/Date/Unit/Week.lean | 3 - src/Std/Time/Date/Unit/Weekday.lean | 60 +-- src/Std/Time/Date/Unit/Year.lean | 3 - src/Std/Time/DateTime.lean | 10 +- src/Std/Time/DateTime/PlainDateTime.lean | 43 ++ src/Std/Time/Format.lean | 53 ++- src/Std/Time/Format/Basic.lean | 66 +--- src/Std/Time/Notation.lean | 1 - src/Std/Time/Notation/Spec.lean | 1 - src/Std/Time/Time/HourMarker.lean | 10 +- src/Std/Time/Time/Unit/Hour.lean | 20 +- src/Std/Time/Time/Unit/Millisecond.lean | 3 - src/Std/Time/Time/Unit/Minute.lean | 20 +- src/Std/Time/Time/Unit/Nanosecond.lean | 3 - src/Std/Time/Time/Unit/Second.lean | 3 - src/Std/Time/Zoned/Database/Leap.lean | 7 +- src/Std/Time/Zoned/Database/Windows.lean | 13 +- src/Std/Time/Zoned/DateTime.lean | 31 +- src/Std/Time/Zoned/Offset.lean | 25 +- src/Std/Time/Zoned/ZoneRules.lean | 2 +- src/Std/Time/Zoned/ZonedDateTime.lean | 7 + src/runtime/io.cpp | 92 +---- tests/lean/run/timeFormats.lean | 474 +++++++++++++++++++++++ 28 files changed, 776 insertions(+), 343 deletions(-) diff --git a/src/CMakeLists.txt b/src/CMakeLists.txt index f18efbebe05b..6efd9836335a 100644 --- a/src/CMakeLists.txt +++ b/src/CMakeLists.txt @@ -252,9 +252,11 @@ if(NOT LEAN_STANDALONE) string(APPEND LEAN_EXTRA_LINKER_FLAGS " ${LIBUV_LIBRARIES}") endif() -set(WINDOWS_SDK_DIR "/c/Program Files (x86)/Windows Kits/10") -string(APPEND LEAN_EXTRA_LINKER_FLAGS " -licu") -string(APPEND CMAKE_CXX_FLAGS " -idirafter \"${WINDOWS_SDK_DIR}/Include/10.0.22621.0/um\"") +if(${CMAKE_SYSTEM_NAME} MATCHES "Windows") + set(WINDOWS_SDK_DIR "/c/Program Files (x86)/Windows Kits/10") + string(APPEND LEAN_EXTRA_LINKER_FLAGS " -licu") + string(APPEND CMAKE_CXX_FLAGS " -idirafter \"${WINDOWS_SDK_DIR}/Include/10.0.22621.0/um\"") +endif() # ccache if(CCACHE AND NOT CMAKE_CXX_COMPILER_LAUNCHER AND NOT CMAKE_C_COMPILER_LAUNCHER) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 62eeefc49da8..cbd4b7920990 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ +prelude import Std.Time.Time import Std.Time.Date import Std.Time.Zoned @@ -18,7 +19,7 @@ namespace Time /-! # Time -The Lean4 API for date, time, and duration functionalities. +The Lean API for date, time, and duration functionalities. # Overview @@ -56,15 +57,15 @@ converted because they use an internal type called `UnitVal`. Ordinal types represent specific bounded values in reference to another unit, e.g., `Day.Ordinal` represents a day in a month, ranging from 1 to 31. Some ordinal types like `Hour.Ordinal` and `Second.Ordinal`, -allow for values beyond the normal range (e.g, 24 hours and 61 seconds) to accomodate special cases -with leap seconds like `24:00:00` that is valid in ISO 8601. +allow for values beyond the normal range (e.g, 60 seconds) to accomodate special cases with leap seconds +like `23:59:60` that is valid in ISO 8601. - Ordinal types: - `Day.Ordinal`: Ranges from 1 to 31. - `Day.Ordinal.OfYear`: Ranges from 1 to (365 or 366). - `Month.Ordinal`: Ranges from 1 to 12. - `WeekOfYear.Ordinal`: Ranges from 1 to 53. - - `Hour.Ordinal`: Ranges from 0 to 24. + - `Hour.Ordinal`: Ranges from 0 to 23. - `Millisecond.Ordinal`: Ranges from 0 to 999. - `Minute.Ordinal`: Ranges from 0 to 59. - `Nanosecond.Ordinal`: Ranges from 0 to 999,999,999. @@ -82,9 +83,13 @@ nanoseconds corresponds to one second. # Date and Time Types -Dates and times are composed of these components. Dates are "absolute" value in contrast with Offsets -that are just shifts in dates and times. Types like `Date` are made using of components such as `Year.Offset`, -`Month.Ordinal`, and `Day.Ordinal`, with a proof of the date's validity. +Dates and times are made up of different parts. an `Ordinal` is an absolute value, like a specific day in a month, +while an `Offset` is a shift forward or backward in time, used in arithmetic operations to add or subtract days, months or years. +Dates use components like `Year.Ordinal`, `Month.Ordinal`, and `Day.Ordinal` to ensure they represent +valid points in time. + +Some types, like `Duration`, include a `Span` to represent ranges over other types, such as `Second.Offset`. +This type can have a fractional nanosecond part that can be negative or positive that is represented as a `Nanosecond.Span`. ## Date These types provide precision down to the day level, useful for representing and manipulating dates. @@ -101,8 +106,8 @@ These types offer precision down to the nanosecond level, useful for representin Combines date and time into a single representation, useful for precise timestamps and scheduling. - **`PlainDateTime`**: Represents both date and time in the format `YYYY-MM-DDTHH:mm:ss,sssssssss`. -- **`Timestamp`**: Represents a point in time with nanosecond-level precision. It starts on the UNIX -epoch and it should be used when you receive or need to send timestamps to other systems. +- **`Timestamp`**: Represents a specific point in time with nanosecond precision. Its zero value corresponds +to the UNIX epoch. This type should be used when sending or receiving timestamps between systems. ## Zoned date and times. Combines date, time and time zones. @@ -117,10 +122,9 @@ Represents spans of time and the difference between two points in time. # Formats -Format strings are used to convert between `String` representations and date/time types, such as `yyyy-MM-dd'T'HH:mm:ss.sssZ`. -The table below outlines the available format specifiers. Some specifiers can be modified by repeating characters to -adjust truncation and offsets. Some of them when a character is repeated `n` times, it truncates the corresponding value to -`n` characters, usually when not specified a quantity. +Format strings are used to convert between `String` representations and date/time types, like `yyyy-MM-dd'T'HH:mm:ss.sssZ`. +The table below shows the available format specifiers. Some specifiers can be repeated to control truncation or offsets. +When a character is repeated `n` times, it usually truncates the value to `n` characters. The supported formats include: - `G`: Represents the era, such as AD (Anno Domini) or BC (Before Christ). @@ -142,24 +146,20 @@ The supported formats include: - `Q`, `QQ`: Displays the quarter as a number (e.g., "3", "03"). - `QQQ` (short): Displays the quarter as an abbreviated text (e.g., "Q3"). - `QQQQ` (full): Displays the full quarter text (e.g., "3rd quarter"). - - `QQQQQ` (narrow): Displays the full quarter text (e.g., "3rd quarter"). + - `QQQQQ` (narrow): Displays the quarter as a short number (e.g., "3"). - `w`: Represents the week of the week-based year (e.g., "27"). - `W`: Represents the week of the month (e.g., "4"). - `E`: Represents the day of the week as text. - - `E`, `EE`, `EEE`: Displays the abbreviated day name (e.g., "Tue"). + - `E`, `EE`, `EEE`: Displays the abbreviated weekday name (e.g., "Tue"). - `EEEE`: Displays the full day name (e.g., "Tuesday"). - `EEEEE`: Displays the narrow day name (e.g., "T" for Tuesday). -- `e`: Represents the localized day of the week. - - `e`, `ee`: Displays the day of the week as a number, starting from 1 (Monday) to 7 (Sunday). - - `eee`, `eeee`, `eeeee`: Displays the localized day of the week as text (same format as `E`). +- `e`: Represents the weekday as number or text. + - `e`, `ee`: Displays the the as a number, starting from 1 (Monday) to 7 (Sunday). + - `eee`, `eeee`, `eeeee`: Displays the weekday as text (same format as `E`). - `F`: Represents the aligned week of the month (e.g., "3"). - `a`: Represents the AM or PM designation of the day. - `a`, `aa`, `aaa`: Displays AM or PM in a concise format (e.g., "PM"). - `aaaa`: Displays the full AM/PM designation (e.g., "PM"). -- `B`: Represents the period of the day (e.g., "in the morning", "at night"). - - `b`, `bb`, `bbb` (short): Displays the abbreviated period of the day (e.g., "morning"). - - `bbbb` (full): Displays the full period of the day (e.g., "in the morning"). - - `bbbbb` (narrow): Displays a narrow version of the period of the day (e.g., "m" for morning). - `h`: Represents the hour of the AM/PM clock (1-12) (e.g., "12"). - `K`: Represents the hour of the AM/PM clock (0-11) (e.g., "0"). - `k`: Represents the hour of the day in a 1-24 format (e.g., "24"). @@ -170,7 +170,7 @@ The supported formats include: - `A`: Represents the millisecond of the day (e.g., "1234"). - `n`: Represents the nanosecond of the second (e.g., "987654321"). - `N`: Represents the nanosecond of the day (e.g., "1234000000"). -- `V`: Represents the time zone ID, which could be a city-based zone (e.g., "America/Los_Angeles"), a UTC marker (`"Z"`), or a specific offset (e.g., "-08:30"). +- `VV`: Represents the time zone ID, which could be a city-based zone (e.g., "America/Los_Angeles"), a UTC marker (`"Z"`), or a specific offset (e.g., "-08:30"). - `z`: Represents the time zone name. - `z`, `zz`, `zzz`: Shows an abbreviated time zone name (e.g., "PST" for Pacific Standard Time). - `zzzz`: Displays the full time zone name (e.g., "Pacific Standard Time"). @@ -200,9 +200,9 @@ The supported formats include: In order to help the user build dates easily, there are a lot of macros available for creating dates. The `.sssssssss` can be ommited in most cases. -- **`date(yyyy-MM-dd)`**: Defines a date in the `YYYY-MM-DD` format. -- **`time(HH:mm:ss.sssssssss)`**: Defines a time in the `HH:mm:ss:sssssssss` format, including fractional seconds. -- **`datetime("yyy-MM-ddTHH:mm:ss.sssssssss")`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss:sssssssss` format. +- **`date("yyyy-MM-dd")`**: Defines a date in the `yyyy-MM-DD` format. +- **`time("HH:mm:ss.sssssssss")`**: Defines a time in the `HH:mm:ss.sssssssss` format, including fractional seconds. +- **`datetime("yyy-MM-ddTHH:mm:ss.sssssssss")`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss.sssssssss` format. - **`offset("+HH:mm")`**: Defines a timezone offset in the format `+HH:mm`. - **`timezone("NAME/ID ZZZ")`**: Defines a timezone with a name and an offset. - **`datespec("format")`**: Defines a date specification format at compile time using the provided format string. diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 36b913f3e244..3baeca1c5ce4 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -46,7 +46,7 @@ valid by adjusting the day to fit within the valid range to fit the given month @[inline] def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := let day := month.clipDay year.isLeap day - PlainDate.mk year month day Month.Ordinal.clipDay_valid + PlainDate.mk year month day Month.Ordinal.valid_clipDay instance : Inhabited PlainDate where default := mk 0 1 1 (by decide) @@ -55,7 +55,7 @@ instance : Inhabited PlainDate where Creates a new `PlainDate` from year, month, and day components. -/ @[inline] -def ofYearMonthDay (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option PlainDate := +def ofYearMonthDay? (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Option PlainDate := if valid : year.Valid month day then some (PlainDate.mk year month day valid) else none @@ -102,10 +102,46 @@ def weekday (date : PlainDate) : Weekday := let j : Bounded.LE (-10) 9 := (Bounded.LE.byMod y 1000 (by decide)).ediv 100 (by decide) let part : Bounded.LE 6 190 := q.addBounds (((m.add 1).mul_pos 13 (by decide)).ediv 5 (by decide)) |>.addBounds k |>.addBounds (k.ediv 4 (by decide)) let h : Bounded.LE (-15) 212 := part.addBounds ((j.ediv 4 (by decide)).addBounds (j.mul_pos 2 (by decide)).neg) - let d := (h.add 6).emod 7 (by decide) + let d := (h.add 5).emod 7 (by decide) .ofOrdinal (d.add 1) +/-- +Determines the week of the year for the given `PlainDate`. +-/ +def weekOfYear (date : PlainDate) : Week.Ordinal := + let res := Month.Ordinal.toOrdinal ⟨⟨date.month, date.day⟩, date.valid⟩ |>.ediv 7 (by decide) |>.add 1 + match date.year.isLeap, res with + | true, res => res + | false, res => res + +/-- +Returns the unaligned week of the month for a `PlainDate` (day divided by 7, plus 1). +-/ +def weekOfMonth (date : PlainDate) : Bounded.LE 1 5 := + date.day.ediv 7 (by decide) |>.add 1 + +/-- +Determines the week of the month for the given `PlainDate`. The week of the month is calculated based +on the day of the month and the weekday. Each week starts on Sunday because the entire library is +based on the Gregorian Calendar. +-/ +def alignedWeekOfMonth (date : PlainDate) : Week.Ordinal.OfMonth := + let weekday := date.weekday.toOrdinal + date.day.addBounds (weekday.sub 1) |>.ediv 7 (by decide) |>.add 1 + +/-- +Determines the quarter of the year for the given `PlainDate`. +-/ +def quarter (date : PlainDate) : Bounded.LE 1 4 := + date.month.sub 1 |>.ediv 3 (by decide) |>.add 1 + +/-- +Transforms a tuple of a `PlainDate` into a `Day.Ordinal.OfYear`. +-/ +def toOrdinal (date : PlainDate) : Day.Ordinal.OfYear date.year.isLeap := + Month.Ordinal.toOrdinal ⟨(date.month, date.day), date.valid⟩ + /-- Determines the era of the given `PlainDate` based on its year. -/ @@ -250,7 +286,7 @@ Creates a new `PlainDate` by adjusting the day of the month to the given `days` out-of-range days clipped to the nearest valid date. -/ @[inline] -def withDaysClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := +def withDayClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := clip dt.year dt.month days /-- @@ -258,7 +294,7 @@ Creates a new `PlainDate` by adjusting the day of the month to the given `days` out-of-range days rolled over to the next month or year as needed. -/ @[inline] -def withDaysRollOver (dt : PlainDate) (days : Day.Ordinal) : PlainDate := +def withDayRollOver (dt : PlainDate) (days : Day.Ordinal) : PlainDate := rollOver dt.year dt.month days /-- diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index 97e23aae649e..e3d9c3fa786a 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -19,9 +19,6 @@ set_option linter.all true def Ordinal := Bounded.LE 1 31 deriving Repr, BEq, LE, LT -instance : ToString Ordinal where - toString x := toString x.val - instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (30 : Nat))) n) @@ -48,46 +45,6 @@ instance {x y : Offset} : Decidable (x ≤ y) := instance {x y : Offset} : Decidable (x < y) := inferInstanceAs (Decidable (x.val < y.val)) -/-- -`Period` is an enumeration representing different times of the day : morning, afternoon, evening, and night. --/ -inductive Period - /-- Represents the morning period. -/ - | morning - - /-- Represents the afternoon period. -/ - | afternoon - - /-- Represents the evening period. -/ - | evening - - /-- Represents the night period. -/ - | night - deriving Repr, BEq, Inhabited - -namespace Period - -/-- -Determines the `Period` of the day based on the given hour. - -- If the hour is between 20 and 4, it returns `night`. -- If the hour is between 17 and 20, it returns `evening`. -- If the hour is between 12 and 17, it returns `afternoon`. -- If the hour is between 5 and 12, it reutrns `morning`. --/ -@[inline] -def fromHour (hour : Hour.Ordinal) : Day.Period := - if hour ≥ 20 ∨ hour ≤ 4 then - .night - else if hour ≥ 17 then - .evening - else if hour ≥ 12 then - .afternoon - else - .morning - -end Period - namespace Ordinal /-- diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index c2fa99d985a7..2fb6b7598035 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -21,9 +21,6 @@ set_option linter.all true def Ordinal := Bounded.LE 1 12 deriving Repr, BEq, LE, LT -instance : ToString Ordinal where - toString x := toString x.val - instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (11 : Nat))) n) @@ -236,11 +233,10 @@ Gets the number of days in a month. def days (leap : Bool) (month : Ordinal) : Day.Ordinal := if month.val = 2 then if leap then 29 else 28 - else by + else let ⟨months, p⟩ := monthSizesNonLeap let index : Fin 12 := (month.sub 1).toFin (by decide) (by decide) - rw [← p] at index - exact months.get index + months.get (index.cast (by rw [p])) theorem days_gt_27 (leap : Bool) (i : Month.Ordinal) : days leap i > 27 := by match i with @@ -262,7 +258,7 @@ def cumulativeDays (leap : Bool) (month : Ordinal) : Day.Offset := by let res := months.get index exact res + (if leap ∧ month.val > 2 then 1 else 0) -theorem cumulativeDays_le_335 (leap : Bool) (month : Month.Ordinal) : cumulativeDays leap month ≥ 0 ∧ cumulativeDays leap month ≤ 334 + (if leap then 1 else 0) := by +theorem cumulativeDays_le (leap : Bool) (month : Month.Ordinal) : cumulativeDays leap month ≥ 0 ∧ cumulativeDays leap month ≤ 334 + (if leap then 1 else 0) := by match month with | ⟨1, _⟩ | ⟨2, _⟩ | ⟨3, _⟩ | ⟨4, _⟩ | ⟨5, _⟩ | ⟨6, _⟩ | ⟨7, _⟩ | ⟨8, _⟩ | ⟨9, _⟩ | ⟨10, _⟩ | ⟨11, _⟩ | ⟨12, _⟩ => simp [cumulativeSizes, Bounded.LE.sub, Bounded.LE.add, Bounded.LE.toFin, cumulativeDays] @@ -280,13 +276,15 @@ theorem difference_eq (p : month.val ≤ 11) : | ⟨12, _⟩ => contradiction /-- -Check if the day is valid in a month and a leap year. +Checks if a given day is valid for the specified month and year. For example, `29/02` is valid only +if the year is a leap year. -/ abbrev Valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := day.val ≤ (days leap month).val /-- -Type for dates and months that are valid within a leap year. +Represents a valid date for a given year, considering whether it is a leap year. Example: `(2, 29)` +is valid only if `leap` is `true`. -/ def ValidDate (leap : Bool) := { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } @@ -298,7 +296,7 @@ Transforms a tuple of a `Month` and a `Day` into a `Day.Ordinal.OfYear`. -/ def toOrdinal (ordinal : ValidDate leap) : Day.Ordinal.OfYear leap := let days := cumulativeDays leap ordinal.val.fst - let proof := cumulativeDays_le_335 leap ordinal.val.fst + let proof := cumulativeDays_le leap ordinal.val.fst let bounded := Bounded.LE.mk days.toInt proof |>.addBounds ordinal.val.snd match leap, bounded with | true, bounded => bounded @@ -361,7 +359,7 @@ def clipDay (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Day.Ordi /-- Proves that every value provided by a clipDay is a valid day in a year. -/ -theorem clipDay_valid : Valid leap month (clipDay leap month day) := by +theorem valid_clipDay : Valid leap month (clipDay leap month day) := by simp [Valid, clipDay] split exact Int.le_refl (days leap month).val diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 4127392bbf11..9d7e49597987 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -21,9 +21,6 @@ set_option linter.all true def Ordinal := Bounded.LE 1 53 deriving Repr, BEq, LE, LT -instance : ToString Ordinal where - toString x := toString x.val - instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 1 (1 + (52 : Nat))) n) diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index 18e94e4d02ad..2df96a8d5001 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -55,49 +55,49 @@ Converts a `Ordinal` representing a day index into a corresponding `Weekday`. Th for mapping numerical representations to days of the week. -/ def ofOrdinal : Ordinal → Weekday - | 1 => .sunday - | 2 => .monday - | 3 => .tuesday - | 4 => .wednesday - | 5 => .thursday - | 6 => .friday - | 7 => .saturday + | 1 => .monday + | 2 => .tuesday + | 3 => .wednesday + | 4 => .thursday + | 5 => .friday + | 6 => .saturday + | 7 => .sunday /-- Converts a `Weekday` to a `Ordinal`. -/ def toOrdinal : Weekday → Ordinal - | .sunday => 1 - | .monday => 2 - | .tuesday => 3 - | .wednesday => 4 - | .thursday => 5 - | .friday => 6 - | .saturday => 7 + | .monday => 1 + | .tuesday => 2 + | .wednesday => 3 + | .thursday => 4 + | .friday => 5 + | .saturday => 6 + | .sunday => 7 /-- Converts a `Weekday` to a `Nat`. -/ def toNat : Weekday → Nat - | .sunday => 1 - | .monday => 2 - | .tuesday => 3 - | .wednesday => 4 - | .thursday => 5 - | .friday => 6 - | .saturday => 7 + | .monday => 1 + | .tuesday => 2 + | .wednesday => 3 + | .thursday => 4 + | .friday => 5 + | .saturday => 6 + | .sunday => 7 /-- Converts a `Nat` to an `Option Weekday`. -/ def ofNat? : Nat → Option Weekday - | 1 => some .sunday - | 2 => some .monday - | 3 => some .tuesday - | 4 => some .wednesday - | 5 => some .thursday - | 6 => some .friday - | 7 => some .saturday + | 1 => some .monday + | 2 => some .tuesday + | 3 => some .wednesday + | 4 => some .thursday + | 5 => some .friday + | 6 => some .saturday + | 7 => some .sunday | _ => none /-- @@ -122,9 +122,9 @@ def next : Weekday → Weekday | .sunday => .monday /-- -Check if it's a Weekend. +Check if it's a weekend. -/ -def weekend : Weekday → Bool +def isWeekend : Weekday → Bool | .saturday => true | .sunday => true | _ => false diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index fb64d6d996f0..b42c9ef8d4ee 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -95,9 +95,6 @@ Checks if the given date is valid for the specified year, month, and day. abbrev Valid (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := day ≤ month.days year.isLeap -instance : Decidable (Valid year month day) := - dite (day ≤ month.days year.isLeap) isTrue isFalse - end Offset end Year end Time diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 6625be2dca6e..2cb13344dc3d 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -46,19 +46,11 @@ def toPlainDateAssumingUTC (timestamp : Timestamp) : PlainDate := let days := Day.Offset.ofSeconds secs PlainDate.ofDaysSinceUNIXEpoch days -/-- -Converts a `PlainTime` to a `Timestamp` --/ -@[inline] -def ofPlainTime (pt : PlainTime) : Timestamp := - let nanos := pt.toNanoseconds - Timestamp.ofNanosecondsSinceUnixEpoch nanos - /-- Converts a `Timestamp` to a `PlainTime` -/ @[inline] -def getTime (timestamp : Timestamp) : PlainTime := +def getTimeAssumingUTC (timestamp : Timestamp) : PlainTime := let nanos := timestamp.toNanosecondsSinceUnixEpoch PlainTime.ofNanoseconds nanos diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 7640f1a5d82a..7d6af0b9cf01 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -381,6 +381,13 @@ Getter for the `Day` inside of a `PlainDateTime`. def day (dt : PlainDateTime) : Day.Ordinal := dt.date.day +/-- +Getter for the `Weekday` inside of a `PlainDateTime`. +-/ +@[inline] +def weekday (dt : PlainDateTime) : Weekday := + dt.date.weekday + /-- Getter for the `Hour` inside of a `PlainDateTime`. -/ @@ -423,6 +430,42 @@ Checks if the `PlainDateTime` is in a leap year. def inLeapYear (date : PlainDateTime) : Bool := date.year.isLeap +/-- +Determines the week of the year for the given `PlainDateTime`. +-/ +@[inline] +def weekOfYear (date : PlainDateTime) : Week.Ordinal := + date.date.weekOfYear + +/-- +Returns the unaligned week of the month for a `PlainDateTime` (day divided by 7, plus 1). +-/ +def weekOfMonth (date : PlainDateTime) : Bounded.LE 1 5 := + date.date.weekOfMonth + +/-- +Determines the week of the month for the given `PlainDateTime`. The week of the month is calculated based +on the day of the month and the weekday. Each week starts on Sunday because the entire library is +based on the Gregorian Calendar. +-/ +@[inline] +def alignedWeekOfMonth (date : PlainDateTime) : Week.Ordinal.OfMonth := + date.date.alignedWeekOfMonth + +/-- +Transforms a tuple of a `PlainDateTime` into a `Day.Ordinal.OfYear`. +-/ +@[inline] +def toOrdinal (date : PlainDateTime) : Day.Ordinal.OfYear date.year.isLeap := + Month.Ordinal.toOrdinal ⟨(date.month, date.day), date.date.valid⟩ + +/-- +Determines the quarter of the year for the given `PlainDateTime`. +-/ +@[inline] +def quarter (date : PlainDateTime) : Bounded.LE 1 4 := + date.date.quarter + instance : HAdd PlainDateTime Day.Offset PlainDateTime where hAdd := addDays diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 7fb9acd57331..e499e947a4da 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -141,8 +141,8 @@ namespace TimeZone Parses a string into a `TimeZone` object. The input string must be in the format `"VV ZZZZZ"`. -/ def fromTimeZone (input : String) : Except String TimeZone := do - let spec : GenericFormat .any := datespec("VV ZZZZZ") - spec.parseBuilder (λid off => some (TimeZone.mk off id "Unknown" false)) input + let spec : GenericFormat .any := datespec("V ZZZZZ") + spec.parseBuilder (fun id off => some (TimeZone.mk off id "Unknown" false)) input namespace Offset @@ -167,10 +167,17 @@ def format (date : PlainDate) (format : String) : String := | .error err => s!"error: {err}" | .ok res => let res := res.formatGeneric fun + | .G _ => some date.era | .y _ => some date.year + | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) + | .Qorq _ => some date.quarter + | .w _ => some date.weekOfYear + | .W _ => some date.weekOfMonth | .MorL _ => some date.month | .d _ => some date.day | .E _ => some date.weekday + | .eorc _ => some date.weekday + | .F _ => some date.alignedWeekOfMonth | _ => none match res with | some res => res @@ -180,7 +187,7 @@ def format (date : PlainDate) (format : String) : String := Parses a date string in the American format (`MM-dd-yyyy`) and returns a `PlainDate`. -/ def fromAmericanDateString (input : String) : Except String PlainDate := do - Formats.americanDate.parseBuilder (fun m d y => PlainDate.ofYearMonthDay y m d) input + Formats.americanDate.parseBuilder (fun m d y => PlainDate.ofYearMonthDay? y m d) input /-- Converts a date in the American format (`MM-dd-yyyy`) into a `String`. @@ -192,7 +199,7 @@ def toAmericanDateString (input : PlainDate) : String := Parses a date string in the SQL format (`yyyy-MM-dd`) and returns a `PlainDate`. -/ def fromSQLDateString (input : String) : Except String PlainDate := do - Formats.sqlDate.parseBuilder (PlainDate.ofYearMonthDay) input + Formats.sqlDate.parseBuilder PlainDate.ofYearMonthDay? input /-- Converts a date in the SQL format (`yyyy-MM-dd`) into a `String`. @@ -204,7 +211,7 @@ def toSQLDateString (input : PlainDate) : String := Parses a date string in the Lean format (`yyyy-MM-dd`) and returns a `PlainDate`. -/ def fromLeanDateString (input : String) : Except String PlainDate := do - Formats.leanDate.parseBuilder (PlainDate.ofYearMonthDay) input + Formats.leanDate.parseBuilder PlainDate.ofYearMonthDay? input /-- Converts a date in the Lean format (`yyyy-MM-dd`) into a `String`. @@ -239,9 +246,16 @@ def format (time : PlainTime) (format : String) : String := | .ok res => let res := res.formatGeneric fun | .H _ => some time.hour + | .k _ => some (time.hour.add 1) | .m _ => some time.minute | .n _ => some time.nano | .s _ => some time.second + | .a _ => some (HourMarker.ofOrdinal time.hour) + | .h _ => some time.hour.toRelative + | .K _ => some (time.hour.emod 12 (by decide)) + | .S _ => some time.nano + | .A _ => some time.toMilliseconds + | .N _ => some time.toNanoseconds | _ => none match res with | some res => res @@ -402,15 +416,28 @@ def format (date : PlainDateTime) (format : String) : String := | .error err => s!"error: {err}" | .ok res => let res := res.formatGeneric fun - | .y _ => some date.date.year - | .MorL _ => some date.date.month - | .d _ => some date.date.day - | .E _ => some date.date.weekday - | .H _ => some date.time.hour - | .m _ => some date.time.minute - | .n _ => some date.time.nano - | .S _ => some date.time.nano + | .G _ => some date.era + | .y _ => some date.year + | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) + | .Qorq _ => some date.quarter + | .w _ => some date.weekOfYear + | .W _ => some date.weekOfMonth + | .MorL _ => some date.month + | .d _ => some date.day + | .E _ => some date.weekday + | .eorc _ => some date.weekday + | .F _ => some date.alignedWeekOfMonth + | .H _ => some date.hour + | .k _ => some (date.hour.add 1) + | .m _ => some date.minute + | .n _ => some date.nanosecond | .s _ => some date.time.second + | .a _ => some (HourMarker.ofOrdinal date.hour) + | .h _ => some date.hour.toRelative + | .K _ => some (date.hour.emod 12 (by decide)) + | .S _ => some date.nanosecond + | .A _ => some date.time.toMilliseconds + | .N _ => some date.time.toNanoseconds | _ => none match res with | some res => res diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 46447af2ec3c..f11a88f88388 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -345,11 +345,6 @@ inductive Modifier -/ | a (presentation : Text) - /-- - `B`: Period of day (e.g., in the morning). - -/ - | B (presentation : Text) - /-- `h`: Clock hour of AM/PM (1-12) (e.g., 12). -/ @@ -463,7 +458,7 @@ private def parseOffsetO (constructor : OffsetO → Modifier) (p : String) : Par parseMod constructor OffsetO.classify p private def parseZoneId (p : String) : Parser Modifier := - if p.length = 2 then pure .V else fail s!"invalid quantity of characters for '{p.get 0}'" + if p.length = 1 then pure .V else fail s!"invalid quantity of characters for '{p.get 0}'" private def parseNumberText (constructor : (Number ⊕ Text) → Modifier) (p : String) : Parser Modifier := parseMod constructor classifyNumberText p @@ -490,7 +485,6 @@ private def parseModifier : Parser Modifier <|> parseNumberText Modifier.eorc =<< many1Chars (pchar 'c') <|> parseNumber Modifier.F =<< many1Chars (pchar 'F') <|> parseText Modifier.a =<< many1Chars (pchar 'a') - <|> parseText Modifier.B =<< many1Chars (pchar 'B') <|> parseNumber Modifier.h =<< many1Chars (pchar 'h') <|> parseNumber Modifier.K =<< many1Chars (pchar 'K') <|> parseNumber Modifier.k =<< many1Chars (pchar 'k') @@ -731,24 +725,6 @@ private def formatMarkerNarrow (marker : HourMarker) : String := | .am => "A" | .pm => "P" -private def formatPeriodOfDayLong : Day.Period → String - | .morning => "in the morning" - | .afternoon => "in the afternoon" - | .evening => "in the evening" - | .night => "at night" - -private def formatPeriodOfDayShort : Day.Period → String - | .morning => "AM" - | .afternoon => "PM" - | .evening => "Eve" - | .night => "Night" - -private def formatPeriodOfDayNarrow : Day.Period → String - | .morning => "M" - | .afternoon => "A" - | .evening => "E" - | .night => "N" - private def toSigned (data : Int) : String := if data < 0 then toString data else "+" ++ toString data @@ -771,12 +747,11 @@ private def TypeFormat : Modifier → Type | .d _ => Day.Ordinal | .Qorq _ => Month.Quarter | .w _ => Week.Ordinal - | .W _ => Week.Ordinal.OfMonth + | .W _ => Bounded.LE 1 5 | .E _ => Weekday | .eorc _ => Weekday | .F _ => Week.Ordinal.OfMonth | .a _ => HourMarker - | .B _ => Day.Period | .h _ => Bounded.LE 1 12 | .K _ => Bounded.LE 0 11 | .k _ => Bounded.LE 1 24 @@ -844,11 +819,6 @@ private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : Strin | .short => formatMarkerShort data | .full => formatMarkerLong data | .narrow => formatMarkerNarrow data - | .B format => - match format with - | .short => formatPeriodOfDayLong data - | .full => formatPeriodOfDayShort data - | .narrow => formatPeriodOfDayNarrow data | .h format => truncate format.padding (data.val % 12) | .K format => truncate format.padding (data.val % 12) | .k format => truncate format.padding (data.val) @@ -872,7 +842,7 @@ private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : Strin | .full => data | .O format => match format with - | .short => s!"GMT{toSigned data.hour.val}" + | .short => s!"GMT{toSigned data.second.toHours.toInt}" | .full => s!"GMT{toIsoString data true false true}" | .X format => if data.second == 0 then @@ -921,9 +891,8 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .W _ => date.weekOfMonth | .E _ => date.weekday | .eorc _ => date.weekday - | .F _ => date.weekOfMonth + | .F _ => date.alignedWeekOfMonth | .a _ => HourMarker.ofOrdinal date.hour - | .B _ => date.period | .h _ => HourMarker.toRelative date.hour |>.fst | .K _ => date.hour.emod 12 (by decide) | .k _ => date.hour.add 1 @@ -1055,24 +1024,6 @@ private def parseMarkerNarrow : Parser HourMarker := pstring "A" *> pure HourMarker.am <|> pstring "P" *> pure HourMarker.pm -private def parsePeriodOfDayLong : Parser Day.Period - := pstring "in the morning" *> pure .morning - <|> pstring "in the afternoon" *> pure .afternoon - <|> pstring "in the evening" *> pure .evening - <|> pstring "at night" *> pure .night - -private def parsePeriodOfDayShort : Parser Day.Period - := pstring "AM" *> pure .morning - <|> pstring "PM" *> pure .afternoon - <|> pstring "Eve" *> pure .evening - <|> pstring "Night" *> pure .night - -private def parsePeriodOfDayNarrow : Parser Day.Period - := pstring "M" *> pure .morning - <|> pstring "A" *> pure .afternoon - <|> pstring "E" *> pure .evening - <|> pstring "N" *> pure .night - private def exactly (parse : Parser α) (size : Nat) : Parser (Array α) := let rec go (acc : Array α) (count : Nat) : Parser (Array α) := if count ≥ size then @@ -1184,11 +1135,6 @@ private def parseWith : (mod : Modifier) → Parser (TypeFormat mod) | .short => parseMarkerShort | .full => parseMarkerLong | .narrow => parseMarkerNarrow - | .B format => - match format with - | .short => parsePeriodOfDayLong - | .full => parsePeriodOfDayShort - | .narrow => parsePeriodOfDayNarrow | .h format => parseNatToBounded (parseNum format.padding) | .K format => parseNatToBounded (parseNum format.padding) | .k format => parseNatToBounded (parseNum format.padding) @@ -1265,12 +1211,11 @@ private structure DateBuilder where d : Option Day.Ordinal := none Qorq : Option Month.Quarter := none w : Option Week.Ordinal := none - W : Option Week.Ordinal.OfMonth := none + W : Option (Bounded.LE 1 5) := none E : Option Weekday := none eorc : Option Weekday := none F : Option Week.Ordinal.OfMonth := none a : Option HourMarker := none - B : Option Day.Period := none h : Option (Bounded.LE 1 12) := none K : Option (Bounded.LE 0 11) := none k : Option (Bounded.LE 1 24) := none @@ -1306,7 +1251,6 @@ private def insert (date : DateBuilder) (modifier : Modifier) (data : TypeFormat | .eorc _ => { date with eorc := some data } | .F _ => { date with F := some data } | .a _ => { date with a := some data } - | .B _ => { date with B := some data } | .h _ => { date with h := some data } | .K _ => { date with K := some data } | .k _ => { date with k := some data } diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 33ece36c104f..25611ae09fbd 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -75,7 +75,6 @@ private def convertModifier : Modifier → MacroM (TSyntax `term) | .inr txt => do `(Std.Time.Modifier.eorc (.inr $(← convertText txt))) | .F p => do `(Std.Time.Modifier.F $(← convertNumber p)) | .a p => do `(Std.Time.Modifier.a $(← convertText p)) - | .B p => do `(Std.Time.Modifier.B $(← convertText p)) | .h p => do `(Std.Time.Modifier.h $(← convertNumber p)) | .K p => do `(Std.Time.Modifier.K $(← convertNumber p)) | .k p => do `(Std.Time.Modifier.k $(← convertNumber p)) diff --git a/src/Std/Time/Notation/Spec.lean b/src/Std/Time/Notation/Spec.lean index 4b49af7bbb2e..d12252164028 100644 --- a/src/Std/Time/Notation/Spec.lean +++ b/src/Std/Time/Notation/Spec.lean @@ -73,7 +73,6 @@ private def convertModifier : Modifier → MacroM (TSyntax `term) | .inr txt => do `(Std.Time.Modifier.eorc (.inr $(← convertText txt))) | .F p => do `(Std.Time.Modifier.F $(← convertNumber p)) | .a p => do `(Std.Time.Modifier.a $(← convertText p)) - | .B p => do `(Std.Time.Modifier.B $(← convertText p)) | .h p => do `(Std.Time.Modifier.h $(← convertNumber p)) | .K p => do `(Std.Time.Modifier.K $(← convertNumber p)) | .k p => do `(Std.Time.Modifier.k $(← convertNumber p)) diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean index 7b0246eeaf58..0a05dd161384 100644 --- a/src/Std/Time/Time/HourMarker.lean +++ b/src/Std/Time/Time/HourMarker.lean @@ -17,9 +17,15 @@ set_option linter.all true 11:59 AM, and `pm` for hours between 12:00 PM and 11:59 PM. -/ inductive HourMarker - /-- Ante meridiem. -/ + + /-- + Ante meridiem. + -/ | am - /-- Post meridiem. -/ + + /-- + Post meridiem. + -/ | pm deriving Repr, BEq diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index e5a6eac664ec..3de289e9a26d 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -23,9 +23,6 @@ set_option linter.all true def Ordinal := Bounded.LE 0 23 deriving Repr, BEq, LE, LT -instance : ToString Ordinal where - toString x := toString x.val - instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (23 : Nat))) n) @@ -78,6 +75,23 @@ def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val end Ordinal +namespace Offset + +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + UnitVal.mk data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + UnitVal.mk data + +end Offset end Hour end Time end Std diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index e900df7ab78d..b01fc897384f 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -22,9 +22,6 @@ set_option linter.all true def Ordinal := Bounded.LE 0 999 deriving Repr, BEq, LE, LT -instance : ToString Ordinal where - toString x := toString x.val - instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (999 : Nat))) n) diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index 8e0447b05a26..822b14e172a9 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -22,9 +22,6 @@ set_option linter.all true def Ordinal := Bounded.LE 0 59 deriving Repr, BEq, LE, LT -instance : ToString Ordinal where - toString x := toString x.val - instance : OfNat Ordinal n := inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) n) @@ -77,6 +74,23 @@ def toOffset (ordinal : Ordinal) : Offset := UnitVal.ofInt ordinal.val end Ordinal +namespace Offset + +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Offset := + UnitVal.mk data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Offset := + UnitVal.mk data + +end Offset end Minute end Time end Std diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index c56bcda46f6c..d009a48e15b7 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -21,9 +21,6 @@ set_option linter.all true def Ordinal := Bounded.LE 0 999999999 deriving Repr, BEq, LE, LT -instance : ToString Ordinal where - toString x := toString x.val - instance : OfNat Ordinal n where ofNat := Bounded.LE.ofFin (Fin.ofNat n) diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 13ddc12d6915..1dc112a8cd41 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -30,9 +30,6 @@ instance : LE (Ordinal leap) where instance : LT (Ordinal leap) where lt x y := LT.lt x.val y.val -instance : ToString (Ordinal leap) where - toString x := toString x.val - instance : Repr (Ordinal l) where reprPrec r := reprPrec r.val diff --git a/src/Std/Time/Zoned/Database/Leap.lean b/src/Std/Time/Zoned/Database/Leap.lean index d6d30683cd69..4cabdc1156b2 100644 --- a/src/Std/Time/Zoned/Database/Leap.lean +++ b/src/Std/Time/Zoned/Database/Leap.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sofia Rodrigues -/ +prelude import Init.Data.Range import Std.Internal.Parsec import Std.Internal.Parsec.ByteArray @@ -52,7 +53,7 @@ private def failWith (opt : Option α) : Parser α := | none => fail "invalid number" | some res => pure res -def parseMonthShort : Parser Month.Ordinal +private def parseMonthShort : Parser Month.Ordinal := pstring "Jan" *> pure ⟨1, by decide⟩ <|> pstring "Feb" *> pure ⟨2, by decide⟩ <|> pstring "Mar" *> pure ⟨3, by decide⟩ @@ -88,11 +89,11 @@ private def parseLeapSecond : Parser LeapSecond := do let stationary ← manyChars (satisfy Char.isAlpha) skipChar '\n' - dbg_trace s!"{year}--{month}--{day}" + dbg_trace s!"{repr year}--{repr month}--{day}" let day ← failWith <| Internal.Bounded.ofInt? day let time : PlainTime ← failWith <| PlainTime.ofHourMinuteSeconds hour minute second - let date : PlainDate ← failWith <| PlainDate.ofYearMonthDay (Year.Offset.ofNat year) month day + let date : PlainDate ← failWith <| PlainDate.ofYearMonthDay? (Year.Offset.ofNat year) month day let pdt := PlainDateTime.mk date time return { timestamp := pdt.toTimestampAssumingUTC, positive, stationary } diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index c5790cdb947a..230c215caf72 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -24,8 +24,8 @@ opaque Windows.getTimeZoneAt : String -> UInt64 -> IO TimeZone /-- Fetches the timezone at a timestamp. -/ -@[extern "lean_get_timezone_offset_at"] -opaque Windows.getLocalTimezoneAt : UInt64 -> IO TimeZone +@[extern "lean_get_windows_local_timezone_id_at"] +opaque Windows.getLocalTimeZoneIdentifierAt : UInt64 → IO String /-- Represents a Time Zone Database that we get from ICU available on Windows SDK. @@ -42,5 +42,10 @@ Returns a default `WindowsDb` instance. def default : WindowsDb := {} instance : Database WindowsDb where - getTimeZoneAt _ id tm := Windows.getTimeZoneAt id (tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64) - getLocalTimeZoneAt _ tm := Windows.getLocalTimezoneAt (tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64) + getTimeZoneAt _ id tm := + Windows.getTimeZoneAt id (tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64) + + getLocalTimeZoneAt _ tm := do + let time := tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64 + let id ← Windows.getLocalTimeZoneIdentifierAt time + Windows.getTimeZoneAt id time diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 3f07f2e2f65f..5d6e9c8e2411 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -15,7 +15,7 @@ open Internal set_option linter.all true /-- -It stores a `Timestamp`, a `PlainDateTime` and a `TimeZone` +Represents a specific point in time associated with a `TimeZone`. -/ structure DateTime (tz : TimeZone) where private mk :: @@ -366,14 +366,6 @@ Gets the `Weekday` of a DateTime. def weekday (dt : DateTime tz) : Weekday := dt.date.get.date.weekday -/-- -Gets the `Period` of a `DateTime`, corresponding to the part of the day (e.g. night, morning, -afternoon, evening) based on the hour. --/ -@[inline] -def period (dt : DateTime tz) : Day.Period := - Day.Period.fromHour dt.hour - /-- Determines the era of the given `DateTime` based on its year. -/ @@ -395,26 +387,31 @@ def toOrdinal (date : DateTime tz) : Day.Ordinal.OfYear date.year.isLeap := /-- Determines the week of the year for the given `DateTime`. -/ +@[inline] def weekOfYear (date : DateTime tz) : Week.Ordinal := - let res := Month.Ordinal.toOrdinal ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ |>.ediv 7 (by decide) |>.add 1 - match date.date.get.date.year.isLeap, res with - | true, res => res - | false, res => res + date.date.get.weekOfYear + +/-- +Returns the unaligned week of the month for a `DateTime` (day divided by 7, plus 1). +-/ +def weekOfMonth (date : DateTime tz) : Bounded.LE 1 5 := + date.date.get.weekOfMonth /-- Determines the week of the month for the given `DateTime`. The week of the month is calculated based on the day of the month and the weekday. Each week starts on Sunday because the entire library is based on the Gregorian Calendar. -/ -def weekOfMonth (date : DateTime tz) : Week.Ordinal.OfMonth := - let weekday := date.weekday.toOrdinal - date.day.addBounds (weekday.sub 1) |>.ediv 7 (by decide) |>.add 1 +@[inline] +def alignedWeekOfMonth (date : DateTime tz) : Week.Ordinal.OfMonth := + date.date.get.alignedWeekOfMonth /-- Determines the quarter of the year for the given `DateTime`. -/ +@[inline] def quarter (date : DateTime tz) : Bounded.LE 1 4 := - date.month.sub 1 |>.ediv 3 (by decide) |>.add 1 + date.date.get.quarter instance : ToTimestamp (DateTime tz) where toTimestamp dt := dt.toUTCTimestamp diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 70a8774c9fa5..7263cb472174 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -18,24 +18,14 @@ Represents a timezone offset with an hour and second component. -/ structure Offset where - /-- - The timezone offset in Hours. - -/ - hour : Hour.Offset - /-- The same timezone offset in seconds. -/ second : Second.Offset - - /-- - The proof that both are equal - -/ - proof : second.toHours = hour deriving Repr instance : Inhabited Offset where - default := ⟨0, 0, rfl⟩ + default := ⟨0⟩ instance : BEq Offset where beq x y := BEq.beq x.second y.second @@ -59,26 +49,25 @@ def toIsoString (offset : Offset) (colon : Bool) : String := A zero `Offset` representing UTC (no offset). -/ def zero : Offset := - { hour := 0, second := 0, proof := rfl } + { second := 0 } /-- Creates an `Offset` from a given number of hour. -/ def ofHours (n : Hour.Offset) : Offset := - mk n n.toSeconds (by simp [Hour.Offset.toSeconds, Second.Offset.toHours, UnitVal.mul, UnitVal.div]; rfl) + mk n.toSeconds /-- -Creates an `Offset` from a given number of hour and minuets. +Creates an `Offset` from a given number of hours and minutes. -/ def ofHoursAndMinutes (n : Hour.Offset) (m : Minute.Offset) : Offset := - let secs := n.toSeconds + m.toSeconds - mk secs.toHours secs rfl + mk (n.toSeconds + m.toSeconds) /-- -Creates an `Offset` from a given number of second. +Creates an `Offset` from a given number of seconds. -/ def ofSeconds (n : Second.Offset) : Offset := - mk n.toHours n rfl + ⟨n⟩ end Offset end TimeZone diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 733a9f604c4b..c3b76a3702b5 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -138,7 +138,7 @@ def createTimeZoneFromTransition (transition : Transition) : TimeZone := Applies the transition to a Timestamp. -/ def apply (timestamp : Timestamp) (transition : Transition) : Timestamp := - let offsetInSeconds := transition.localTimeType.gmtOffset.hour.mul 3600 |>.add transition.localTimeType.gmtOffset.second + let offsetInSeconds := transition.localTimeType.gmtOffset.second |>.add transition.localTimeType.gmtOffset.second timestamp.addSeconds offsetInSeconds /-- diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index ce8d437d56a1..13c4b6da0007 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -157,6 +157,13 @@ Returns the weekday. def weekday (zdt : ZonedDateTime) : Weekday := zdt.snd.weekday +/-- +Transforms a tuple of a `ZonedDateTime` into a `Day.Ordinal.OfYear`. +-/ +@[inline] +def toOrdinal (date : ZonedDateTime) : Day.Ordinal.OfYear date.year.isLeap := + Month.Ordinal.toOrdinal ⟨(date.month, date.day), date.snd.date.get.date.valid⟩ + /-- Add `Day.Offset` to a `ZonedDateTime`. -/ diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index bf087a70d222..9a34f3e783f9 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -6,8 +6,6 @@ Authors: Leonardo de Moura, Sebastian Ullrich */ #if defined(LEAN_WINDOWS) #include -#include -#include #include #include #define NOMINMAX // prevent ntdef.h from defining min/max macros @@ -541,8 +539,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } -/* Std.Time.TimeZone.getWindowsTimeZoneAt : String -> UInt64 -> IO Timezone */ -extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str, uint64_t timestamp_secs, obj_arg /* w */) { +/* Std.Time.Database.Windows.getLocalTimeZoneIdentifier : IO String */ +extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; const char* dst_name = lean_string_cstr(timezone_str); @@ -593,11 +591,14 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); zone_offset += dst_offset; - + if (U_FAILURE(status)) { + ucal_close(cal); return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); } + ucal_close(cal); + int offset_hour = zone_offset / 3600000; int offset_seconds = zone_offset / 1000; int is_dst = dst_offset != 0; @@ -611,19 +612,18 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked(dst_name)); lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked(display_name_str)); lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); - + return lean_io_result_mk_ok(lean_tz); #else return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get timezone, its windows only."))); #endif } -/* Std.Time.TimeZone.getCurrentTimezone : IO Timezone */ -extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset_at(uint64_t timestamp_secs, obj_arg /* w */) { - using namespace std::chrono; +/* SStd.Time.Database.Windows.getLocalTimeZoneIdentifierAt : UInt64 → IO String */ +extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(uint64_t timestamp_secs, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; - UCalendar *cal = ucal_open(NULL, -1, NULL, UCAL_GREGORIAN, &status); + UCalendar* cal = ucal_open(NULL, -1, NULL, UCAL_GREGORIAN, &status); if (U_FAILURE(status)) { return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); @@ -636,80 +636,24 @@ extern "C" LEAN_EXPORT obj_res lean_get_timezone_offset_at(uint64_t timestamp_se } UChar tzId[256]; - int32_t tzIdLength = ucal_getTimeZoneID(cal, tzId, sizeof(tzId)/sizeof(tzId[0]), &status); - - if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); - } + int32_t tzIdLength = ucal_getTimeZoneID(cal, tzId, sizeof(tzId) / sizeof(tzId[0]), &status); + ucal_close(cal); - char dst_name[256]; - u_strToUTF8(dst_name, sizeof(dst_name), NULL, tzId, tzIdLength, &status); if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get timezone ID"))); } - UChar display_name[32]; - int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); + char tzIdStr[256]; + u_strToUTF8(tzIdStr, sizeof(tzIdStr), NULL, tzId, tzIdLength, &status); if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to read abbreaviation"))); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to convert timezone ID to UTF-8"))); } - char display_name_str[32]; - u_strToUTF8(display_name_str, sizeof(display_name_str), NULL, display_name, display_name_len, &status); - - if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get abbreviation to cstr"))); - } - - int32_t zone_offset = ucal_get(cal, UCAL_ZONE_OFFSET, &status); - - if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); - } - - int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); - - zone_offset += dst_offset; - - if (U_FAILURE(status)) { - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); - } - - int offset_hour = zone_offset / 3600000; - int offset_seconds = zone_offset / 1000; - int is_dst = dst_offset != 0; + return lean_io_result_mk_ok(lean_mk_ascii_string_unchecked(tzIdStr)); #else - std::time_t input_time = static_cast(timestamp_secs); - std::tm tm_info; - struct tm *tm_ptr = localtime_r(&input_time, &tm_info); - - if (tm_ptr == NULL) { - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string(""))); - } - - int offset_hour = tm_info.tm_gmtoff / 3600; - int offset_seconds = tm_info.tm_gmtoff; - int is_dst = tm_info.tm_isdst; - char dst_name[] = "Unknown"; - char display_name_str[32] = "Unknown"; + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("timezone retrieval is Windows-only"))); #endif - lean_object *lean_offset = lean_alloc_ctor(0, 2, 0); - lean_ctor_set(lean_offset, 0, lean_int_to_int(offset_hour)); - lean_ctor_set(lean_offset, 1, lean_int_to_int(offset_seconds)); - - lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); - lean_ctor_set(lean_tz, 0, lean_offset); - lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked(dst_name)); - lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked(display_name_str)); - lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); - - return lean_io_result_mk_ok(lean_tz); } /* monoMsNow : BaseIO Nat */ diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 4a61533921fb..452499636a17 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -269,3 +269,477 @@ info: "14-13-12" -/ #guard_msgs in #eval zoned("2002-07-14T14:13:12+09:00").format "HH-mm-ss" + +/- +Format +-/ + +def time₄ := time("23:13:12.324354679") +def date₄ := date("2002-07-14") +def datetime₄ := datetime("2002-07-14T23:13:12.324354679") +def zoned₄ := zoned("2002-07-14T23:13:12.324354679+09:00") + +/-- +info: "CE CE CE Common Era C" +-/ +#guard_msgs in +#eval zoned₄.format "G GG GGG GGGG GGGGG" + +/-- +info: "02 2002 000002002" +-/ +#guard_msgs in +#eval zoned₄.format "yy yyyy yyyyyyyyy" + +/-- +info: "5 95 195" +-/ +#guard_msgs in +#eval zoned₄.format "D DD DDD" + +/-- +info: "7 07 Jul J" +-/ +#guard_msgs in +#eval zoned₄.format "M MM MMM MMMMM" + +/-- +info: "4 14 014 0014 00014" +-/ +#guard_msgs in +#eval zoned₄.format "d dd ddd dddd ddddd" + +/-- +info: "7 07 Jul July J" +-/ +#guard_msgs in +#eval zoned₄.format "M MM MMM MMMM MMMMM" + +/-- +info: "4 14 0014 0014" +-/#guard_msgs in +#eval zoned₄.format "d dd dddd dddd" + +/-- +info: "3 03 3rd quarter 3" +-/ +#guard_msgs in +#eval zoned₄.format "Q QQ QQQQ QQQQQ" + +/-- +info: "8 28 028 0028" +-/ +#guard_msgs in +#eval zoned₄.format "w ww www wwww" + +/-- +info: "3 03 003 0003" +-/ +#guard_msgs in +#eval zoned₄.format "W WW WWW WWWW" + +/-- +info: "Sun Sun Sun Sunday S" +-/ +#guard_msgs in +#eval zoned₄.format "E EE EEE EEEE EEEEE" + +/-- +info: "7 07 Sun Sunday S" +-/ +#guard_msgs in +#eval zoned₄.format "e ee eee eeee eeeee" + +/-- +info: "3 03 003 0003" +-/ +#guard_msgs in +#eval zoned₄.format "F FF FFF FFFF" + +/-- +info: "1 11 011 0011 0011" +-/ +#guard_msgs in +#eval zoned₄.format "h hh hhh hhhh hhhh" + +/-- +info: "1 11 011 0011 000011" +-/ +#guard_msgs in +#eval zoned₄.format "K KK KKK KKKK KKKKKK" + +/-- +info: "4 24 024 0024 000024" +-/ +#guard_msgs in +#eval zoned₄.format "k kk kkk kkkk kkkkkk" + +/-- +info: "3 23 023 0023 00023" +-/ +#guard_msgs in +#eval zoned₄.format "H HH HHH HHHH HHHHH" + +/-- +info: "3 13 013 0013 00013" +-/ +#guard_msgs in +#eval zoned₄.format "m mm mmm mmmm mmmmm" + +/-- +info: "2 12 012 0012 00012" +-/ +#guard_msgs in +#eval zoned₄.format "s ss sss ssss sssss" + + +/-- +info: "9 79 679 4679 54679" +-/#guard_msgs in +#eval zoned₄.format "S SS SSS SSSS SSSSS" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval zoned₄.format "S SS SSS SSSS SSSSSSSSS" + +/-- +info: "4 24 324 2324 083592324" +-/ +#guard_msgs in +#eval zoned₄.format "A AA AAA AAAA AAAAAAAAA" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval zoned₄.format "n nn nnn nnnn nnnnnnnnn" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval zoned₄.format "N NN NNN NNNN NNNNNNNNN" + +/-- +info: "Unknown" +-/ +#guard_msgs in +#eval zoned₄.format "V" + +/-- +info: "Unknown Unknown Unknown Unknown" +-/ +#guard_msgs in +#eval zoned₄.format "z zz zzzz zzzz" + +/-- +info: "GMT+9 GMT+09:00" +-/ +#guard_msgs in +#eval zoned₄.format "O OOOO" + +/-- +info: "+09 +0900 +09:00 +090000 +09:00:00" +-/ +#guard_msgs in +#eval zoned₄.format "X XX XXX XXXX XXXXX" + +/-- +info: "+09 +0900 +09:00 +0900 +09:00:00" +-/ +#guard_msgs in +#eval zoned₄.format "x xx xxx xxxx xxxxx" + +/-- +info: "+0900 +0900 +0900 GMT+09:00 +09:00" +-/ +#guard_msgs in +#eval zoned₄.format "Z ZZ ZZZ ZZZZ ZZZZZ" + +/-- +info: "CE CE CE Common Era C" +-/ +#guard_msgs in +#eval datetime₄.format "G GG GGG GGGG GGGGG" + +/-- +info: "02 2002 000002002" +-/ +#guard_msgs in +#eval datetime₄.format "yy yyyy yyyyyyyyy" + +/-- +info: "5 95 195" +-/ +#guard_msgs in +#eval datetime₄.format "D DD DDD" + +/-- +info: "7 07 Jul J" +-/ +#guard_msgs in +#eval datetime₄.format "M MM MMM MMMMM" + +/-- +info: "4 14 014 0014 00014" +-/ +#guard_msgs in +#eval datetime₄.format "d dd ddd dddd ddddd" + +/-- +info: "7 07 Jul July J" +-/ +#guard_msgs in +#eval datetime₄.format "M MM MMM MMMM MMMMM" + +/-- +info: "4 14 0014 0014" +-/#guard_msgs in +#eval datetime₄.format "d dd dddd dddd" + +/-- +info: "3 03 3rd quarter 3" +-/ +#guard_msgs in +#eval datetime₄.format "Q QQ QQQQ QQQQQ" + +/-- +info: "8 28 028 0028" +-/ +#guard_msgs in +#eval datetime₄.format "w ww www wwww" + +/-- +info: "3 03 003 0003" +-/ +#guard_msgs in +#eval datetime₄.format "W WW WWW WWWW" + +/-- +info: "Sun Sun Sun Sunday S" +-/ +#guard_msgs in +#eval datetime₄.format "E EE EEE EEEE EEEEE" + +/-- +info: "7 07 Sun Sunday S" +-/ +#guard_msgs in +#eval datetime₄.format "e ee eee eeee eeeee" + +/-- +info: "3 03 003 0003" +-/ +#guard_msgs in +#eval datetime₄.format "F FF FFF FFFF" + +/-- +info: "1 11 011 0011 0011" +-/ +#guard_msgs in +#eval datetime₄.format "h hh hhh hhhh hhhh" + +/-- +info: "1 11 011 0011 000011" +-/ +#guard_msgs in +#eval datetime₄.format "K KK KKK KKKK KKKKKK" + +/-- +info: "4 24 024 0024 000024" +-/ +#guard_msgs in +#eval datetime₄.format "k kk kkk kkkk kkkkkk" + +/-- +info: "3 23 023 0023 00023" +-/ +#guard_msgs in +#eval datetime₄.format "H HH HHH HHHH HHHHH" + +/-- +info: "3 13 013 0013 00013" +-/ +#guard_msgs in +#eval datetime₄.format "m mm mmm mmmm mmmmm" + +/-- +info: "2 12 012 0012 00012" +-/ +#guard_msgs in +#eval datetime₄.format "s ss sss ssss sssss" + + +/-- +info: "9 79 679 4679 54679" +-/#guard_msgs in +#eval datetime₄.format "S SS SSS SSSS SSSSS" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval datetime₄.format "S SS SSS SSSS SSSSSSSSS" + +/-- +info: "4 24 324 2324 083592324" +-/ +#guard_msgs in +#eval datetime₄.format "A AA AAA AAAA AAAAAAAAA" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval datetime₄.format "n nn nnn nnnn nnnnnnnnn" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval datetime₄.format "N NN NNN NNNN NNNNNNNNN" + +/-- +info: "1 11 011 0011 0011" +-/ +#guard_msgs in +#eval time₄.format "h hh hhh hhhh hhhh" + +/-- +info: "1 11 011 0011 000011" +-/ +#guard_msgs in +#eval time₄.format "K KK KKK KKKK KKKKKK" + +/-- +info: "4 24 024 0024 000024" +-/ +#guard_msgs in +#eval time₄.format "k kk kkk kkkk kkkkkk" + +/-- +info: "3 23 023 0023 00023" +-/ +#guard_msgs in +#eval time₄.format "H HH HHH HHHH HHHHH" + +/-- +info: "3 13 013 0013 00013" +-/ +#guard_msgs in +#eval time₄.format "m mm mmm mmmm mmmmm" + +/-- +info: "2 12 012 0012 00012" +-/ +#guard_msgs in +#eval time₄.format "s ss sss ssss sssss" + + +/-- +info: "9 79 679 4679 54679" +-/#guard_msgs in +#eval time₄.format "S SS SSS SSSS SSSSS" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval time₄.format "S SS SSS SSSS SSSSSSSSS" + +/-- +info: "4 24 324 2324 083592324" +-/ +#guard_msgs in +#eval time₄.format "A AA AAA AAAA AAAAAAAAA" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval time₄.format "n nn nnn nnnn nnnnnnnnn" + +/-- +info: "9 79 679 4679 324354679" +-/ +#guard_msgs in +#eval time₄.format "N NN NNN NNNN NNNNNNNNN" + +/-- +info: "CE CE CE Common Era C" +-/ +#guard_msgs in +#eval date₄.format "G GG GGG GGGG GGGGG" + +/-- +info: "02 2002 000002002" +-/ +#guard_msgs in +#eval date₄.format "yy yyyy yyyyyyyyy" + +/-- +info: "5 95 195" +-/ +#guard_msgs in +#eval date₄.format "D DD DDD" + +/-- +info: "7 07 Jul J" +-/ +#guard_msgs in +#eval date₄.format "M MM MMM MMMMM" + +/-- +info: "4 14 014 0014 00014" +-/ +#guard_msgs in +#eval date₄.format "d dd ddd dddd ddddd" + +/-- +info: "7 07 Jul July J" +-/ +#guard_msgs in +#eval date₄.format "M MM MMM MMMM MMMMM" + +/-- +info: "4 14 0014 0014" +-/#guard_msgs in +#eval date₄.format "d dd dddd dddd" + +/-- +info: "3 03 3rd quarter 3" +-/ +#guard_msgs in +#eval date₄.format "Q QQ QQQQ QQQQQ" + +/-- +info: "8 28 028 0028" +-/ +#guard_msgs in +#eval date₄.format "w ww www wwww" + +/-- +info: "3 03 003 0003" +-/ +#guard_msgs in +#eval date₄.format "W WW WWW WWWW" + +/-- +info: "Sun Sun Sun Sunday S" +-/ +#guard_msgs in +#eval date₄.format "E EE EEE EEEE EEEEE" + +/-- +info: "7 07 Sun Sunday S" +-/ +#guard_msgs in +#eval date₄.format "e ee eee eeee eeeee" + +/-- +info: "3 03 003 0003" +-/ +#guard_msgs in +#eval date₄.format "F FF FFF FFFF" From 0dd6c267c9385faec9e01a7c82cedca8e589e465 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 23 Oct 2024 00:51:53 -0300 Subject: [PATCH 076/118] feat: wip refactor of zoneddatetime --- src/Std/Time.lean | 26 ++- src/Std/Time/Date/Basic.lean | 1 + src/Std/Time/Date/PlainDate.lean | 111 +++++----- src/Std/Time/Date/Unit/Month.lean | 64 ------ src/Std/Time/Date/Unit/Week.lean | 14 ++ src/Std/Time/Date/Unit/Year.lean | 21 ++ src/Std/Time/Date/ValidDate.lean | 88 ++++++++ src/Std/Time/DateTime.lean | 42 ---- src/Std/Time/DateTime/PlainDateTime.lean | 22 +- src/Std/Time/DateTime/Timestamp.lean | 14 -- src/Std/Time/Format.lean | 14 +- src/Std/Time/Format/Basic.lean | 43 +++- src/Std/Time/Internal/Bounded.lean | 36 +++- src/Std/Time/Notation.lean | 3 +- src/Std/Time/Notation/Spec.lean | 1 + src/Std/Time/Zoned.lean | 79 ++----- src/Std/Time/Zoned/Database.lean | 21 +- src/Std/Time/Zoned/Database/Basic.lean | 8 +- src/Std/Time/Zoned/Database/TZdb.lean | 16 +- src/Std/Time/Zoned/Database/Windows.lean | 9 +- src/Std/Time/Zoned/DateTime.lean | 5 +- src/Std/Time/Zoned/ZoneRules.lean | 30 ++- src/Std/Time/Zoned/ZonedDateTime.lean | 254 ++++++++++++----------- tests/lean/run/timeFormats.lean | 60 +++++- tests/lean/run/timeOperations.lean | 34 ++- tests/lean/run/timeParse.lean | 44 ++-- tests/lean/run/timeSet.lean | 37 ++-- 27 files changed, 614 insertions(+), 483 deletions(-) create mode 100644 src/Std/Time/Date/ValidDate.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index cbd4b7920990..077ec66c069b 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -159,41 +159,51 @@ The supported formats include: - `F`: Represents the aligned week of the month (e.g., "3"). - `a`: Represents the AM or PM designation of the day. - `a`, `aa`, `aaa`: Displays AM or PM in a concise format (e.g., "PM"). - - `aaaa`: Displays the full AM/PM designation (e.g., "PM"). + - `aaaa`: Displays the full AM/PM designation (e.g., "Post Meridium"). - `h`: Represents the hour of the AM/PM clock (1-12) (e.g., "12"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `K`: Represents the hour of the AM/PM clock (0-11) (e.g., "0"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `k`: Represents the hour of the day in a 1-24 format (e.g., "24"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `H`: Represents the hour of the day in a 0-23 format (e.g., "0"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `m`: Represents the minute of the hour (e.g., "30"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `s`: Represents the second of the minute (e.g., "55"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `S`: Represents a fraction of a second, typically displayed as a decimal number (e.g., "978" for milliseconds). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `A`: Represents the millisecond of the day (e.g., "1234"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `n`: Represents the nanosecond of the second (e.g., "987654321"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `N`: Represents the nanosecond of the day (e.g., "1234000000"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `VV`: Represents the time zone ID, which could be a city-based zone (e.g., "America/Los_Angeles"), a UTC marker (`"Z"`), or a specific offset (e.g., "-08:30"). + - One or more repetitions of the character indicates the truncation of the value to the specified number of characters. - `z`: Represents the time zone name. - `z`, `zz`, `zzz`: Shows an abbreviated time zone name (e.g., "PST" for Pacific Standard Time). - `zzzz`: Displays the full time zone name (e.g., "Pacific Standard Time"). - `O`: Represents the localized zone offset in the format "GMT" followed by the time difference from UTC. - `O`: Displays the GMT offset in a simple format (e.g., "GMT+8"). - - `O`: Displays the full GMT offset, including hours and minutes (e.g., "GMT+08:00"). -- `X`: Represents the zone offset, using 'Z' for UTC and specifying the offset. + - `OOOO`: Displays the full GMT offset, including hours and minutes (e.g., "GMT+08:00"). +- `X`: Represents the zone offset. It uses 'Z' for UTC and can represent any offset (positive or negative). - `X`: Displays the hour offset (e.g., "-08"). - `XX`: Displays the hour and minute offset without a colon (e.g., "-0830"). - `XXX`: Displays the hour and minute offset with a colon (e.g., "-08:30"). - `XXXX`: Displays the hour, minute, and second offset without a colon (e.g., "-083045"). - `XXXXX`: Displays the hour, minute, and second offset with a colon (e.g., "-08:30:45"). - - It also uses `Z` to represent UTC without any offset. -- `x`: Represents the zone offset without using 'Z' for zero offsets. +- `x`: Represents the zone offset. Similar to X, but does not display 'Z' for UTC and focuses only on positive offsets. - `x`: Displays the hour offset (e.g., "+08"). - `xx`: Displays the hour and minute offset without a colon (e.g., "+0830"). - `xxx`: Displays the hour and minute offset with a colon (e.g., "+08:30"). - `xxxx`: Displays the hour, minute, and second offset without a colon (e.g., "+083045"). - `xxxxx`: Displays the hour, minute, and second offset with a colon (e.g., "+08:30:45"). -- `Z`: Represents the zone offset, with 'Z' for UTC and an optional time offset. +- `Z`: Always includes an hour and minute offset and may use 'Z' for UTC, providing clear differentiation between UTC and other time zones. - `Z`: Displays the hour and minute offset without a colon (e.g., "+0800"). - - `ZZ`: Displays "GMT" followed by the time offset (e.g., "GMT+08:00" or Z). - - `ZZZ`: Displays the full hour, minute, and second offset with a colon (e.g., "+08:30:45" or Z). + - `ZZ`: Displays "GMT" followed by the time offset (e.g., "GMT+08:00" or "Z"). + - `ZZZ`: Displays the full hour, minute, and second offset with a colon (e.g., "+08:30:45" or "Z"). # Macros diff --git a/src/Std/Time/Date/Basic.lean b/src/Std/Time/Date/Basic.lean index e50fbc17d028..e4b124774810 100644 --- a/src/Std/Time/Date/Basic.lean +++ b/src/Std/Time/Date/Basic.lean @@ -5,3 +5,4 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Date.Unit.Basic +import Std.Time.Date.ValidDate diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 3baeca1c5ce4..dadeddd8c9ee 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -11,6 +11,7 @@ import Std.Internal.Rat namespace Std namespace Time open Std.Internal +open Std.Time open Internal open Lean @@ -65,7 +66,7 @@ Creates a `PlainDate` from a year and a day ordinal within that year. -/ @[inline] def ofYearOrdinal (year : Year.Offset) (ordinal : Day.Ordinal.OfYear year.isLeap) : PlainDate := - let ⟨⟨month, day⟩, proof⟩ := Month.Ordinal.ofOrdinal ordinal + let ⟨⟨month, day⟩, proof⟩ := ValidDate.ofOrdinal ordinal ⟨year, month, day, proof⟩ /-- @@ -84,52 +85,12 @@ def ofDaysSinceUNIXEpoch (day : Day.Offset) : PlainDate := let y := y + (if m <= 2 then 1 else 0) .clip y (.clip m (by decide)) (.clip d (by decide)) -/-- -Calculates the `Weekday` of a given `PlainDate` using Zeller's Congruence for the Gregorian calendar. --/ -def weekday (date : PlainDate) : Weekday := - let q := date.day - let m := date.month - let y := date.year.toInt - - let y := if m.val < 2 then y - 1 else y - - let m : Bounded.LE 1 13 := if h : m.val ≤ 1 - then (m.truncateTop h |>.add 12 |>.expandBottom (by decide)) - else m.expandTop (by decide) - - let k := Bounded.LE.byEmod y 100 (by decide) - let j : Bounded.LE (-10) 9 := (Bounded.LE.byMod y 1000 (by decide)).ediv 100 (by decide) - let part : Bounded.LE 6 190 := q.addBounds (((m.add 1).mul_pos 13 (by decide)).ediv 5 (by decide)) |>.addBounds k |>.addBounds (k.ediv 4 (by decide)) - let h : Bounded.LE (-15) 212 := part.addBounds ((j.ediv 4 (by decide)).addBounds (j.mul_pos 2 (by decide)).neg) - let d := (h.add 5).emod 7 (by decide) - - .ofOrdinal (d.add 1) - -/-- -Determines the week of the year for the given `PlainDate`. --/ -def weekOfYear (date : PlainDate) : Week.Ordinal := - let res := Month.Ordinal.toOrdinal ⟨⟨date.month, date.day⟩, date.valid⟩ |>.ediv 7 (by decide) |>.add 1 - match date.year.isLeap, res with - | true, res => res - | false, res => res - /-- Returns the unaligned week of the month for a `PlainDate` (day divided by 7, plus 1). -/ def weekOfMonth (date : PlainDate) : Bounded.LE 1 5 := date.day.ediv 7 (by decide) |>.add 1 -/-- -Determines the week of the month for the given `PlainDate`. The week of the month is calculated based -on the day of the month and the weekday. Each week starts on Sunday because the entire library is -based on the Gregorian Calendar. --/ -def alignedWeekOfMonth (date : PlainDate) : Week.Ordinal.OfMonth := - let weekday := date.weekday.toOrdinal - date.day.addBounds (weekday.sub 1) |>.ediv 7 (by decide) |>.add 1 - /-- Determines the quarter of the year for the given `PlainDate`. -/ @@ -140,7 +101,7 @@ def quarter (date : PlainDate) : Bounded.LE 1 4 := Transforms a tuple of a `PlainDate` into a `Day.Ordinal.OfYear`. -/ def toOrdinal (date : PlainDate) : Day.Ordinal.OfYear date.year.isLeap := - Month.Ordinal.toOrdinal ⟨(date.month, date.day), date.valid⟩ + ValidDate.toOrdinal ⟨(date.month, date.day), date.valid⟩ /-- Determines the era of the given `PlainDate` based on its year. @@ -231,6 +192,14 @@ Creates a `PlainDate` by rolling over the extra days to the next month. def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := clip year month 1 |>.addDays (day.toOffset - 1) +/-- +Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day remain unchanged, +and any invalid days for the new year will be handled according to the `clip` behavior. +-/ +@[inline] +def withYearClip (dt : PlainDate) (year : Year.Offset) : PlainDate := + clip year dt.month dt.day + /-- Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day are rolled over to the next valid month and day if necessary. @@ -314,12 +283,60 @@ def withMonthRollOver (dt : PlainDate) (month : Month.Ordinal) : PlainDate := rollOver dt.year month dt.day /-- -Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day remain unchanged, -and any invalid days for the new year will be handled according to the `clip` behavior. +Calculates the `Weekday` of a given `PlainDate` using Zeller's Congruence for the Gregorian calendar. -/ -@[inline] -def withYearClip (dt : PlainDate) (year : Year.Offset) : PlainDate := - clip year dt.month dt.day +def weekday (date : PlainDate) : Weekday := + let days := date.toDaysSinceUNIXEpoch.val + let res := if days ≥ -4 then (days + 4) % 7 else (days + 5) % 7 + 6 + .ofOrdinal (Bounded.LE.ofNatWrapping res (by decide)) + +/-- +Determines the week of the month for the given `PlainDate`. The week of the month is calculated based +on the day of the month and the weekday. Each week starts on Sunday because the entire library is +based on the Gregorian Calendar. +-/ +def alignedWeekOfMonth (date : PlainDate) : Week.Ordinal.OfMonth := + let weekday := date.withDayClip 1 |>.weekday |>.toOrdinal |>.sub 1 + let days := date.day |>.sub 1 |>.addBounds weekday + days |>.ediv 7 (by decide) |>.add 1 + +/-- +Sets the date to the specified `desiredWeekday`. +-/ +def setWeekday (date : PlainDate) (desiredWeekday : Weekday) : PlainDate := + let weekday := date |>.weekday |>.toOrdinal + let offset := desiredWeekday.toOrdinal |>.subBounds weekday + + let offset : Bounded.LE 0 6 := + if h : offset.val < 0 then + offset.truncateTop (Int.le_sub_one_of_lt h) |>.addBounds (.exact 7) + |>.expandBottom (by decide) + else + offset.truncateBottom (Int.not_lt.mp h) + |>.expandTop (by decide) + + date.addDays (Day.Offset.ofInt offset.toInt) + +/-- +Calculates the starting Monday of the ISO week-based year for a given year. +-/ +def weekOfYear (date : PlainDate) : Week.Ordinal := + let y := date.year + + let w := Bounded.LE.exact 10 + |>.addBounds date.toOrdinal + |>.subBounds date.weekday.toOrdinal + |>.ediv 7 (by decide) + + if h : w.val < 1 then + (y-1).weeks |>.expandBottom (by decide) + else if h₁ : w.val > y.weeks.val then + .ofNat' 1 (by decide) + else + let h := Int.not_lt.mp h + let h₁ := Int.not_lt.mp h₁ + let w := w.truncateBottom h |>.truncateTop (Int.le_trans h₁ y.weeks.property.right) + w instance : HAdd PlainDate Day.Offset PlainDate where hAdd := addDays diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 2fb6b7598035..c292bd092fb6 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -282,70 +282,6 @@ if the year is a leap year. abbrev Valid (leap : Bool) (month : Month.Ordinal) (day : Day.Ordinal) : Prop := day.val ≤ (days leap month).val -/-- -Represents a valid date for a given year, considering whether it is a leap year. Example: `(2, 29)` -is valid only if `leap` is `true`. --/ -def ValidDate (leap : Bool) := { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } - -instance : Inhabited (ValidDate l) where - default := ⟨⟨1, 1⟩, (by cases l <;> decide)⟩ - -/-- -Transforms a tuple of a `Month` and a `Day` into a `Day.Ordinal.OfYear`. --/ -def toOrdinal (ordinal : ValidDate leap) : Day.Ordinal.OfYear leap := - let days := cumulativeDays leap ordinal.val.fst - let proof := cumulativeDays_le leap ordinal.val.fst - let bounded := Bounded.LE.mk days.toInt proof |>.addBounds ordinal.val.snd - match leap, bounded with - | true, bounded => bounded - | false, bounded => bounded - -/-- -Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. --/ -def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : ValidDate leap := - let rec go (idx : Month.Ordinal) (acc : Int) (h : ordinal.val > acc) (p : acc = (cumulativeDays leap idx).val) : ValidDate leap := - let monthDays := days leap idx - if h₁ : ordinal.val ≤ acc + monthDays.val then - let bounded := Bounded.LE.mk ordinal.val (And.intro h h₁) |>.sub acc - let bounded : Bounded.LE 1 monthDays.val := bounded.cast (by omega) (by omega) - let days₁ : Day.Ordinal := ⟨bounded.val, And.intro bounded.property.left (Int.le_trans bounded.property.right monthDays.property.right)⟩ - ⟨⟨idx, days₁⟩, Int.le_trans bounded.property.right (by simp)⟩ - else - let h₂ := Int.not_le.mp h₁ - if h₃ : idx.val > 11 then by - have h₅ := ordinal.property.right - let eq := Int.eq_iff_le_and_ge.mpr (And.intro idx.property.right h₃) - simp [monthDays, days, eq] at h₂ - simp [cumulativeDays, eq] at p - simp [p] at h₂ - cases leap - all_goals (simp at h₂; simp_all) - · have h₂ : 365 < ordinal.val := h₂ - omega - · have h₂ : 366 < ordinal.val := h₂ - omega - else by - let h₃ := Int.not_le.mp h₃ - let idx₂ := idx.truncateTop (Int.le_sub_one_of_lt h₃) |>.addTop 1 (by decide) - refine go idx₂ (acc + monthDays.val) h₂ ?_ - simp [monthDays, p] - rw [difference_eq (Int.le_of_lt_add_one h₃)] - termination_by 12 - idx.val.toNat - decreasing_by - simp_wf - simp [Bounded.LE.addTop] - let gt0 : idx.val ≥ 0 := Int.le_trans (by decide) idx.property.left - refine Nat.sub_lt_sub_left (Int.toNat_lt gt0 |>.mpr h₃) ?_ - let toNat_lt_lt {n z : Int} (h : 0 ≤ z) (h₁ : 0 ≤ n) : z.toNat < n.toNat ↔ z < n := by - rw [← Int.not_le, ← Nat.not_le, ← Int.ofNat_le, Int.toNat_of_nonneg h, Int.toNat_of_nonneg h₁] - rw [toNat_lt_lt (by omega) (by omega)] - omega - - go 1 0 (Int.le_trans (by decide) ordinal.property.left) (by cases leap <;> decide) - /-- Clips the day to be within the valid range. -/ diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 9d7e49597987..af6860b4c9ab 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -73,6 +73,20 @@ def toOffset (ordinal : Ordinal) : Offset := end Ordinal namespace Offset +/-- +Creates an `Offset` from a natural number. +-/ +@[inline] +def ofNat (data : Nat) : Week.Offset := + UnitVal.mk data + +/-- +Creates an `Offset` from an integer. +-/ +@[inline] +def ofInt (data : Int) : Week.Offset := + UnitVal.mk data + /-- Convert `Week.Offset` into `Millisecond.Offset`. -/ diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index b42c9ef8d4ee..b512e339ad58 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -88,6 +88,27 @@ def era (year : Offset) : Era := then .ce else .bce +/-- +Calculates the number of days in the specified `year`. +-/ +def days (year : Offset) : Bounded.LE 365 366 := + if year.isLeap + then .ofNatWrapping 366 (by decide) + else .ofNatWrapping 355 (by decide) + +/-- +Calculates the number of weeks in the specified `year`. +-/ +def weeks (year : Offset) : Bounded.LE 52 53 := + let p (year : Offset) := Bounded.LE.byEmod (year.toInt + year.toInt/4 - year.toInt/100 + year.toInt/400) 7 (by decide) + + let add : Bounded.LE 0 1 := + if (p year).val = 4 ∨ (p (year - 1)).val = 3 + then Bounded.LE.ofNat 1 (by decide) + else Bounded.LE.ofNat 0 (by decide) + + Bounded.LE.exact 52 |>.addBounds add + /-- Checks if the given date is valid for the specified year, month, and day. -/ diff --git a/src/Std/Time/Date/ValidDate.lean b/src/Std/Time/Date/ValidDate.lean new file mode 100644 index 000000000000..450b1db53378 --- /dev/null +++ b/src/Std/Time/Date/ValidDate.lean @@ -0,0 +1,88 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sofia Rodrigues +-/ +prelude +import Std.Internal.Rat +import Std.Time.Date.Unit.Day +import Std.Time.Date.Unit.Month + +namespace Std +namespace Time +open Std.Internal +open Internal +open Month.Ordinal + +set_option linter.all true + +/-- +Represents a valid date for a given year, considering whether it is a leap year. Example: `(2, 29)` +is valid only if `leap` is `true`. +-/ +def ValidDate (leap : Bool) := { val : Month.Ordinal × Day.Ordinal // Valid leap (Prod.fst val) (Prod.snd val) } + +instance : Inhabited (ValidDate l) where + default := ⟨⟨1, 1⟩, (by cases l <;> decide)⟩ + +namespace ValidDate + +/-- +Transforms a tuple of a `Month` and a `Day` into a `Day.Ordinal.OfYear`. +-/ +def toOrdinal (ordinal : ValidDate leap) : Day.Ordinal.OfYear leap := + let days := cumulativeDays leap ordinal.val.fst + let proof := cumulativeDays_le leap ordinal.val.fst + let bounded := Bounded.LE.mk days.toInt proof |>.addBounds ordinal.val.snd + match leap, bounded with + | true, bounded => bounded + | false, bounded => bounded + +/-- +Transforms a `Day.Ordinal.OfYear` into a tuple of a `Month` and a `Day`. +-/ +def ofOrdinal (ordinal : Day.Ordinal.OfYear leap) : ValidDate leap := + let rec go (idx : Month.Ordinal) (acc : Int) (h : ordinal.val > acc) (p : acc = (cumulativeDays leap idx).val) : ValidDate leap := + let monthDays := days leap idx + if h₁ : ordinal.val ≤ acc + monthDays.val then + let bounded := Bounded.LE.mk ordinal.val (And.intro h h₁) |>.sub acc + let bounded : Bounded.LE 1 monthDays.val := bounded.cast (by omega) (by omega) + let days₁ : Day.Ordinal := ⟨bounded.val, And.intro bounded.property.left (Int.le_trans bounded.property.right monthDays.property.right)⟩ + ⟨⟨idx, days₁⟩, Int.le_trans bounded.property.right (by simp)⟩ + else by + let h₂ := Int.not_le.mp h₁ + + have h₃ : idx.val < 12 := Int.not_le.mp <| λh₃ => by + have h₅ := ordinal.property.right + let eq := Int.eq_iff_le_and_ge.mpr (And.intro idx.property.right h₃) + simp [monthDays, days, eq] at h₂ + simp [cumulativeDays, eq] at p + simp [p] at h₂ + cases leap + all_goals (simp at h₂; simp_all) + · have h₂ : 365 < ordinal.val := h₂ + omega + · have h₂ : 366 < ordinal.val := h₂ + omega + + let idx₂ := idx.truncateTop (Int.le_sub_one_of_lt h₃) |>.addTop 1 (by decide) + refine go idx₂ (acc + monthDays.val) h₂ ?_ + simp [monthDays, p] + rw [difference_eq (Int.le_of_lt_add_one h₃)] + + termination_by 12 - idx.val.toNat + decreasing_by + simp_wf + simp [Bounded.LE.addTop] + let gt0 : idx.val ≥ 0 := Int.le_trans (by decide) idx.property.left + refine Nat.sub_lt_sub_left (Int.toNat_lt gt0 |>.mpr h₃) ?_ + let toNat_lt_lt {n z : Int} (h : 0 ≤ z) (h₁ : 0 ≤ n) : z.toNat < n.toNat ↔ z < n := by + rw [← Int.not_le, ← Nat.not_le, ← Int.ofNat_le, Int.toNat_of_nonneg h, Int.toNat_of_nonneg h₁] + rw [toNat_lt_lt (by omega) (by omega)] + omega + + go 1 0 (Int.le_trans (by decide) ordinal.property.left) (by cases leap <;> decide) + +end ValidDate +end Time +end Std diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index 2cb13344dc3d..dc08064d5100 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -64,24 +64,6 @@ Converts a `PlainDate` to a `Timestamp` def toTimestampAssumingUTC (pdt : PlainDate) : Timestamp := Timestamp.ofPlainDateAssumingUTC pdt -/-- -Calculates the duration between a given `PlainDate` and a specified date. - -## Example - -```lean -def example : Duration := - let startDate := date("2023-1-1") - let endDate := date("2023-3-15") - endDate.since startDate -``` --/ -@[inline] -def since [ToTimestamp α] (date : PlainDate) (since : α) : Duration := - let date := date.toTimestampAssumingUTC - let since := ToTimestamp.toTimestamp since - Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch - instance : HSub PlainDate PlainDate Duration where hSub x y := x.toTimestampAssumingUTC - y.toTimestampAssumingUTC @@ -116,30 +98,6 @@ Converts a `PlainDateTime` to a `PlainTime` def toPlainTime (pdt : PlainDateTime) : PlainTime := pdt.time -instance : ToTimestamp PlainDateTime where - toTimestamp := Timestamp.ofPlainDateTimeAssumingUTC - -instance : ToTimestamp PlainDate where - toTimestamp := Timestamp.ofPlainDateAssumingUTC - -/-- -Calculates the duration between a given `PlainDateTime` and a specified date. - -## Example - -```lean -example : Duration := - let startDate := datetime("2023-1-1:05:10:20") - let endDate := datetime("2023-3-15:05:10:20") - endDate.since startDate -``` --/ -@[inline] -def since [ToTimestamp α] (date : PlainDateTime) (since : α) : Duration := - let date := date.toTimestampAssumingUTC - let since := ToTimestamp.toTimestamp since - Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch - instance : HSub PlainDateTime PlainDateTime Duration where hSub x y := x.toTimestampAssumingUTC - y.toTimestampAssumingUTC diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 7d6af0b9cf01..3b7421a80452 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -457,7 +457,7 @@ Transforms a tuple of a `PlainDateTime` into a `Day.Ordinal.OfYear`. -/ @[inline] def toOrdinal (date : PlainDateTime) : Day.Ordinal.OfYear date.year.isLeap := - Month.Ordinal.toOrdinal ⟨(date.month, date.day), date.date.valid⟩ + ValidDate.toOrdinal ⟨(date.month, date.day), date.date.valid⟩ /-- Determines the quarter of the year for the given `PlainDateTime`. @@ -506,5 +506,25 @@ instance : HAdd PlainDateTime Duration PlainDateTime where hAdd x y := addNanoseconds x y.toNanoseconds end PlainDateTime +namespace PlainDate + +/-- +Combines a `PlainDate` and `PlainTime` into a `PlainDateTime`. +-/ +@[inline] +def atTime : PlainDate → PlainTime → PlainDateTime := + PlainDateTime.mk + +end PlainDate +namespace PlainTime + +/-- +Combines a `PlainTime` and `PlainDate` into a `PlainDateTime`. +-/ +@[inline] +def atDate (time: PlainTime) (date: PlainDate) : PlainDateTime := + PlainDateTime.mk date time + +end PlainTime end Time end Std diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 7095f29877a9..18954fbe4227 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -42,20 +42,6 @@ instance : ToString Timestamp where instance : Repr Timestamp where reprPrec s := reprPrec (toString s) -/-- -Type class to transform to `Timestamp`. It's used internally to generate timestamps out of -some absolute date types. --/ -class ToTimestamp (α : Type) where - - /-- - Transforms into a `Timestamp`. - -/ - toTimestamp : α → Timestamp - -instance : ToTimestamp Timestamp where - toTimestamp := id - namespace Timestamp /-- diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index e499e947a4da..d40ed3582525 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -169,6 +169,7 @@ def format (date : PlainDate) (format : String) : String := let res := res.formatGeneric fun | .G _ => some date.era | .y _ => some date.year + | .u _ => some date.year | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) | .Qorq _ => some date.quarter | .w _ => some date.weekOfYear @@ -326,7 +327,7 @@ def format (data: ZonedDateTime) (format : String) : String := let format : Except String (GenericFormat .any) := GenericFormat.spec format match format with | .error err => s!"error: {err}" - | .ok res => res.format data.snd + | .ok res => res.format data.toDateTime /-- Parses a `String` in the `ISO8601` format and returns a `ZonedDateTime`. @@ -338,7 +339,7 @@ def fromISO8601String (input : String) : Except String ZonedDateTime := Formats a `ZonedDateTime` value into an ISO8601 string. -/ def toISO8601String (date : ZonedDateTime) : String := - Formats.iso8601.format date.snd + Formats.iso8601.format date.toDateTime /-- Parses a `String` in the `RFC822` format and returns a `ZonedDateTime`. @@ -350,7 +351,7 @@ def fromRFC822String (input : String) : Except String ZonedDateTime := Formats a `ZonedDateTime` value into an RFC822 format string. -/ def toRFC822String (date : ZonedDateTime) : String := - Formats.rfc822.format date.snd + Formats.rfc822.format date.toDateTime /-- Parses a `String` in the `RFC850` format and returns a `ZonedDateTime`. @@ -362,7 +363,7 @@ def fromRFC850String (input : String) : Except String ZonedDateTime := Formats a `ZonedDateTime` value into an RFC850 format string. -/ def toRFC850String (date : ZonedDateTime) : String := - Formats.rfc850.format date.snd + Formats.rfc850.format date.toDateTime /-- Parses a `String` in the `DateTimeWithZone` format and returns a `DateTime` object in the GMT time zone. @@ -374,7 +375,7 @@ def fromDateTimeWithZoneString (input : String) : Except String ZonedDateTime := Formats a `DateTime` value into a simple date time with timezone string. -/ def toDateTimeWithZoneString (pdt : ZonedDateTime) : String := - Formats.dateTimeWithZone.format pdt.snd + Formats.dateTimeWithZone.format pdt.toDateTime /-- Parses a `String` in the `DateTimeWithZone` format and returns a `DateTime` object in the GMT time zone. @@ -387,7 +388,7 @@ def fromLeanDateTimeWithZoneString (input : String) : Except String ZonedDateTim Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notationg. -/ def toLeanDateTimeWithZoneString (zdt : ZonedDateTime) : String := - Formats.leanDateTimeWithZone.formatBuilder zdt.year zdt.month zdt.day zdt.hour zdt.minute zdt.snd.date.get.time.second zdt.nanosecond zdt.fst.offset + Formats.leanDateTimeWithZone.formatBuilder zdt.year zdt.month zdt.day zdt.hour zdt.minute zdt.date.get.time.second zdt.nanosecond zdt.offset /-- Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. @@ -418,6 +419,7 @@ def format (date : PlainDateTime) (format : String) : String := let res := res.formatGeneric fun | .G _ => some date.era | .y _ => some date.year + | .u _ => some date.year | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) | .Qorq _ => some date.quarter | .w _ => some date.weekOfYear diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index f11a88f88388..98bd29a2ebc1 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -291,10 +291,15 @@ inductive Modifier | G (presentation : Text) /-- - `y`: Year of era (e.g., 2004, 04). + `y`: Year of era (e.g., 2004, 04, 0002, 2). -/ | y (presentation : Year) + /-- + `u`: Year (e.g., 2004, 04, -0001, -1). + -/ + | u (presentation : Year) + /-- `D`: Day of year (e.g., 189). -/ @@ -472,6 +477,7 @@ private def parseZoneName (constructor : ZoneName → Modifier) (p : String) : P private def parseModifier : Parser Modifier := (parseText Modifier.G =<< many1Chars (pchar 'G')) <|> parseYear Modifier.y =<< many1Chars (pchar 'y') + <|> parseYear Modifier.u =<< many1Chars (pchar 'u') <|> parseNumber Modifier.D =<< many1Chars (pchar 'D') <|> parseNumberText Modifier.MorL =<< many1Chars (pchar 'M') <|> parseNumberText Modifier.MorL =<< many1Chars (pchar 'L') @@ -742,6 +748,7 @@ private def toIsoString (offset : Offset) (withMinutes : Bool) (withSeconds : Bo private def TypeFormat : Modifier → Type | .G _ => Year.Era | .y _ => Year.Offset + | .u _ => Year.Offset | .D _ => Sigma Day.Ordinal.OfYear | .MorL _ => Month.Ordinal | .d _ => Day.Ordinal @@ -777,6 +784,13 @@ private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : Strin | .full => formatEraLong data | .narrow => formatEraNarrow data | .y format => + let info := data.toInt + let info := if info ≤ 0 then -info + 1 else info + match format with + | .twoDigit => truncate 2 (info % 100) + | .fourDigit => truncate 4 info + | .extended n => truncate n info + | .u format => match format with | .twoDigit => truncate 2 (data.toInt % 100) | .fourDigit => truncate 4 data.toInt @@ -883,6 +897,7 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := match modifier with | .G _ => date.era | .y _ => date.year + | .u _ => date.year | .D _ => Sigma.mk _ date.toOrdinal | .MorL _ => date.month | .d _ => date.day @@ -1098,6 +1113,11 @@ private def parseWith : (mod : Modifier) → Parser (TypeFormat mod) | .full => parseEraLong | .narrow => parseEraNarrow | .y format => + match format with + | .twoDigit => (2000 + ·) <$> (Int.ofNat <$> parseNum 2) + | .fourDigit => Int.ofNat <$> parseNum 4 + | .extended n => Int.ofNat <$> parseNum n + | .u format => match format with | .twoDigit => (2000 + ·) <$> (parseSigned <| parseNum 2) | .fourDigit => parseSigned <| parseNum 4 @@ -1206,6 +1226,7 @@ namespace GenericFormat private structure DateBuilder where G : Option Year.Era := none y : Option Year.Offset := none + u : Option Year.Offset := none D : Option (Sigma Day.Ordinal.OfYear) := none MorL : Option Month.Ordinal := none d : Option Day.Ordinal := none @@ -1239,8 +1260,8 @@ namespace DateBuilder private def insert (date : DateBuilder) (modifier : Modifier) (data : TypeFormat modifier) : DateBuilder := match modifier with | .G _ => { date with G := some data } - | .y .twoDigit => { date with y := some (Year.Offset.ofInt (data.toInt + 2000)) } | .y _ => { date with y := some data } + | .u _ => { date with u := some data } | .D _ => { date with D := some data } | .MorL _ => { date with MorL := some data } | .d _ => { date with d := some data } @@ -1269,6 +1290,10 @@ private def insert (date : DateBuilder) (modifier : Modifier) (data : TypeFormat | .x _ => { date with x := some data } | .Z _ => { date with Z := some data } +private def convertYearAndEra (year : Year.Offset) : Year.Era → Year.Offset + | .ce => year + | .bce => -(year + 1) + private def build (builder : DateBuilder) (aw : Awareness) : Option aw.type := let tz : TimeZone := { offset := builder.O <|> builder.X <|> builder.x <|> builder.Z |>.getD Offset.zero, @@ -1277,16 +1302,14 @@ private def build (builder : DateBuilder) (aw : Awareness) : Option aw.type := isDST := false, } - let year := builder.y |>.getD 0 let month := builder.MorL |>.getD 0 let day := builder.d |>.getD 0 + let era := (builder.G.getD .ce) - let era := builder.G |>.getD .ce - - let year := - match era with - | .ce => year - | .bce => -(year + 1) + let year + := builder.u + <|> ((convertYearAndEra · era) <$> builder.y) + |>.getD 0 let hour : Option (Bounded.LE 0 23) := if let some marker := builder.a then @@ -1320,7 +1343,7 @@ private def build (builder : DateBuilder) (aw : Awareness) : Option aw.type := match aw with | .only newTz => (DateTime.ofLocalDateTime · newTz) <$> datetime - | .any => (ZonedDateTime.ofLocalDateTime · tz) <$> datetime + | .any => (ZonedDateTime.ofPlainDateTime · (ZoneRules.ofTimeZone tz)) <$> datetime end DateBuilder diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index b1e9df28aaae..7b4bc514c0b2 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -114,6 +114,13 @@ Creates a new `Bounded` integer that the relation is less-equal. def mk (val : Int) (proof : lo ≤ val ∧ val ≤ hi) : Bounded.LE lo hi := ⟨val, proof⟩ +/-- +Creates a new `Bounded` integer that the relation is less-equal. +-/ +@[inline] +def exact (val : Nat) : Bounded.LE val val := + ⟨val, by simp⟩ + /-- Creates a new `Bounded` integer. -/ @@ -341,10 +348,8 @@ def sub (bounded : Bounded.LE n m) (num : Int) : Bounded.LE (n - num) (m - num) Adds two `Bounded` and adjust the boundaries. -/ @[inline] -def subBounds (bounded : Bounded.LE n m) (bounded₂ : Bounded.LE n m) : Bounded.LE (n - m) (m - n) := by - refine ⟨bounded.val - bounded₂.val, And.intro ?_ ?_⟩ - · exact Int.sub_le_sub bounded.property.left bounded₂.property.right - · exact Int.sub_le_sub bounded.property.right bounded₂.property.left +def subBounds (bounded : Bounded.LE n m) (bounded₂ : Bounded.LE i j) : Bounded.LE (n - j) (m - i) := + addBounds bounded bounded₂.neg /-- Adjust the bounds of a `Bounded` by applying the emod operation constraining the lower bound to 0 and @@ -439,6 +444,29 @@ def abs (bo : Bounded.LE (-i) i) : Bounded.LE 0 i := rw [Int.neg_neg] at r exact r +/-- +Returns the maximum between a number and the bounded. +-/ +def max (bounded : Bounded.LE n m) (val : Int) : Bounded.LE (Max.max n val) (Max.max m val) := by + let ⟨left, right⟩ := bounded.property + refine ⟨Max.max bounded.val val, And.intro ?_ ?_⟩ + + all_goals + simp [Int.max_def] + split <;> split + + next h => simp [h, Int.le_trans left h] + next h h₁ => exact Int.le_of_lt <| Int.not_le.mp h₁ + next h => simp [h, Int.le_trans left h] + next h h₁ => exact left + next h h₁ => simp [h, Int.le_trans left h] + next h h₁ => exact Int.le_of_lt <| Int.not_le.mp h₁ + next h h₁ => + let h₃ := Int.lt_of_lt_of_le (Int.not_le.mp h) right + let h₄ := Int.not_le.mpr h₃ h₁ + contradiction + next h h₁ => exact right + end LE end Bounded end Internal diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 25611ae09fbd..cbc87c516ea6 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -56,6 +56,7 @@ private def convertOffsetZ : OffsetZ → MacroM (TSyntax `term) private def convertModifier : Modifier → MacroM (TSyntax `term) | .G p => do `(Std.Time.Modifier.G $(← convertText p)) | .y p => do `(Std.Time.Modifier.y $(← convertYear p)) + | .u p => do `(Std.Time.Modifier.u $(← convertYear p)) | .D p => do `(Std.Time.Modifier.D $(← convertNumber p)) | .MorL p => match p with @@ -131,7 +132,7 @@ private def convertPlainDateTime (d : Std.Time.PlainDateTime) : MacroM (TSyntax `(Std.Time.PlainDateTime.mk $(← convertPlainDate d.date) $(← convertPlainTime d.time)) private def convertZonedDateTime (d : Std.Time.ZonedDateTime) : MacroM (TSyntax `term) := do - `(Std.Time.ZonedDateTime.mk $(← convertTimezone d.timezone) (DateTime.ofLocalDateTime $(← convertPlainDateTime d.snd.date.get) $(← convertTimezone d.timezone))) + `(Std.Time.ZonedDateTime.ofPlainDateTime $(← convertPlainDateTime d.toPlainDateTime) (Std.Time.TimeZone.ZoneRules.ofTimeZone $(← convertTimezone d.timezone))) /-- Defines a syntax for zoned datetime values. It expects a string representing a datetime with diff --git a/src/Std/Time/Notation/Spec.lean b/src/Std/Time/Notation/Spec.lean index d12252164028..cd77dd4948e3 100644 --- a/src/Std/Time/Notation/Spec.lean +++ b/src/Std/Time/Notation/Spec.lean @@ -54,6 +54,7 @@ private def convertOffsetZ : OffsetZ → MacroM (TSyntax `term) private def convertModifier : Modifier → MacroM (TSyntax `term) | .G p => do `(Std.Time.Modifier.G $(← convertText p)) | .y p => do `(Std.Time.Modifier.y $(← convertYear p)) + | .u p => do `(Std.Time.Modifier.u $(← convertYear p)) | .D p => do `(Std.Time.Modifier.D $(← convertNumber p)) | .MorL p => match p with diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 2c5becfa172b..c3ab8cc17ce2 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -30,8 +30,10 @@ Get the current time. @[inline] def now : IO PlainDateTime := do let tm ← Timestamp.now - let tz ← Database.defaultGetCurrentTimeZone - return PlainDateTime.ofTimestamp tm tz + let rules ← Database.defaultGetLocalZoneRulesAt + let transition ← rules.findTransitionForTimestamp tm |>.elim (throw <| IO.userError "cannot find timezone") pure + + return PlainDateTime.ofTimestamp tm transition.localTimeType.getTimeZone end PlainDateTime @@ -51,13 +53,6 @@ Converts a `DateTime` to a `PlainDate` def toPlainDate (dt : DateTime tz) : PlainDate := Timestamp.toPlainDateAssumingUTC dt.toUTCTimestamp -/-- -Converts a `PlainTime` and a `TimeZone` to a `DateTime` --/ -@[inline] -def ofPlainTime (pt : PlainTime) (tz : TimeZone) : DateTime tz := - DateTime.ofTimestamp (Timestamp.ofPlainTime pt) tz - /-- Converts a `DateTime` to a `PlainTime` -/ @@ -65,24 +60,6 @@ Converts a `DateTime` to a `PlainTime` def toPlainTime (dt : DateTime tz) : PlainTime := dt.date.get.time -/-- -Calculates the duration between a given `DateTime` and a specified date. - -## Example - -```lean -example : Duration := - let startDate := date("2023-1-1:05:10:20UTC") - let endDate := date("2023-3-15:05:10:20UTC") - endDate.since startDate -``` --/ -@[inline] -def since [ToTimestamp α] (date : DateTime tz) (since : α) : Duration := - let date := date.toUTCTimestamp - let since := ToTimestamp.toTimestamp since - Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch - end DateTime namespace ZonedDateTime @@ -91,61 +68,37 @@ Gets the current `ZonedDateTime`. -/ @[inline] def now : IO ZonedDateTime := do - let date ← Timestamp.now - let tz ← Database.defaultGetCurrentTimeZone - return ofTimestamp date tz + let tm ← Timestamp.now + let rules ← Database.defaultGetLocalZoneRulesAt + return ZonedDateTime.ofTimestamp tm rules /-- -Converts a `PlainDate` to a `ZonedDateTime` +Converts a `PlainDate` to a `ZonedDateTime`. -/ @[inline] -def ofPlainDate (pd : PlainDate) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainDateAssumingUTC pd) tz⟩ +def ofPlainDate (pd : PlainDate) (zr : TimeZone.ZoneRules) : ZonedDateTime := + ZonedDateTime.ofPlainDateTime (pd.atTime PlainTime.midnight) zr /-- Converts a `ZonedDateTime` to a `PlainDate` -/ @[inline] def toPlainDate (dt : ZonedDateTime) : PlainDate := - DateTime.toPlainDate dt.snd - -/-- -Converts a `PlainTime` to a `ZonedDateTime` --/ -@[inline] -def ofPlainTime (pt : PlainTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainTime pt) tz⟩ - -/-- -Converts a `PlainDateTime` to a `ZonedDateTime` assuming the Plain Date is Local. --/ -@[inline] -def ofLocalDateTime (pd : PlainDateTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofLocalDateTime pd tz⟩ + dt.toPlainDateTime.date /-- Converts a `ZonedDateTime` to a `PlainTime` -/ @[inline] def toPlainTime (dt : ZonedDateTime) : PlainTime := - DateTime.toPlainTime dt.snd + dt.toPlainDateTime.time /-- -Calculates the duration between a given `ZonedDateTime` and a specified date. - -## Example - -```lean -def example : Duration := - let startDate := date% 2023-1-1:05:10:20UTC - let endDate := date% 2023-3-15:05:10:20UTC - endDate.since startDate -``` +Creates a new `ZonedDateTime` out of a `PlainDateTime` and a time zone identifier. -/ @[inline] -def since [ToTimestamp α] (date : ZonedDateTime) (since : α) : Duration := - let date := date.toUTCTimestamp - let since := ToTimestamp.toTimestamp since - Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch +def of (pdt : PlainDateTime) (id : String) : IO ZonedDateTime := do + let zr ← Database.defaultGetZoneRulesAt id + return ZonedDateTime.ofPlainDateTime pdt zr end ZonedDateTime diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index 72af7352a13f..fd6940fb4fad 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -20,24 +20,15 @@ set_option linter.all true /-- Gets the time zone for a given location and timestamp, handling Windows and non-Windows platforms. -/ -def defaultGetTimeZoneAt : String → Timestamp → IO TimeZone := +def defaultGetZoneRulesAt : String → IO TimeZone.ZoneRules := if System.Platform.isWindows - then Database.getTimeZoneAt WindowsDb.default - else Database.getTimeZoneAt TZdb.default + then getZoneRulesAt WindowsDb.default + else getZoneRulesAt TZdb.default /-- Gets the local time zone for a specific timestamp, accounting for platform differences. -/ -def defaultGetLocalTimeZoneAt : Timestamp → IO TimeZone := +def defaultGetLocalZoneRulesAt : IO TimeZone.ZoneRules := if System.Platform.isWindows - then Database.getLocalTimeZoneAt WindowsDb.default - else Database.getLocalTimeZoneAt TZdb.default - -/-- -Retrieves the current local time zone based on the system platform and the current timestamp. --/ -def defaultGetCurrentTimeZone : IO TimeZone := do - let now <- Timestamp.now - if System.Platform.isWindows - then Database.getLocalTimeZoneAt WindowsDb.default now - else Database.getLocalTimeZoneAt TZdb.default now + then getLocalZoneRulesAt WindowsDb.default + else getLocalZoneRulesAt TZdb.default diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 9bf6208b7418..69a5d9d683f4 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -19,14 +19,14 @@ A timezone database from which we can read the `ZoneRules` of some area by it's class Database (α : Type) where /-- - Loads a `ZoneRules` by its id. + Retrieves the timezone information (`ZoneRules`) for a given area at a specific point in time. -/ - getTimeZoneAt : α → String → Timestamp → IO TimeZone + getZoneRulesAt : α → String → IO TimeZone.ZoneRules /-- - Loads a `ZoneRules` that is set by the local machine. + Retrieves the local timezone information (`ZoneRules`) at a given timestamp. -/ - getLocalTimeZoneAt : α → Timestamp → IO TimeZone + getLocalZoneRulesAt : α → IO TimeZone.ZoneRules namespace TimeZone diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 6c0d3a1a7d4b..3b2ab241db72 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -93,18 +93,6 @@ def readLeapsFromDisk (path : System.FilePath) : IO (Array LeapSecond) := do then return res else throw (IO.userError "cannot read the id of the path.") -/-- -Gets the `TimeZone` at the local timezone. --/ -def getLocalTimeZoneAt (db : TZdb) (tm : Timestamp) : IO TimeZone := do - (IO.ofExcept <| Transition.timezoneAt ·.transitions tm) =<< localRules db.localPath - -/-- -Gets the TimeZone at a timezone using a `Database`. --/ -def getTimeZoneAt (db : TZdb) (id : String) (tm : Timestamp) : IO TimeZone := do - (IO.ofExcept <| Transition.timezoneAt ·.transitions tm) =<< readRulesFromDisk db.zonesPath id - instance : Database TZdb where - getLocalTimeZoneAt := TZdb.getLocalTimeZoneAt - getTimeZoneAt := TZdb.getTimeZoneAt + getLocalZoneRulesAt db := localRules db.localPath + getZoneRulesAt db id := readRulesFromDisk db.zonesPath id diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index 230c215caf72..ecf0bed4111e 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -42,10 +42,5 @@ Returns a default `WindowsDb` instance. def default : WindowsDb := {} instance : Database WindowsDb where - getTimeZoneAt _ id tm := - Windows.getTimeZoneAt id (tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64) - - getLocalTimeZoneAt _ tm := do - let time := tm.toSecondsSinceUnixEpoch |>.toInt |>.toNat |>.toUInt64 - let id ← Windows.getLocalTimeZoneIdentifierAt time - Windows.getTimeZoneAt id time + getZoneRulesAt _ _ := pure <| ZoneRules.mk #[] #[] + getLocalZoneRulesAt _ := pure Inhabited.default diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 5d6e9c8e2411..952953e02e41 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -382,7 +382,7 @@ def inLeapYear (date : DateTime tz) : Bool := Determines the ordinal day of the year for the given `DateTime`. -/ def toOrdinal (date : DateTime tz) : Day.Ordinal.OfYear date.year.isLeap := - Month.Ordinal.toOrdinal ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ + ValidDate.toOrdinal ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ /-- Determines the week of the year for the given `DateTime`. @@ -413,9 +413,6 @@ Determines the quarter of the year for the given `DateTime`. def quarter (date : DateTime tz) : Bounded.LE 1 4 := date.date.get.quarter -instance : ToTimestamp (DateTime tz) where - toTimestamp dt := dt.toUTCTimestamp - instance : HAdd (DateTime tz) (Day.Offset) (DateTime tz) where hAdd := addDays diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index c3b76a3702b5..7ee15bd61baf 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -147,14 +147,40 @@ If the timestamp falls between two transitions, it returns the most recent trans -/ def findTransitionForTimestamp (transitions : Array Transition) (timestamp : Timestamp) : Option Transition := let value := timestamp.toSecondsSinceUnixEpoch - if let some idx := transitions.findIdx? (fun t => t.time.val > value.val) + if let some idx := transitions.findIdx? (fun t => t.time.val ≥ value.val) then transitions.get? (idx - 1) else transitions.back? /-- Find the current `TimeZone` out of a `Transition` in a `Array Transition` -/ -def timezoneAt(transitions : Array Transition) (tm : Timestamp) : Except String TimeZone := +def timezoneAt (transitions : Array Transition) (tm : Timestamp) : Except String TimeZone := if let some transition := findTransitionForTimestamp transitions tm then .ok transition.createTimeZoneFromTransition else .error "cannot find local timezone." + +end Transition +namespace ZoneRules + +/-- +Finds the transition corresponding to a given timestamp in `ZoneRules`. +If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. +-/ +@[inline] +def findTransitionForTimestamp (zr : ZoneRules) (timestamp : Timestamp) : Option Transition := + Transition.findTransitionForTimestamp zr.transitions timestamp + +/-- +Find the current `TimeZone` out of a `Transition` in a `ZoneRules` +-/ +@[inline] +def timezoneAt (zr : ZoneRules) (tm : Timestamp) : Except String TimeZone := + Transition.timezoneAt zr.transitions tm + +/-- +Creates `ZoneRules` for the given `TimeZone`. +-/ +def ofTimeZone (tz : TimeZone) : ZoneRules := + ZoneRules.mk #[] #[Transition.mk tz.toSeconds (LocalTimeType.mk tz.offset tz.isDST tz.name .wall .local tz.abbreviation)] + +end ZoneRules diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 13c4b6da0007..ecc08f77c8a0 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -13,298 +13,306 @@ namespace Time set_option linter.all true /-- -An existential version of `DateTime` that encapsulates a `DateTime` value without explicitly storing -the time zone in the type. +Represents a date and time with timezone information. -/ -def ZonedDateTime := Sigma DateTime +structure ZonedDateTime where + private mk:: -instance : BEq ZonedDateTime where - beq x y := x.fst == y.fst && x.snd.timestamp == y.snd.timestamp + /-- + The plain datetime component, evaluated lazily. + -/ + date : Thunk PlainDateTime -instance : CoeDep ZonedDateTime d (DateTime d.fst) where - coe := d.snd + /-- + The corresponding timestamp for the datetime. + -/ + timestamp : Timestamp + + /-- + The timezone rules applied to this datetime. + -/ + rules : TimeZone.ZoneRules + + /-- + The timezone associated with this datetime. + -/ + timezone : TimeZone instance : Inhabited ZonedDateTime where - default := ⟨Inhabited.default, Inhabited.default⟩ + default := ⟨Thunk.mk Inhabited.default, Inhabited.default, Inhabited.default, Inhabited.default⟩ namespace ZonedDateTime open DateTime /-- -Creates a new `ZonedDateTime` out of a `DateTime` and `TimeZone`. +Creates a new `ZonedDateTime` out of a `Timestamp` and a `TimeZone`. -/ @[inline] -def mk (tz : TimeZone) (datetime : DateTime tz) : ZonedDateTime := - ⟨tz, datetime⟩ +def ofTimestamp (tm : Timestamp) (rules : TimeZone.ZoneRules) : ZonedDateTime := + let tz := TimeZone.Transition.timezoneAt rules.transitions tm |>.toOption |>.getD TimeZone.UTC + ZonedDateTime.mk (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds) |>.toPlainDateTimeAssumingUTC) tm rules tz /-- -Creates a new `ZonedDateTime` out of a `Timestamp` and a `TimeZone`. +Creates a new `ZonedDateTime` out of a `PlainDateTime` and a `TimeZone`. -/ @[inline] -def ofTimestamp (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofTimestamp tm tz⟩ +def ofPlainDateTime (pdt : PlainDateTime) (zr : TimeZone.ZoneRules) : ZonedDateTime := + let tm := pdt.toTimestampAssumingUTC + let tz := zr.findTransitionForTimestamp tm |>.map (·.localTimeType.getTimeZone) |>.getD TimeZone.UTC + let tm := tm.subSeconds tz.toSeconds + ZonedDateTime.mk (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds) |>.toPlainDateTimeAssumingUTC) tm zr tz /-- Creates a new UTC `Timestamp` out of a `ZonedDateTime`. -/ @[inline] def toUTCTimestamp (date : ZonedDateTime) : Timestamp := - date.snd.toUTCTimestamp + date.timestamp /-- -Creates a new `ZonedDateTime` out of a `Timestamp` and `ZoneRules`. +Changes the `ZoleRules` to a new one. -/ @[inline] -def ofZoneRules (tm : Timestamp) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do - let transition ← TimeZone.Transition.findTransitionForTimestamp rules.transitions tm - return ofTimestamp tm transition.localTimeType.getTimeZone - -/-- -Converts a `PlainDateTime` to a `ZonedDateTime` using the given `ZoneRules`. --/ -@[inline] -def ofLocalDateTimeWithZoneRules (pdt : PlainDateTime) (rules : TimeZone.ZoneRules) : Option ZonedDateTime := do - ofZoneRules (pdt.toTimestampAssumingUTC) rules +def convertZoneRules (date : ZonedDateTime) (tz₁ : TimeZone.ZoneRules) : ZonedDateTime := + ofTimestamp date.toUTCTimestamp tz₁ /-- -Changes the `TimeZone` to a new one. +Creates a new `ZonedDateTime` out of a `PlainDateTime` -/ @[inline] -def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := - ofTimestamp date.toUTCTimestamp tz₁ +def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone.ZoneRules) : ZonedDateTime := + ofTimestamp date.toTimestampAssumingUTC tz /-- -Creates a new `ZonedDateTime` out of a `PlainDateTime` +Converts a `ZonedDateTime` to a `PlainDateTime` -/ @[inline] -def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : ZonedDateTime := - ⟨tz, DateTime.ofPlainDateTimeAssumingUTC date tz⟩ +def toPlainDateTime (dt : ZonedDateTime) : PlainDateTime := + dt.date.get /-- Converts a `ZonedDateTime` to a `PlainDateTime` -/ @[inline] -def toPlainDateTime (dt : ZonedDateTime) : PlainDateTime := - DateTime.toPlainDateTime dt.snd +def toDateTime (dt : ZonedDateTime) : DateTime dt.timezone := + DateTime.ofTimestamp dt.timestamp dt.timezone /-- Getter for the `Year` inside of a `ZonedDateTime` -/ @[inline] def year (zdt : ZonedDateTime) : Year.Offset := - zdt.snd.year + zdt.date.get.year /-- Getter for the `Month` inside of a `ZonedDateTime` -/ @[inline] def month (zdt : ZonedDateTime) : Month.Ordinal := - zdt.snd.month + zdt.date.get.month /-- Getter for the `Day` inside of a `ZonedDateTime` -/ @[inline] def day (zdt : ZonedDateTime) : Day.Ordinal := - zdt.snd.day + zdt.date.get.day /-- Getter for the `Hour` inside of a `ZonedDateTime` -/ @[inline] def hour (zdt : ZonedDateTime) : Hour.Ordinal := - zdt.snd.date.get.time.hour + zdt.date.get.time.hour /-- Getter for the `Minute` inside of a `ZonedDateTime` -/ @[inline] def minute (zdt : ZonedDateTime) : Minute.Ordinal := - zdt.snd.minute + zdt.date.get.minute /-- Getter for the `Second` inside of a `ZonedDateTime` -/ @[inline] -def second (zdt : ZonedDateTime) : Second.Ordinal zdt.snd.date.get.time.second.fst := - zdt.snd.date.get.time.second.snd +def second (zdt : ZonedDateTime) : Second.Ordinal zdt.date.get.time.second.fst := + zdt.date.get.time.second.snd /-- Getter for the `Nanosecond` inside of a `ZonedDateTime` -/ @[inline] def nanosecond (zdt : ZonedDateTime) : Nanosecond.Ordinal := - zdt.snd.date.get.time.nano + zdt.date.get.time.nano /-- Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` -/ @[inline] def offset (zdt : ZonedDateTime) : TimeZone.Offset := - zdt.fst.offset - -/-- -Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` --/ -@[inline] -def timezone (zdt : ZonedDateTime) : TimeZone := - zdt.fst + zdt.timezone.offset /-- Returns the weekday. -/ @[inline] def weekday (zdt : ZonedDateTime) : Weekday := - zdt.snd.weekday + zdt.date.get.weekday /-- Transforms a tuple of a `ZonedDateTime` into a `Day.Ordinal.OfYear`. -/ @[inline] def toOrdinal (date : ZonedDateTime) : Day.Ordinal.OfYear date.year.isLeap := - Month.Ordinal.toOrdinal ⟨(date.month, date.day), date.snd.date.get.date.valid⟩ + ValidDate.toOrdinal ⟨(date.month, date.day), date.date.get.date.valid⟩ /-- Add `Day.Offset` to a `ZonedDateTime`. -/ def addDays (dt : ZonedDateTime) (days : Day.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addDays days) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addDays days).toTimestampAssumingUTC dt.rules /-- -Subtracts `Day.Offset` to a `ZonedDateTime`. +Subtract `Day.Offset` from a `ZonedDateTime`. -/ -@[inline] def subDays (dt : ZonedDateTime) (days : Day.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subDays days) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subDays days).toTimestampAssumingUTC dt.rules /-- Add `Week.Offset` to a `ZonedDateTime`. -/ def addWeeks (dt : ZonedDateTime) (weeks : Week.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addWeeks weeks) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addWeeks weeks).toTimestampAssumingUTC dt.rules /-- -Subtracts `Week.Offset` to a `ZonedDateTime`. +Subtract `Week.Offset` from a `ZonedDateTime`. -/ -@[inline] def subWeeks (dt : ZonedDateTime) (weeks : Week.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subWeeks weeks) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subWeeks weeks).toTimestampAssumingUTC dt.rules /-- -Add `Month.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. +Add `Month.Offset` to a `ZonedDateTime`, clipping to the last valid day. -/ def addMonthsClip (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addMonthsClip months) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addMonthsClip months).toTimestampAssumingUTC dt.rules /-- -Subtracts `Month.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. +Subtract `Month.Offset` from a `ZonedDateTime`, clipping to the last valid day. -/ -@[inline] def subMonthsClip (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subMonthsClip months) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subMonthsClip months).toTimestampAssumingUTC dt.rules /-- -Add `Month.Offset` to a `ZonedDateTime`, this function rolls over any excess days into the following -month. +Add `Month.Offset` to a `ZonedDateTime`, rolling over excess days. -/ def addMonthsRollOver (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addMonthsRollOver months) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addMonthsRollOver months).toTimestampAssumingUTC dt.rules /-- -Add `Month.Offset` to a `ZonedDateTime`, this function rolls over any excess days into the following -month. +Subtract `Month.Offset` from a `ZonedDateTime`, rolling over excess days. -/ -@[inline] def subMonthsRollOver (dt : ZonedDateTime) (months : Month.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subMonthsRollOver months) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subMonthsRollOver months).toTimestampAssumingUTC dt.rules /-- -Add `Year.Offset` to a `ZonedDateTime`, this function rolls over any excess days into the following -month. +Add `Year.Offset` to a `ZonedDateTime`, rolling over excess days. -/ -@[inline] def addYearsRollOver (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addYearsRollOver years) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addYearsRollOver years).toTimestampAssumingUTC dt.rules /-- -Add `Year.Offset` to a `ZonedDateTime`, it clips the day to the last valid day of that month. +Add `Year.Offset` to a `ZonedDateTime`, clipping to the last valid day. -/ -@[inline] def addYearsClip (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addYearsClip years) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addYearsClip years).toTimestampAssumingUTC dt.rules /-- -Subtract `Year.Offset` from a `ZonedDateTime`, this function clips the day to the last valid day of that month. +Subtract `Year.Offset` from a `ZonedDateTime`, clipping to the last valid day. -/ -@[inline] def subYearsClip (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subYearsClip years) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subYearsClip years).toTimestampAssumingUTC dt.rules /-- -Subtract `Year.Offset` from a `ZonedDateTime`, this function rolls over any excess days into the previous month. +Subtract `Year.Offset` from a `ZonedDateTime`, rolling over excess days. -/ -@[inline] def subYearsRollOver (dt : ZonedDateTime) (years : Year.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subYearsRollOver years) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subYearsRollOver years).toTimestampAssumingUTC dt.rules /-- -Add `Hour.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +Add `Hour.Offset` to a `ZonedDateTime`. -/ def addHours (dt : ZonedDateTime) (hours : Hour.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addHours hours) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addHours hours).toTimestampAssumingUTC dt.rules /-- -Subtract `Hour.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +Subtract `Hour.Offset` from a `ZonedDateTime`. -/ -@[inline] def subHours (dt : ZonedDateTime) (hours : Hour.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subHours hours) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subHours hours).toTimestampAssumingUTC dt.rules /-- -Add `Minute.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +Add `Minute.Offset` to a `ZonedDateTime`. -/ def addMinutes (dt : ZonedDateTime) (minutes : Minute.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addMinutes minutes) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addMinutes minutes).toTimestampAssumingUTC dt.rules /-- -Subtract `Minute.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +Subtract `Minute.Offset` from a `ZonedDateTime`. -/ -@[inline] def subMinutes (dt : ZonedDateTime) (minutes : Minute.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subMinutes minutes) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subMinutes minutes).toTimestampAssumingUTC dt.rules /-- -Add `Second.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +Add `Second.Offset` to a `ZonedDateTime`. -/ def addSeconds (dt : ZonedDateTime) (seconds : Second.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addSeconds seconds) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addSeconds seconds).toTimestampAssumingUTC dt.rules /-- -Subtract `Second.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +Subtract `Second.Offset` from a `ZonedDateTime`. -/ -@[inline] def subSeconds (dt : ZonedDateTime) (seconds : Second.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subSeconds seconds) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subSeconds seconds).toTimestampAssumingUTC dt.rules /-- -Add `Nanosecond.Offset` to a `ZonedDateTime`, adjusting the date if necessary. +Add `Nanosecond.Offset` to a `ZonedDateTime`. -/ def addNanoseconds (dt : ZonedDateTime) (nanoseconds : Nanosecond.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.addNanoseconds nanoseconds) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addNanoseconds nanoseconds).toTimestampAssumingUTC dt.rules /-- -Subtract `Nanosecond.Offset` from a `ZonedDateTime`, adjusting the date if necessary. +Subtract `Nanosecond.Offset` from a `ZonedDateTime`. -/ -@[inline] def subNanoseconds (dt : ZonedDateTime) (nanoseconds : Nanosecond.Offset) : ZonedDateTime := - Sigma.mk dt.fst (dt.snd.subNanoseconds nanoseconds) + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subNanoseconds nanoseconds).toTimestampAssumingUTC dt.rules /-- Determines the era of the given `ZonedDateTime` based on its year. -/ @[inline] def era (date : ZonedDateTime) : Year.Era := - date.snd.era + date.date.get.era /-- Creates a new `ZonedDateTime` by adjusting the day of the month to the given `days` value, with any @@ -312,7 +320,8 @@ out-of-range days clipped to the nearest valid date. -/ @[inline] def withDaysClip (dt : ZonedDateTime) (days : Day.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withDaysClip days⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withDaysClip days) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the day of the month to the given `days` value, with any @@ -320,7 +329,8 @@ out-of-range days rolled over to the next month or year as needed. -/ @[inline] def withDaysRollOver (dt : ZonedDateTime) (days : Day.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withDaysRollOver days⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withDaysRollOver days) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the month to the given `month` value. @@ -328,7 +338,8 @@ The day remains unchanged, and any invalid days for the new month will be handle -/ @[inline] def withMonthClip (dt : ZonedDateTime) (month : Month.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withMonthClip month⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withMonthClip month) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the month to the given `month` value. @@ -336,7 +347,8 @@ The day is rolled over to the next valid month if necessary. -/ @[inline] def withMonthRollOver (dt : ZonedDateTime) (month : Month.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withMonthRollOver month⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withMonthRollOver month) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the year to the given `year` value. The month and day remain unchanged, @@ -344,7 +356,9 @@ and any invalid days for the new year will be handled according to the `clip` be -/ @[inline] def withYearClip (dt : ZonedDateTime) (year : Year.Offset) : ZonedDateTime := - ⟨dt.fst, dt.snd.withYearClip year⟩ + let date := dt.date.get + + ZonedDateTime.ofPlainDateTime (date.withYearClip year) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the year to the given `year` value. The month and day are rolled @@ -352,35 +366,40 @@ over to the next valid month and day if necessary. -/ @[inline] def withYearRollOver (dt : ZonedDateTime) (year : Year.Offset) : ZonedDateTime := - ⟨dt.fst, dt.snd.withYearRollOver year⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withYearRollOver year) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the `hour` component. -/ @[inline] def withHours (dt : ZonedDateTime) (hour : Hour.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withHours hour⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withHours hour) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the `minute` component. -/ @[inline] def withMinutes (dt : ZonedDateTime) (minute : Minute.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withMinutes minute⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withMinutes minute) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the `second` component. -/ @[inline] def withSeconds (dt : ZonedDateTime) (second : Sigma Second.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withSeconds second⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withSeconds second) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the `nano` component. -/ @[inline] def withNanoseconds (dt : ZonedDateTime) (nano : Nanosecond.Ordinal) : ZonedDateTime := - ⟨dt.fst, dt.snd.withNanoseconds nano⟩ + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withNanoseconds nano) dt.rules /-- Checks if the `ZonedDateTime` is in a leap year. @@ -388,9 +407,6 @@ Checks if the `ZonedDateTime` is in a leap year. def inLeapYear (date : ZonedDateTime) : Bool := date.year.isLeap -instance : ToTimestamp ZonedDateTime where - toTimestamp dt := dt.toUTCTimestamp - instance : HAdd ZonedDateTime Day.Offset ZonedDateTime where hAdd := addDays diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 452499636a17..bf0ed501a9a2 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -11,6 +11,7 @@ def Time24Hour : GenericFormat .any := datespec("HH:mm:ss") def Time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM yyyy HH:mm") +def EraDate : GenericFormat .any := datespec("MM D, yyyy G") -- Dates @@ -26,9 +27,9 @@ def time₂ := time("03:11:01") info: "Monday, June 16, 2014 03:03:03 -0300" -/ #guard_msgs in -#eval FullDayTimeZone.format date₁.snd +#eval FullDayTimeZone.format date₁.toDateTime -def tm := date₁.snd.timestamp +def tm := date₁.toUTCTimestamp def date₂ := DateTime.ofTimestamp tm brTZ /-- @@ -205,31 +206,31 @@ info: "0053-06-19" #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-700000⟩) .UTC) /-- -info: "-0002-09-16" +info: "0003-09-16" -/ #guard_msgs in #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-720000⟩) .UTC) /-- -info: "-0084-07-28" +info: "0085-07-28" -/ #guard_msgs in #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-750000⟩) .UTC) /-- -info: "-0221-09-04" +info: "0222-09-04" -/ #guard_msgs in #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) /-- -info: date("-0221-09-04") +info: date("0222-09-04") -/ #guard_msgs in #eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ /-- -info: date("-0221-09-04") +info: date("0222-09-04") -/ #guard_msgs in #eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ @@ -276,6 +277,7 @@ Format def time₄ := time("23:13:12.324354679") def date₄ := date("2002-07-14") +def datetime₅ := PlainDateTime.mk (PlainDate.clip (-2000) 3 4) (PlainTime.mk 12 23 ⟨false, 12⟩ 0) def datetime₄ := datetime("2002-07-14T23:13:12.324354679") def zoned₄ := zoned("2002-07-14T23:13:12.324354679+09:00") @@ -291,6 +293,12 @@ info: "02 2002 000002002" #guard_msgs in #eval zoned₄.format "yy yyyy yyyyyyyyy" +/-- +info: "02 2002 000002002" +-/ +#guard_msgs in +#eval zoned₄.format "uu uuuu uuuuuuuuu" + /-- info: "5 95 195" -/ @@ -351,7 +359,7 @@ info: "7 07 Sun Sunday S" #eval zoned₄.format "e ee eee eeee eeeee" /-- -info: "3 03 003 0003" +info: "2 02 002 0002" -/ #guard_msgs in #eval zoned₄.format "F FF FFF FFFF" @@ -470,6 +478,12 @@ info: "02 2002 000002002" #guard_msgs in #eval datetime₄.format "yy yyyy yyyyyyyyy" +/-- +info: "02 2002 000002002" +-/ +#guard_msgs in +#eval datetime₄.format "uu uuuu uuuuuuuuu" + /-- info: "5 95 195" -/ @@ -530,7 +544,7 @@ info: "7 07 Sun Sunday S" #eval datetime₄.format "e ee eee eeee eeeee" /-- -info: "3 03 003 0003" +info: "2 02 002 0002" -/ #guard_msgs in #eval datetime₄.format "F FF FFF FFFF" @@ -679,6 +693,12 @@ info: "02 2002 000002002" #guard_msgs in #eval date₄.format "yy yyyy yyyyyyyyy" +/-- +info: "02 2002 000002002" +-/ +#guard_msgs in +#eval date₄.format "uu uuuu uuuuuuuuu" + /-- info: "5 95 195" -/ @@ -739,7 +759,27 @@ info: "7 07 Sun Sunday S" #eval date₄.format "e ee eee eeee eeeee" /-- -info: "3 03 003 0003" +info: "2 02 002 0002" -/ #guard_msgs in #eval date₄.format "F FF FFF FFFF" + +/-- +info: "-2000 2001 BCE" +-/ +#guard_msgs in +#eval datetime₅.format "uuuu yyyy G" + +/-- +info: "2002 2002 CE" +-/ +#guard_msgs in +#eval datetime₄.format "uuuu yyyy G" + +/-- +info: 3 +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := .ofPlainDateTime datetime("2018-12-31T12:00:00") (TimeZone.ZoneRules.ofTimeZone TimeZone.UTC) + IO.println s!"{t.format "w"}" diff --git a/tests/lean/run/timeOperations.lean b/tests/lean/run/timeOperations.lean index 9f4eba8a858c..2aee656f9121 100644 --- a/tests/lean/run/timeOperations.lean +++ b/tests/lean/run/timeOperations.lean @@ -274,34 +274,46 @@ info: zoned("2000-01-20T06:01:01.000000000-03:00") #guard_msgs in #eval datetimetz.subSeconds 60 +def now := zoned("2024-08-29T10:56:43.276801081+02:00") + /-- -info: "86400s" +info: zoned("2024-08-29T10:56:43.276801081+02:00") -/ #guard_msgs in -#eval (date("2000-01-20").since (date("2000-01-19"))) +#eval now /-- -info: "86399s" +info: zoned("2024-08-30T10:56:43.276801081+02:00") -/ #guard_msgs in -#eval (datetime("2000-01-20T00:00:00").since (datetime("2000-01-19T00:00:01"))) +#eval now.addDays 1 /-- -info: "86399s" +info: zoned("2000-01-20T06:01:01.000000000-03:00") -/ #guard_msgs in -#eval (zoned("2000-01-20T00:00:00Z").since (zoned("2000-01-19T00:00:01Z"))) +#eval datetimetz.subSeconds 60 -def now := zoned("2024-08-29T10:56:43.276801081+02:00") +/-- +info: 3 +-/ +#guard_msgs in +#eval date("2024-11-17").alignedWeekOfMonth /-- -info: zoned("2024-08-29T10:56:43.276801081+02:00") +info: 4 -/ #guard_msgs in -#eval now +#eval date("2024-11-18").alignedWeekOfMonth /-- -info: zoned("2024-08-30T10:56:43.276801081+02:00") +info: 3 -/ #guard_msgs in -#eval now.addDays 1 +#eval date("2024-01-21").alignedWeekOfMonth + +/-- +info: 4 +-/ +#guard_msgs in +#eval date("2024-01-22").alignedWeekOfMonth diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 4b8c4586a1a1..d242f3c7d9a3 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -31,9 +31,9 @@ info: "2014-06-16T03:03:03.000000100-03:00" #guard_msgs in #eval let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03.000000100-03:00" - ISO8601UTC.format t.snd + ISO8601UTC.format t.toDateTime -def tm := date₁.snd.timestamp +def tm := date₁.toUTCTimestamp def date₂ := DateTime.ofTimestamp tm brTZ /-- @@ -42,9 +42,9 @@ info: "2014-06-16T03:03:03.000000000-03:00" #guard_msgs in #eval let t : ZonedDateTime := RFC1123.parse! "Mon, 16 Jun 2014 03:03:03 -0300" - ISO8601UTC.format t.snd + ISO8601UTC.format t.toDateTime -def tm₃ := date₁.toTimestamp +def tm₃ := date₁.toUTCTimestamp def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- @@ -53,7 +53,7 @@ info: "2014-06-16T00:00:00.000000000Z" #guard_msgs in #eval let t : ZonedDateTime := ShortDate.parse! "06/16/2014" - ISO8601UTC.format t.snd + ISO8601UTC.format t.toDateTime -- the timestamp is always related to UTC. @@ -74,7 +74,7 @@ info: "2024-08-15T13:28:12.000000000-03:00" #guard_msgs in #eval let t := FullDayTimeZone.parse! "Thursday, August 15, 2024 13:28:12 -0300" - ISO8601UTC.format t.snd + ISO8601UTC.format t.toDateTime /-- info: "2024-08-16T01:28:00.000000000Z" @@ -82,23 +82,23 @@ info: "2024-08-16T01:28:00.000000000Z" #guard_msgs in #eval let t : ZonedDateTime := LongDateTime.parse! "August 16, 2024 01:28 AM" - ISO8601UTC.format t.snd + ISO8601UTC.format t.toDateTime /-- -info: "0000-12-30T22:28:12.000000000+09:00" +info: "0001-12-30T22:28:12.000000000+09:00" -/ #guard_msgs in #eval let t : ZonedDateTime := Time24Hour.parse! "13:28:12" - ISO8601UTC.format (t.snd.convertTimeZone jpTZ) + ISO8601UTC.format (t.toDateTime.convertTimeZone jpTZ) /-- -info: "0000-12-29T21:28:12.000000000-03:00" +info: "0001-12-29T21:28:12.000000000-03:00" -/ #guard_msgs in #eval let t1 : ZonedDateTime := Time12Hour.parse! "12:28:12 AM" - ISO8601UTC.format (t1.snd.convertTimeZone brTZ) + ISO8601UTC.format (t1.toDateTime.convertTimeZone brTZ) /-- info: "Thu 15 Aug 2024 16:28" @@ -106,7 +106,7 @@ info: "Thu 15 Aug 2024 16:28" #guard_msgs in #eval let t2 : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 16:28:12 -0000" - CustomDayTime.format t2.snd + CustomDayTime.format t2.toDateTime /-- info: "2024-08-16T13:28:00.000000000Z" @@ -114,7 +114,7 @@ info: "2024-08-16T13:28:00.000000000Z" #guard_msgs in #eval let t5 : ZonedDateTime := CustomDayTime.parse! "Thu 16 Aug 2024 13:28" - ISO8601UTC.format t5.snd + ISO8601UTC.format t5.toDateTime /-- info: "2024-08-16T01:28:12.000000000+09:00" @@ -122,7 +122,7 @@ info: "2024-08-16T01:28:12.000000000+09:00" #guard_msgs in #eval let t6 : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" - ISO8601UTC.format (t6.snd.convertTimeZone jpTZ) + ISO8601UTC.format (t6.toDateTime.convertTimeZone jpTZ) /-- info: "2024-08-16T01:28:12.000000000+09:00" @@ -130,7 +130,7 @@ info: "2024-08-16T01:28:12.000000000+09:00" #guard_msgs in #eval let t7 : ZonedDateTime := FullDayTimeZone.parse! "Friday, August 16, 2024 01:28:12 +0900" - ISO8601UTC.format (t7.snd.convertTimeZone jpTZ) + ISO8601UTC.format (t7.toDateTime.convertTimeZone jpTZ) /-- TM: 1723730627 @@ -154,7 +154,7 @@ info: "2024-08-15T14:03:47.000000000-03:00" #guard_msgs in #eval let t : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 14:03:47 -0300" - ISO8601UTC.format t.snd + ISO8601UTC.format t.toDateTime /-- info: "2024-08-15T14:03:47.000000000+09:00" @@ -162,7 +162,7 @@ info: "2024-08-15T14:03:47.000000000+09:00" #guard_msgs in #eval let t1 : ZonedDateTime := FullDayTimeZone.parse! "Thursday, August 15, 2024 14:03:47 +0900" - ISO8601UTC.format t1.snd + ISO8601UTC.format t1.toDateTime /-- info: "2014-06-16T03:03:03.000000000-03:00" @@ -170,7 +170,7 @@ info: "2014-06-16T03:03:03.000000000-03:00" #guard_msgs in #eval let t2 : ZonedDateTime := FullDayTimeZone.parse! "Monday, June 16, 2014 03:03:03 -0300" - ISO8601UTC.format t2.snd + ISO8601UTC.format t2.toDateTime /-- info: Except.ok "1993-05-10T10:30:23.000000000+03:00" @@ -178,7 +178,7 @@ info: Except.ok "1993-05-10T10:30:23.000000000+03:00" #guard_msgs in #eval let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 AM +03:00" - (ISO8601UTC.format ·.snd) <$> t2 + (ISO8601UTC.format ·.toDateTime) <$> t2 /-- info: Except.ok "1993-05-10T22:30:23.000000000+03:00" @@ -186,7 +186,7 @@ info: Except.ok "1993-05-10T22:30:23.000000000+03:00" #guard_msgs in #eval let t2 := Full12HourWrong.parse "05/10/1993 10:30:23 PM +03:00" - (ISO8601UTC.format ·.snd) <$> t2 + (ISO8601UTC.format ·.toDateTime) <$> t2 /-- info: Except.error "offset 13: need a natural number in the interval of 1 to 12" @@ -194,7 +194,7 @@ info: Except.error "offset 13: need a natural number in the interval of 1 to 12" #guard_msgs in #eval let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 AM +03:00" - (ISO8601UTC.format ·.snd) <$> t2 + (ISO8601UTC.format ·.toDateTime) <$> t2 /-- info: Except.error "offset 13: need a natural number in the interval of 1 to 12" @@ -202,4 +202,4 @@ info: Except.error "offset 13: need a natural number in the interval of 1 to 12" #guard_msgs in #eval let t2 := Full12HourWrong.parse "05/10/1993 20:30:23 PM +03:00" - (ISO8601UTC.format ·.snd) <$> t2 + (ISO8601UTC.format ·.toDateTime) <$> t2 diff --git a/tests/lean/run/timeSet.lean b/tests/lean/run/timeSet.lean index 36020cf2d137..60b6748e5487 100644 --- a/tests/lean/run/timeSet.lean +++ b/tests/lean/run/timeSet.lean @@ -19,75 +19,82 @@ def Full12HourWrong : GenericFormat .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX def brTZ : TimeZone := timezone("America/Sao_Paulo -03:00") def jpTZ : TimeZone := timezone("Asia/Tokyo +09:00") -def date₁ := zoned("2014-06-16T03:03:03-03:00") +def date₁ := zoned("2014-06-16T10:03:03-03:00") def time₁ := time("14:11:01") def time₂ := time("03:11:01") /-- -info: "2014-06-16T03:03:03.000000100-03:00" +info: "2014-06-16T10:03:03.000000100-03:00" -/ #guard_msgs in #eval - let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03.000000100-03:00" - ISO8601UTC.format t.snd + let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T10:03:03.000000100-03:00" + ISO8601UTC.format t.toDateTime /-- -info: zoned("2014-06-30T03:03:03.000000000-03:00") +info: zoned("2014-06-30T10:03:03.000000000-03:00") -/ #guard_msgs in #eval date₁.withDaysClip 31 /-- -info: zoned("2014-07-01T03:03:03.000000000-03:00") +info: zoned("2014-07-01T10:03:03.000000000-03:00") -/ #guard_msgs in #eval date₁.withDaysRollOver 31 /-- -info: zoned("2014-05-16T03:03:03.000000000-03:00") +info: zoned("2014-05-16T10:03:03.000000000-03:00") -/ #guard_msgs in #eval date₁.withMonthClip 5 /-- -info: zoned("2014-05-16T03:03:03.000000000-03:00") +info: zoned("2014-05-16T10:03:03.000000000-03:00") -/ #guard_msgs in #eval date₁.withMonthRollOver 5 /-- -info: zoned("2016-06-16T03:03:03.000000000-03:00") +info: zoned("2016-06-16T10:03:03.000000000-03:00") -/ #guard_msgs in #eval date₁.withYearClip 2016 /-- -info: zoned("2016-06-16T03:03:03.000000000-03:00") +info: zoned("2016-06-16T10:03:03.000000000-03:00") -/ #guard_msgs in #eval date₁.withYearRollOver 2016 /-- -info: zoned("2014-06-16T19:03:03.000000000-03:00") +info: zoned("2014-06-16T10:03:03.000000000-03:00") -/ #guard_msgs in -#eval date₁.withHours 19 +#eval date₁.withDaysClip 16 /-- -info: zoned("2014-06-16T03:45:03.000000000-03:00") +info: zoned("2014-06-16T10:45:03.000000000-03:00") -/ #guard_msgs in #eval date₁.withMinutes 45 + +/-- +info: zoned("2014-06-16T10:03:03.000000000-03:00") +-/ +#guard_msgs in +#eval date₁.withHours 10 + /-- -info: zoned("2014-06-16T03:03:59.000000000-03:00") +info: zoned("2014-06-16T10:03:59.000000000-03:00") -/ #guard_msgs in #eval date₁.withSeconds ⟨true, 59⟩ /-- -info: zoned("2014-06-16T03:03:03.000000002-03:00") +info: zoned("2014-06-16T10:03:03.000000002-03:00") -/ #guard_msgs in #eval date₁.withNanoseconds 2 From 25c4aab4409c8a296182827acf6f488cfde41a94 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 23 Oct 2024 00:53:47 -0300 Subject: [PATCH 077/118] fix: remove dbg_trace --- src/Std/Time/Zoned/Database/Leap.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/Std/Time/Zoned/Database/Leap.lean b/src/Std/Time/Zoned/Database/Leap.lean index 4cabdc1156b2..d3852567d00f 100644 --- a/src/Std/Time/Zoned/Database/Leap.lean +++ b/src/Std/Time/Zoned/Database/Leap.lean @@ -89,8 +89,6 @@ private def parseLeapSecond : Parser LeapSecond := do let stationary ← manyChars (satisfy Char.isAlpha) skipChar '\n' - dbg_trace s!"{repr year}--{repr month}--{day}" - let day ← failWith <| Internal.Bounded.ofInt? day let time : PlainTime ← failWith <| PlainTime.ofHourMinuteSeconds hour minute second let date : PlainDate ← failWith <| PlainDate.ofYearMonthDay? (Year.Offset.ofNat year) month day From 26adc207e3f0a634f87a9a4f020ad06968e976af Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 23 Oct 2024 03:57:15 -0300 Subject: [PATCH 078/118] feat: wip c++ next_transition --- src/Std/Time/Zoned/Database/Windows.lean | 4 ++-- src/runtime/io.cpp | 27 ++++++++++++++++++------ 2 files changed, 22 insertions(+), 9 deletions(-) diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index ecf0bed4111e..a83b82f581ec 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -18,8 +18,8 @@ set_option linter.all true /-- Fetches the timezone information from the current windows machine. -/ -@[extern "lean_get_windows_timezone_at"] -opaque Windows.getTimeZoneAt : String -> UInt64 -> IO TimeZone +@[extern "lean_get_windows_next_transition"] +opaque Windows.getNextTransition : String -> @&Int -> IO (Option (Int × TimeZone)) /-- Fetches the timezone at a timestamp. diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index a1be24b9dda4..ac3956511938 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -648,8 +648,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } -/* Std.Time.Database.Windows.getLocalTimeZoneIdentifier : IO String */ -extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str, obj_arg /* w */) { +/* Std.Time.Database.Windows.getNextTransition : String -> @&Int -> IO (Option (Int × TimeZone)) */ +extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone_str, obj_arg tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; const char* dst_name = lean_string_cstr(timezone_str); @@ -668,12 +668,20 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); } + int64_t timestamp_secs = lean_scalar_to_int64(tm_obj); + ucal_setMillis(cal, timestamp_secs * 1000, &status); if (U_FAILURE(status)) { ucal_close(cal); return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); } + UDate nextTransition; + if (!ucal_getTimeZoneTransitionDate(cal, UCAL_TZ_TRANSITION_NEXT, &nextTransition, &status)) { + ucal_close(cal); + return io_result_mk_ok(mk_option_none()); + } + UChar display_name[32]; int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); @@ -706,23 +714,28 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_timezone_at(obj_arg timezone_str return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); } + int64_t tm = (int64_t)(nextTransition / 1000.0); ucal_close(cal); - int offset_hour = zone_offset / 3600000; int offset_seconds = zone_offset / 1000; int is_dst = dst_offset != 0; - lean_object *lean_offset = lean_alloc_ctor(0, 2, 0); - lean_ctor_set(lean_offset, 0, lean_int_to_int(offset_hour)); + lean_object *lean_offset = lean_alloc_ctor(0, 1, 0); lean_ctor_set(lean_offset, 1, lean_int_to_int(offset_seconds)); lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); - lean_ctor_set(lean_tz, 0, lean_offset); + lean_ctor_set(lean_tz, 0, lean_int_to_int(offset_seconds)); lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked(dst_name)); lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked(display_name_str)); lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); + + lean_object *lean_pair = lean_alloc_ctor(0, 2, 0); + lean_ctor_set(lean_pair, 0, lean_int64_to_int(tm)); + lean_ctor_set(lean_pair, 1, lean_tz); + + - return lean_io_result_mk_ok(lean_tz); + return lean_io_result_mk_ok(mk_option_some(lean_pair)); #else return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get timezone, its windows only."))); #endif From fcef2e0e3af7708420ee8ebbc39b3c2a8cd0c58b Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 23 Oct 2024 04:07:05 -0300 Subject: [PATCH 079/118] feat: add getZoneRules --- src/Std/Time/Zoned/Database/Windows.lean | 35 ++++++++++++++++++++---- src/runtime/io.cpp | 4 ++- 2 files changed, 33 insertions(+), 6 deletions(-) diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index a83b82f581ec..bf21625f3edc 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -15,17 +15,42 @@ namespace Database set_option linter.all true +namespace Windows + /-- -Fetches the timezone information from the current windows machine. +Fetches the next timezone transition for a given timezone identifier and timestamp. -/ @[extern "lean_get_windows_next_transition"] -opaque Windows.getNextTransition : String -> @&Int -> IO (Option (Int × TimeZone)) +opaque getNextTransition : String -> @&Int -> IO (Option (Int × TimeZone)) /-- Fetches the timezone at a timestamp. -/ @[extern "lean_get_windows_local_timezone_id_at"] -opaque Windows.getLocalTimeZoneIdentifierAt : UInt64 → IO String +opaque getLocalTimeZoneIdentifierAt : Int → IO String + +/-- +Retrieves the timezone rules, including all transitions, for a given timezone identifier. +-/ +partial def getZoneRules (id : String) : IO TimeZone.ZoneRules := do + let mut start := -2147483647 + let mut transitions := #[] + + while true do + let result ← getNextTransition id start + if let some res := result then + transitions := transitions.push { time := Second.Offset.ofInt res.fst, localTimeType := { + gmtOffset := res.snd.offset, + abbreviation := res.snd.abbreviation, + identifier := res.snd.name, + isDst := res.snd.isDST, + wall := .wall, + utLocal := .local + }} + + return { transitions, localTimes := #[] } + +end Windows /-- Represents a Time Zone Database that we get from ICU available on Windows SDK. @@ -42,5 +67,5 @@ Returns a default `WindowsDb` instance. def default : WindowsDb := {} instance : Database WindowsDb where - getZoneRulesAt _ _ := pure <| ZoneRules.mk #[] #[] - getLocalZoneRulesAt _ := pure Inhabited.default + getZoneRulesAt _ id := Windows.getZoneRules id + getLocalZoneRulesAt _ := Windows.getZoneRules =<< Windows.getLocalTimeZoneIdentifierAt (-2147483647) diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index ac3956511938..7edbe10ce821 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -742,7 +742,7 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone } /* SStd.Time.Database.Windows.getLocalTimeZoneIdentifierAt : UInt64 → IO String */ -extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(uint64_t timestamp_secs, obj_arg /* w */) { +extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(obj_arg tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; UCalendar* cal = ucal_open(NULL, -1, NULL, UCAL_GREGORIAN, &status); @@ -751,7 +751,9 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(uint64_t ti return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); } + int64_t timestamp_secs = lean_scalar_to_int64(tm_obj); ucal_setMillis(cal, timestamp_secs * 1000, &status); + if (U_FAILURE(status)) { ucal_close(cal); return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); From aa02d6e2665197cfb90e6420d48422fda8b829ef Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 23 Oct 2024 13:48:36 -0300 Subject: [PATCH 080/118] feat: finish get next transition --- src/Std/Time/Zoned/Database/Windows.lean | 15 ++++---- src/runtime/io.cpp | 46 ++++++++++++++---------- 2 files changed, 37 insertions(+), 24 deletions(-) diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index bf21625f3edc..436a2c31f430 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -21,23 +21,23 @@ namespace Windows Fetches the next timezone transition for a given timezone identifier and timestamp. -/ @[extern "lean_get_windows_next_transition"] -opaque getNextTransition : String -> @&Int -> IO (Option (Int × TimeZone)) +opaque getNextTransition : @&String -> @&Int -> IO (Option (Int × TimeZone)) /-- Fetches the timezone at a timestamp. -/ @[extern "lean_get_windows_local_timezone_id_at"] -opaque getLocalTimeZoneIdentifierAt : Int → IO String +opaque getLocalTimeZoneIdentifierAt : @&Int → IO String /-- Retrieves the timezone rules, including all transitions, for a given timezone identifier. -/ -partial def getZoneRules (id : String) : IO TimeZone.ZoneRules := do - let mut start := -2147483647 +def getZoneRules (id : String) : IO TimeZone.ZoneRules := do + let mut start := -2147483648 let mut transitions := #[] while true do - let result ← getNextTransition id start + let result ← Windows.getNextTransition id start if let some res := result then transitions := transitions.push { time := Second.Offset.ofInt res.fst, localTimeType := { gmtOffset := res.snd.offset, @@ -47,6 +47,9 @@ partial def getZoneRules (id : String) : IO TimeZone.ZoneRules := do wall := .wall, utLocal := .local }} + start := res.fst + else + break return { transitions, localTimes := #[] } @@ -68,4 +71,4 @@ def default : WindowsDb := {} instance : Database WindowsDb where getZoneRulesAt _ id := Windows.getZoneRules id - getLocalZoneRulesAt _ := Windows.getZoneRules =<< Windows.getLocalTimeZoneIdentifierAt (-2147483647) + getLocalZoneRulesAt _ := Windows.getZoneRules =<< Windows.getLocalTimeZoneIdentifierAt (-2147483648) diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 7edbe10ce821..baf83699e409 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -652,10 +652,10 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone_str, obj_arg tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; - const char* dst_name = lean_string_cstr(timezone_str); + const char* dst_name_id = lean_string_cstr(timezone_str); UChar tzID[256]; - u_strFromUTF8(tzID, sizeof(tzID) / sizeof(tzID[0]), NULL, dst_name, strlen(dst_name), &status); + u_strFromUTF8(tzID, sizeof(tzID) / sizeof(tzID[0]), NULL, dst_name_id, strlen(dst_name_id), &status); if (U_FAILURE(status)) { return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to read identifier"))); @@ -675,6 +675,15 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone ucal_close(cal); return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); } + + int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); + } + + int is_dst = dst_offset != 0; UDate nextTransition; if (!ucal_getTimeZoneTransitionDate(cal, UCAL_TZ_TRANSITION_NEXT, &nextTransition, &status)) { @@ -682,8 +691,23 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone return io_result_mk_ok(mk_option_none()); } + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get next transation"))); + } + + int32_t tzIDLength = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_DST : UCAL_STANDARD, "en_US", tzID, 32, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to timezone identifier"))); + } + + char dst_name[256]; + u_strToUTF8(dst_name, sizeof(dst_name), NULL, tzID, tzIDLength, &status); + UChar display_name[32]; - int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); + int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_SHORT_DST : UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); if (U_FAILURE(status)) { ucal_close(cal); @@ -699,29 +723,17 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone } int32_t zone_offset = ucal_get(cal, UCAL_ZONE_OFFSET, &status); - - if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); - } - - int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); - zone_offset += dst_offset; if (U_FAILURE(status)) { ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); } int64_t tm = (int64_t)(nextTransition / 1000.0); ucal_close(cal); int offset_seconds = zone_offset / 1000; - int is_dst = dst_offset != 0; - - lean_object *lean_offset = lean_alloc_ctor(0, 1, 0); - lean_ctor_set(lean_offset, 1, lean_int_to_int(offset_seconds)); lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); lean_ctor_set(lean_tz, 0, lean_int_to_int(offset_seconds)); @@ -732,8 +744,6 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone lean_object *lean_pair = lean_alloc_ctor(0, 2, 0); lean_ctor_set(lean_pair, 0, lean_int64_to_int(tm)); lean_ctor_set(lean_pair, 1, lean_tz); - - return lean_io_result_mk_ok(mk_option_some(lean_pair)); #else From 70c82254c064758033a7332d235d961198f2c654 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 28 Oct 2024 10:11:22 -0300 Subject: [PATCH 081/118] fix: a bunch of stuff --- script/prepare-llvm-mingw.sh | 1 - src/Std/Time.lean | 8 +- src/Std/Time/Format.lean | 76 ++++++------ src/Std/Time/Format/Basic.lean | 16 +-- src/Std/Time/Time/Unit/Hour.lean | 9 +- src/Std/Time/Zoned.lean | 6 +- src/Std/Time/Zoned/Database/Basic.lean | 12 +- src/Std/Time/Zoned/Database/TZdb.lean | 2 +- src/Std/Time/Zoned/Database/Windows.lean | 11 +- src/Std/Time/Zoned/DateTime.lean | 6 +- src/Std/Time/Zoned/ZoneRules.lean | 19 ++- src/Std/Time/Zoned/ZonedDateTime.lean | 16 +-- src/runtime/io.cpp | 6 +- tests/lean/run/timeFormats.lean | 150 ++++++++++------------- tests/lean/run/timeLocalDateTime.lean | 4 +- tests/lean/run/timeParse.lean | 26 ++-- tests/lean/run/timeSet.lean | 18 +-- tests/lean/run/timeTzifParse.lean | 8 +- 18 files changed, 203 insertions(+), 191 deletions(-) diff --git a/script/prepare-llvm-mingw.sh b/script/prepare-llvm-mingw.sh index 6f1dd937d30f..f0ca954664af 100644 --- a/script/prepare-llvm-mingw.sh +++ b/script/prepare-llvm-mingw.sh @@ -40,7 +40,6 @@ cp /clang64/lib/{crtbegin,crtend,crt2,dllcrt2}.o stage1/lib/ cp /clang64/lib/lib{m,bcrypt,mingw32,moldname,mingwex,msvcrt,pthread,advapi32,shell32,user32,kernel32,ucrtbase,psapi,iphlpapi,userenv,ws2_32,dbghelp,ole32,icu}.* /clang64/lib/libgmp.a /clang64/lib/libuv.a llvm/lib/lib{c++,c++abi,unwind}.a stage1/lib/ echo -n " -DLEAN_STANDALONE=ON" echo -n " -DCMAKE_C_COMPILER=$PWD/stage1/bin/clang.exe -DCMAKE_C_COMPILER_WORKS=1 -DCMAKE_CXX_COMPILER=$PWD/llvm/bin/clang++.exe -DCMAKE_CXX_COMPILER_WORKS=1 -DLEAN_CXX_STDLIB='-lc++ -lc++abi'" -echo -n " -DSTAGE0_CMAKE_C_COMPILER=clang -DSTAGE0_CMAKE_CXX_COMPILER=clang++" echo -n " -DLEAN_EXTRA_CXX_FLAGS='-idirafter \"${WINDOWS_SDK_DIR}/Include/10.0.22621.0/um\" --sysroot $PWD/llvm -idirafter /clang64/include/'" echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang.exe" echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -static-libgcc -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld'" diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 077ec66c069b..1b025f91700a 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -135,6 +135,10 @@ The supported formats include: - `yy`: Displays the year in a two-digit format, showing the last two digits (e.g., "04" for 2004). - `yyyy`: Displays the year in a four-digit format (e.g., "2004"). - `yyyy+`: Extended format for years with more than four digits. +- `u`: Represents the *year* (aligned with ISO chronology). + - `uu`: Two-digit year format, showing the last two digits (e.g., "04" for 2004). + - `uuuu`: Displays the year in a four-digit format (e.g., "2004" or "-1000"). + - `uuuu+`: Extended format for handling years with more than four digits (e.g., "12345" or "-12345"). Useful for historical dates far into the past or future! - `D`: Represents the day of the year. - `M`: Represents the month of the year, displayed as either a number or text. - `M`, `MM`: Displays the month as a number, with `MM` zero-padded (e.g., "7" for July, "07" for July with padding). @@ -210,9 +214,9 @@ The supported formats include: In order to help the user build dates easily, there are a lot of macros available for creating dates. The `.sssssssss` can be ommited in most cases. -- **`date("yyyy-MM-dd")`**: Defines a date in the `yyyy-MM-DD` format. +- **`date("uuuu-MM-dd")`**: Defines a date in the `uuuu-MM-DD` format. - **`time("HH:mm:ss.sssssssss")`**: Defines a time in the `HH:mm:ss.sssssssss` format, including fractional seconds. -- **`datetime("yyy-MM-ddTHH:mm:ss.sssssssss")`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss.sssssssss` format. +- **`datetime("uuuu-MM-ddTHH:mm:ss.sssssssss")`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss.sssssssss` format. - **`offset("+HH:mm")`**: Defines a timezone offset in the format `+HH:mm`. - **`timezone("NAME/ID ZZZ")`**: Defines a timezone with a name and an offset. - **`datespec("format")`**: Defines a date specification format at compile time using the provided format string. diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index d40ed3582525..62d17b2a3e20 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -17,19 +17,19 @@ set_option linter.all true /-- The ISO8601 format, which is always 24 or 27 characters long, used for representing date and time in -a standardized format. The format follows the pattern `yyyy-MM-dd'T'HH:mm:ssZ`. +a standardized format. The format follows the pattern `uuuu-MM-dd'T'HH:mm:ssZ`. -/ -def iso8601 : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm.ssZ") +def iso8601 : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm.ssZ") /-- -The americanDate format, which follows the pattern `MM-dd-yyyy`. +The americanDate format, which follows the pattern `MM-dd-uuuu`. -/ -def americanDate : GenericFormat .any := datespec("MM-dd-yyyy") +def americanDate : GenericFormat .any := datespec("MM-dd-uuuu") /-- -The europeanDate format, which follows the pattern `dd-MM-yyyy`. +The europeanDate format, which follows the pattern `dd-MM-uuuu`. -/ -def europeanDate : GenericFormat .any := datespec("dd-MM-yyyy") +def europeanDate : GenericFormat .any := datespec("dd-MM-uuuu") /-- The time12Hour format, which follows the pattern `hh:mm:ss aa` for representing time @@ -44,16 +44,16 @@ in a 24-hour clock format. def time24Hour : GenericFormat .any := datespec("HH:mm:ss") /-- -The DateTimeZone24Hour format, which follows the pattern `yyyy-MM-dd:HH:mm:ss.SSSSSSSSS` for +The DateTimeZone24Hour format, which follows the pattern `uuuu-MM-dd:HH:mm:ss.SSSSSSSSS` for representing date, time, and time zone. -/ -def dateTime24Hour : GenericFormat (.only .GMT) := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSS") +def dateTime24Hour : GenericFormat (.only .GMT) := datespec("uuuu-MM-dd:HH:mm:ss.SSSSSSSSS") /-- -The DateTimeWithZone format, which follows the pattern `yyyy-MM-dd:HH:mm:ss.SSSSSSSSSZZZ` +The DateTimeWithZone format, which follows the pattern `uuuu-MM-dd:HH:mm:ss.SSSSSSSSSZZZ` for representing date, time, and time zone. -/ -def dateTimeWithZone : GenericFormat .any := datespec("yyyy-MM-dd:HH:mm:ss.SSSSSSSSSZZZ") +def dateTimeWithZone : GenericFormat .any := datespec("uuuu-MM-dd:HH:mm:ss.SSSSSSSSSZZZ") /-- The leanTime24Hour format, which follows the pattern `HH:mm:ss.SSSSSSSSS` for representing time @@ -70,68 +70,68 @@ notation of dates. def leanTime24HourNoNanos : GenericFormat .any := datespec("HH:mm:ss") /-- -The leanDateTime24Hour format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSS` for +The leanDateTime24Hour format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSS` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTime24Hour : GenericFormat (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSS") +def leanDateTime24Hour : GenericFormat (.only .GMT) := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSS") /-- -The leanDateTime24HourNoNanos format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ss` for +The leanDateTime24HourNoNanos format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTime24HourNoNanos : GenericFormat (.only .GMT) := datespec("yyyy-MM-dd'T'HH:mm:ss") +def leanDateTime24HourNoNanos : GenericFormat (.only .GMT) := datespec("uuuu-MM-dd'T'HH:mm:ss") /-- -The DateTimeWithZone format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ` +The DateTimeWithZone format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithZone : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ") +def leanDateTimeWithZone : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ") /-- -The DateTimeWithZoneNoNanos format, which follows the pattern `yyyy-MM-dd'T'HH:mm:ssZZZZZ` +The DateTimeWithZoneNoNanos format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ssZZZZZ` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithZoneNoNanos : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ssZZZZZ") +def leanDateTimeWithZoneNoNanos : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ssZZZZZ") /-- -The Lean Date format, which follows the pattern `yyyy-MM-dd`. It uses the default value that can be parsed with the +The Lean Date format, which follows the pattern `uuuu-MM-dd`. It uses the default value that can be parsed with the notation of dates. -/ -def leanDate : GenericFormat .any := datespec("yyyy-MM-dd") +def leanDate : GenericFormat .any := datespec("uuuu-MM-dd") /-- -The SQLDate format, which follows the pattern `yyyy-MM-dd` and is commonly used +The SQLDate format, which follows the pattern `uuuu-MM-dd` and is commonly used in SQL databases to represent dates. -/ -def sqlDate : GenericFormat .any := datespec("yyyy-MM-dd") +def sqlDate : GenericFormat .any := datespec("uuuu-MM-dd") /-- -The LongDateFormat, which follows the pattern `EEEE, MMMM D, yyyy HH:mm:ss` for +The LongDateFormat, which follows the pattern `EEEE, MMMM D, uuuu HH:mm:ss` for representing a full date and time with the day of the week and month name. -/ -def longDateFormat : GenericFormat (.only .GMT) := datespec("EEEE, MMMM D, yyyy HH:mm:ss") +def longDateFormat : GenericFormat (.only .GMT) := datespec("EEEE, MMMM D, uuuu HH:mm:ss") /-- -The AscTime format, which follows the pattern `EEE MMM d HH:mm:ss yyyy`. This format +The AscTime format, which follows the pattern `EEE MMM d HH:mm:ss uuuu`. This format is often used in older systems for logging and time-stamping events. -/ -def ascTime : GenericFormat (.only .GMT) := datespec("EEE MMM d HH:mm:ss yyyy") +def ascTime : GenericFormat (.only .GMT) := datespec("EEE MMM d HH:mm:ss uuuu") /-- -The RFC822 format, which follows the pattern `eee, dd MMM yyyy HH:mm:ss ZZZ`. +The RFC822 format, which follows the pattern `eee, dd MMM uuuu HH:mm:ss ZZZ`. This format is used in email headers and HTTP headers. -/ -def rfc822 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") +def rfc822 : GenericFormat .any := datespec("eee, dd MMM uuuu HH:mm:ss ZZZ") /-- The RFC850 format, which follows the pattern `eee, dd-MMM-YY HH:mm:ss ZZZ`. This format is an older standard for representing date and time in headers. -/ -def rfc850 : GenericFormat .any := datespec("eee, dd-MM-yyyy HH:mm:ss ZZZ") +def rfc850 : GenericFormat .any := datespec("eee, dd-MM-uuuu HH:mm:ss ZZZ") end Formats @@ -141,7 +141,7 @@ namespace TimeZone Parses a string into a `TimeZone` object. The input string must be in the format `"VV ZZZZZ"`. -/ def fromTimeZone (input : String) : Except String TimeZone := do - let spec : GenericFormat .any := datespec("V ZZZZZ") + let spec : GenericFormat .any := datespec("VV ZZZZZ") spec.parseBuilder (fun id off => some (TimeZone.mk off id "Unknown" false)) input namespace Offset @@ -185,37 +185,37 @@ def format (date : PlainDate) (format : String) : String := | none => "invalid time" /-- -Parses a date string in the American format (`MM-dd-yyyy`) and returns a `PlainDate`. +Parses a date string in the American format (`MM-dd-uuuu`) and returns a `PlainDate`. -/ def fromAmericanDateString (input : String) : Except String PlainDate := do Formats.americanDate.parseBuilder (fun m d y => PlainDate.ofYearMonthDay? y m d) input /-- -Converts a date in the American format (`MM-dd-yyyy`) into a `String`. +Converts a date in the American format (`MM-dd-uuuu`) into a `String`. -/ def toAmericanDateString (input : PlainDate) : String := Formats.americanDate.formatBuilder input.month input.day input.year /-- -Parses a date string in the SQL format (`yyyy-MM-dd`) and returns a `PlainDate`. +Parses a date string in the SQL format (`uuuu-MM-dd`) and returns a `PlainDate`. -/ def fromSQLDateString (input : String) : Except String PlainDate := do Formats.sqlDate.parseBuilder PlainDate.ofYearMonthDay? input /-- -Converts a date in the SQL format (`yyyy-MM-dd`) into a `String`. +Converts a date in the SQL format (`uuuu-MM-dd`) into a `String`. -/ def toSQLDateString (input : PlainDate) : String := Formats.sqlDate.formatBuilder input.year input.month input.day /-- -Parses a date string in the Lean format (`yyyy-MM-dd`) and returns a `PlainDate`. +Parses a date string in the Lean format (`uuuu-MM-dd`) and returns a `PlainDate`. -/ def fromLeanDateString (input : String) : Except String PlainDate := do Formats.leanDate.parseBuilder PlainDate.ofYearMonthDay? input /-- -Converts a date in the Lean format (`yyyy-MM-dd`) into a `String`. +Converts a date in the Lean format (`uuuu-MM-dd`) into a `String`. -/ def toLeanDateString (input : PlainDate) : String := Formats.leanDate.formatBuilder input.year input.month input.day @@ -247,7 +247,7 @@ def format (time : PlainTime) (format : String) : String := | .ok res => let res := res.formatGeneric fun | .H _ => some time.hour - | .k _ => some (time.hour.add 1) + | .k _ => some (time.hour.shiftTo1BasedHour) | .m _ => some time.minute | .n _ => some time.nano | .s _ => some time.second @@ -430,7 +430,7 @@ def format (date : PlainDateTime) (format : String) : String := | .eorc _ => some date.weekday | .F _ => some date.alignedWeekOfMonth | .H _ => some date.hour - | .k _ => some (date.hour.add 1) + | .k _ => some date.hour.shiftTo1BasedHour | .m _ => some date.minute | .n _ => some date.nanosecond | .s _ => some date.time.second diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 98bd29a2ebc1..f50f89460524 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -463,7 +463,7 @@ private def parseOffsetO (constructor : OffsetO → Modifier) (p : String) : Par parseMod constructor OffsetO.classify p private def parseZoneId (p : String) : Parser Modifier := - if p.length = 1 then pure .V else fail s!"invalid quantity of characters for '{p.get 0}'" + if p.length = 2 then pure .V else fail s!"invalid quantity of characters for '{p.get 0}'" private def parseNumberText (constructor : (Number ⊕ Text) → Modifier) (p : String) : Parser Modifier := parseMod constructor classifyNumberText p @@ -599,21 +599,21 @@ private def leftPad (n : Nat) (a : Char) (s : String) : String := private def rightPad (n : Nat) (a : Char) (s : String) : String := s ++ "".pushn a (n - s.length) -private def truncate (size : Nat) (n : Int) : String := +private def truncate (size : Nat) (n : Int) (cut : Bool := false) : String := let (sign, n) := if n < 0 then ("-", -n) else ("", n) let numStr := toString n if numStr.length > size then - sign ++ numStr.drop (numStr.length - size) + sign ++ if cut then numStr.drop (numStr.length - size) else numStr else sign ++ leftPad size '0' numStr -private def rightTruncate (size : Nat) (n : Int) : String := +private def rightTruncate (size : Nat) (n : Int) (cut : Bool := false) : String := let (sign, n) := if n < 0 then ("-", -n) else ("", n) let numStr := toString n if numStr.length > size then - sign ++ numStr.drop (numStr.length - size) + sign ++ if cut then numStr.take size else numStr else sign ++ rightPad size '0' numStr @@ -735,7 +735,7 @@ private def toSigned (data : Int) : String := if data < 0 then toString data else "+" ++ toString data private def toIsoString (offset : Offset) (withMinutes : Bool) (withSeconds : Bool) (colon : Bool) : String := - let (sign, time) := if offset.second.val > 0 then ("+", offset.second) else ("-", -offset.second) + let (sign, time) := if offset.second.val ≥ 0 then ("+", offset.second) else ("-", -offset.second) let time := PlainTime.ofSeconds time let pad := leftPad 2 '0' ∘ toString @@ -842,7 +842,7 @@ private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : Strin | .S format => match format with | .nano => truncate 9 data.val - | .truncated n => rightTruncate n data.val + | .truncated n => rightTruncate n data.val (cut := true) | .A format => truncate format.padding data.val | .n format => @@ -910,7 +910,7 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .a _ => HourMarker.ofOrdinal date.hour | .h _ => HourMarker.toRelative date.hour |>.fst | .K _ => date.hour.emod 12 (by decide) - | .k _ => date.hour.add 1 + | .k _ => date.hour.shiftTo1BasedHour | .H _ => date.hour | .m _ => date.minute | .s _ => date.date.get.time.second diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 3de289e9a26d..5be3976c7dbd 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -51,9 +51,16 @@ namespace Ordinal Converts an `Ordinal` into a relative hour in the range of 1 to 12. -/ def toRelative (ordinal : Ordinal) : Bounded.LE 1 12 := - (ordinal.add 11).emod 12 (by decide) |>.add 1 + ordinal.emod 12 (by decide) |>.add 1 /-- +Converts an Ordinal into a 1-based hour representation within the range of 1 to 24. +-/ +def shiftTo1BasedHour (ordinal : Ordinal) : Bounded.LE 1 24 := + if h : ordinal.val < 1 + then Internal.Bounded.LE.ofNatWrapping 24 (by decide) + else ordinal.truncateBottom (Int.not_lt.mp h) |>.expandTop (by decide) +/-- Creates an `Ordinal` from a natural number, ensuring the value is within the valid bounds for hours. -/ @[inline] diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index c3ab8cc17ce2..556fa3728967 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -31,9 +31,9 @@ Get the current time. def now : IO PlainDateTime := do let tm ← Timestamp.now let rules ← Database.defaultGetLocalZoneRulesAt - let transition ← rules.findTransitionForTimestamp tm |>.elim (throw <| IO.userError "cannot find timezone") pure + let ltt := rules.findLocalTimeTypeForTimestamp tm - return PlainDateTime.ofTimestamp tm transition.localTimeType.getTimeZone + return PlainDateTime.ofTimestamp tm ltt.getTimeZone end PlainDateTime @@ -51,7 +51,7 @@ Converts a `DateTime` to a `PlainDate` -/ @[inline] def toPlainDate (dt : DateTime tz) : PlainDate := - Timestamp.toPlainDateAssumingUTC dt.toUTCTimestamp + Timestamp.toPlainDateAssumingUTC dt.toTimestamp /-- Converts a `DateTime` to a `PlainTime` diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 69a5d9d683f4..27388196f8cd 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -16,7 +16,7 @@ set_option linter.all true /-- A timezone database from which we can read the `ZoneRules` of some area by it's id. -/ -class Database (α : Type) where +protected class Database (α : Type) where /-- Retrieves the timezone information (`ZoneRules`) for a given area at a specific point in time. @@ -87,9 +87,15 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := for i in [0:tz.transitionTimes.size] do if let some result := convertTransition times i tz then transitions := transitions.push result - else .error s!"cannot convert transtiion {i} of the file" + else .error s!"cannot convert transition {i} of the file" - .ok { transitions, localTimes := times } + let first ← + if let some res := transitions.get? 0 + then .ok res + else .error "empty transitions" + + + .ok { transitions, initialLocalTimeType := first.localTimeType } /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 3b2ab241db72..0935ae522962 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -93,6 +93,6 @@ def readLeapsFromDisk (path : System.FilePath) : IO (Array LeapSecond) := do then return res else throw (IO.userError "cannot read the id of the path.") -instance : Database TZdb where +instance : Std.Time.Database TZdb where getLocalZoneRulesAt db := localRules db.localPath getZoneRulesAt db id := readRulesFromDisk db.zonesPath id diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index 436a2c31f430..11de9493de50 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -34,7 +34,7 @@ Retrieves the timezone rules, including all transitions, for a given timezone id -/ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do let mut start := -2147483648 - let mut transitions := #[] + let mut transitions : Array TimeZone.Transition := #[] while true do let result ← Windows.getNextTransition id start @@ -51,7 +51,12 @@ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do else break - return { transitions, localTimes := #[] } + let initialLocalTimeType ← do + if let some res := transitions.get? 0 + then pure res.localTimeType + else throw (IO.userError "cannot find first transition in zone rules") + + return { transitions, initialLocalTimeType } end Windows @@ -69,6 +74,6 @@ Returns a default `WindowsDb` instance. @[inline] def default : WindowsDb := {} -instance : Database WindowsDb where +instance : Std.Time.Database WindowsDb where getZoneRulesAt _ id := Windows.getZoneRules id getLocalZoneRulesAt _ := Windows.getZoneRules =<< Windows.getLocalTimeZoneIdentifierAt (-2147483648) diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 952953e02e41..e4fa3f6901cc 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -47,10 +47,10 @@ def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := DateTime.mk tm (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds).toPlainDateTimeAssumingUTC) /-- -Creates a UTC `Timestamp` out of a `DateTime`. +Creates a `Timestamp` out of a `DateTime`. -/ @[inline] -def toUTCTimestamp (date : DateTime tz) : Timestamp := +def toTimestamp (date : DateTime tz) : Timestamp := date.timestamp /-- @@ -450,7 +450,7 @@ instance : HSub (DateTime tz) (Nanosecond.Offset) (DateTime tz) where hSub := subNanoseconds instance : HSub (DateTime tz) (DateTime tz₁) Duration where - hSub x y := x.toUTCTimestamp - y.toUTCTimestamp + hSub x y := x.toTimestamp - y.toTimestamp end DateTime end Time diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 7ee15bd61baf..49e02468aefb 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -112,9 +112,9 @@ Represents the rules for a time zone. structure ZoneRules where /-- - The array of local time types for the time zone. + The first `LocalTimeType` used as a fallback when no matching transition is found. -/ - localTimes : Array LocalTimeType + initialLocalTimeType : LocalTimeType /-- The array of transitions for the time zone. @@ -163,24 +163,31 @@ end Transition namespace ZoneRules /-- -Finds the transition corresponding to a given timestamp in `ZoneRules`. +Finds the `LocalTimeType` corresponding to a given `Timestamp` in `ZoneRules`. If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. +If no transition is found, it falls back to `initialLocalTimeType`. -/ @[inline] -def findTransitionForTimestamp (zr : ZoneRules) (timestamp : Timestamp) : Option Transition := +def findLocalTimeTypeForTimestamp (zr : ZoneRules) (timestamp : Timestamp) : LocalTimeType := Transition.findTransitionForTimestamp zr.transitions timestamp + |>.map (·.localTimeType) + |>.getD zr.initialLocalTimeType /-- Find the current `TimeZone` out of a `Transition` in a `ZoneRules` -/ @[inline] -def timezoneAt (zr : ZoneRules) (tm : Timestamp) : Except String TimeZone := +def timezoneAt (zr : ZoneRules) (tm : Timestamp) : TimeZone := Transition.timezoneAt zr.transitions tm + |>.toOption + |>.getD (zr.initialLocalTimeType |>.getTimeZone) /-- Creates `ZoneRules` for the given `TimeZone`. -/ +@[inline] def ofTimeZone (tz : TimeZone) : ZoneRules := - ZoneRules.mk #[] #[Transition.mk tz.toSeconds (LocalTimeType.mk tz.offset tz.isDST tz.name .wall .local tz.abbreviation)] + let ltt := LocalTimeType.mk tz.offset tz.isDST tz.name .wall .local tz.abbreviation + ZoneRules.mk ltt #[] end ZoneRules diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index ecc08f77c8a0..6283692b8861 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -45,28 +45,28 @@ namespace ZonedDateTime open DateTime /-- -Creates a new `ZonedDateTime` out of a `Timestamp` and a `TimeZone`. +Creates a new `ZonedDateTime` out of a `Timestamp` and a `ZoneRules`. -/ @[inline] def ofTimestamp (tm : Timestamp) (rules : TimeZone.ZoneRules) : ZonedDateTime := - let tz := TimeZone.Transition.timezoneAt rules.transitions tm |>.toOption |>.getD TimeZone.UTC + let tz := rules.timezoneAt tm ZonedDateTime.mk (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds) |>.toPlainDateTimeAssumingUTC) tm rules tz /-- -Creates a new `ZonedDateTime` out of a `PlainDateTime` and a `TimeZone`. +Creates a new `ZonedDateTime` out of a `PlainDateTime` and a `ZoneRules`. -/ @[inline] def ofPlainDateTime (pdt : PlainDateTime) (zr : TimeZone.ZoneRules) : ZonedDateTime := let tm := pdt.toTimestampAssumingUTC - let tz := zr.findTransitionForTimestamp tm |>.map (·.localTimeType.getTimeZone) |>.getD TimeZone.UTC + let tz := zr.findLocalTimeTypeForTimestamp tm |>.getTimeZone let tm := tm.subSeconds tz.toSeconds ZonedDateTime.mk (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds) |>.toPlainDateTimeAssumingUTC) tm zr tz /-- -Creates a new UTC `Timestamp` out of a `ZonedDateTime`. +Creates a new `Timestamp` out of a `ZonedDateTime`. -/ @[inline] -def toUTCTimestamp (date : ZonedDateTime) : Timestamp := +def toTimestamp (date : ZonedDateTime) : Timestamp := date.timestamp /-- @@ -74,7 +74,7 @@ Changes the `ZoleRules` to a new one. -/ @[inline] def convertZoneRules (date : ZonedDateTime) (tz₁ : TimeZone.ZoneRules) : ZonedDateTime := - ofTimestamp date.toUTCTimestamp tz₁ + ofTimestamp date.toTimestamp tz₁ /-- Creates a new `ZonedDateTime` out of a `PlainDateTime` @@ -444,7 +444,7 @@ instance : HSub ZonedDateTime Nanosecond.Offset ZonedDateTime where hSub := subNanoseconds instance : HSub ZonedDateTime ZonedDateTime Duration where - hSub x y := x.toUTCTimestamp - y.toUTCTimestamp + hSub x y := x.toTimestamp - y.toTimestamp end ZonedDateTime end Time diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index baf83699e409..6a823204ab28 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -648,8 +648,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } -/* Std.Time.Database.Windows.getNextTransition : String -> @&Int -> IO (Option (Int × TimeZone)) */ -extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone_str, obj_arg tm_obj, obj_arg /* w */) { +/* Std.Time.Database.Windows.getNextTransition : @&String -> @&Int -> IO (Option (Int × TimeZone)) */ +extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(b_obj_arg timezone_str, b_obj_arg tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; const char* dst_name_id = lean_string_cstr(timezone_str); @@ -751,7 +751,7 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(obj_arg timezone #endif } -/* SStd.Time.Database.Windows.getLocalTimeZoneIdentifierAt : UInt64 → IO String */ +/* Std.Time.Database.Windows.getLocalTimeZoneIdentifierAt : Int → IO String */ extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(obj_arg tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index bf0ed501a9a2..49a967fc461b 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -1,17 +1,17 @@ import Std.Time open Std.Time -def ISO8601UTC : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSZ") -def RFC1123 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") -def ShortDate : GenericFormat .any := datespec("MM/dd/yyyy") -def LongDate : GenericFormat .any := datespec("MMMM D, yyyy") -def ShortDateTime : GenericFormat .any := datespec("MM/dd/yyyy HH:mm:ss") -def LongDateTime : GenericFormat .any := datespec("MMMM D, yyyy h:mm aa") +def ISO8601UTC : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZ") +def RFC1123 : GenericFormat .any := datespec("eee, dd MMM uuuu HH:mm:ss ZZZ") +def ShortDate : GenericFormat .any := datespec("MM/dd/uuuu") +def LongDate : GenericFormat .any := datespec("MMMM D, uuuu") +def ShortDateTime : GenericFormat .any := datespec("MM/dd/uuuu HH:mm:ss") +def LongDateTime : GenericFormat .any := datespec("MMMM D, uuuu h:mm aa") def Time24Hour : GenericFormat .any := datespec("HH:mm:ss") def Time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") -def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") -def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM yyyy HH:mm") -def EraDate : GenericFormat .any := datespec("MM D, yyyy G") +def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, uuuu HH:mm:ss ZZZ") +def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM uuuu HH:mm") +def EraDate : GenericFormat .any := datespec("MM D, uuuu G") -- Dates @@ -29,7 +29,7 @@ info: "Monday, June 16, 2014 03:03:03 -0300" #guard_msgs in #eval FullDayTimeZone.format date₁.toDateTime -def tm := date₁.toUTCTimestamp +def tm := date₁.toTimestamp def date₂ := DateTime.ofTimestamp tm brTZ /-- @@ -38,7 +38,7 @@ info: "Monday, June 16, 2014 03:03:03 -0300" #guard_msgs in #eval FullDayTimeZone.format date₂ -def tm₃ := date₁.toUTCTimestamp +def tm₃ := date₁.toTimestamp def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- @@ -87,19 +87,19 @@ info: "Thursday, August 15, 2024 13:28:12 -0300" #eval FullDayTimeZone.format (dateJP.convertTimeZone brTZ) /-- -info: "Thursday, August 15, 2024 16:28:12 -0000" +info: "Thursday, August 15, 2024 16:28:12 +0000" -/ #guard_msgs in #eval FullDayTimeZone.format dateUTC /-- -info: "Thursday, August 15, 2024 16:28:12 -0000" +info: "Thursday, August 15, 2024 16:28:12 +0000" -/ #guard_msgs in #eval FullDayTimeZone.format (dateBR.convertTimeZone .UTC) /-- -info: "Thursday, August 15, 2024 16:28:12 -0000" +info: "Thursday, August 15, 2024 16:28:12 +0000" -/ #guard_msgs in #eval FullDayTimeZone.format (dateJP.convertTimeZone .UTC) @@ -125,7 +125,7 @@ info: "Friday, August 16, 2024 01:28:12 +0900" /-- TM: 1723730627 GMT: Thursday, 15 August 2024 14:03:47 -Your time zone: 15 Aguust 2024 11:03:47 GMT-03:00 +Your time zone: 15 August 2024 11:03:47 GMT-03:00 -/ def localTm : Second.Offset := 1723730627 @@ -206,31 +206,31 @@ info: "0053-06-19" #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-700000⟩) .UTC) /-- -info: "0003-09-16" +info: "-0002-09-16" -/ #guard_msgs in #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-720000⟩) .UTC) /-- -info: "0085-07-28" +info: "-0084-07-28" -/ #guard_msgs in #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-750000⟩) .UTC) /-- -info: "0222-09-04" +info: "-0221-09-04" -/ #guard_msgs in #eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) /-- -info: date("0222-09-04") +info: date("-0221-09-04") -/ #guard_msgs in #eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ /-- -info: date("0222-09-04") +info: date("-0221-09-04") -/ #guard_msgs in #eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ @@ -263,7 +263,7 @@ info: zoned("2002-07-14T14:13:12.000000000+09:00") info: "2002-07-14" -/ #guard_msgs in -#eval zoned("2002-07-14T14:13:12+09:00").format "yyyy-MM-dd" +#eval zoned("2002-07-14T14:13:12+09:00").format "uuuu-MM-dd" /-- info: "14-13-12" @@ -300,19 +300,13 @@ info: "02 2002 000002002" #eval zoned₄.format "uu uuuu uuuuuuuuu" /-- -info: "5 95 195" +info: "195 195 195" -/ #guard_msgs in #eval zoned₄.format "D DD DDD" /-- -info: "7 07 Jul J" --/ -#guard_msgs in -#eval zoned₄.format "M MM MMM MMMMM" - -/-- -info: "4 14 014 0014 00014" +info: "14 14 014 0014 00014" -/ #guard_msgs in #eval zoned₄.format "d dd ddd dddd ddddd" @@ -323,11 +317,6 @@ info: "7 07 Jul July J" #guard_msgs in #eval zoned₄.format "M MM MMM MMMM MMMMM" -/-- -info: "4 14 0014 0014" --/#guard_msgs in -#eval zoned₄.format "d dd dddd dddd" - /-- info: "3 03 3rd quarter 3" -/ @@ -335,7 +324,7 @@ info: "3 03 3rd quarter 3" #eval zoned₄.format "Q QQ QQQQ QQQQQ" /-- -info: "8 28 028 0028" +info: "28 28 028 0028" -/ #guard_msgs in #eval zoned₄.format "w ww www wwww" @@ -365,67 +354,61 @@ info: "2 02 002 0002" #eval zoned₄.format "F FF FFF FFFF" /-- -info: "1 11 011 0011 0011" +info: "11 11 011 0011 0011" -/ #guard_msgs in #eval zoned₄.format "h hh hhh hhhh hhhh" /-- -info: "1 11 011 0011 000011" +info: "11 11 011 0011 000011" -/ #guard_msgs in #eval zoned₄.format "K KK KKK KKKK KKKKKK" /-- -info: "4 24 024 0024 000024" +info: "24 24 024 0024 000024" -/ #guard_msgs in #eval zoned₄.format "k kk kkk kkkk kkkkkk" /-- -info: "3 23 023 0023 00023" +info: "23 23 023 0023 00023" -/ #guard_msgs in #eval zoned₄.format "H HH HHH HHHH HHHHH" /-- -info: "3 13 013 0013 00013" +info: "13 13 013 0013 00013" -/ #guard_msgs in #eval zoned₄.format "m mm mmm mmmm mmmmm" /-- -info: "2 12 012 0012 00012" +info: "12 12 012 0012 00012" -/ #guard_msgs in #eval zoned₄.format "s ss sss ssss sssss" /-- -info: "9 79 679 4679 54679" +info: "3 32 324 3243 32435" -/#guard_msgs in #eval zoned₄.format "S SS SSS SSSS SSSSS" /-- -info: "9 79 679 4679 324354679" --/ -#guard_msgs in -#eval zoned₄.format "S SS SSS SSSS SSSSSSSSS" - -/-- -info: "4 24 324 2324 083592324" +info: "83592324 83592324 83592324 83592324 083592324" -/ #guard_msgs in #eval zoned₄.format "A AA AAA AAAA AAAAAAAAA" /-- -info: "9 79 679 4679 324354679" +info: "324354679 324354679 324354679 324354679 324354679" -/ #guard_msgs in #eval zoned₄.format "n nn nnn nnnn nnnnnnnnn" /-- -info: "9 79 679 4679 324354679" +info: "83592324354679 83592324354679 83592324354679 83592324354679 83592324354679" -/ #guard_msgs in #eval zoned₄.format "N NN NNN NNNN NNNNNNNNN" @@ -434,7 +417,7 @@ info: "9 79 679 4679 324354679" info: "Unknown" -/ #guard_msgs in -#eval zoned₄.format "V" +#eval zoned₄.format "VV" /-- info: "Unknown Unknown Unknown Unknown" @@ -485,7 +468,7 @@ info: "02 2002 000002002" #eval datetime₄.format "uu uuuu uuuuuuuuu" /-- -info: "5 95 195" +info: "195 195 195" -/ #guard_msgs in #eval datetime₄.format "D DD DDD" @@ -497,7 +480,7 @@ info: "7 07 Jul J" #eval datetime₄.format "M MM MMM MMMMM" /-- -info: "4 14 014 0014 00014" +info: "14 14 014 0014 00014" -/ #guard_msgs in #eval datetime₄.format "d dd ddd dddd ddddd" @@ -509,7 +492,7 @@ info: "7 07 Jul July J" #eval datetime₄.format "M MM MMM MMMM MMMMM" /-- -info: "4 14 0014 0014" +info: "14 14 0014 0014" -/#guard_msgs in #eval datetime₄.format "d dd dddd dddd" @@ -520,7 +503,7 @@ info: "3 03 3rd quarter 3" #eval datetime₄.format "Q QQ QQQQ QQQQQ" /-- -info: "8 28 028 0028" +info: "28 28 028 0028" -/ #guard_msgs in #eval datetime₄.format "w ww www wwww" @@ -550,133 +533,134 @@ info: "2 02 002 0002" #eval datetime₄.format "F FF FFF FFFF" /-- -info: "1 11 011 0011 0011" +info: "11 11 011 0011 0011" -/ #guard_msgs in #eval datetime₄.format "h hh hhh hhhh hhhh" /-- -info: "1 11 011 0011 000011" +info: "11 11 011 0011 000011" -/ #guard_msgs in #eval datetime₄.format "K KK KKK KKKK KKKKKK" /-- -info: "4 24 024 0024 000024" +info: "23 23 023 0023 000023" -/ #guard_msgs in #eval datetime₄.format "k kk kkk kkkk kkkkkk" /-- -info: "3 23 023 0023 00023" +info: "23 23 023 0023 00023" -/ #guard_msgs in #eval datetime₄.format "H HH HHH HHHH HHHHH" /-- -info: "3 13 013 0013 00013" +info: "13 13 013 0013 00013" -/ #guard_msgs in #eval datetime₄.format "m mm mmm mmmm mmmmm" /-- -info: "2 12 012 0012 00012" +info: "12 12 012 0012 00012" -/ #guard_msgs in #eval datetime₄.format "s ss sss ssss sssss" /-- -info: "9 79 679 4679 54679" +info: "3 32 324 3243 32435" -/#guard_msgs in #eval datetime₄.format "S SS SSS SSSS SSSSS" /-- -info: "9 79 679 4679 324354679" +info: "3 32 324 3243 324354679" -/ #guard_msgs in #eval datetime₄.format "S SS SSS SSSS SSSSSSSSS" /-- -info: "4 24 324 2324 083592324" +info: "83592324 83592324 83592324 83592324 083592324" -/ #guard_msgs in #eval datetime₄.format "A AA AAA AAAA AAAAAAAAA" /-- -info: "9 79 679 4679 324354679" +info: "324354679 324354679 324354679 324354679 324354679" -/ #guard_msgs in #eval datetime₄.format "n nn nnn nnnn nnnnnnnnn" /-- -info: "9 79 679 4679 324354679" +info: "83592324354679 83592324354679 83592324354679 83592324354679 83592324354679" -/ #guard_msgs in #eval datetime₄.format "N NN NNN NNNN NNNNNNNNN" /-- -info: "1 11 011 0011 0011" +info: "11 11 011 0011 0011" -/ #guard_msgs in #eval time₄.format "h hh hhh hhhh hhhh" /-- -info: "1 11 011 0011 000011" +info: "11 11 011 0011 000011" -/ #guard_msgs in #eval time₄.format "K KK KKK KKKK KKKKKK" /-- -info: "4 24 024 0024 000024" +info: "24 24 024 0024 000024" + -/ #guard_msgs in #eval time₄.format "k kk kkk kkkk kkkkkk" /-- -info: "3 23 023 0023 00023" +info: "23 23 023 0023 00023" -/ #guard_msgs in #eval time₄.format "H HH HHH HHHH HHHHH" /-- -info: "3 13 013 0013 00013" +info: "13 13 013 0013 00013" -/ #guard_msgs in #eval time₄.format "m mm mmm mmmm mmmmm" /-- -info: "2 12 012 0012 00012" +info: "12 12 012 0012 00012" -/ #guard_msgs in #eval time₄.format "s ss sss ssss sssss" /-- -info: "9 79 679 4679 54679" +info: "3 32 324 3243 32435" -/#guard_msgs in #eval time₄.format "S SS SSS SSSS SSSSS" /-- -info: "9 79 679 4679 324354679" +info: "3 32 324 3243 324354679" -/ #guard_msgs in #eval time₄.format "S SS SSS SSSS SSSSSSSSS" /-- -info: "4 24 324 2324 083592324" +info: "83592324 83592324 83592324 83592324 083592324" -/ #guard_msgs in #eval time₄.format "A AA AAA AAAA AAAAAAAAA" /-- -info: "9 79 679 4679 324354679" +info: "324354679 324354679 324354679 324354679 324354679" -/ #guard_msgs in #eval time₄.format "n nn nnn nnnn nnnnnnnnn" /-- -info: "9 79 679 4679 324354679" +info: "83592324354679 83592324354679 83592324354679 83592324354679 83592324354679" -/ #guard_msgs in #eval time₄.format "N NN NNN NNNN NNNNNNNNN" @@ -700,7 +684,7 @@ info: "02 2002 000002002" #eval date₄.format "uu uuuu uuuuuuuuu" /-- -info: "5 95 195" +info: "195 195 195" -/ #guard_msgs in #eval date₄.format "D DD DDD" @@ -712,7 +696,7 @@ info: "7 07 Jul J" #eval date₄.format "M MM MMM MMMMM" /-- -info: "4 14 014 0014 00014" +info: "14 14 014 0014 00014" -/ #guard_msgs in #eval date₄.format "d dd ddd dddd ddddd" @@ -724,7 +708,7 @@ info: "7 07 Jul July J" #eval date₄.format "M MM MMM MMMM MMMMM" /-- -info: "4 14 0014 0014" +info: "14 14 0014 0014" -/#guard_msgs in #eval date₄.format "d dd dddd dddd" @@ -735,7 +719,7 @@ info: "3 03 3rd quarter 3" #eval date₄.format "Q QQ QQQQ QQQQQ" /-- -info: "8 28 028 0028" +info: "28 28 028 0028" -/ #guard_msgs in #eval date₄.format "w ww www wwww" @@ -777,7 +761,7 @@ info: "2002 2002 CE" #eval datetime₄.format "uuuu yyyy G" /-- -info: 3 +info: 1 -/ #guard_msgs in #eval diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 60bf46e59313..3ab2d18eb356 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -2,8 +2,8 @@ import Std.Time import Init open Std.Time -def ShortDateTime : GenericFormat .any := datespec("dd/MM/yyyy HH:mm:ss") -def ShortDate : GenericFormat .any := datespec("dd/MM/yyyy") +def ShortDateTime : GenericFormat .any := datespec("dd/MM/uuuu HH:mm:ss") +def ShortDate : GenericFormat .any := datespec("dd/MM/uuuu") def format (PlainDate : PlainDateTime) : String := ShortDateTime.formatBuilder PlainDate.day PlainDate.month PlainDate.year PlainDate.time.hour PlainDate.minute PlainDate.time.second def format₂ (PlainDate : PlainDate) : String := ShortDate.formatBuilder PlainDate.day PlainDate.month PlainDate.year diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index d242f3c7d9a3..6d1b7884084e 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -1,18 +1,18 @@ import Std.Time open Std.Time -def ISO8601UTC : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") -def RFC1123 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") -def ShortDate : GenericFormat .any := datespec("MM/dd/yyyy") -def LongDate : GenericFormat .any := datespec("MMMM D, yyyy") -def ShortDateTime : GenericFormat .any := datespec("MM/dd/yyyy HH:mm:ss") -def LongDateTime : GenericFormat .any := datespec("MMMM dd, yyyy hh:mm aa") +def ISO8601UTC : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") +def RFC1123 : GenericFormat .any := datespec("eee, dd MMM uuuu HH:mm:ss ZZZ") +def ShortDate : GenericFormat .any := datespec("MM/dd/uuuu") +def LongDate : GenericFormat .any := datespec("MMMM D, uuuu") +def ShortDateTime : GenericFormat .any := datespec("MM/dd/uuuu HH:mm:ss") +def LongDateTime : GenericFormat .any := datespec("MMMM dd, uuuu hh:mm aa") def Time24Hour : GenericFormat .any := datespec("HH:mm:ss") def Time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") -def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") -def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM yyyy HH:mm") +def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, uuuu HH:mm:ss ZZZ") +def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM uuuu HH:mm") -def Full12HourWrong : GenericFormat .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") +def Full12HourWrong : GenericFormat .any := datespec("MM/dd/uuuu hh:mm:ss aa XXX") -- Dates @@ -33,7 +33,7 @@ info: "2014-06-16T03:03:03.000000100-03:00" let t : ZonedDateTime := ISO8601UTC.parse! "2014-06-16T03:03:03.000000100-03:00" ISO8601UTC.format t.toDateTime -def tm := date₁.toUTCTimestamp +def tm := date₁.toTimestamp def date₂ := DateTime.ofTimestamp tm brTZ /-- @@ -44,7 +44,7 @@ info: "2014-06-16T03:03:03.000000000-03:00" let t : ZonedDateTime := RFC1123.parse! "Mon, 16 Jun 2014 03:03:03 -0300" ISO8601UTC.format t.toDateTime -def tm₃ := date₁.toUTCTimestamp +def tm₃ := date₁.toTimestamp def date₃ := DateTime.ofTimestamp tm₃ brTZ /-- @@ -85,7 +85,7 @@ info: "2024-08-16T01:28:00.000000000Z" ISO8601UTC.format t.toDateTime /-- -info: "0001-12-30T22:28:12.000000000+09:00" +info: "0000-12-30T22:28:12.000000000+09:00" -/ #guard_msgs in #eval @@ -93,7 +93,7 @@ info: "0001-12-30T22:28:12.000000000+09:00" ISO8601UTC.format (t.toDateTime.convertTimeZone jpTZ) /-- -info: "0001-12-29T21:28:12.000000000-03:00" +info: "0000-12-29T21:28:12.000000000-03:00" -/ #guard_msgs in #eval diff --git a/tests/lean/run/timeSet.lean b/tests/lean/run/timeSet.lean index 60b6748e5487..4f179d506132 100644 --- a/tests/lean/run/timeSet.lean +++ b/tests/lean/run/timeSet.lean @@ -1,18 +1,18 @@ import Std.Time open Std.Time -def ISO8601UTC : GenericFormat .any := datespec("yyyy-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") -def RFC1123 : GenericFormat .any := datespec("eee, dd MMM yyyy HH:mm:ss ZZZ") -def ShortDate : GenericFormat .any := datespec("MM/dd/yyyy") -def LongDate : GenericFormat .any := datespec("MMMM D, yyyy") -def ShortDateTime : GenericFormat .any := datespec("MM/dd/yyyy HH:mm:ss") -def LongDateTime : GenericFormat .any := datespec("MMMM dd, yyyy hh:mm aa") +def ISO8601UTC : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSXXX") +def RFC1123 : GenericFormat .any := datespec("eee, dd MMM uuuu HH:mm:ss ZZZ") +def ShortDate : GenericFormat .any := datespec("MM/dd/uuuu") +def LongDate : GenericFormat .any := datespec("MMMM D, uuuu") +def ShortDateTime : GenericFormat .any := datespec("MM/dd/uuuu HH:mm:ss") +def LongDateTime : GenericFormat .any := datespec("MMMM dd, uuuu hh:mm aa") def Time24Hour : GenericFormat .any := datespec("HH:mm:ss") def Time12Hour : GenericFormat .any := datespec("hh:mm:ss aa") -def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, yyyy HH:mm:ss ZZZ") -def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM yyyy HH:mm") +def FullDayTimeZone : GenericFormat .any := datespec("EEEE, MMMM dd, uuuu HH:mm:ss ZZZ") +def CustomDayTime : GenericFormat .any := datespec("EEE dd MMM uuuu HH:mm") -def Full12HourWrong : GenericFormat .any := datespec("MM/dd/yyyy hh:mm:ss aa XXX") +def Full12HourWrong : GenericFormat .any := datespec("MM/dd/uuuu hh:mm:ss aa XXX") -- Dates diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index 10ac2ca911d6..a9b7765219a2 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -59,13 +59,13 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: some (zoned("1969-12-30T21:00:00.000000000-03:00")) +info: zoned("1969-12-30T21:00:00.000000000-03:00") -/ #guard_msgs in -#eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules +#eval ZonedDateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch 0) rules /-- -info: some (zoned("2012-12-10T00:35:47.000000000-02:00")) +info: zoned("2012-12-10T00:35:47.000000000-02:00") -/ #guard_msgs in -#eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules +#eval ZonedDateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules From 6086487f5aa7a6cc79c94511ee631f51d55bc331 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 28 Oct 2024 12:54:38 -0300 Subject: [PATCH 082/118] fix: comments and remove leap file --- src/Std/Time.lean | 14 ++-- src/Std/Time/Date/PlainDate.lean | 2 +- src/Std/Time/Format.lean | 10 +-- src/Std/Time/Format/Basic.lean | 22 ++--- src/Std/Time/Time/Unit/Hour.lean | 2 +- src/Std/Time/Zoned.lean | 9 ++ src/Std/Time/Zoned/Database/Basic.lean | 6 +- src/Std/Time/Zoned/Database/Leap.lean | 110 ------------------------- src/Std/Time/Zoned/Database/TZdb.lean | 10 --- src/Std/Time/Zoned/Offset.lean | 2 +- src/Std/Time/Zoned/ZonedDateTime.lean | 14 ++++ tests/lean/run/timeFormats.lean | 37 +++++++-- 12 files changed, 83 insertions(+), 155 deletions(-) delete mode 100644 src/Std/Time/Zoned/Database/Leap.lean diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 1b025f91700a..dafa2a6b8c5d 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -113,7 +113,11 @@ to the UNIX epoch. This type should be used when sending or receiving timestamps Combines date, time and time zones. - **`DateTime`**: Represents both date and time but with a time zone in the type constructor. -- **`ZonedDateTime`**: An existential version of the `DateTime`. +- **`ZonedDateTime`**: Is a way to represent date and time that includes `ZoneRules`, which consider +Daylight Saving Time (DST). This means it can handle local time changes throughout the year better +than a regular `DateTime`. If you want to use a specific time zone without worrying about DST, you can +use the `ofTimestampWithZone` function, which gives you a `ZonedDateTime` based only on that time zone, +without considering the zone rules, otherwise you can use `ofTimestamp` or `ofTimestampWithIdentifier`. ## Duration Represents spans of time and the difference between two points in time. @@ -135,7 +139,7 @@ The supported formats include: - `yy`: Displays the year in a two-digit format, showing the last two digits (e.g., "04" for 2004). - `yyyy`: Displays the year in a four-digit format (e.g., "2004"). - `yyyy+`: Extended format for years with more than four digits. -- `u`: Represents the *year* (aligned with ISO chronology). +- `u`: Represents the year. - `uu`: Two-digit year format, showing the last two digits (e.g., "04" for 2004). - `uuuu`: Displays the year in a four-digit format (e.g., "2004" or "-1000"). - `uuuu+`: Extended format for handling years with more than four digits (e.g., "12345" or "-12345"). Useful for historical dates far into the past or future! @@ -151,8 +155,8 @@ The supported formats include: - `QQQ` (short): Displays the quarter as an abbreviated text (e.g., "Q3"). - `QQQQ` (full): Displays the full quarter text (e.g., "3rd quarter"). - `QQQQQ` (narrow): Displays the quarter as a short number (e.g., "3"). -- `w`: Represents the week of the week-based year (e.g., "27"). -- `W`: Represents the week of the month (e.g., "4"). +- `w`: Represents the week of the week-based year, each week starts on Monday (e.g., "27"). +- `W`: Represents the week of the month, each week starts on Monday (e.g., "4"). - `E`: Represents the day of the week as text. - `E`, `EE`, `EEE`: Displays the abbreviated weekday name (e.g., "Tue"). - `EEEE`: Displays the full day name (e.g., "Tuesday"). @@ -160,7 +164,7 @@ The supported formats include: - `e`: Represents the weekday as number or text. - `e`, `ee`: Displays the the as a number, starting from 1 (Monday) to 7 (Sunday). - `eee`, `eeee`, `eeeee`: Displays the weekday as text (same format as `E`). -- `F`: Represents the aligned week of the month (e.g., "3"). +- `F`: Represents the week of the month that the first week starts on the first day of the month (e.g., "3"). - `a`: Represents the AM or PM designation of the day. - `a`, `aa`, `aaa`: Displays AM or PM in a concise format (e.g., "PM"). - `aaaa`: Displays the full AM/PM designation (e.g., "Post Meridium"). diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index dadeddd8c9ee..9d024da8112d 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -89,7 +89,7 @@ def ofDaysSinceUNIXEpoch (day : Day.Offset) : PlainDate := Returns the unaligned week of the month for a `PlainDate` (day divided by 7, plus 1). -/ def weekOfMonth (date : PlainDate) : Bounded.LE 1 5 := - date.day.ediv 7 (by decide) |>.add 1 + date.day.sub 1 |>.ediv 7 (by decide) |>.add 1 /-- Determines the quarter of the year for the given `PlainDate`. diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 62d17b2a3e20..f174b03243c4 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -142,7 +142,7 @@ Parses a string into a `TimeZone` object. The input string must be in the format -/ def fromTimeZone (input : String) : Except String TimeZone := do let spec : GenericFormat .any := datespec("VV ZZZZZ") - spec.parseBuilder (fun id off => some (TimeZone.mk off id "Unknown" false)) input + spec.parseBuilder (fun id off => some (TimeZone.mk off id (off.toIsoString true) false)) input namespace Offset @@ -173,12 +173,12 @@ def format (date : PlainDate) (format : String) : String := | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) | .Qorq _ => some date.quarter | .w _ => some date.weekOfYear - | .W _ => some date.weekOfMonth + | .W _ => some date.alignedWeekOfMonth | .MorL _ => some date.month | .d _ => some date.day | .E _ => some date.weekday | .eorc _ => some date.weekday - | .F _ => some date.alignedWeekOfMonth + | .F _ => some date.weekOfMonth | _ => none match res with | some res => res @@ -423,12 +423,12 @@ def format (date : PlainDateTime) (format : String) : String := | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) | .Qorq _ => some date.quarter | .w _ => some date.weekOfYear - | .W _ => some date.weekOfMonth + | .W _ => some date.alignedWeekOfMonth | .MorL _ => some date.month | .d _ => some date.day | .E _ => some date.weekday | .eorc _ => some date.weekday - | .F _ => some date.alignedWeekOfMonth + | .F _ => some date.weekOfMonth | .H _ => some date.hour | .k _ => some date.hour.shiftTo1BasedHour | .m _ => some date.minute diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index f50f89460524..d681c1fb3e31 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -741,7 +741,7 @@ private def toIsoString (offset : Offset) (withMinutes : Bool) (withSeconds : Bo let data := s!"{sign}{pad time.hour.val}" let data := if withMinutes then s!"{data}{if colon then ":" else ""}{pad time.minute.val}" else data - let data := if withSeconds then s!"{data}{if colon then ":" else ""}{pad time.second.snd.val}" else data + let data := if withSeconds ∧ time.second.snd.val ≠ 0 then s!"{data}{if colon then ":" else ""}{pad time.second.snd.val}" else data data @@ -754,10 +754,10 @@ private def TypeFormat : Modifier → Type | .d _ => Day.Ordinal | .Qorq _ => Month.Quarter | .w _ => Week.Ordinal - | .W _ => Bounded.LE 1 5 + | .W _ => Week.Ordinal.OfMonth | .E _ => Weekday | .eorc _ => Weekday - | .F _ => Week.Ordinal.OfMonth + | .F _ => Bounded.LE 1 5 | .a _ => HourMarker | .h _ => Bounded.LE 1 12 | .K _ => Bounded.LE 0 11 @@ -903,10 +903,10 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .d _ => date.day | .Qorq _ => date.quarter | .w _ => date.weekOfYear - | .W _ => date.weekOfMonth + | .W _ => date.alignedWeekOfMonth | .E _ => date.weekday | .eorc _ => date.weekday - | .F _ => date.alignedWeekOfMonth + | .F _ => date.weekOfMonth | .a _ => HourMarker.ofOrdinal date.hour | .h _ => HourMarker.toRelative date.hour |>.fst | .K _ => date.hour.emod 12 (by decide) @@ -1232,10 +1232,10 @@ private structure DateBuilder where d : Option Day.Ordinal := none Qorq : Option Month.Quarter := none w : Option Week.Ordinal := none - W : Option (Bounded.LE 1 5) := none + W : Option Week.Ordinal.OfMonth := none E : Option Weekday := none eorc : Option Weekday := none - F : Option Week.Ordinal.OfMonth := none + F : Option (Bounded.LE 1 5) := none a : Option HourMarker := none h : Option (Bounded.LE 1 12) := none K : Option (Bounded.LE 0 11) := none @@ -1295,10 +1295,12 @@ private def convertYearAndEra (year : Year.Offset) : Year.Era → Year.Offset | .bce => -(year + 1) private def build (builder : DateBuilder) (aw : Awareness) : Option aw.type := + let offset := builder.O <|> builder.X <|> builder.x <|> builder.Z |>.getD Offset.zero + let tz : TimeZone := { - offset := builder.O <|> builder.X <|> builder.x <|> builder.Z |>.getD Offset.zero, - name := builder.V <|> builder.z |>.getD "Unknown", - abbreviation := builder.zabbrev |>.getD "Unknown", + offset, + name := builder.V <|> builder.z |>.getD (offset.toIsoString true), + abbreviation := builder.zabbrev |>.getD (offset.toIsoString true), isDST := false, } diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 5be3976c7dbd..056fd26c619b 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -51,7 +51,7 @@ namespace Ordinal Converts an `Ordinal` into a relative hour in the range of 1 to 12. -/ def toRelative (ordinal : Ordinal) : Bounded.LE 1 12 := - ordinal.emod 12 (by decide) |>.add 1 + (ordinal.add 11).emod 12 (by decide) |>.add 1 /-- Converts an Ordinal into a 1-based hour representation within the range of 1 to 24. diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 556fa3728967..d8ed4b7bc937 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -72,6 +72,15 @@ def now : IO ZonedDateTime := do let rules ← Database.defaultGetLocalZoneRulesAt return ZonedDateTime.ofTimestamp tm rules +/-- +Gets the current `ZonedDateTime` using the identifier of a time zone. +-/ +@[inline] +def nowAt (id : String) : IO ZonedDateTime := do + let tm ← Timestamp.now + let rules ← Database.defaultGetZoneRulesAt id + return ZonedDateTime.ofTimestamp tm rules + /-- Converts a `PlainDate` to a `ZonedDateTime`. -/ diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 27388196f8cd..5b7ee9e4bfc1 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -5,7 +5,6 @@ Authors: Sofia Rodrigues -/ prelude import Std.Time.Zoned.ZoneRules -import Std.Time.Zoned.Database.Leap import Std.Time.Zoned.Database.TzIf namespace Std @@ -49,12 +48,13 @@ Converts a given time index into a `LocalTimeType` by using a time zone (`tz`) a -/ def convertLocalTimeType (index : Nat) (tz : TZif.TZifV1) (identifier : String) : Option LocalTimeType := do let localType ← tz.localTimeTypes.get? index - let abbreviation ← tz.abbreviations.getD index "Unknown" + let offset := Offset.ofSeconds <| .ofInt localType.gmtOffset + let abbreviation ← tz.abbreviations.getD index (offset.toIsoString true) let wallflag := convertWall (tz.stdWallIndicators.getD index true) let utLocal := convertUt (tz.utLocalIndicators.getD index true) return { - gmtOffset := Offset.ofSeconds <| .ofInt localType.gmtOffset + gmtOffset := offset isDst := localType.isDst abbreviation wall := wallflag diff --git a/src/Std/Time/Zoned/Database/Leap.lean b/src/Std/Time/Zoned/Database/Leap.lean deleted file mode 100644 index d3852567d00f..000000000000 --- a/src/Std/Time/Zoned/Database/Leap.lean +++ /dev/null @@ -1,110 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sofia Rodrigues --/ -prelude -import Init.Data.Range -import Std.Internal.Parsec -import Std.Internal.Parsec.ByteArray - -import Std.Time.Internal.Bounded -import Std.Time.Time -import Std.Time.Date -import Std.Time.DateTime - -namespace Std -namespace Time -namespace TimeZone -open Std.Internal.Parsec Std.Internal.Parsec.String - -set_option linter.all true - -/-- -Represents a leap second entry with details such as the year, month, day, time, whether it's a positive leap second, -and a stationary string (presumably to capture additional information). --/ -structure LeapSecond where - - /-- - The timestamp when the leap second occurs. - -/ - timestamp : Timestamp - - /-- - Indicates whether the leap second is positive (`true` for positive, `false` for negative). - -/ - positive : Bool - - /-- - A string representing the stationary field, which could be used for additional information - or metadata about the leap second. - -/ - stationary : String - -private def skipComment : Parser Unit := do - skipChar '#' - discard <| many (satisfy (· ≠ '\n')) - discard ∘ optional <| skipChar '\n' - return () - -private def failWith (opt : Option α) : Parser α := - match opt with - | none => fail "invalid number" - | some res => pure res - -private def parseMonthShort : Parser Month.Ordinal - := pstring "Jan" *> pure ⟨1, by decide⟩ - <|> pstring "Feb" *> pure ⟨2, by decide⟩ - <|> pstring "Mar" *> pure ⟨3, by decide⟩ - <|> pstring "Apr" *> pure ⟨4, by decide⟩ - <|> pstring "May" *> pure ⟨5, by decide⟩ - <|> pstring "Jun" *> pure ⟨6, by decide⟩ - <|> pstring "Jul" *> pure ⟨7, by decide⟩ - <|> pstring "Aug" *> pure ⟨8, by decide⟩ - <|> pstring "Sep" *> pure ⟨9, by decide⟩ - <|> pstring "Oct" *> pure ⟨10, by decide⟩ - <|> pstring "Nov" *> pure ⟨11, by decide⟩ - <|> pstring "Dec" *> pure ⟨12, by decide⟩ - -private def parseLeapSecond : Parser LeapSecond := do - skipString "Leap" - ws - let year ← digits - ws - let month ← parseMonthShort - ws - let day ← digits - ws - - let hour : Hour.Ordinal ← failWith =<< Internal.Bounded.ofInt? <$> digits - skipChar ':' - let minute : Minute.Ordinal ← failWith =<< Internal.Bounded.ofInt? <$> digits - skipChar ':' - let second : Second.Ordinal true ← failWith =<< Internal.Bounded.ofInt? <$> digits - - ws - let positive ← (pchar '+' *> pure true) - ws - let stationary ← manyChars (satisfy Char.isAlpha) - skipChar '\n' - - let day ← failWith <| Internal.Bounded.ofInt? day - let time : PlainTime ← failWith <| PlainTime.ofHourMinuteSeconds hour minute second - let date : PlainDate ← failWith <| PlainDate.ofYearMonthDay? (Year.Offset.ofNat year) month day - let pdt := PlainDateTime.mk date time - - return { timestamp := pdt.toTimestampAssumingUTC, positive, stationary } - -private def parseComments : Parser (Array Unit) := - many1 (ws *> skipComment) - -/-- -Parses a sequence of leap second entries. --/ -def parseLeapSeconds : Parser (Array LeapSecond) := do - discard <| many (ws *> skipComment) - let res ← many parseLeapSecond - discard <| many (ws *> skipComment) - eof - return res diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 0935ae522962..1d6c7dacc26e 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -83,16 +83,6 @@ Reads timezone rules from disk based on the provided file path and timezone ID. def readRulesFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := do parseTZIfFromDisk (System.FilePath.join path id) id -/-- -Reads leap seconds form from disk based on the provided file path and timezone ID. --/ -def readLeapsFromDisk (path : System.FilePath) : IO (Array LeapSecond) := do - let res ← parseLeapSeconds.run <$> IO.FS.readFile path - - if let .ok res := res - then return res - else throw (IO.userError "cannot read the id of the path.") - instance : Std.Time.Database TZdb where getLocalZoneRulesAt db := localRules db.localPath getZoneRulesAt db id := readRulesFromDisk db.zonesPath id diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 7263cb472174..5e954b55e62c 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -37,7 +37,7 @@ Converts an `Offset` to a string in ISO 8601 format. The `colon` parameter deter and minute components are separated by a colon (e.g., "+01:00" or "+0100"). -/ def toIsoString (offset : Offset) (colon : Bool) : String := - let (sign, time) := if offset.second.val > 0 then ("+", offset.second) else ("-", -offset.second) + let (sign, time) := if offset.second.val ≥ 0 then ("+", offset.second) else ("-", -offset.second) let hour : Hour.Offset := time.ediv 3600 let minute := Int.ediv (Int.tmod time.val 3600) 60 let hourStr := if hour.val < 10 then s!"0{hour.val}" else toString hour.val diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 6283692b8861..5e191b31638b 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -62,6 +62,20 @@ def ofPlainDateTime (pdt : PlainDateTime) (zr : TimeZone.ZoneRules) : ZonedDateT let tm := tm.subSeconds tz.toSeconds ZonedDateTime.mk (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds) |>.toPlainDateTimeAssumingUTC) tm zr tz +/-- +Creates a new `ZonedDateTime` out of a `Timestamp` and a `TimeZone`. +-/ +@[inline] +def ofTimestampWithZone (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := + ofTimestamp tm (TimeZone.ZoneRules.ofTimeZone tz) + +/-- +Creates a new `ZonedDateTime` out of a `PlainDateTime` and a `TimeZone`. +-/ +@[inline] +def ofPlainDateTimeWithZone (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := + ofTimestamp tm (TimeZone.ZoneRules.ofTimeZone tz) + /-- Creates a new `Timestamp` out of a `ZonedDateTime`. -/ diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index 49a967fc461b..d3a8ed39fdd7 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -280,6 +280,7 @@ def date₄ := date("2002-07-14") def datetime₅ := PlainDateTime.mk (PlainDate.clip (-2000) 3 4) (PlainTime.mk 12 23 ⟨false, 12⟩ 0) def datetime₄ := datetime("2002-07-14T23:13:12.324354679") def zoned₄ := zoned("2002-07-14T23:13:12.324354679+09:00") +def zoned₅ := zoned("2002-07-14T23:13:12.324354679+00:00") /-- info: "CE CE CE Common Era C" @@ -330,7 +331,7 @@ info: "28 28 028 0028" #eval zoned₄.format "w ww www wwww" /-- -info: "3 03 003 0003" +info: "2 02 002 0002" -/ #guard_msgs in #eval zoned₄.format "W WW WWW WWWW" @@ -366,7 +367,7 @@ info: "11 11 011 0011 000011" #eval zoned₄.format "K KK KKK KKKK KKKKKK" /-- -info: "24 24 024 0024 000024" +info: "23 23 023 0023 000023" -/ #guard_msgs in #eval zoned₄.format "k kk kkk kkkk kkkkkk" @@ -414,17 +415,23 @@ info: "83592324354679 83592324354679 83592324354679 83592324354679 8359232435467 #eval zoned₄.format "N NN NNN NNNN NNNNNNNNN" /-- -info: "Unknown" +info: "+09:00" -/ #guard_msgs in #eval zoned₄.format "VV" /-- -info: "Unknown Unknown Unknown Unknown" +info: "+09:00 +09:00 +09:00 +09:00" -/ #guard_msgs in #eval zoned₄.format "z zz zzzz zzzz" +/-- +info: "+00:00 +00:00 +00:00 +00:00" +-/ +#guard_msgs in +#eval zoned₅.format "z zz zzzz zzzz" + /-- info: "GMT+9 GMT+09:00" -/ @@ -432,17 +439,29 @@ info: "GMT+9 GMT+09:00" #eval zoned₄.format "O OOOO" /-- -info: "+09 +0900 +09:00 +090000 +09:00:00" +info: "+09 +0900 +09:00 +0900 +09:00" -/ #guard_msgs in #eval zoned₄.format "X XX XXX XXXX XXXXX" /-- -info: "+09 +0900 +09:00 +0900 +09:00:00" +info: "+09 +0900 +09:00 +0900 +09:00" -/ #guard_msgs in #eval zoned₄.format "x xx xxx xxxx xxxxx" +/-- +info: "Z Z Z Z Z" +-/ +#guard_msgs in +#eval zoned₅.format "X XX XXX XXXX XXXXX" + +/-- +info: "+00 +0000 +00:00 +0000 +00:00" +-/ +#guard_msgs in +#eval zoned₅.format "x xx xxx xxxx xxxxx" + /-- info: "+0900 +0900 +0900 GMT+09:00 +09:00" -/ @@ -509,7 +528,7 @@ info: "28 28 028 0028" #eval datetime₄.format "w ww www wwww" /-- -info: "3 03 003 0003" +info: "2 02 002 0002" -/ #guard_msgs in #eval datetime₄.format "W WW WWW WWWW" @@ -611,7 +630,7 @@ info: "11 11 011 0011 000011" #eval time₄.format "K KK KKK KKKK KKKKKK" /-- -info: "24 24 024 0024 000024" +info: "23 23 023 0023 000023" -/ #guard_msgs in @@ -725,7 +744,7 @@ info: "28 28 028 0028" #eval date₄.format "w ww www wwww" /-- -info: "3 03 003 0003" +info: "2 02 002 0002" -/ #guard_msgs in #eval date₄.format "W WW WWW WWWW" From 946aff67b354726abeed4122deff8c0a5908f0ab Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 28 Oct 2024 13:18:42 -0300 Subject: [PATCH 083/118] fix: nix test --- tests/lean/run/timeClassOperations.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index 2c92b96813d4..8d4c48a2bd9d 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -167,7 +167,7 @@ info: "3600s" def now := PlainDateTime.ofUTCTimestamp ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ def after := PlainDateTime.ofUTCTimestamp ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ -def duration := after.since now +def duration := after - now /-- info: "15:40:38.907328169" From 81b70351bdadcfb10268fda8c8329500643d604e Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 29 Oct 2024 12:06:49 -0300 Subject: [PATCH 084/118] fix: multiple fixes --- src/Std/Time.lean | 17 ++++++++---- src/Std/Time/DateTime.lean | 2 +- src/Std/Time/DateTime/PlainDateTime.lean | 2 +- src/Std/Time/Format.lean | 35 +++++++++++++++++++----- src/Std/Time/Format/Basic.lean | 3 +- src/Std/Time/Notation.lean | 30 ++++++++++++++++++-- src/Std/Time/Zoned.lean | 9 +----- src/Std/Time/Zoned/Database/TZdb.lean | 8 ++++-- src/Std/Time/Zoned/ZoneRules.lean | 2 +- src/Std/Time/Zoned/ZonedDateTime.lean | 4 +-- tests/lean/run/timeClassOperations.lean | 4 +-- tests/lean/run/timeFormats.lean | 8 ++++++ tests/lean/run/timeLocalDateTime.lean | 2 +- 13 files changed, 91 insertions(+), 35 deletions(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index dafa2a6b8c5d..f85d37cef44b 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -218,10 +218,15 @@ The supported formats include: In order to help the user build dates easily, there are a lot of macros available for creating dates. The `.sssssssss` can be ommited in most cases. -- **`date("uuuu-MM-dd")`**: Defines a date in the `uuuu-MM-DD` format. -- **`time("HH:mm:ss.sssssssss")`**: Defines a time in the `HH:mm:ss.sssssssss` format, including fractional seconds. -- **`datetime("uuuu-MM-ddTHH:mm:ss.sssssssss")`**: Defines a datetime in the `YYYY-MM-DD:HH:mm:ss.sssssssss` format. -- **`offset("+HH:mm")`**: Defines a timezone offset in the format `+HH:mm`. -- **`timezone("NAME/ID ZZZ")`**: Defines a timezone with a name and an offset. -- **`datespec("format")`**: Defines a date specification format at compile time using the provided format string. + +- **`date("uuuu-MM-dd")`**: Represents a date in the `uuuu-MM-dd` format, where `uuuu` refers to the year. +- **`time("HH:mm:ss.sssssssss")`**: Represents a time in the format `HH:mm:ss.sssssssss`, including optional support for nanoseconds. +- **`datetime("uuuu-MM-ddTHH:mm:ss.sssssssss")`**: Represents a datetime value in the `uuuu-MM-ddTHH:mm:ss.sssssssss` format, with optional nanoseconds. +- **`offset("+HH:mm")`**: Represents a timezone offset in the format `+HH:mm`, where `+` or `-` indicates the direction from UTC. +- **`timezone("NAME/ID ZZZ")`**: Specifies a timezone using a region-based name or ID, along with its associated offset. +- **`datespec("FORMAT")`**: Defines a compile-time date format based on the provided string. +- **`zoned("uuuu-MM-ddTHH:mm:ss.sssssssssZZZ")`**: Represents a `ZonedDateTime` with a fixed timezone and optional nanosecond precision. +- **`zoned("uuuu-MM-ddTHH:mm:ss.sssssssss[IDENTIFIER]")`**: Defines an `IO ZonedDateTime`, where the timezone identifier is dynamically retrieved from the default timezone database. +- **`zoned("uuuu-MM-ddTHH:mm:ss.sssssssss, timezone")`**: Represents an `IO ZonedDateTime`, using a specified `timezone` term and allowing optional nanoseconds. + -/ diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index dc08064d5100..eab45c7d8b5c 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -26,7 +26,7 @@ Converts a `Timestamp` to a `PlainDateTime` -/ @[inline] def toPlainDateTimeAssumingUTC (timestamp : Timestamp) : PlainDateTime := - PlainDateTime.ofUTCTimestamp timestamp + PlainDateTime.ofTimestamp timestamp /-- Converts a `PlainDate` to a `Timestamp` diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 3b7421a80452..f2fcb554f8db 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -46,7 +46,7 @@ def toTimestampAssumingUTC (dt : PlainDateTime) : Timestamp := /-- Converts a UNIX `Timestamp` to a `PlainDateTime`. -/ -def ofUTCTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do +def ofTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do let leapYearEpoch := 11017 let daysPer400Y := 365 * 400 + 97 let daysPer100Y := 365 * 100 + 24 diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index f174b03243c4..d5dbcfe17e83 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -84,19 +84,33 @@ notation of dates. def leanDateTime24HourNoNanos : GenericFormat (.only .GMT) := datespec("uuuu-MM-dd'T'HH:mm:ss") /-- -The DateTimeWithZone format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ` +The leanDateTimeWithZone format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ def leanDateTimeWithZone : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZZZ") /-- -The DateTimeWithZoneNoNanos format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ssZZZZZ` +The leanDateTimeWithZoneNoNanos format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ssZZZZZ` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ def leanDateTimeWithZoneNoNanos : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ssZZZZZ") +/-- +The leanDateTimeWithIdentifier format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss[z]` +for representing date, time, and time zone. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanDateTimeWithIdentifier : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss'['z']'") + +/-- +The leanDateTimeWithIdentifierAndNanos format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSS'[z]'` +for representing date, time, and time zone. It uses the default value that can be parsed with the +notation of dates. +-/ +def leanDateTimeWithIdentifierAndNanos : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSS'['z']'") + /-- The Lean Date format, which follows the pattern `uuuu-MM-dd`. It uses the default value that can be parsed with the notation of dates. @@ -342,7 +356,7 @@ def toISO8601String (date : ZonedDateTime) : String := Formats.iso8601.format date.toDateTime /-- -Parses a `String` in the `RFC822` format and returns a `ZonedDateTime`. +Parses a `String` in the rfc822 format and returns a `ZonedDateTime`. -/ def fromRFC822String (input : String) : Except String ZonedDateTime := Formats.rfc822.parse input @@ -354,7 +368,7 @@ def toRFC822String (date : ZonedDateTime) : String := Formats.rfc822.format date.toDateTime /-- -Parses a `String` in the `RFC850` format and returns a `ZonedDateTime`. +Parses a `String` in the rfc850 format and returns a `ZonedDateTime`. -/ def fromRFC850String (input : String) : Except String ZonedDateTime := Formats.rfc850.parse input @@ -366,24 +380,31 @@ def toRFC850String (date : ZonedDateTime) : String := Formats.rfc850.format date.toDateTime /-- -Parses a `String` in the `DateTimeWithZone` format and returns a `DateTime` object in the GMT time zone. +Parses a `String` in the dateTimeWithZone format and returns a `ZonedDateTime` object in the GMT time zone. -/ def fromDateTimeWithZoneString (input : String) : Except String ZonedDateTime := Formats.dateTimeWithZone.parse input /-- -Formats a `DateTime` value into a simple date time with timezone string. +Formats a `ZonedDateTime` value into a simple date time with timezone string. -/ def toDateTimeWithZoneString (pdt : ZonedDateTime) : String := Formats.dateTimeWithZone.format pdt.toDateTime /-- -Parses a `String` in the `DateTimeWithZone` format and returns a `DateTime` object in the GMT time zone. +Parses a `String` in the lean date time format with timezone format and returns a `ZonedDateTime` object. -/ def fromLeanDateTimeWithZoneString (input : String) : Except String ZonedDateTime := Formats.leanDateTimeWithZone.parse input <|> Formats.leanDateTimeWithZoneNoNanos.parse input +/-- +Parses a `String` in the lean date time format with identifier and returns a `ZonedDateTime` object. +-/ +def fromLeanDateTimeWithIdentifierString (input : String) : Except String ZonedDateTime := + Formats.leanDateTimeWithIdentifier.parse input + <|> Formats.leanDateTimeWithIdentifierAndNanos.parse input + /-- Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notationg. -/ diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index d681c1fb3e31..fb06c3bd66af 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -919,7 +919,8 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .n _ => date.nanoseconds | .N _ => date.date.get.time.toNanoseconds | .V => tz.name - | .z _ => tz.abbreviation + | .z .short => tz.abbreviation + | .z .full => tz.name | .O _ => tz.offset | .X _ => tz.offset | .x _ => tz.offset diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index cbc87c516ea6..522cd3d0d47a 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -131,8 +131,13 @@ private def convertPlainTime (d : Std.Time.PlainTime) : MacroM (TSyntax `term) : private def convertPlainDateTime (d : Std.Time.PlainDateTime) : MacroM (TSyntax `term) := do `(Std.Time.PlainDateTime.mk $(← convertPlainDate d.date) $(← convertPlainTime d.time)) -private def convertZonedDateTime (d : Std.Time.ZonedDateTime) : MacroM (TSyntax `term) := do - `(Std.Time.ZonedDateTime.ofPlainDateTime $(← convertPlainDateTime d.toPlainDateTime) (Std.Time.TimeZone.ZoneRules.ofTimeZone $(← convertTimezone d.timezone))) +private def convertZonedDateTime (d : Std.Time.ZonedDateTime) (identifier := false) : MacroM (TSyntax `term) := do + let plain ← convertPlainDateTime d.toPlainDateTime + + if identifier then + `(Std.Time.ZonedDateTime.ofPlainDateTime $plain <$> Std.Time.Database.defaultGetZoneRulesAt $(Syntax.mkStrLit d.timezone.name)) + else + `(Std.Time.ZonedDateTime.ofPlainDateTime $plain (Std.Time.TimeZone.ZoneRules.ofTimeZone $(← convertTimezone d.timezone))) /-- Defines a syntax for zoned datetime values. It expects a string representing a datetime with @@ -143,6 +148,16 @@ Example: -/ syntax "zoned(" str ")" : term +/-- +Defines a syntax for zoned datetime values. It expects a string representing a datetime and a +timezone information as a term. + +Example: +`zoned("2024-10-13T15:00:00", timezone)` +-/ +syntax "zoned(" str "," term ")" : term + + /-- Defines a syntax for datetime values without timezone. The input should be a string in an ISO8601-like format. @@ -191,8 +206,17 @@ syntax "timezone(" str ")" : term macro_rules | `(zoned( $date:str )) => do match ZonedDateTime.fromLeanDateTimeWithZoneString date.getString with + | .ok res => do return ← convertZonedDateTime res + | .error _ => + match ZonedDateTime.fromLeanDateTimeWithIdentifierString date.getString with + | .ok res => do return ← convertZonedDateTime res (identifier := true) + | .error res => Macro.throwErrorAt date s!"error: {res}" + + | `(zoned( $date:str, $timezone )) => do + match PlainDateTime.fromLeanDateTimeString date.getString with | .ok res => do - return ← convertZonedDateTime res + let plain ← convertPlainDateTime res + `(Std.Time.ZonedDateTime.ofPlainDateTime $plain $timezone) | .error res => Macro.throwErrorAt date s!"error: {res}" | `(datetime( $date:str )) => do diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index d8ed4b7bc937..6e800570b52d 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -17,13 +17,6 @@ namespace PlainDateTime set_option linter.all true -/-- -Creaates a new `PlainDateTime` out of a `Timestamp` and a `TimeZone`. --/ -def ofTimestamp (stamp : Timestamp) (tz : TimeZone) : PlainDateTime := - let stamp := stamp.addSeconds tz.toSeconds - PlainDateTime.ofUTCTimestamp stamp - /-- Get the current time. -/ @@ -33,7 +26,7 @@ def now : IO PlainDateTime := do let rules ← Database.defaultGetLocalZoneRulesAt let ltt := rules.findLocalTimeTypeForTimestamp tm - return PlainDateTime.ofTimestamp tm ltt.getTimeZone + return PlainDateTime.ofTimestamp tm |>.addSeconds ltt.getTimeZone.toSeconds end PlainDateTime diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index 1d6c7dacc26e..ec6242eb1caa 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -52,7 +52,7 @@ def parseTZif (bin : ByteArray) (id : String) : Except String ZoneRules := do Reads a TZif file from disk and retrieves the zone rules for the specified timezone ID. -/ def parseTZIfFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := do - let binary ← IO.FS.readBinFile path + let binary ← try IO.FS.readBinFile path catch _ => throw <| IO.userError s!"cannot find {id} in the local timezone database" IO.ofExcept (parseTZif binary id) /-- @@ -71,7 +71,11 @@ def idFromPath (path : System.FilePath) : Option String := do Retrieves the timezone rules from the local timezone data file. -/ def localRules (path : System.FilePath) : IO ZoneRules := do - let localTimePath ← IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } + let localTimePath ← + try + IO.Process.run { cmd := "readlink", args := #["-f", path.toString] } + catch _ => + throw <| IO.userError "cannot find the local timezone database" if let some id := idFromPath localTimePath then parseTZIfFromDisk path id diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 49e02468aefb..8d3856047957 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -187,7 +187,7 @@ Creates `ZoneRules` for the given `TimeZone`. -/ @[inline] def ofTimeZone (tz : TimeZone) : ZoneRules := - let ltt := LocalTimeType.mk tz.offset tz.isDST tz.name .wall .local tz.abbreviation + let ltt := LocalTimeType.mk tz.offset tz.isDST tz.abbreviation .wall .local tz.name ZoneRules.mk ltt #[] end ZoneRules diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 5e191b31638b..65041ffa5cac 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -50,7 +50,7 @@ Creates a new `ZonedDateTime` out of a `Timestamp` and a `ZoneRules`. @[inline] def ofTimestamp (tm : Timestamp) (rules : TimeZone.ZoneRules) : ZonedDateTime := let tz := rules.timezoneAt tm - ZonedDateTime.mk (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds) |>.toPlainDateTimeAssumingUTC) tm rules tz + ZonedDateTime.mk (Thunk.mk fun _ => tm.toPlainDateTimeAssumingUTC.addSeconds tz.toSeconds) tm rules tz /-- Creates a new `ZonedDateTime` out of a `PlainDateTime` and a `ZoneRules`. @@ -60,7 +60,7 @@ def ofPlainDateTime (pdt : PlainDateTime) (zr : TimeZone.ZoneRules) : ZonedDateT let tm := pdt.toTimestampAssumingUTC let tz := zr.findLocalTimeTypeForTimestamp tm |>.getTimeZone let tm := tm.subSeconds tz.toSeconds - ZonedDateTime.mk (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds) |>.toPlainDateTimeAssumingUTC) tm zr tz + ZonedDateTime.mk (Thunk.mk fun _ => tm.toPlainDateTimeAssumingUTC.addSeconds tz.toSeconds) tm zr tz /-- Creates a new `ZonedDateTime` out of a `Timestamp` and a `TimeZone`. diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index 8d4c48a2bd9d..de4ca9c399f3 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -165,8 +165,8 @@ info: "3600s" #guard_msgs in #eval zoned("2000-12-20T00:00:00-03:00") - zoned("2000-12-20T00:00:00-02:00") -def now := PlainDateTime.ofUTCTimestamp ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ -def after := PlainDateTime.ofUTCTimestamp ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ +def now := PlainDateTime.ofTimestamp ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ +def after := PlainDateTime.ofTimestamp ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ def duration := after - now /-- diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index d3a8ed39fdd7..fe2be5ac23c8 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -281,6 +281,8 @@ def datetime₅ := PlainDateTime.mk (PlainDate.clip (-2000) 3 4) (PlainTime.mk 1 def datetime₄ := datetime("2002-07-14T23:13:12.324354679") def zoned₄ := zoned("2002-07-14T23:13:12.324354679+09:00") def zoned₅ := zoned("2002-07-14T23:13:12.324354679+00:00") +def tz : TimeZone := { offset := { second := -3600 }, name := "America/Sao_Paulo", abbreviation := "BRT", isDST := false} +def zoned₆ := ZonedDateTime.ofPlainDateTime (zoned₄.toPlainDateTime) (TimeZone.ZoneRules.ofTimeZone tz) /-- info: "CE CE CE Common Era C" @@ -779,6 +781,12 @@ info: "2002 2002 CE" #guard_msgs in #eval datetime₄.format "uuuu yyyy G" +/-- +info: "BRT BRT BRT America/Sao_Paulo America/Sao_Paulo" +-/ +#guard_msgs in +#eval zoned₆.format "z zz zzz zzzz zzzz" + /-- info: 1 -/ diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 3ab2d18eb356..3d2123e641d8 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -44,7 +44,7 @@ info: 736952399 info: "09/05/1993 12:59:59" -/ #guard_msgs in -#eval PlainDateTime.ofUTCTimestamp 736952399 |> format +#eval PlainDateTime.ofTimestamp 736952399 |> format /-- info: 736952399 From 6c7ad97861081425feaf3862f0cabaf77a920ea6 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 31 Oct 2024 02:17:48 -0300 Subject: [PATCH 085/118] feat: add functions that are lacking, fix other names --- src/Std/Time/Date/PlainDate.lean | 24 ++++---- src/Std/Time/Date/ValidDate.lean | 2 +- src/Std/Time/DateTime/PlainDateTime.lean | 73 +++++++++++++++++++++--- src/Std/Time/Format.lean | 10 ++-- src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Notation.lean | 2 +- src/Std/Time/Time/PlainTime.lean | 26 +++++++-- src/Std/Time/Zoned.lean | 25 +++++++- src/Std/Time/Zoned/DateTime.lean | 8 +-- src/Std/Time/Zoned/ZonedDateTime.lean | 56 +++++++++++++++++- 10 files changed, 184 insertions(+), 44 deletions(-) diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 9d024da8112d..0170c5f7bfc8 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -35,6 +35,9 @@ structure PlainDate where /-- Validates the date by ensuring that the year, month, and day form a correct and valid date. -/ valid : year.Valid month day +instance : Inhabited PlainDate where + default := ⟨1, 1, 1, by decide⟩ + instance : BEq PlainDate where beq x y := x.day == y.day && x.month == y.month && x.year == y.year @@ -100,8 +103,8 @@ def quarter (date : PlainDate) : Bounded.LE 1 4 := /-- Transforms a tuple of a `PlainDate` into a `Day.Ordinal.OfYear`. -/ -def toOrdinal (date : PlainDate) : Day.Ordinal.OfYear date.year.isLeap := - ValidDate.toOrdinal ⟨(date.month, date.day), date.valid⟩ +def dayOfYear (date : PlainDate) : Day.Ordinal.OfYear date.year.isLeap := + ValidDate.dayOfYear ⟨(date.month, date.day), date.valid⟩ /-- Determines the era of the given `PlainDate` based on its year. @@ -131,13 +134,6 @@ def toDaysSinceUNIXEpoch (date : PlainDate) : Day.Offset := .ofInt (era * 146097 + doe - 719468) -/-- -Calculates the difference in years between a `PlainDate` and a given year. --/ -@[inline] -def yearsSince (date : PlainDate) (year : Year.Offset) : Year.Offset := - date.year - year - /-- Adds a given number of days to a `PlainDate`. -/ @@ -255,7 +251,7 @@ Creates a new `PlainDate` by adjusting the day of the month to the given `days` out-of-range days clipped to the nearest valid date. -/ @[inline] -def withDayClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := +def withDaysClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := clip dt.year dt.month days /-- @@ -263,7 +259,7 @@ Creates a new `PlainDate` by adjusting the day of the month to the given `days` out-of-range days rolled over to the next month or year as needed. -/ @[inline] -def withDayRollOver (dt : PlainDate) (days : Day.Ordinal) : PlainDate := +def withDaysRollOver (dt : PlainDate) (days : Day.Ordinal) : PlainDate := rollOver dt.year dt.month days /-- @@ -296,14 +292,14 @@ on the day of the month and the weekday. Each week starts on Sunday because the based on the Gregorian Calendar. -/ def alignedWeekOfMonth (date : PlainDate) : Week.Ordinal.OfMonth := - let weekday := date.withDayClip 1 |>.weekday |>.toOrdinal |>.sub 1 + let weekday := date.withDaysClip 1 |>.weekday |>.toOrdinal |>.sub 1 let days := date.day |>.sub 1 |>.addBounds weekday days |>.ediv 7 (by decide) |>.add 1 /-- Sets the date to the specified `desiredWeekday`. -/ -def setWeekday (date : PlainDate) (desiredWeekday : Weekday) : PlainDate := +def withWeekday (date : PlainDate) (desiredWeekday : Weekday) : PlainDate := let weekday := date |>.weekday |>.toOrdinal let offset := desiredWeekday.toOrdinal |>.subBounds weekday @@ -324,7 +320,7 @@ def weekOfYear (date : PlainDate) : Week.Ordinal := let y := date.year let w := Bounded.LE.exact 10 - |>.addBounds date.toOrdinal + |>.addBounds date.dayOfYear |>.subBounds date.weekday.toOrdinal |>.ediv 7 (by decide) diff --git a/src/Std/Time/Date/ValidDate.lean b/src/Std/Time/Date/ValidDate.lean index 450b1db53378..bd38c6520d26 100644 --- a/src/Std/Time/Date/ValidDate.lean +++ b/src/Std/Time/Date/ValidDate.lean @@ -30,7 +30,7 @@ namespace ValidDate /-- Transforms a tuple of a `Month` and a `Day` into a `Day.Ordinal.OfYear`. -/ -def toOrdinal (ordinal : ValidDate leap) : Day.Ordinal.OfYear leap := +def dayOfYear (ordinal : ValidDate leap) : Day.Ordinal.OfYear leap := let days := cumulativeDays leap ordinal.val.fst let proof := cumulativeDays_le leap ordinal.val.fst let bounded := Bounded.LE.mk days.toInt proof |>.addBounds ordinal.val.snd diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index f2fcb554f8db..3cf479ea93d8 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -40,7 +40,7 @@ Converts a `PlainDateTime` to a `Timestamp` def toTimestampAssumingUTC (dt : PlainDateTime) : Timestamp := let days := dt.date.toDaysSinceUNIXEpoch let nanos := days.toSeconds + dt.time.toSeconds |>.mul 1000000000 - let nanos := nanos.val + dt.time.nano.val + let nanos := nanos.val + dt.time.nanosecond.val Timestamp.ofNanosecondsSinceUnixEpoch (Nanosecond.Offset.ofInt nanos) /-- @@ -126,6 +126,12 @@ def ofTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do time := PlainTime.ofHourMinuteSecondsNano (leap := false) (hour.expandTop (by decide)) minute second nano } +/-- +Sets the `PlainDateTime` to the specified `desiredWeekday`. +-/ +def withWeekday (dt : PlainDateTime) (desiredWeekday : Weekday) : PlainDateTime := + { dt with date := PlainDate.withWeekday dt.date desiredWeekday } + /-- Creates a new `PlainDateTime` by adjusting the day of the month to the given `days` value, with any out-of-range days clipped to the nearest valid date. @@ -195,12 +201,19 @@ Creates a new `PlainDateTime` by adjusting the `second` component of its `time` def withSeconds (dt : PlainDateTime) (second : Sigma Second.Ordinal) : PlainDateTime := { dt with time := { dt.time with second := second } } +/-- +Creates a new `PlainDateTime` by adjusting the milliseconds component inside the `nano` component of its `time` to the given value. +-/ +@[inline] +def withMillisecond (dt : PlainDateTime) (millis : Millisecond.Ordinal) : PlainDateTime := + { dt with time := dt.time.withMillisecond millis } + /-- Creates a new `PlainDateTime` by adjusting the `nano` component of its `time` to the given value. -/ @[inline] def withNanoseconds (dt : PlainDateTime) (nano : Nanosecond.Ordinal) : PlainDateTime := - { dt with time := { dt.time with nano := nano } } + { dt with time := { dt.time with nanosecond := nano } } /-- Adds a `Day.Offset` to a `PlainDateTime`. @@ -341,17 +354,34 @@ Subtracts a `Second.Offset` from a `PlainDateTime`, adjusting the minute, hour, def subSeconds (dt : PlainDateTime) (seconds : Second.Offset) : PlainDateTime := addSeconds dt (-seconds) +/-- +Adds a `Millisecond.Offset` to a `PlainDateTime`, adjusting the second, minute, hour, and date if the milliseconds overflow. +-/ +@[inline] +def addMilliseconds (dt : PlainDateTime) (milliseconds : Millisecond.Offset) : PlainDateTime := + let totalMilliseconds := dt.time.toMilliseconds + milliseconds + let days := totalMilliseconds.ediv 86400000 -- 86400000 ms in a day + let newTime := dt.time.addMilliseconds milliseconds + { dt with date := dt.date.addDays days, time := newTime } + +/-- +Subtracts a `Millisecond.Offset` from a `PlainDateTime`, adjusting the second, minute, hour, and date if the milliseconds underflow. +-/ +@[inline] +def subMilliseconds (dt : PlainDateTime) (milliseconds : Millisecond.Offset) : PlainDateTime := + addMilliseconds dt (-milliseconds) + /-- Adds a `Nanosecond.Offset` to a `PlainDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds overflow. -/ @[inline] def addNanoseconds (dt : PlainDateTime) (nanos : Nanosecond.Offset) : PlainDateTime := - let nano := Nanosecond.Offset.ofInt dt.time.nano.val + let nano := Nanosecond.Offset.ofInt dt.time.nanosecond.val let totalNanos := nano + nanos let extraSeconds := totalNanos.ediv 1000000000 - let nano := Bounded.LE.byEmod totalNanos.val 1000000000 (by decide) + let nanosecond := Bounded.LE.byEmod totalNanos.val 1000000000 (by decide) let newTime := dt.time.addSeconds extraSeconds - { dt with time := { newTime with nano } } + { dt with time := { newTime with nanosecond } } /-- Subtracts a `Nanosecond.Offset` from a `PlainDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds underflow. @@ -402,6 +432,13 @@ Getter for the `Minute` inside of a `PlainDateTime`. def minute (dt : PlainDateTime) : Minute.Ordinal := dt.time.minute +/-- +Getter for the `Millisecond` inside of a `PlainDateTime`. +-/ +@[inline] +def millisecond (dt : PlainDateTime) : Millisecond.Ordinal := + dt.time.millisecond + /-- Getter for the `Second` inside of a `PlainDateTime`. -/ @@ -414,7 +451,7 @@ Getter for the `Nanosecond.Ordinal` inside of a `PlainDateTime`. -/ @[inline] def nanosecond (dt : PlainDateTime) : Nanosecond.Ordinal := - dt.time.nano + dt.time.nanosecond /-- Determines the era of the given `PlainDateTime` based on its year. @@ -456,8 +493,8 @@ def alignedWeekOfMonth (date : PlainDateTime) : Week.Ordinal.OfMonth := Transforms a tuple of a `PlainDateTime` into a `Day.Ordinal.OfYear`. -/ @[inline] -def toOrdinal (date : PlainDateTime) : Day.Ordinal.OfYear date.year.isLeap := - ValidDate.toOrdinal ⟨(date.month, date.day), date.date.valid⟩ +def dayOfYear (date : PlainDateTime) : Day.Ordinal.OfYear date.year.isLeap := + ValidDate.dayOfYear ⟨(date.month, date.day), date.date.valid⟩ /-- Determines the quarter of the year for the given `PlainDateTime`. @@ -466,6 +503,20 @@ Determines the quarter of the year for the given `PlainDateTime`. def quarter (date : PlainDateTime) : Bounded.LE 1 4 := date.date.quarter +/-- +Combines a `PlainDate` and `PlainTime` into a `PlainDateTime`. +-/ +@[inline] +def atTime : PlainDate → PlainTime → PlainDateTime := + PlainDateTime.mk + +/-- +Combines a `PlainTime` and `PlainDate` into a `PlainDateTime`. +-/ +@[inline] +def atDate (time: PlainTime) (date: PlainDate) : PlainDateTime := + PlainDateTime.mk date time + instance : HAdd PlainDateTime Day.Offset PlainDateTime where hAdd := addDays @@ -490,6 +541,12 @@ instance : HAdd PlainDateTime Minute.Offset PlainDateTime where instance : HSub PlainDateTime Minute.Offset PlainDateTime where hSub := subMinutes +instance : HAdd PlainDateTime Millisecond.Offset PlainDateTime where + hAdd := addMilliseconds + +instance : HSub PlainDateTime Millisecond.Offset PlainDateTime where + hSub := addMilliseconds + instance : HAdd PlainDateTime Second.Offset PlainDateTime where hAdd := addSeconds diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index d5dbcfe17e83..c374540e09de 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -184,7 +184,7 @@ def format (date : PlainDate) (format : String) : String := | .G _ => some date.era | .y _ => some date.year | .u _ => some date.year - | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) + | .D _ => some (Sigma.mk date.year.isLeap date.dayOfYear) | .Qorq _ => some date.quarter | .w _ => some date.weekOfYear | .W _ => some date.alignedWeekOfMonth @@ -263,12 +263,12 @@ def format (time : PlainTime) (format : String) : String := | .H _ => some time.hour | .k _ => some (time.hour.shiftTo1BasedHour) | .m _ => some time.minute - | .n _ => some time.nano + | .n _ => some time.nanosecond | .s _ => some time.second | .a _ => some (HourMarker.ofOrdinal time.hour) | .h _ => some time.hour.toRelative | .K _ => some (time.hour.emod 12 (by decide)) - | .S _ => some time.nano + | .S _ => some time.nanosecond | .A _ => some time.toMilliseconds | .N _ => some time.toNanoseconds | _ => none @@ -299,7 +299,7 @@ def fromLeanTime24Hour (input : String) : Except String PlainTime := Formats a `PlainTime` value into a 24-hour format string (`HH:mm:ss.SSSSSSSSS`). -/ def toLeanTime24Hour (input : PlainTime) : String := - Formats.leanTime24Hour.formatBuilder input.hour input.minute input.second input.nano + Formats.leanTime24Hour.formatBuilder input.hour input.minute input.second input.nanosecond /-- Parses a time string in the 12-hour format (`hh:mm:ss aa`) and returns a `PlainTime`. @@ -441,7 +441,7 @@ def format (date : PlainDateTime) (format : String) : String := | .G _ => some date.era | .y _ => some date.year | .u _ => some date.year - | .D _ => some (Sigma.mk date.year.isLeap date.toOrdinal) + | .D _ => some (Sigma.mk date.year.isLeap date.dayOfYear) | .Qorq _ => some date.quarter | .w _ => some date.weekOfYear | .W _ => some date.alignedWeekOfMonth diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index fb06c3bd66af..8bdfcea96e4b 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -898,7 +898,7 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .G _ => date.era | .y _ => date.year | .u _ => date.year - | .D _ => Sigma.mk _ date.toOrdinal + | .D _ => Sigma.mk _ date.dayOfYear | .MorL _ => date.month | .d _ => date.day | .Qorq _ => date.quarter diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 522cd3d0d47a..67d4e62587f5 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -126,7 +126,7 @@ private def convertPlainDate (d : Std.Time.PlainDate) : MacroM (TSyntax `term) : `(Std.Time.PlainDate.clip $(← syntaxInt d.year) $(← syntaxBounded d.month.val) $(← syntaxBounded d.day.val)) private def convertPlainTime (d : Std.Time.PlainTime) : MacroM (TSyntax `term) := do - `(Std.Time.PlainTime.mk $(← syntaxBounded d.hour.val) $(← syntaxBounded d.minute.val) ⟨true, $(← syntaxBounded d.second.snd.val)⟩ $(← syntaxBounded d.nano.val)) + `(Std.Time.PlainTime.mk $(← syntaxBounded d.hour.val) $(← syntaxBounded d.minute.val) ⟨true, $(← syntaxBounded d.second.snd.val)⟩ $(← syntaxBounded d.nanosecond.val)) private def convertPlainDateTime (d : Std.Time.PlainDateTime) : MacroM (TSyntax `term) := do `(Std.Time.PlainDateTime.mk $(← convertPlainDate d.date) $(← convertPlainTime d.time)) diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 232651963c81..551d6076d05b 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -35,7 +35,7 @@ structure PlainTime where /-- `Nanoseconds` component of the `PlainTime` -/ - nano : Nanosecond.Ordinal + nanosecond : Nanosecond.Ordinal deriving Repr instance : Inhabited PlainTime where @@ -43,7 +43,7 @@ instance : Inhabited PlainTime where instance : BEq PlainTime where beq x y := x.hour.val == y.hour.val && x.minute == y.minute - && x.second.snd.val == y.second.snd.val && x.nano == y.nano + && x.second.snd.val == y.second.snd.val && x.nanosecond == y.nanosecond namespace PlainTime @@ -74,7 +74,7 @@ def toMilliseconds (time : PlainTime) : Millisecond.Offset := time.hour.toOffset.toMilliseconds + time.minute.toOffset.toMilliseconds + time.second.snd.toOffset.toMilliseconds + - time.nano.toOffset.toMilliseconds + time.nanosecond.toOffset.toMilliseconds /-- Converts a `PlainTime` value to the total number of nanoseconds. @@ -83,7 +83,7 @@ def toNanoseconds (time : PlainTime) : Nanosecond.Offset := time.hour.toOffset.toNanoseconds + time.minute.toOffset.toNanoseconds + time.second.snd.toOffset.toNanoseconds + - time.nano.toOffset + time.nanosecond.toOffset /-- Converts a `PlainTime` value to the total number of seconds. @@ -232,12 +232,21 @@ Creates a new `PlainTime` by adjusting the `minute` component to the given value def withMinutes (pt : PlainTime) (minute : Minute.Ordinal) : PlainTime := { pt with minute := minute } +/-- +Creates a new `PlainTime` by adjusting the milliseconds component inside the `nano` component of its `time` to the given value. +-/ +@[inline] +def withMillisecond (pt : PlainTime) (millis : Millisecond.Ordinal) : PlainTime := + let minorPart := pt.nanosecond.emod 1000 (by decide) + let majorPart := millis.mul_pos 1000000 (by decide) |>.addBounds minorPart + { pt with nanosecond := majorPart |>.expandTop (by decide) } + /-- Creates a new `PlainTime` by adjusting the `nano` component to the given value. -/ @[inline] def withNanoseconds (pt : PlainTime) (nano : Nanosecond.Ordinal) : PlainTime := - { pt with nano := nano } + { pt with nanosecond := nano } /-- Creates a new `PlainTime` by adjusting the `hour` component to the given value. @@ -246,6 +255,13 @@ Creates a new `PlainTime` by adjusting the `hour` component to the given value. def withHours (pt : PlainTime) (hour : Hour.Ordinal) : PlainTime := { pt with hour := hour } +/-- +`Millisecond` component of the `PlainTime` +-/ +@[inline] +def millisecond (pt : PlainTime) : Millisecond.Ordinal := + pt.nanosecond.ediv 1000000 (by decide) + instance : HAdd PlainTime Nanosecond.Offset PlainTime where hAdd := addNanoseconds diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 6e800570b52d..97aafebf131b 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -13,10 +13,10 @@ import Std.Time.Zoned.Database namespace Std namespace Time -namespace PlainDateTime - set_option linter.all true +namespace PlainDateTime + /-- Get the current time. -/ @@ -30,6 +30,27 @@ def now : IO PlainDateTime := do end PlainDateTime +namespace PlainDate + +/-- +Get the current date. +-/ +@[inline] +def now : IO PlainDate := + PlainDateTime.date <$> PlainDateTime.now + +end PlainDate +namespace PlainTime + +/-- +Get the current time. +-/ +@[inline] +def now : IO PlainTime := + PlainDateTime.time <$> PlainDateTime.now + +end PlainTime + namespace DateTime /-- diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index e4fa3f6901cc..081685978d6b 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -350,14 +350,14 @@ Getter for the `Milliseconds` inside of a `DateTime` -/ @[inline] def milliseconds (dt : DateTime tz) : Millisecond.Ordinal := - dt.date.get.time.nano.emod 1000 (by decide) + dt.date.get.time.nanosecond.emod 1000 (by decide) /-- Getter for the `Nanosecond` inside of a `DateTime` -/ @[inline] def nanoseconds (dt : DateTime tz) : Nanosecond.Ordinal := - dt.date.get.time.nano + dt.date.get.time.nanosecond /-- Gets the `Weekday` of a DateTime. @@ -381,8 +381,8 @@ def inLeapYear (date : DateTime tz) : Bool := /-- Determines the ordinal day of the year for the given `DateTime`. -/ -def toOrdinal (date : DateTime tz) : Day.Ordinal.OfYear date.year.isLeap := - ValidDate.toOrdinal ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ +def dayOfYear (date : DateTime tz) : Day.Ordinal.OfYear date.year.isLeap := + ValidDate.dayOfYear ⟨⟨date.month, date.day⟩, date.date.get.date.valid⟩ /-- Determines the week of the year for the given `DateTime`. diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 65041ffa5cac..cbc68649a419 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -111,6 +111,13 @@ Converts a `ZonedDateTime` to a `PlainDateTime` def toDateTime (dt : ZonedDateTime) : DateTime dt.timezone := DateTime.ofTimestamp dt.timestamp dt.timezone +/-- +Getter for the `PlainTime` inside of a `ZonedDateTime` +-/ +@[inline] +def time (zdt : ZonedDateTime) : PlainTime := + zdt.date.get.time + /-- Getter for the `Year` inside of a `ZonedDateTime` -/ @@ -153,12 +160,19 @@ Getter for the `Second` inside of a `ZonedDateTime` def second (zdt : ZonedDateTime) : Second.Ordinal zdt.date.get.time.second.fst := zdt.date.get.time.second.snd +/-- +Getter for the `Millisecond` inside of a `ZonedDateTime`. +-/ +@[inline] +def millisecond (dt : ZonedDateTime) : Millisecond.Ordinal := + dt.date.get.time.millisecond + /-- Getter for the `Nanosecond` inside of a `ZonedDateTime` -/ @[inline] def nanosecond (zdt : ZonedDateTime) : Nanosecond.Ordinal := - zdt.date.get.time.nano + zdt.date.get.time.nanosecond /-- Getter for the `TimeZone.Offset` inside of a `ZonedDateTime` @@ -178,8 +192,28 @@ def weekday (zdt : ZonedDateTime) : Weekday := Transforms a tuple of a `ZonedDateTime` into a `Day.Ordinal.OfYear`. -/ @[inline] -def toOrdinal (date : ZonedDateTime) : Day.Ordinal.OfYear date.year.isLeap := - ValidDate.toOrdinal ⟨(date.month, date.day), date.date.get.date.valid⟩ +def dayOfYear (date : ZonedDateTime) : Day.Ordinal.OfYear date.year.isLeap := + ValidDate.dayOfYear ⟨(date.month, date.day), date.date.get.date.valid⟩ + +/-- +Determines the week of the year for the given `ZonedDateTime`. +-/ +@[inline] +def weekOfYear (date : ZonedDateTime) : Week.Ordinal := + date.date.get.weekOfYear + +/-- +Returns the unaligned week of the month for a `ZonedDateTime` (day divided by 7, plus 1). +-/ +def weekOfMonth (date : ZonedDateTime) : Internal.Bounded.LE 1 5 := + date.date.get.weekOfMonth + +/-- +Determines the quarter of the year for the given `ZonedDateTime`. +-/ +@[inline] +def quarter (date : ZonedDateTime) : Internal.Bounded.LE 1 4 := + date.date.get.quarter /-- Add `Day.Offset` to a `ZonedDateTime`. @@ -328,6 +362,13 @@ Determines the era of the given `ZonedDateTime` based on its year. def era (date : ZonedDateTime) : Year.Era := date.date.get.era +/-- +Sets the `ZonedDateTime` to the specified `desiredWeekday`. +-/ +def withWeekday (dt : ZonedDateTime) (desiredWeekday : Weekday) : ZonedDateTime := + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withWeekday desiredWeekday) dt.rules + /-- Creates a new `ZonedDateTime` by adjusting the day of the month to the given `days` value, with any out-of-range days clipped to the nearest valid date. @@ -407,6 +448,15 @@ def withSeconds (dt : ZonedDateTime) (second : Sigma Second.Ordinal) : ZonedDate let date := dt.date.get ZonedDateTime.ofPlainDateTime (date.withSeconds second) dt.rules +/-- +Creates a new `ZonedDateTime` by adjusting the `nano` component with a new `millis` that will set +in the millisecond scale. +-/ +@[inline] +def withMillisecond (dt : ZonedDateTime) (millis : Millisecond.Ordinal) : ZonedDateTime := + let date := dt.date.get + ZonedDateTime.ofPlainDateTime (date.withMillisecond millis) dt.rules + /-- Creates a new `ZonedDateTime` by adjusting the `nano` component. -/ From 646c8d3c34c1858ec67c6e9b5a8bb4f5d265cb58 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 31 Oct 2024 02:38:04 -0300 Subject: [PATCH 086/118] fix: some functions and tests --- src/Std/Time/DateTime/PlainDateTime.lean | 6 +-- src/Std/Time/Time/PlainTime.lean | 14 +++--- src/Std/Time/Zoned/ZonedDateTime.lean | 4 +- tests/lean/run/nicerNestedDos.lean | 57 ++++-------------------- tests/lean/run/timeIO.lean | 0 tests/lean/run/timeParse.lean | 1 - 6 files changed, 20 insertions(+), 62 deletions(-) create mode 100644 tests/lean/run/timeIO.lean diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 3cf479ea93d8..c9527dd1ae68 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -205,15 +205,15 @@ def withSeconds (dt : PlainDateTime) (second : Sigma Second.Ordinal) : PlainDate Creates a new `PlainDateTime` by adjusting the milliseconds component inside the `nano` component of its `time` to the given value. -/ @[inline] -def withMillisecond (dt : PlainDateTime) (millis : Millisecond.Ordinal) : PlainDateTime := - { dt with time := dt.time.withMillisecond millis } +def withMilliseconds (dt : PlainDateTime) (millis : Millisecond.Ordinal) : PlainDateTime := + { dt with time := dt.time.withMilliseconds millis } /-- Creates a new `PlainDateTime` by adjusting the `nano` component of its `time` to the given value. -/ @[inline] def withNanoseconds (dt : PlainDateTime) (nano : Nanosecond.Ordinal) : PlainDateTime := - { dt with time := { dt.time with nanosecond := nano } } + { dt with time := dt.time.withNanoseconds nano } /-- Adds a `Day.Offset` to a `PlainDateTime`. diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean index 551d6076d05b..3c0b2d3e0061 100644 --- a/src/Std/Time/Time/PlainTime.lean +++ b/src/Std/Time/Time/PlainTime.lean @@ -152,8 +152,8 @@ Adds seconds to a `PlainTime`. -/ @[inline] def addSeconds (time : PlainTime) (secondsToAdd : Second.Offset) : PlainTime := - let totalSeconds := time.toSeconds + secondsToAdd - ofSeconds totalSeconds + let totalSeconds := time.toNanoseconds + secondsToAdd.toNanoseconds + ofNanoseconds totalSeconds /-- Subtracts seconds from a `PlainTime`. @@ -167,8 +167,8 @@ Adds minutes to a `PlainTime`. -/ @[inline] def addMinutes (time : PlainTime) (minutesToAdd : Minute.Offset) : PlainTime := - let total := time.toSeconds + minutesToAdd.toSeconds - ofSeconds total + let total := time.toNanoseconds + minutesToAdd.toNanoseconds + ofNanoseconds total /-- Subtracts minutes from a `PlainTime`. @@ -182,8 +182,8 @@ Adds hours to a `PlainTime`. -/ @[inline] def addHours (time : PlainTime) (hoursToAdd : Hour.Offset) : PlainTime := - let total := time.toSeconds + hoursToAdd.toSeconds - ofSeconds total + let total := time.toNanoseconds + hoursToAdd.toNanoseconds + ofNanoseconds total /-- Subtracts hours from a `PlainTime`. @@ -236,7 +236,7 @@ def withMinutes (pt : PlainTime) (minute : Minute.Ordinal) : PlainTime := Creates a new `PlainTime` by adjusting the milliseconds component inside the `nano` component of its `time` to the given value. -/ @[inline] -def withMillisecond (pt : PlainTime) (millis : Millisecond.Ordinal) : PlainTime := +def withMilliseconds (pt : PlainTime) (millis : Millisecond.Ordinal) : PlainTime := let minorPart := pt.nanosecond.emod 1000 (by decide) let majorPart := millis.mul_pos 1000000 (by decide) |>.addBounds minorPart { pt with nanosecond := majorPart |>.expandTop (by decide) } diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index cbc68649a419..106c02fce201 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -453,9 +453,9 @@ Creates a new `ZonedDateTime` by adjusting the `nano` component with a new `mill in the millisecond scale. -/ @[inline] -def withMillisecond (dt : ZonedDateTime) (millis : Millisecond.Ordinal) : ZonedDateTime := +def withMilliseconds (dt : ZonedDateTime) (millis : Millisecond.Ordinal) : ZonedDateTime := let date := dt.date.get - ZonedDateTime.ofPlainDateTime (date.withMillisecond millis) dt.rules + ZonedDateTime.ofPlainDateTime (date.withMilliseconds millis) dt.rules /-- Creates a new `ZonedDateTime` by adjusting the `nano` component. diff --git a/tests/lean/run/nicerNestedDos.lean b/tests/lean/run/nicerNestedDos.lean index d952a57568d1..09ff22f0bb9a 100644 --- a/tests/lean/run/nicerNestedDos.lean +++ b/tests/lean/run/nicerNestedDos.lean @@ -1,49 +1,8 @@ -def f (x : Nat) : IO Nat := do - IO.println "hello" - if x > 5 then - IO.println ("x: " ++ toString x) - IO.println "done" - pure (x + 1) - -/-- -info: hello ---- -info: 3 --/ -#guard_msgs in -#eval f 2 - -/-- -info: hello -x: 10 -done ---- -info: 11 --/ -#guard_msgs in -#eval f 10 - -def g (x : Nat) : StateT Nat Id Unit := do - if x > 10 then - let s ← get - set (s + x) - pure () - -theorem ex1 : (g 10).run 1 = ((), 1) := -rfl - -theorem ex2 : (g 20).run 1 = ((), 21) := -rfl - -def h (x : Nat) : StateT Nat Id Unit := do -if x > 10 then { - let s ← get; -set (s + x) -- we don't need to respect indentation when `{` `}` are used -} -pure () - -theorem ex3 : (h 10).run 1 = ((), 1) := -rfl - -theorem ex4 : (h 20).run 1 = ((), 21) := -rfl +import Std.Time +import Init +open Std.Time + +#eval do + let res ← Database.defaultGetZoneRulesAt "America/Sao_Paulo" + if res.transitions.size < 1 then + throw <| IO.userError "invalid quantity for America/Sao_Paulo" diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean new file mode 100644 index 000000000000..e69de29bb2d1 diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index 6d1b7884084e..a16c3df1cbfc 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -16,7 +16,6 @@ def Full12HourWrong : GenericFormat .any := datespec("MM/dd/uuuu hh:mm:ss aa XXX -- Dates - def brTZ : TimeZone := timezone("America/Sao_Paulo -03:00") def jpTZ : TimeZone := timezone("Asia/Tokyo +09:00") From b2fc40ac03a664d03bf94763b365df4da041b292 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 31 Oct 2024 03:03:32 -0300 Subject: [PATCH 087/118] fix: tests in tzif parse --- tests/lean/run/timeTzifParse.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index a9b7765219a2..e0f5f2e492be 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -59,7 +59,7 @@ info: 0 #eval code.v1.utLocalIndicators.size /-- -info: zoned("1969-12-30T21:00:00.000000000-03:00") +info: zoned("1969-12-31T21:00:00.000000000-03:00") -/ #guard_msgs in #eval ZonedDateTime.ofTimestamp (Timestamp.ofSecondsSinceUnixEpoch 0) rules From c2392aa57c00b8f07f7f3752189f8eb695d50e4d Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 1 Nov 2024 13:27:56 -0300 Subject: [PATCH 088/118] fix: memory leak in getZoneRules and tests --- src/Std/Time/Notation.lean | 2 +- src/Std/Time/Zoned.lean | 8 ++-- src/Std/Time/Zoned/Database.lean | 20 +++++---- src/Std/Time/Zoned/Database/Basic.lean | 8 ++-- src/Std/Time/Zoned/Database/TZdb.lean | 4 +- src/Std/Time/Zoned/Database/Windows.lean | 8 +++- tests/lean/run/nicerNestedDos.lean | 57 ++++++++++++++++++++---- tests/lean/run/timeIO.lean | 13 ++++++ 8 files changed, 91 insertions(+), 29 deletions(-) diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 67d4e62587f5..08520c18ae72 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -135,7 +135,7 @@ private def convertZonedDateTime (d : Std.Time.ZonedDateTime) (identifier := fal let plain ← convertPlainDateTime d.toPlainDateTime if identifier then - `(Std.Time.ZonedDateTime.ofPlainDateTime $plain <$> Std.Time.Database.defaultGetZoneRulesAt $(Syntax.mkStrLit d.timezone.name)) + `(Std.Time.ZonedDateTime.ofPlainDateTime $plain <$> Std.Time.Database.defaultGetZoneRules $(Syntax.mkStrLit d.timezone.name)) else `(Std.Time.ZonedDateTime.ofPlainDateTime $plain (Std.Time.TimeZone.ZoneRules.ofTimeZone $(← convertTimezone d.timezone))) diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 97aafebf131b..dfe54dc5eee0 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -23,7 +23,7 @@ Get the current time. @[inline] def now : IO PlainDateTime := do let tm ← Timestamp.now - let rules ← Database.defaultGetLocalZoneRulesAt + let rules ← Database.defaultGetLocalZoneRules let ltt := rules.findLocalTimeTypeForTimestamp tm return PlainDateTime.ofTimestamp tm |>.addSeconds ltt.getTimeZone.toSeconds @@ -83,7 +83,7 @@ Gets the current `ZonedDateTime`. @[inline] def now : IO ZonedDateTime := do let tm ← Timestamp.now - let rules ← Database.defaultGetLocalZoneRulesAt + let rules ← Database.defaultGetLocalZoneRules return ZonedDateTime.ofTimestamp tm rules /-- @@ -92,7 +92,7 @@ Gets the current `ZonedDateTime` using the identifier of a time zone. @[inline] def nowAt (id : String) : IO ZonedDateTime := do let tm ← Timestamp.now - let rules ← Database.defaultGetZoneRulesAt id + let rules ← Database.defaultGetZoneRules id return ZonedDateTime.ofTimestamp tm rules /-- @@ -121,7 +121,7 @@ Creates a new `ZonedDateTime` out of a `PlainDateTime` and a time zone identifie -/ @[inline] def of (pdt : PlainDateTime) (id : String) : IO ZonedDateTime := do - let zr ← Database.defaultGetZoneRulesAt id + let zr ← Database.defaultGetZoneRules id return ZonedDateTime.ofPlainDateTime pdt zr end ZonedDateTime diff --git a/src/Std/Time/Zoned/Database.lean b/src/Std/Time/Zoned/Database.lean index fd6940fb4fad..66e206a998db 100644 --- a/src/Std/Time/Zoned/Database.lean +++ b/src/Std/Time/Zoned/Database.lean @@ -18,17 +18,21 @@ open TimeZone.ZoneRules set_option linter.all true /-- -Gets the time zone for a given location and timestamp, handling Windows and non-Windows platforms. +Gets the zone rules for a specific time zone identifier, handling Windows and non-Windows platforms. +In windows it uses the current `icu.h` in Windows SDK. If it's linux or macos then it will use the `tzdata` +files. -/ -def defaultGetZoneRulesAt : String → IO TimeZone.ZoneRules := +def defaultGetZoneRules : String → IO TimeZone.ZoneRules := if System.Platform.isWindows - then getZoneRulesAt WindowsDb.default - else getZoneRulesAt TZdb.default + then getZoneRules WindowsDb.default + else getZoneRules TZdb.default /-- -Gets the local time zone for a specific timestamp, accounting for platform differences. +Gets the local zone rules, accounting for platform differences. +In windows it uses the current `icu.h` in Windows SDK. If it's linux or macos then it will use the `tzdata` +files. -/ -def defaultGetLocalZoneRulesAt : IO TimeZone.ZoneRules := +def defaultGetLocalZoneRules : IO TimeZone.ZoneRules := if System.Platform.isWindows - then getLocalZoneRulesAt WindowsDb.default - else getLocalZoneRulesAt TZdb.default + then getLocalZoneRules WindowsDb.default + else getLocalZoneRules TZdb.default diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 5b7ee9e4bfc1..349e9fadb926 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -18,14 +18,14 @@ A timezone database from which we can read the `ZoneRules` of some area by it's protected class Database (α : Type) where /-- - Retrieves the timezone information (`ZoneRules`) for a given area at a specific point in time. + Retrieves the zone rules information (`ZoneRules`) for a given area at a specific point in time. -/ - getZoneRulesAt : α → String → IO TimeZone.ZoneRules + getZoneRules : α → String → IO TimeZone.ZoneRules /-- - Retrieves the local timezone information (`ZoneRules`) at a given timestamp. + Retrieves the local zone rules information (`ZoneRules`) at a given timestamp. -/ - getLocalZoneRulesAt : α → IO TimeZone.ZoneRules + getLocalZoneRules : α → IO TimeZone.ZoneRules namespace TimeZone diff --git a/src/Std/Time/Zoned/Database/TZdb.lean b/src/Std/Time/Zoned/Database/TZdb.lean index ec6242eb1caa..40fe0f7fd13e 100644 --- a/src/Std/Time/Zoned/Database/TZdb.lean +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -88,5 +88,5 @@ def readRulesFromDisk (path : System.FilePath) (id : String) : IO ZoneRules := d parseTZIfFromDisk (System.FilePath.join path id) id instance : Std.Time.Database TZdb where - getLocalZoneRulesAt db := localRules db.localPath - getZoneRulesAt db id := readRulesFromDisk db.zonesPath id + getLocalZoneRules db := localRules db.localPath + getZoneRules db id := readRulesFromDisk db.zonesPath id diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index 11de9493de50..241738c967bd 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -47,6 +47,10 @@ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do wall := .wall, utLocal := .local }} + + if res.fst ≤ start then + break + start := res.fst else break @@ -75,5 +79,5 @@ Returns a default `WindowsDb` instance. def default : WindowsDb := {} instance : Std.Time.Database WindowsDb where - getZoneRulesAt _ id := Windows.getZoneRules id - getLocalZoneRulesAt _ := Windows.getZoneRules =<< Windows.getLocalTimeZoneIdentifierAt (-2147483648) + getZoneRules _ id := Windows.getZoneRules id + getLocalZoneRules _ := Windows.getZoneRules =<< Windows.getLocalTimeZoneIdentifierAt (-2147483648) diff --git a/tests/lean/run/nicerNestedDos.lean b/tests/lean/run/nicerNestedDos.lean index 09ff22f0bb9a..d952a57568d1 100644 --- a/tests/lean/run/nicerNestedDos.lean +++ b/tests/lean/run/nicerNestedDos.lean @@ -1,8 +1,49 @@ -import Std.Time -import Init -open Std.Time - -#eval do - let res ← Database.defaultGetZoneRulesAt "America/Sao_Paulo" - if res.transitions.size < 1 then - throw <| IO.userError "invalid quantity for America/Sao_Paulo" +def f (x : Nat) : IO Nat := do + IO.println "hello" + if x > 5 then + IO.println ("x: " ++ toString x) + IO.println "done" + pure (x + 1) + +/-- +info: hello +--- +info: 3 +-/ +#guard_msgs in +#eval f 2 + +/-- +info: hello +x: 10 +done +--- +info: 11 +-/ +#guard_msgs in +#eval f 10 + +def g (x : Nat) : StateT Nat Id Unit := do + if x > 10 then + let s ← get + set (s + x) + pure () + +theorem ex1 : (g 10).run 1 = ((), 1) := +rfl + +theorem ex2 : (g 20).run 1 = ((), 21) := +rfl + +def h (x : Nat) : StateT Nat Id Unit := do +if x > 10 then { + let s ← get; +set (s + x) -- we don't need to respect indentation when `{` `}` are used +} +pure () + +theorem ex3 : (h 10).run 1 = ((), 1) := +rfl + +theorem ex4 : (h 20).run 1 = ((), 21) := +rfl diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index e69de29bb2d1..6adbf9db23ae 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -0,0 +1,13 @@ +import Std.Time +import Init +open Std.Time + +#eval do + let res ← Database.defaultGetZoneRules "America/Sao_Paulo" + if res.transitions.size < 1 then + throw <| IO.userError "invalid quantity for America/Sao_Paulo" + +#eval do + let res ← Database.defaultGetLocalZoneRules + if res.transitions.size < 1 then + throw <| IO.userError "invalid quantity for the local timezone" From bb2459741c7d76036f339a7f549b9f363acd9e23 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 1 Nov 2024 13:41:35 -0300 Subject: [PATCH 089/118] chore: small fixes in the documentation --- src/Std/Time/Duration.lean | 2 +- src/Std/Time/Zoned/DateTime.lean | 6 +++--- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index b2d2d945df82..f5ad1ebc3a8c 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -14,7 +14,7 @@ open Internal set_option linter.all true /-- -Represents a point in time relative to the UNIX Epoch, with nanoseconds precision. +Represents a time interval with nanoseconds precision. -/ structure Duration where diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 081685978d6b..4e06b1f55567 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -44,7 +44,7 @@ Creates a new `DateTime` out of a `Timestamp` that is in a `TimeZone`. -/ @[inline] def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := - DateTime.mk tm (Thunk.mk fun _ => (tm.addSeconds tz.toSeconds).toPlainDateTimeAssumingUTC) + DateTime.mk tm (Thunk.mk fun _ => tm.toPlainDateTimeAssumingUTC |>.addSeconds tz.toSeconds) /-- Creates a `Timestamp` out of a `DateTime`. @@ -75,8 +75,8 @@ is relative to the given `TimeZone`. -/ @[inline] def ofLocalDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := - let tm := Timestamp.ofPlainDateTimeAssumingUTC date - DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk fun _ => date) + let tm := date.subSeconds tz.toSeconds + DateTime.mk (Timestamp.ofPlainDateTimeAssumingUTC tm) (Thunk.mk fun _ => date) /-- Add `Hour.Offset` to a `DateTime`. From 13bd57803c229835cc3d14fd55d4b1f2bfa15c07 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 4 Nov 2024 19:20:04 -0300 Subject: [PATCH 090/118] fix: nix ci --- nix/bootstrap.nix | 4 ++-- script/prepare-llvm-mingw.sh | 3 ++- 2 files changed, 4 insertions(+), 3 deletions(-) diff --git a/nix/bootstrap.nix b/nix/bootstrap.nix index 661fa34230e7..170b807494df 100644 --- a/nix/bootstrap.nix +++ b/nix/bootstrap.nix @@ -1,5 +1,5 @@ { src, debug ? false, stage0debug ? false, extraCMakeFlags ? [], - stdenv, lib, cmake, gmp, libuv, cadical, git, gnumake, bash, buildLeanPackage, writeShellScriptBin, runCommand, symlinkJoin, lndir, perl, gnused, darwin, llvmPackages, linkFarmFromDrvs, + stdenv, lib, cmake, gmp, libuv, cadical, git, tzdata, gnumake, bash, buildLeanPackage, writeShellScriptBin, runCommand, symlinkJoin, lndir, perl, gnused, darwin, llvmPackages, linkFarmFromDrvs, ... } @ args: with builtins; lib.warn "The Nix-based build is deprecated" rec { @@ -159,7 +159,7 @@ lib.warn "The Nix-based build is deprecated" rec { test = buildCMake { name = "lean-test-${desc}"; realSrc = lib.sourceByRegex src [ "src.*" "tests.*" ]; - buildInputs = [ gmp libuv perl git cadical ]; + buildInputs = [ gmp libuv perl git cadical tzdata ]; preConfigure = '' cd src ''; diff --git a/script/prepare-llvm-mingw.sh b/script/prepare-llvm-mingw.sh index f0ca954664af..f3e1b0a65a20 100644 --- a/script/prepare-llvm-mingw.sh +++ b/script/prepare-llvm-mingw.sh @@ -40,7 +40,8 @@ cp /clang64/lib/{crtbegin,crtend,crt2,dllcrt2}.o stage1/lib/ cp /clang64/lib/lib{m,bcrypt,mingw32,moldname,mingwex,msvcrt,pthread,advapi32,shell32,user32,kernel32,ucrtbase,psapi,iphlpapi,userenv,ws2_32,dbghelp,ole32,icu}.* /clang64/lib/libgmp.a /clang64/lib/libuv.a llvm/lib/lib{c++,c++abi,unwind}.a stage1/lib/ echo -n " -DLEAN_STANDALONE=ON" echo -n " -DCMAKE_C_COMPILER=$PWD/stage1/bin/clang.exe -DCMAKE_C_COMPILER_WORKS=1 -DCMAKE_CXX_COMPILER=$PWD/llvm/bin/clang++.exe -DCMAKE_CXX_COMPILER_WORKS=1 -DLEAN_CXX_STDLIB='-lc++ -lc++abi'" -echo -n " -DLEAN_EXTRA_CXX_FLAGS='-idirafter \"${WINDOWS_SDK_DIR}/Include/10.0.22621.0/um\" --sysroot $PWD/llvm -idirafter /clang64/include/'" +echo -n " -DSTAGE0_CMAKE_C_COMPILER=clang -DSTAGE0_CMAKE_CXX_COMPILER=clang++" +echo -n " -DLEAN_EXTRA_CXX_FLAGS='--sysroot $PWD/llvm -idirafter /clang64/include/'" echo -n " -DLEANC_INTERNAL_FLAGS='--sysroot ROOT -nostdinc -isystem ROOT/include/clang' -DLEANC_CC=ROOT/bin/clang.exe" echo -n " -DLEANC_INTERNAL_LINKER_FLAGS='-L ROOT/lib -static-libgcc -Wl,-Bstatic -lgmp $(pkg-config --static --libs libuv) -lunwind -Wl,-Bdynamic -fuse-ld=lld'" # when not using the above flags, link GMP dynamically/as usual. Always link ICU dynamically. From 596588d30b3c96e9f9f0e75af40e42ef5e982bfc Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 4 Nov 2024 20:02:46 -0300 Subject: [PATCH 091/118] fix: trying to add nix ci again --- nix/buildLeanPackage.nix | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/nix/buildLeanPackage.nix b/nix/buildLeanPackage.nix index c8b82ebde648..8cd4015fd5b6 100644 --- a/nix/buildLeanPackage.nix +++ b/nix/buildLeanPackage.nix @@ -1,5 +1,5 @@ { lean, lean-leanDeps ? lean, lean-final ? lean, leanc, - stdenv, lib, coreutils, gnused, writeShellScriptBin, bash, substituteAll, symlinkJoin, linkFarmFromDrvs, + stdenv, lib, coreutils, tzdata, gnused, writeShellScriptBin, bash, substituteAll, symlinkJoin, linkFarmFromDrvs, runCommand, darwin, mkShell, ... }: let lean-final' = lean-final; in lib.makeOverridable ( @@ -45,7 +45,7 @@ with builtins; let name = lib.strings.sanitizeDerivationName args.name; stdenv = bareStdenv; inherit (stdenv) system; - buildInputs = (args.buildInputs or []) ++ [ coreutils ]; + buildInputs = (args.buildInputs or []) ++ [ coreutils tzdata ]; builder = stdenv.shell; args = [ "-c" '' source $stdenv/setup From 0fef367c9c5d22a744daba589f28a144bdc90ea3 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 5 Nov 2024 22:15:35 -0300 Subject: [PATCH 092/118] Update src/Std/Time/DateTime/PlainDateTime.lean Co-authored-by: Markus Himmel --- src/Std/Time/DateTime/PlainDateTime.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index c9527dd1ae68..c87ce51b2fff 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -44,7 +44,7 @@ def toTimestampAssumingUTC (dt : PlainDateTime) : Timestamp := Timestamp.ofNanosecondsSinceUnixEpoch (Nanosecond.Offset.ofInt nanos) /-- -Converts a UNIX `Timestamp` to a `PlainDateTime`. +Converts a `Timestamp` to a `PlainDateTime`. -/ def ofTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do let leapYearEpoch := 11017 From 35c014a5c210d019c3c5eb052142750cffbe120a Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 5 Nov 2024 23:03:14 -0300 Subject: [PATCH 093/118] feat: all the functions in the sheet --- src/Std/Time/Date/Basic.lean | 174 +++++++++++++++++++++++ src/Std/Time/Date/PlainDate.lean | 16 +-- src/Std/Time/Date/Unit/Month.lean | 8 -- src/Std/Time/Date/Unit/Week.lean | 4 +- src/Std/Time/DateTime/PlainDateTime.lean | 8 +- src/Std/Time/DateTime/Timestamp.lean | 20 +++ src/Std/Time/Duration.lean | 23 +++ src/Std/Time/Notation.lean | 2 +- src/Std/Time/Zoned.lean | 44 +++++- src/Std/Time/Zoned/DateTime.lean | 26 ++++ src/Std/Time/Zoned/ZoneRules.lean | 11 +- src/Std/Time/Zoned/ZonedDateTime.lean | 32 ++++- tests/lean/run/timeFormats.lean | 2 +- 13 files changed, 342 insertions(+), 28 deletions(-) diff --git a/src/Std/Time/Date/Basic.lean b/src/Std/Time/Date/Basic.lean index e4b124774810..5c3cca0a97e6 100644 --- a/src/Std/Time/Date/Basic.lean +++ b/src/Std/Time/Date/Basic.lean @@ -6,3 +6,177 @@ Authors: Sofia Rodrigues prelude import Std.Time.Date.Unit.Basic import Std.Time.Date.ValidDate +import Std.Time.Time.Basic + +namespace Std +namespace Time + +namespace Nanosecond +namespace Offset + +/-- +Convert `Nanosecond.Offset` into `Day.Offset`. +-/ +@[inline] +def toDays (nanoseconds : Nanosecond.Offset) : Day.Offset := + nanoseconds.div 86400000000000 + +/-- +Convert `Day.Offset` into `Nanosecond.Offset`. +-/ +@[inline] +def ofDays (days : Day.Offset) : Nanosecond.Offset := + days.mul 86400000000000 + +/-- +Convert `Nanosecond.Offset` into `Week.Offset`. +-/ +@[inline] +def toWeeks (nanoseconds : Nanosecond.Offset) : Week.Offset := + nanoseconds.div 604800000000000 + +/-- +Convert `Week.Offset` into `Nanosecond.Offset`. +-/ +@[inline] +def ofWeeks (weeks : Week.Offset) : Nanosecond.Offset := + weeks.mul 604800000000000 + +end Offset +end Nanosecond + +namespace Millisecond +namespace Offset + +/-- +Convert `Millisecond.Offset` into `Day.Offset`. +-/ +@[inline] +def toDays (milliseconds : Millisecond.Offset) : Day.Offset := + milliseconds.div 86400000 + +/-- +Convert `Day.Offset` into `Millisecond.Offset`. +-/ +@[inline] +def ofDays (days : Day.Offset) : Millisecond.Offset := + days.mul 86400000 + +/-- +Convert `Millisecond.Offset` into `Week.Offset`. +-/ +@[inline] +def toWeeks (milliseconds : Millisecond.Offset) : Week.Offset := + milliseconds.div 604800000 + +/-- +Convert `Week.Offset` into `Millisecond.Offset`. +-/ +@[inline] +def ofWeeks (weeks : Week.Offset) : Millisecond.Offset := + weeks.mul 604800000 + +end Offset +end Millisecond + +namespace Second +namespace Offset + +/-- +Convert `Second.Offset` into `Day.Offset`. +-/ +@[inline] +def toDays (seconds : Second.Offset) : Day.Offset := + seconds.div 86400 + +/-- +Convert `Day.Offset` into `Second.Offset`. +-/ +@[inline] +def ofDays (days : Day.Offset) : Second.Offset := + days.mul 86400 + +/-- +Convert `Second.Offset` into `Week.Offset`. +-/ +@[inline] +def toWeeks (seconds : Second.Offset) : Week.Offset := + seconds.div 604800 + +/-- +Convert `Week.Offset` into `Second.Offset`. +-/ +@[inline] +def ofWeeks (weeks : Week.Offset) : Second.Offset := + weeks.mul 604800 + +end Offset +end Second + +namespace Minute +namespace Offset + +/-- +Convert `Minute.Offset` into `Day.Offset`. +-/ +@[inline] +def toDays (minutes : Minute.Offset) : Day.Offset := + minutes.div 1440 + +/-- +Convert `Day.Offset` into `Minute.Offset`. +-/ +@[inline] +def ofDays (days : Day.Offset) : Minute.Offset := + days.mul 1440 + +/-- +Convert `Minute.Offset` into `Week.Offset`. +-/ +@[inline] +def toWeeks (minutes : Minute.Offset) : Week.Offset := + minutes.div 10080 + +/-- +Convert `Week.Offset` into `Minute.Offset`. +-/ +@[inline] +def ofWeeks (weeks : Week.Offset) : Minute.Offset := + weeks.mul 10080 + +end Offset +end Minute + +namespace Hour +namespace Offset + +/-- +Convert `Hour.Offset` into `Day.Offset`. +-/ +@[inline] +def toDays (hours : Hour.Offset) : Day.Offset := + hours.div 24 + +/-- +Convert `Day.Offset` into `Hour.Offset`. +-/ +@[inline] +def ofDays (days : Day.Offset) : Hour.Offset := + days.mul 24 + +/-- +Convert `Hour.Offset` into `Week.Offset`. +-/ +@[inline] +def toWeeks (hours : Hour.Offset) : Week.Offset := + hours.div 168 + +/-- +Convert `Week.Offset` into `Hour.Offset`. +-/ +@[inline] +def ofWeeks (weeks : Week.Offset) : Hour.Offset := + weeks.mul 168 + +end Offset +end Hour diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 0170c5f7bfc8..0d8a6f5ecc6f 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -48,7 +48,7 @@ Creates a `PlainDate` by clipping the day to ensure validity. This function forc valid by adjusting the day to fit within the valid range to fit the given month and year. -/ @[inline] -def clip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := +def ofYearMonthDayClip (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := let day := month.clipDay year.isLeap day PlainDate.mk year month day Month.Ordinal.valid_clipDay @@ -86,7 +86,7 @@ def ofDaysSinceUNIXEpoch (day : Day.Offset) : PlainDate := let d := doy - (153 * mp + 2).tdiv 5 + 1 let m := mp + (if mp < 10 then 3 else -9) let y := y + (if m <= 2 then 1 else 0) - .clip y (.clip m (by decide)) (.clip d (by decide)) + .ofYearMonthDayClip y (.clip m (by decide)) (.clip d (by decide)) /-- Returns the unaligned week of the month for a `PlainDate` (day divided by 7, plus 1). @@ -173,7 +173,7 @@ def addMonthsClip (date : PlainDate) (months : Month.Offset) : PlainDate := let totalMonths : Int := totalMonths let wrappedMonths := Bounded.LE.byEmod totalMonths 12 (by decide) |>.add 1 let yearsOffset := totalMonths / 12 - PlainDate.clip (date.year.add yearsOffset) wrappedMonths date.day + PlainDate.ofYearMonthDayClip (date.year.add yearsOffset) wrappedMonths date.day /-- Subtracts `Month.Offset` from a `PlainDate`, it clips the day to the last valid day of that month. @@ -186,7 +186,7 @@ def subMonthsClip (date : PlainDate) (months : Month.Offset) : PlainDate := Creates a `PlainDate` by rolling over the extra days to the next month. -/ def rollOver (year : Year.Offset) (month : Month.Ordinal) (day : Day.Ordinal) : PlainDate := - clip year month 1 |>.addDays (day.toOffset - 1) + ofYearMonthDayClip year month 1 |>.addDays (day.toOffset - 1) /-- Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day remain unchanged, @@ -194,7 +194,7 @@ and any invalid days for the new year will be handled according to the `clip` be -/ @[inline] def withYearClip (dt : PlainDate) (year : Year.Offset) : PlainDate := - clip year dt.month dt.day + ofYearMonthDayClip year dt.month dt.day /-- Creates a new `PlainDate` by adjusting the year to the given `year` value. The month and day are rolled @@ -208,7 +208,7 @@ def withYearRollOver (dt : PlainDate) (year : Year.Offset) : PlainDate := Adds a given number of months to a `PlainDate`, rolling over any excess days into the following month. -/ def addMonthsRollOver (date : PlainDate) (months : Month.Offset) : PlainDate := - addMonthsClip (clip date.year date.month 1) months + addMonthsClip (ofYearMonthDayClip date.year date.month 1) months |>.addDays (date.day.toOffset - 1) /-- @@ -252,7 +252,7 @@ out-of-range days clipped to the nearest valid date. -/ @[inline] def withDaysClip (dt : PlainDate) (days : Day.Ordinal) : PlainDate := - clip dt.year dt.month days + ofYearMonthDayClip dt.year dt.month days /-- Creates a new `PlainDate` by adjusting the day of the month to the given `days` value, with any @@ -268,7 +268,7 @@ The day remains unchanged, and any invalid days for the new month will be handle -/ @[inline] def withMonthClip (dt : PlainDate) (month : Month.Ordinal) : PlainDate := - clip dt.year month dt.day + ofYearMonthDayClip dt.year month dt.day /-- Creates a new `PlainDate` by adjusting the month to the given `month` value. diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index c292bd092fb6..b607ccee96e2 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -62,14 +62,6 @@ end Quarter namespace Offset -/-- -Creates an `Ordinal` from a `Fin`, ensuring the value is within bounds, if its 0 then its converted -to 1. --/ -@[inline] -def ofFin (data : Fin 13) : Ordinal := - Bounded.LE.ofFin' data (by decide) - /-- Creates an `Offset` from a natural number. -/ diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index af6860b4c9ab..1f6311feb97f 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -168,8 +168,8 @@ def toDays (weeks : Week.Offset) : Day.Offset := Convert `Day.Offset` into `Week.Offset`. -/ @[inline] -def ofDays (hours : Day.Offset) : Week.Offset := - hours.ediv 7 +def ofDays (days : Day.Offset) : Week.Offset := + days.ediv 7 end Offset end Week diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index c9527dd1ae68..bf7b1dd9ef8d 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -122,7 +122,7 @@ def ofTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do let nano : Bounded.LE 0 999999999 := Bounded.LE.byEmod nanos.val 1000000000 (by decide) return { - date := PlainDate.clip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) + date := PlainDate.ofYearMonthDayClip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) time := PlainTime.ofHourMinuteSecondsNano (leap := false) (hour.expandTop (by decide)) minute second nano } @@ -138,7 +138,7 @@ out-of-range days clipped to the nearest valid date. -/ @[inline] def withDaysClip (dt : PlainDateTime) (days : Day.Ordinal) : PlainDateTime := - { dt with date := PlainDate.clip dt.date.year dt.date.month days } + { dt with date := PlainDate.ofYearMonthDayClip dt.date.year dt.date.month days } /-- Creates a new `PlainDateTime` by adjusting the day of the month to the given `days` value, with any @@ -154,7 +154,7 @@ out-of-range days clipped to the nearest valid date. -/ @[inline] def withMonthClip (dt : PlainDateTime) (month : Month.Ordinal) : PlainDateTime := - { dt with date := PlainDate.clip dt.date.year month dt.date.day } + { dt with date := PlainDate.ofYearMonthDayClip dt.date.year month dt.date.day } /-- Creates a new `PlainDateTime` by adjusting the month to the given `month` value. @@ -170,7 +170,7 @@ remain unchanged, with any out-of-range days clipped to the nearest valid date. -/ @[inline] def withYearClip (dt : PlainDateTime) (year : Year.Offset) : PlainDateTime := - { dt with date := PlainDate.clip year dt.date.month dt.date.day } + { dt with date := PlainDate.ofYearMonthDayClip year dt.date.month dt.date.day } /-- Creates a new `PlainDateTime` by adjusting the year to the given `year` value. The month and day are rolled diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 18954fbe4227..e56b076a94e0 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -109,6 +109,20 @@ Returns the `Duration` represented by the `Timestamp` since the Unix epoch. def toDurationSinceUnixEpoch (tm : Timestamp) : Duration := tm.val +/-- +Adds a `Millisecond.Offset` to the given `Timestamp`. +-/ +@[inline] +def addMilliseconds (t : Timestamp) (s : Millisecond.Offset) : Timestamp := + ⟨t.val + s⟩ + +/-- +Subtracts a `Millisecond.Offset` from the given `Timestamp`. +-/ +@[inline] +def subMilliseconds (t : Timestamp) (s : Millisecond.Offset) : Timestamp := + ⟨t.val - s⟩ + /-- Adds a `Nanosecond.Offset` to the given `Timestamp`. -/ @@ -243,6 +257,12 @@ instance : HAdd Timestamp Second.Offset Timestamp where instance : HSub Timestamp Second.Offset Timestamp where hSub := subSeconds +instance : HAdd Timestamp Millisecond.Offset Timestamp where + hAdd := addMilliseconds + +instance : HSub Timestamp Millisecond.Offset Timestamp where + hSub := subMilliseconds + instance : HAdd Timestamp Nanosecond.Offset Timestamp where hAdd := addNanoseconds diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index f5ad1ebc3a8c..7549f0d4a13d 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -180,6 +180,20 @@ Adds a `Nanosecond.Offset` to a `Duration` def addNanoseconds (t : Duration) (s : Nanosecond.Offset) : Duration := t.add (ofNanoseconds s) +/-- +Adds a `Millisecond.Offset` to a `Duration` +-/ +@[inline] +def addMilliseconds (t : Duration) (s : Millisecond.Offset) : Duration := + t.add (ofNanoseconds s.toNanoseconds) + +/-- +Adds a `Millisecond.Offset` to a `Duration` +-/ +@[inline] +def subMilliseconds (t : Duration) (s : Millisecond.Offset) : Duration := + t.sub (ofNanoseconds s.toNanoseconds) + /-- Adds a `Nanosecond.Offset` to a `Duration` -/ @@ -301,6 +315,12 @@ instance : HAdd Duration Nanosecond.Offset Duration where instance : HSub Duration Nanosecond.Offset Duration where hSub := subNanoseconds +instance : HAdd Duration Millisecond.Offset Duration where + hAdd := addMilliseconds + +instance : HSub Duration Millisecond.Offset Duration where + hSub := subMilliseconds + instance : HSub Duration Duration Duration where hSub := sub @@ -334,6 +354,9 @@ instance : HMul Duration Int Duration where instance : HAdd PlainTime Duration PlainTime where hAdd pt d := PlainTime.ofNanoseconds (d.toNanoseconds + pt.toNanoseconds) +instance : HSub PlainTime Duration PlainTime where + hSub pt d := PlainTime.ofNanoseconds (d.toNanoseconds - pt.toNanoseconds) + end Duration end Time end Std diff --git a/src/Std/Time/Notation.lean b/src/Std/Time/Notation.lean index 08520c18ae72..b8b5e35bcfdb 100644 --- a/src/Std/Time/Notation.lean +++ b/src/Std/Time/Notation.lean @@ -123,7 +123,7 @@ private def convertTimezone (tz : Std.Time.TimeZone) : MacroM (TSyntax `term) := `(Std.Time.TimeZone.mk $(← convertOffset tz.offset) $(Syntax.mkStrLit tz.name) $(Syntax.mkStrLit tz.abbreviation) false) private def convertPlainDate (d : Std.Time.PlainDate) : MacroM (TSyntax `term) := do - `(Std.Time.PlainDate.clip $(← syntaxInt d.year) $(← syntaxBounded d.month.val) $(← syntaxBounded d.day.val)) + `(Std.Time.PlainDate.ofYearMonthDayClip $(← syntaxInt d.year) $(← syntaxBounded d.month.val) $(← syntaxBounded d.day.val)) private def convertPlainTime (d : Std.Time.PlainTime) : MacroM (TSyntax `term) := do `(Std.Time.PlainTime.mk $(← syntaxBounded d.hour.val) $(← syntaxBounded d.minute.val) ⟨true, $(← syntaxBounded d.second.snd.val)⟩ $(← syntaxBounded d.nanosecond.val)) diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index dfe54dc5eee0..4125c2d776ce 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -7,7 +7,6 @@ prelude import Std.Time.Zoned.DateTime import Std.Time.Zoned.ZoneRules import Std.Time.Zoned.ZonedDateTime -import Std.Time.Zoned.ZonedDateTime import Std.Time.Zoned.Database namespace Std @@ -102,6 +101,13 @@ Converts a `PlainDate` to a `ZonedDateTime`. def ofPlainDate (pd : PlainDate) (zr : TimeZone.ZoneRules) : ZonedDateTime := ZonedDateTime.ofPlainDateTime (pd.atTime PlainTime.midnight) zr +/-- +Converts a `PlainDate` to a `ZonedDateTime` using `TimeZone`. +-/ +@[inline] +def ofPlainDateWithZone (pd : PlainDate) (zr : TimeZone) : ZonedDateTime := + ZonedDateTime.ofPlainDateTime (pd.atTime PlainTime.midnight) (TimeZone.ZoneRules.ofTimeZone zr) + /-- Converts a `ZonedDateTime` to a `PlainDate` -/ @@ -125,3 +131,39 @@ def of (pdt : PlainDateTime) (id : String) : IO ZonedDateTime := do return ZonedDateTime.ofPlainDateTime pdt zr end ZonedDateTime + +namespace PlainDateTime + +/-- +Converts a `PlainDateTime` to a `Timestamp` using the `ZoneRules`. +-/ +@[inline] +def toTimestamp (pdt : PlainDateTime) (zr : TimeZone.ZoneRules) : Timestamp := + ZonedDateTime.ofPlainDateTime pdt zr |>.toTimestamp + +/-- +Converts a `PlainDateTime` to a `Timestamp` using the `TimeZone`. +-/ +@[inline] +def toTimestampWithZone (pdt : PlainDateTime) (tz : TimeZone) : Timestamp := + ZonedDateTime.ofPlainDateTimeWithZone pdt tz |>.toTimestamp + +end PlainDateTime + +namespace PlainDate + +/-- +Converts a `PlainDate` to a `Timestamp` using the `ZoneRules`. +-/ +@[inline] +def toTimestamp (dt : PlainDate) (zr : TimeZone.ZoneRules) : Timestamp := + ZonedDateTime.ofPlainDate dt zr |>.toTimestamp + +/-- +Converts a `PlainDate` to a `Timestamp` using the `TimeZone`. +-/ +@[inline] +def toTimestampWithZone (dt : PlainDate) (tz : TimeZone) : Timestamp := + ZonedDateTime.ofPlainDateWithZone dt tz |>.toTimestamp + +end PlainDate diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 4e06b1f55567..0d13c4bf88b9 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -120,6 +120,20 @@ Subtract `Second.Offset` from a `DateTime`. def subSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := ofLocalDateTime (dt.date.get.subSeconds seconds) tz +/-- +Add `Millisecond.Offset` to a `DateTime`. +-/ +@[inline] +def addMilliseconds (dt : DateTime tz) (milliseconds : Millisecond.Offset) : DateTime tz := + ofLocalDateTime (dt.date.get.addMilliseconds milliseconds) tz + +/-- +Subtract `Millisecond.Offset` from a `DateTime`. +-/ +@[inline] +def subMilliseconds (dt : DateTime tz) (milliseconds : Millisecond.Offset) : DateTime tz := + ofLocalDateTime (dt.date.get.subMilliseconds milliseconds) tz + /-- Add `Nanosecond.Offset` to a `DateTime`. -/ @@ -443,6 +457,12 @@ instance : HAdd (DateTime tz) (Second.Offset) (DateTime tz) where instance : HSub (DateTime tz) (Second.Offset) (DateTime tz) where hSub := subSeconds +instance : HAdd (DateTime tz) (Millisecond.Offset) (DateTime tz) where + hAdd := addMilliseconds + +instance : HSub (DateTime tz) (Millisecond.Offset) (DateTime tz) where + hSub := subMilliseconds + instance : HAdd (DateTime tz) (Nanosecond.Offset) (DateTime tz) where hAdd := addNanoseconds @@ -452,6 +472,12 @@ instance : HSub (DateTime tz) (Nanosecond.Offset) (DateTime tz) where instance : HSub (DateTime tz) (DateTime tz₁) Duration where hSub x y := x.toTimestamp - y.toTimestamp +instance : HAdd (DateTime tz) Duration (DateTime tz) where + hAdd x y := x.addNanoseconds y.toNanoseconds + +instance : HSub (DateTime tz) Duration (DateTime tz) where + hSub x y := x.subNanoseconds y.toNanoseconds + end DateTime end Time end Std diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 8d3856047957..72253c6d05f2 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -145,9 +145,18 @@ def apply (timestamp : Timestamp) (transition : Transition) : Timestamp := Finds the transition corresponding to a given timestamp in `Array Transition`. If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. -/ -def findTransitionForTimestamp (transitions : Array Transition) (timestamp : Timestamp) : Option Transition := +def findTransitionIndexForTimestamp (transitions : Array Transition) (timestamp : Timestamp) : Option Nat := let value := timestamp.toSecondsSinceUnixEpoch if let some idx := transitions.findIdx? (fun t => t.time.val ≥ value.val) + then some (idx - 1) + else none + +/-- +Finds the transition corresponding to a given timestamp in `Array Transition`. +If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. +-/ +def findTransitionForTimestamp (transitions : Array Transition) (timestamp : Timestamp) : Option Transition := + if let some idx := findTransitionIndexForTimestamp transitions timestamp then transitions.get? (idx - 1) else transitions.back? diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 106c02fce201..420e944c03bb 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -73,8 +73,8 @@ def ofTimestampWithZone (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := Creates a new `ZonedDateTime` out of a `PlainDateTime` and a `TimeZone`. -/ @[inline] -def ofPlainDateTimeWithZone (tm : Timestamp) (tz : TimeZone) : ZonedDateTime := - ofTimestamp tm (TimeZone.ZoneRules.ofTimeZone tz) +def ofPlainDateTimeWithZone (tm : PlainDateTime) (tz : TimeZone) : ZonedDateTime := + ofPlainDateTime tm (TimeZone.ZoneRules.ofTimeZone tz) /-- Creates a new `Timestamp` out of a `ZonedDateTime`. @@ -327,6 +327,22 @@ def subMinutes (dt : ZonedDateTime) (minutes : Minute.Offset) : ZonedDateTime := let date := dt.timestamp.toPlainDateTimeAssumingUTC ZonedDateTime.ofTimestamp (date.subMinutes minutes).toTimestampAssumingUTC dt.rules +/-- +Add `Millisecond.Offset` to a `DateTime`. +-/ +@[inline] +def addMilliseconds (dt : ZonedDateTime) (milliseconds : Millisecond.Offset) : ZonedDateTime := + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.addMilliseconds milliseconds).toTimestampAssumingUTC dt.rules + +/-- +Subtract `Millisecond.Offset` from a `DateTime`. +-/ +@[inline] +def subMilliseconds (dt : ZonedDateTime) (milliseconds : Millisecond.Offset) : ZonedDateTime := + let date := dt.timestamp.toPlainDateTimeAssumingUTC + ZonedDateTime.ofTimestamp (date.subMilliseconds milliseconds).toTimestampAssumingUTC dt.rules + /-- Add `Second.Offset` to a `ZonedDateTime`. -/ @@ -501,6 +517,12 @@ instance : HAdd ZonedDateTime Second.Offset ZonedDateTime where instance : HSub ZonedDateTime Second.Offset ZonedDateTime where hSub := subSeconds +instance : HAdd ZonedDateTime Millisecond.Offset ZonedDateTime where + hAdd := addMilliseconds + +instance : HSub ZonedDateTime Millisecond.Offset ZonedDateTime where + hSub := subMilliseconds + instance : HAdd ZonedDateTime Nanosecond.Offset ZonedDateTime where hAdd := addNanoseconds @@ -510,6 +532,12 @@ instance : HSub ZonedDateTime Nanosecond.Offset ZonedDateTime where instance : HSub ZonedDateTime ZonedDateTime Duration where hSub x y := x.toTimestamp - y.toTimestamp +instance : HAdd ZonedDateTime Duration ZonedDateTime where + hAdd x y := x.addNanoseconds y.toNanoseconds + +instance : HSub ZonedDateTime Duration ZonedDateTime where + hSub x y := x.subNanoseconds y.toNanoseconds + end ZonedDateTime end Time end Std diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index fe2be5ac23c8..d71ea8b70a19 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -277,7 +277,7 @@ Format def time₄ := time("23:13:12.324354679") def date₄ := date("2002-07-14") -def datetime₅ := PlainDateTime.mk (PlainDate.clip (-2000) 3 4) (PlainTime.mk 12 23 ⟨false, 12⟩ 0) +def datetime₅ := PlainDateTime.mk (PlainDate.ofYearMonthDayClip (-2000) 3 4) (PlainTime.mk 12 23 ⟨false, 12⟩ 0) def datetime₄ := datetime("2002-07-14T23:13:12.324354679") def zoned₄ := zoned("2002-07-14T23:13:12.324354679+09:00") def zoned₅ := zoned("2002-07-14T23:13:12.324354679+00:00") From d6e2024ff06c4a2106a31bc5be1254a9969b3fcb Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 5 Nov 2024 23:03:36 -0300 Subject: [PATCH 094/118] fix: problem with index in the timezone --- src/Std/Time/Zoned/ZoneRules.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 72253c6d05f2..e196a3f84dbb 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -148,7 +148,7 @@ If the timestamp falls between two transitions, it returns the most recent trans def findTransitionIndexForTimestamp (transitions : Array Transition) (timestamp : Timestamp) : Option Nat := let value := timestamp.toSecondsSinceUnixEpoch if let some idx := transitions.findIdx? (fun t => t.time.val ≥ value.val) - then some (idx - 1) + then some idx else none /-- From 9ed4dc15abe15306b33cadb9a4ec18b72360d21e Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 6 Nov 2024 03:48:55 -0300 Subject: [PATCH 095/118] feat: add all the arithmetic operations between offsets --- src/Std/Time/Date/Basic.lean | 294 +++++++++++ tests/lean/run/timeOperationsOffset.lean | 596 +++++++++++++++++++++++ 2 files changed, 890 insertions(+) create mode 100644 tests/lean/run/timeOperationsOffset.lean diff --git a/src/Std/Time/Date/Basic.lean b/src/Std/Time/Date/Basic.lean index 5c3cca0a97e6..4d4755901faf 100644 --- a/src/Std/Time/Date/Basic.lean +++ b/src/Std/Time/Date/Basic.lean @@ -180,3 +180,297 @@ def ofWeeks (weeks : Week.Offset) : Hour.Offset := end Offset end Hour + +instance : HAdd Nanosecond.Offset Nanosecond.Offset Nanosecond.Offset where + hAdd x y := x.add y + +instance : HAdd Nanosecond.Offset Millisecond.Offset Nanosecond.Offset where + hAdd x y := x.add y.toNanoseconds + +instance : HAdd Nanosecond.Offset Second.Offset Nanosecond.Offset where + hAdd x y := x.add y.toNanoseconds + +instance : HAdd Nanosecond.Offset Minute.Offset Nanosecond.Offset where + hAdd x y := x.add y.toNanoseconds + +instance : HAdd Nanosecond.Offset Hour.Offset Nanosecond.Offset where + hAdd x y := x.add y.toNanoseconds + +instance : HAdd Nanosecond.Offset Day.Offset Nanosecond.Offset where + hAdd x y := x.add y.toNanoseconds + +instance : HAdd Nanosecond.Offset Week.Offset Nanosecond.Offset where + hAdd x y := x.add y.toNanoseconds + +instance : HAdd Millisecond.Offset Nanosecond.Offset Nanosecond.Offset where + hAdd x y := x.toNanoseconds.add y + +instance : HAdd Millisecond.Offset Millisecond.Offset Millisecond.Offset where + hAdd x y := x.add y + +instance : HAdd Millisecond.Offset Second.Offset Millisecond.Offset where + hAdd x y := x.add y.toMilliseconds + +instance : HAdd Millisecond.Offset Minute.Offset Millisecond.Offset where + hAdd x y := x.add y.toMilliseconds + +instance : HAdd Millisecond.Offset Hour.Offset Millisecond.Offset where + hAdd x y := x.add y.toMilliseconds + +instance : HAdd Millisecond.Offset Day.Offset Millisecond.Offset where + hAdd x y := x.add y.toMilliseconds + +instance : HAdd Millisecond.Offset Week.Offset Millisecond.Offset where + hAdd x y := x.add y.toMilliseconds + +instance : HAdd Second.Offset Nanosecond.Offset Nanosecond.Offset where + hAdd x y := x.toNanoseconds.add y + +instance : HAdd Second.Offset Millisecond.Offset Millisecond.Offset where + hAdd x y := x.toMilliseconds.add y + +instance : HAdd Second.Offset Second.Offset Second.Offset where + hAdd x y := x.add y + +instance : HAdd Second.Offset Minute.Offset Second.Offset where + hAdd x y := x.add y.toSeconds + +instance : HAdd Second.Offset Hour.Offset Second.Offset where + hAdd x y := x.add y.toSeconds + +instance : HAdd Second.Offset Day.Offset Second.Offset where + hAdd x y := x.add y.toSeconds + +instance : HAdd Second.Offset Week.Offset Second.Offset where + hAdd x y := x.add y.toSeconds + +instance : HAdd Minute.Offset Nanosecond.Offset Nanosecond.Offset where + hAdd x y := x.toNanoseconds.add y + +instance : HAdd Minute.Offset Millisecond.Offset Millisecond.Offset where + hAdd x y := x.toMilliseconds.add y + +instance : HAdd Minute.Offset Second.Offset Second.Offset where + hAdd x y := x.toSeconds.add y + +instance : HAdd Minute.Offset Minute.Offset Minute.Offset where + hAdd x y := x.add y + +instance : HAdd Minute.Offset Hour.Offset Minute.Offset where + hAdd x y := x.add y.toMinutes + +instance : HAdd Minute.Offset Day.Offset Minute.Offset where + hAdd x y := x.add y.toMinutes + +instance : HAdd Minute.Offset Week.Offset Minute.Offset where + hAdd x y := x.add y.toMinutes + +instance : HAdd Hour.Offset Nanosecond.Offset Nanosecond.Offset where + hAdd x y := x.toNanoseconds.add y + +instance : HAdd Hour.Offset Millisecond.Offset Millisecond.Offset where + hAdd x y := x.toMilliseconds.add y + +instance : HAdd Hour.Offset Second.Offset Second.Offset where + hAdd x y := x.toSeconds.add y + +instance : HAdd Hour.Offset Minute.Offset Minute.Offset where + hAdd x y := x.toMinutes.add y + +instance : HAdd Hour.Offset Hour.Offset Hour.Offset where + hAdd x y := x.add y + +instance : HAdd Hour.Offset Day.Offset Hour.Offset where + hAdd x y := x.add y.toHours + +instance : HAdd Hour.Offset Week.Offset Hour.Offset where + hAdd x y := x.add y.toHours + +instance : HAdd Day.Offset Nanosecond.Offset Nanosecond.Offset where + hAdd x y := x.toNanoseconds.add y + +instance : HAdd Day.Offset Millisecond.Offset Millisecond.Offset where + hAdd x y := x.toMilliseconds.add y + +instance : HAdd Day.Offset Second.Offset Second.Offset where + hAdd x y := x.toSeconds.add y + +instance : HAdd Day.Offset Minute.Offset Minute.Offset where + hAdd x y := x.toMinutes.add y + +instance : HAdd Day.Offset Hour.Offset Hour.Offset where + hAdd x y := x.toHours.add y + +instance : HAdd Day.Offset Day.Offset Day.Offset where + hAdd x y := x.add y + +instance : HAdd Day.Offset Week.Offset Day.Offset where + hAdd x y := x.add y.toDays + +instance : HAdd Week.Offset Nanosecond.Offset Nanosecond.Offset where + hAdd x y := x.toNanoseconds.add y + +instance : HAdd Week.Offset Millisecond.Offset Millisecond.Offset where + hAdd x y := x.toMilliseconds.add y + +instance : HAdd Week.Offset Second.Offset Second.Offset where + hAdd x y := x.toSeconds.add y + +instance : HAdd Week.Offset Minute.Offset Minute.Offset where + hAdd x y := x.toMinutes.add y + +instance : HAdd Week.Offset Hour.Offset Hour.Offset where + hAdd x y := x.toHours.add y + +instance : HAdd Week.Offset Day.Offset Day.Offset where + hAdd x y := x.toDays.add y + +instance : HAdd Week.Offset Week.Offset Week.Offset where + hAdd x y := x.add y + +instance : HSub Nanosecond.Offset Nanosecond.Offset Nanosecond.Offset where + hSub x y := x.sub y + +instance : HSub Nanosecond.Offset Millisecond.Offset Nanosecond.Offset where + hSub x y := x.sub y.toNanoseconds + +instance : HSub Nanosecond.Offset Second.Offset Nanosecond.Offset where + hSub x y := x.sub y.toNanoseconds + +instance : HSub Nanosecond.Offset Minute.Offset Nanosecond.Offset where + hSub x y := x.sub y.toNanoseconds + +instance : HSub Nanosecond.Offset Hour.Offset Nanosecond.Offset where + hSub x y := x.sub y.toNanoseconds + +instance : HSub Nanosecond.Offset Day.Offset Nanosecond.Offset where + hSub x y := x.sub y.toNanoseconds + +instance : HSub Nanosecond.Offset Week.Offset Nanosecond.Offset where + hSub x y := x.sub y.toNanoseconds + +instance : HSub Millisecond.Offset Nanosecond.Offset Nanosecond.Offset where + hSub x y := x.toNanoseconds.sub y + +instance : HSub Millisecond.Offset Millisecond.Offset Millisecond.Offset where + hSub x y := x.sub y + +instance : HSub Millisecond.Offset Second.Offset Millisecond.Offset where + hSub x y := x.sub y.toMilliseconds + +instance : HSub Millisecond.Offset Minute.Offset Millisecond.Offset where + hSub x y := x.sub y.toMilliseconds + +instance : HSub Millisecond.Offset Hour.Offset Millisecond.Offset where + hSub x y := x.sub y.toMilliseconds + +instance : HSub Millisecond.Offset Day.Offset Millisecond.Offset where + hSub x y := x.sub y.toMilliseconds + +instance : HSub Millisecond.Offset Week.Offset Millisecond.Offset where + hSub x y := x.sub y.toMilliseconds + +instance : HSub Second.Offset Nanosecond.Offset Nanosecond.Offset where + hSub x y := x.toNanoseconds.sub y + +instance : HSub Second.Offset Millisecond.Offset Millisecond.Offset where + hSub x y := x.toMilliseconds.sub y + +instance : HSub Second.Offset Second.Offset Second.Offset where + hSub x y := x.sub y + +instance : HSub Second.Offset Minute.Offset Second.Offset where + hSub x y := x.sub y.toSeconds + +instance : HSub Second.Offset Hour.Offset Second.Offset where + hSub x y := x.sub y.toSeconds + +instance : HSub Second.Offset Day.Offset Second.Offset where + hSub x y := x.sub y.toSeconds + +instance : HSub Second.Offset Week.Offset Second.Offset where + hSub x y := x.sub y.toSeconds + +instance : HSub Minute.Offset Nanosecond.Offset Nanosecond.Offset where + hSub x y := x.toNanoseconds.sub y + +instance : HSub Minute.Offset Millisecond.Offset Millisecond.Offset where + hSub x y := x.toMilliseconds.sub y + +instance : HSub Minute.Offset Second.Offset Second.Offset where + hSub x y := x.toSeconds.sub y + +instance : HSub Minute.Offset Minute.Offset Minute.Offset where + hSub x y := x.sub y + +instance : HSub Minute.Offset Hour.Offset Minute.Offset where + hSub x y := x.sub y.toMinutes + +instance : HSub Minute.Offset Day.Offset Minute.Offset where + hSub x y := x.sub y.toMinutes + +instance : HSub Minute.Offset Week.Offset Minute.Offset where + hSub x y := x.sub y.toMinutes + +instance : HSub Hour.Offset Nanosecond.Offset Nanosecond.Offset where + hSub x y := x.toNanoseconds.sub y + +instance : HSub Hour.Offset Millisecond.Offset Millisecond.Offset where + hSub x y := x.toMilliseconds.sub y + +instance : HSub Hour.Offset Second.Offset Second.Offset where + hSub x y := x.toSeconds.sub y + +instance : HSub Hour.Offset Minute.Offset Minute.Offset where + hSub x y := x.toMinutes.sub y + +instance : HSub Hour.Offset Hour.Offset Hour.Offset where + hSub x y := x.sub y + +instance : HSub Hour.Offset Day.Offset Hour.Offset where + hSub x y := x.sub y.toHours + +instance : HSub Hour.Offset Week.Offset Hour.Offset where + hSub x y := x.sub y.toHours + +instance : HSub Day.Offset Nanosecond.Offset Nanosecond.Offset where + hSub x y := x.toNanoseconds.sub y + +instance : HSub Day.Offset Millisecond.Offset Millisecond.Offset where + hSub x y := x.toMilliseconds.sub y + +instance : HSub Day.Offset Second.Offset Second.Offset where + hSub x y := x.toSeconds.sub y + +instance : HSub Day.Offset Minute.Offset Minute.Offset where + hSub x y := x.toMinutes.sub y + +instance : HSub Day.Offset Hour.Offset Hour.Offset where + hSub x y := x.toHours.sub y + +instance : HSub Day.Offset Day.Offset Day.Offset where + hSub x y := x.sub y + +instance : HSub Day.Offset Week.Offset Day.Offset where + hSub x y := x.sub y.toDays + +instance : HSub Week.Offset Nanosecond.Offset Nanosecond.Offset where + hSub x y := x.toNanoseconds.sub y + +instance : HSub Week.Offset Millisecond.Offset Millisecond.Offset where + hSub x y := x.toMilliseconds.sub y + +instance : HSub Week.Offset Second.Offset Second.Offset where + hSub x y := x.toSeconds.sub y + +instance : HSub Week.Offset Minute.Offset Minute.Offset where + hSub x y := x.toMinutes.sub y + +instance : HSub Week.Offset Hour.Offset Hour.Offset where + hSub x y := x.toHours.sub y + +instance : HSub Week.Offset Day.Offset Day.Offset where + hSub x y := x.toDays.sub y + +instance : HSub Week.Offset Week.Offset Week.Offset where + hSub x y := x.sub y diff --git a/tests/lean/run/timeOperationsOffset.lean b/tests/lean/run/timeOperationsOffset.lean new file mode 100644 index 000000000000..53810de730d7 --- /dev/null +++ b/tests/lean/run/timeOperationsOffset.lean @@ -0,0 +1,596 @@ +import Std.Time +open Std.Time + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 1000001 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Millisecond.Offset) + +/-- +info: 1000000001 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Second.Offset) + +/-- +info: 60000000001 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Minute.Offset) + +/-- +info: 3600000000001 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Hour.Offset) + +/-- +info: 86400000000001 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Day.Offset) + +/-- +info: 604800000000001 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) + (1 : Week.Offset) + +/-- +info: 1000001 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) + (1 : Millisecond.Offset) + +/-- +info: 1001 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) + (1 : Second.Offset) + +/-- +info: 60001 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) + (1 : Minute.Offset) + +/-- +info: 3600001 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) + (1 : Hour.Offset) + +/-- +info: 86400001 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) + (1 : Day.Offset) + +/-- +info: 604800001 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) + (1 : Week.Offset) + +/-- +info: 1000000001 +-/ +#guard_msgs in +#eval (1 : Second.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 1001 +-/ +#guard_msgs in +#eval (1 : Second.Offset) + (1 : Millisecond.Offset) + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Second.Offset) + (1 : Second.Offset) + +/-- +info: 61 +-/ +#guard_msgs in +#eval (1 : Second.Offset) + (1 : Minute.Offset) + +/-- +info: 3601 +-/ +#guard_msgs in +#eval (1 : Second.Offset) + (1 : Hour.Offset) + +/-- +info: 86401 +-/ +#guard_msgs in +#eval (1 : Second.Offset) + (1 : Day.Offset) + +/-- +info: 604801 +-/ +#guard_msgs in +#eval (1 : Second.Offset) + (1 : Week.Offset) + +/-- +info: 60000000001 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 60001 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) + (1 : Millisecond.Offset) + +/-- +info: 61 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) + (1 : Second.Offset) + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) + (1 : Minute.Offset) + +/-- +info: 61 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) + (1 : Hour.Offset) + +/-- +info: 1441 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) + (1 : Day.Offset) + +/-- +info: 10081 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) + (1 : Week.Offset) + +/-- +info: 3600000000001 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 3600001 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) + (1 : Millisecond.Offset) + +/-- +info: 3601 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) + (1 : Second.Offset) + +/-- +info: 61 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) + (1 : Minute.Offset) + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) + (1 : Hour.Offset) + +/-- +info: 25 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) + (1 : Day.Offset) + +/-- +info: 169 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) + (1 : Week.Offset) + +/-- +info: 86400000000001 +-/ +#guard_msgs in +#eval (1 : Day.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 86400001 +-/ +#guard_msgs in +#eval (1 : Day.Offset) + (1 : Millisecond.Offset) + +/-- +info: 86401 +-/ +#guard_msgs in +#eval (1 : Day.Offset) + (1 : Second.Offset) + +/-- +info: 1441 +-/ +#guard_msgs in +#eval (1 : Day.Offset) + (1 : Minute.Offset) + +/-- +info: 25 +-/ +#guard_msgs in +#eval (1 : Day.Offset) + (1 : Hour.Offset) + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Day.Offset) + (1 : Day.Offset) + +/-- +info: 8 +-/ +#guard_msgs in +#eval (1 : Day.Offset) + (1 : Week.Offset) + +/-- +info: 604800000000001 +-/ +#guard_msgs in +#eval (1 : Week.Offset) + (1 : Nanosecond.Offset) + +/-- +info: 604800001 +-/ +#guard_msgs in +#eval (1 : Week.Offset) + (1 : Millisecond.Offset) + +/-- +info: 604801 +-/ +#guard_msgs in +#eval (1 : Week.Offset) + (1 : Second.Offset) + +/-- +info: 10081 +-/ +#guard_msgs in +#eval (1 : Week.Offset) + (1 : Minute.Offset) + +/-- +info: 169 +-/ +#guard_msgs in +#eval (1 : Week.Offset) + (1 : Hour.Offset) + +/-- +info: 8 +-/ +#guard_msgs in +#eval (1 : Week.Offset) + (1 : Day.Offset) + +/-- +info: 2 +-/ +#guard_msgs in +#eval (1 : Week.Offset) + (1 : Week.Offset) + +/-- +info: 0 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) - (1 : Nanosecond.Offset) + +/-- +info: -999999 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) - (1 : Millisecond.Offset) + +/-- +info: -999999999 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) - (1 : Second.Offset) + +/-- +info: -59999999999 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) - (1 : Minute.Offset) + +/-- +info: -3599999999999 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) - (1 : Hour.Offset) + +/-- +info: -86399999999999 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) - (1 : Day.Offset) + +/-- +info: -604799999999999 +-/ +#guard_msgs in +#eval (1 : Nanosecond.Offset) - (1 : Week.Offset) + +/-- +info: 999999 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) - (1 : Nanosecond.Offset) + +/-- +info: 0 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) - (1 : Millisecond.Offset) + +/-- +info: -999 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) - (1 : Second.Offset) + +/-- +info: -59999 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) - (1 : Minute.Offset) + +/-- +info: -3599999 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) - (1 : Hour.Offset) + +/-- +info: -86399999 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) - (1 : Day.Offset) + +/-- +info: -604799999 +-/ +#guard_msgs in +#eval (1 : Millisecond.Offset) - (1 : Week.Offset) + +/-- +info: 999999999 +-/ +#guard_msgs in +#eval (1 : Second.Offset) - (1 : Nanosecond.Offset) + +/-- +info: 999 +-/ +#guard_msgs in +#eval (1 : Second.Offset) - (1 : Millisecond.Offset) + +/-- +info: 0 +-/ +#guard_msgs in +#eval (1 : Second.Offset) - (1 : Second.Offset) + +/-- +info: -59 +-/ +#guard_msgs in +#eval (1 : Second.Offset) - (1 : Minute.Offset) + +/-- +info: -3599 +-/ +#guard_msgs in +#eval (1 : Second.Offset) - (1 : Hour.Offset) + +/-- +info: -86399 +-/ +#guard_msgs in +#eval (1 : Second.Offset) - (1 : Day.Offset) + +/-- +info: -604799 +-/ +#guard_msgs in +#eval (1 : Second.Offset) - (1 : Week.Offset) + +/-- +info: 59999999999 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) - (1 : Nanosecond.Offset) + +/-- +info: 59999 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) - (1 : Millisecond.Offset) + +/-- +info: 59 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) - (1 : Second.Offset) + +/-- +info: 0 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) - (1 : Minute.Offset) + +/-- +info: -59 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) - (1 : Hour.Offset) + +/-- +info: -1439 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) - (1 : Day.Offset) + +/-- +info: -10079 +-/ +#guard_msgs in +#eval (1 : Minute.Offset) - (1 : Week.Offset) + +/-- +info: 3599999999999 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) - (1 : Nanosecond.Offset) + +/-- +info: 3599999 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) - (1 : Millisecond.Offset) + +/-- +info: 3599 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) - (1 : Second.Offset) + +/-- +info: 59 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) - (1 : Minute.Offset) + +/-- +info: 0 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) - (1 : Hour.Offset) + +/-- +info: -23 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) - (1 : Day.Offset) + +/-- +info: -167 +-/ +#guard_msgs in +#eval (1 : Hour.Offset) - (1 : Week.Offset) + +/-- +info: 86399999999999 +-/ +#guard_msgs in +#eval (1 : Day.Offset) - (1 : Nanosecond.Offset) + +/-- +info: 86399999 +-/ +#guard_msgs in +#eval (1 : Day.Offset) - (1 : Millisecond.Offset) + +/-- +info: 86399 +-/ +#guard_msgs in +#eval (1 : Day.Offset) - (1 : Second.Offset) + +/-- +info: 1439 +-/ +#guard_msgs in +#eval (1 : Day.Offset) - (1 : Minute.Offset) + +/-- +info: 23 +-/ +#guard_msgs in +#eval (1 : Day.Offset) - (1 : Hour.Offset) + +/-- +info: 0 +-/ +#guard_msgs in +#eval (1 : Day.Offset) - (1 : Day.Offset) + +/-- +info: -6 +-/ +#guard_msgs in +#eval (1 : Day.Offset) - (1 : Week.Offset) + +/-- +info: 604799999999999 +-/ +#guard_msgs in +#eval (1 : Week.Offset) - (1 : Nanosecond.Offset) + +/-- +info: 604799999 +-/ +#guard_msgs in +#eval (1 : Week.Offset) - (1 : Millisecond.Offset) + +/-- +info: 604799 +-/ +#guard_msgs in +#eval (1 : Week.Offset) - (1 : Second.Offset) + +/-- +info: 10079 +-/ +#guard_msgs in +#eval (1 : Week.Offset) - (1 : Minute.Offset) + +/-- +info: 167 +-/ +#guard_msgs in +#eval (1 : Week.Offset) - (1 : Hour.Offset) + +/-- +info: 6 +-/ +#guard_msgs in +#eval (1 : Week.Offset) - (1 : Day.Offset) + +/-- +info: 0 +-/ +#guard_msgs in +#eval (1 : Week.Offset) - (1 : Week.Offset) From 8ec73c26548c186230063d89cebc795c84a5dc9a Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 6 Nov 2024 09:14:01 -0300 Subject: [PATCH 096/118] fix: nix ci --- nix/bootstrap.nix | 6 +++--- nix/buildLeanPackage.nix | 4 ++-- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/nix/bootstrap.nix b/nix/bootstrap.nix index 170b807494df..3c2b9b9151c9 100644 --- a/nix/bootstrap.nix +++ b/nix/bootstrap.nix @@ -1,5 +1,5 @@ { src, debug ? false, stage0debug ? false, extraCMakeFlags ? [], - stdenv, lib, cmake, gmp, libuv, cadical, git, tzdata, gnumake, bash, buildLeanPackage, writeShellScriptBin, runCommand, symlinkJoin, lndir, perl, gnused, darwin, llvmPackages, linkFarmFromDrvs, + stdenv, lib, cmake, gmp, libuv, cadical, git, gnumake, bash, buildLeanPackage, writeShellScriptBin, runCommand, symlinkJoin, lndir, perl, gnused, darwin, llvmPackages, linkFarmFromDrvs, ... } @ args: with builtins; lib.warn "The Nix-based build is deprecated" rec { @@ -159,7 +159,7 @@ lib.warn "The Nix-based build is deprecated" rec { test = buildCMake { name = "lean-test-${desc}"; realSrc = lib.sourceByRegex src [ "src.*" "tests.*" ]; - buildInputs = [ gmp libuv perl git cadical tzdata ]; + buildInputs = [ gmp libuv perl git cadical ]; preConfigure = '' cd src ''; @@ -170,7 +170,7 @@ lib.warn "The Nix-based build is deprecated" rec { ln -sf ${lean-all}/* . ''; buildPhase = '' - ctest --output-junit test-results.xml --output-on-failure -E 'leancomptest_(doc_example|foreign)|leanlaketest_reverse-ffi' -j$NIX_BUILD_CORES + ctest --output-junit test-results.xml --output-on-failure -E 'leancomptest_(doc_example|foreign)|leanlaketest_reverse-ffi|leanruntest_timeIO' -j$NIX_BUILD_CORES ''; installPhase = '' mkdir $out diff --git a/nix/buildLeanPackage.nix b/nix/buildLeanPackage.nix index 8cd4015fd5b6..c8b82ebde648 100644 --- a/nix/buildLeanPackage.nix +++ b/nix/buildLeanPackage.nix @@ -1,5 +1,5 @@ { lean, lean-leanDeps ? lean, lean-final ? lean, leanc, - stdenv, lib, coreutils, tzdata, gnused, writeShellScriptBin, bash, substituteAll, symlinkJoin, linkFarmFromDrvs, + stdenv, lib, coreutils, gnused, writeShellScriptBin, bash, substituteAll, symlinkJoin, linkFarmFromDrvs, runCommand, darwin, mkShell, ... }: let lean-final' = lean-final; in lib.makeOverridable ( @@ -45,7 +45,7 @@ with builtins; let name = lib.strings.sanitizeDerivationName args.name; stdenv = bareStdenv; inherit (stdenv) system; - buildInputs = (args.buildInputs or []) ++ [ coreutils tzdata ]; + buildInputs = (args.buildInputs or []) ++ [ coreutils ]; builder = stdenv.shell; args = [ "-c" '' source $stdenv/setup From a69dca171178ea2b03c9bef3f2c8f7828c8f243f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 6 Nov 2024 11:46:19 -0300 Subject: [PATCH 097/118] fix: problems with int64 --- src/Std/Time/Duration.lean | 3 +-- src/Std/Time/Zoned/Database/Windows.lean | 6 +++--- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 7549f0d4a13d..36a28d86f456 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -56,8 +56,7 @@ instance : Inhabited Duration where instance : OfNat Duration n where ofNat := by refine ⟨.ofInt n, ⟨0, by decide⟩, ?_⟩ - simp <;> exact Int.le_total s.val 0 |>.symm - exact Int.le_total 0 n + simp <;> exact Int.le_total n 0 |>.symm namespace Duration diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index 241738c967bd..0ae44837832a 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -21,13 +21,13 @@ namespace Windows Fetches the next timezone transition for a given timezone identifier and timestamp. -/ @[extern "lean_get_windows_next_transition"] -opaque getNextTransition : @&String -> @&Int -> IO (Option (Int × TimeZone)) +opaque getNextTransition : @&String -> @&Int64 -> IO (Option (Int64 × TimeZone)) /-- Fetches the timezone at a timestamp. -/ @[extern "lean_get_windows_local_timezone_id_at"] -opaque getLocalTimeZoneIdentifierAt : @&Int → IO String +opaque getLocalTimeZoneIdentifierAt : @&Int64 → IO String /-- Retrieves the timezone rules, including all transitions, for a given timezone identifier. @@ -39,7 +39,7 @@ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do while true do let result ← Windows.getNextTransition id start if let some res := result then - transitions := transitions.push { time := Second.Offset.ofInt res.fst, localTimeType := { + transitions := transitions.push { time := Second.Offset.ofInt res.fst.toInt, localTimeType := { gmtOffset := res.snd.offset, abbreviation := res.snd.abbreviation, identifier := res.snd.name, From 5b7a8fd07aa65216f08989848dc46a0199af54bb Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Thu, 7 Nov 2024 11:19:27 -0300 Subject: [PATCH 098/118] Update src/Std/Time.lean Co-authored-by: Markus Himmel --- src/Std/Time.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index f85d37cef44b..96e67b4f4891 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -83,7 +83,7 @@ nanoseconds corresponds to one second. # Date and Time Types -Dates and times are made up of different parts. an `Ordinal` is an absolute value, like a specific day in a month, +Dates and times are made up of different parts. An `Ordinal` is an absolute value, like a specific day in a month, while an `Offset` is a shift forward or backward in time, used in arithmetic operations to add or subtract days, months or years. Dates use components like `Year.Ordinal`, `Month.Ordinal`, and `Day.Ordinal` to ensure they represent valid points in time. From 0236e9a52eb0eafbeda8cbe8e80c0b6971f41010 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 00:27:12 -0300 Subject: [PATCH 099/118] Update src/Std/Time.lean Co-authored-by: Markus Himmel --- src/Std/Time.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index 96e67b4f4891..b54c634e2775 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -57,7 +57,7 @@ converted because they use an internal type called `UnitVal`. Ordinal types represent specific bounded values in reference to another unit, e.g., `Day.Ordinal` represents a day in a month, ranging from 1 to 31. Some ordinal types like `Hour.Ordinal` and `Second.Ordinal`, -allow for values beyond the normal range (e.g, 60 seconds) to accomodate special cases with leap seconds +allow for values beyond the normal range (e.g, 60 seconds) to accommodate special cases with leap seconds like `23:59:60` that is valid in ISO 8601. - Ordinal types: From 3feda98ac5510dd0a7045d3d8a02af32da76fd7b Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 00:32:19 -0300 Subject: [PATCH 100/118] fix: all the comments --- src/Std/Time.lean | 3 +- src/Std/Time/Date/PlainDate.lean | 1 + src/Std/Time/DateTime.lean | 2 +- src/Std/Time/DateTime/PlainDateTime.lean | 4 +- src/Std/Time/DateTime/Timestamp.lean | 2 +- src/Std/Time/Duration.lean | 8 ++-- src/Std/Time/Internal/UnitVal.lean | 7 +++ src/Std/Time/Zoned.lean | 2 +- src/Std/Time/Zoned/Database/Windows.lean | 2 +- src/Std/Time/Zoned/ZoneRules.lean | 2 +- src/Std/Time/Zoned/ZonedDateTime.lean | 23 +++++++++- src/runtime/io.cpp | 12 ++--- tests/lean/run/timeClassOperations.lean | 4 +- tests/lean/run/timeIO.lean | 57 ++++++++++++++++++++++++ tests/lean/run/timeLocalDateTime.lean | 2 +- 15 files changed, 108 insertions(+), 23 deletions(-) diff --git a/src/Std/Time.lean b/src/Std/Time.lean index f85d37cef44b..4e9c2d3fd0dc 100644 --- a/src/Std/Time.lean +++ b/src/Std/Time.lean @@ -44,7 +44,7 @@ converted because they use an internal type called `UnitVal`. - Types with a correspondence to seconds: - `Day.Offset` - `Hour.Offset` - - `WeekOfYear.Offset` + - `Week.Offset` - `Millisecond.Offset` - `Nanosecond.Offset` - `Second.Offset` @@ -95,7 +95,6 @@ This type can have a fractional nanosecond part that can be negative or positive These types provide precision down to the day level, useful for representing and manipulating dates. - **`PlainDate`:** Represents a calendar date in the format `YYYY-MM-DD`. -- **`WeekDate`:** Uses the `YYYY-Www` format with week level precision. ## Time These types offer precision down to the nanosecond level, useful for representing and manipulating time of day. diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 0d8a6f5ecc6f..c5a6898fc88f 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -34,6 +34,7 @@ structure PlainDate where /-- Validates the date by ensuring that the year, month, and day form a correct and valid date. -/ valid : year.Valid month day + deriving Repr instance : Inhabited PlainDate where default := ⟨1, 1, 1, by decide⟩ diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean index eab45c7d8b5c..ec83e83f0790 100644 --- a/src/Std/Time/DateTime.lean +++ b/src/Std/Time/DateTime.lean @@ -26,7 +26,7 @@ Converts a `Timestamp` to a `PlainDateTime` -/ @[inline] def toPlainDateTimeAssumingUTC (timestamp : Timestamp) : PlainDateTime := - PlainDateTime.ofTimestamp timestamp + PlainDateTime.ofTimestampAssumingUTC timestamp /-- Converts a `PlainDate` to a `Timestamp` diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index fd40748d3853..83715f8dad53 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -30,7 +30,7 @@ structure PlainDateTime where -/ time : PlainTime - deriving Inhabited, BEq + deriving Inhabited, BEq, Repr namespace PlainDateTime @@ -46,7 +46,7 @@ def toTimestampAssumingUTC (dt : PlainDateTime) : Timestamp := /-- Converts a `Timestamp` to a `PlainDateTime`. -/ -def ofTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do +def ofTimestampAssumingUTC (stamp : Timestamp) : PlainDateTime := Id.run do let leapYearEpoch := 11017 let daysPer400Y := 365 * 400 + 97 let daysPer100Y := 365 * 100 + 24 diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index e56b076a94e0..7a32102323aa 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -17,7 +17,7 @@ open Internal set_option linter.all true /-- -`Timestamp` represents a specific point in time since the UNIX Epoch. +Represents an exact point in time as a UNIX Epoch timestamp. -/ structure Timestamp where diff --git a/src/Std/Time/Duration.lean b/src/Std/Time/Duration.lean index 36a28d86f456..54d6d0364340 100644 --- a/src/Std/Time/Duration.lean +++ b/src/Std/Time/Duration.lean @@ -82,10 +82,10 @@ def ofSeconds (s : Second.Offset) : Duration := by Creates a new `Duration` out of `Nanosecond.Offset`. -/ def ofNanoseconds (s : Nanosecond.Offset) : Duration := by - refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ - cases Int.le_total s.val 0 - next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) - next n => exact Or.inl (And.intro (Int.ediv_nonneg n (by decide)) (Int.tmod_nonneg 1000000000 n)) + refine ⟨s.ediv 1000000000, Bounded.LE.byMod s.val 1000000000 (by decide), ?_⟩ + cases Int.le_total s.val 0 + next n => exact Or.inr (And.intro (Int.ediv_le_ediv (by decide) n) (mod_nonpos 1000000000 n (by decide))) + next n => exact Or.inl (And.intro (Int.ediv_nonneg n (by decide)) (Int.tmod_nonneg 1000000000 n)) where mod_nonpos : ∀ {a : Int} (b : Int), (a ≤ 0) → (b ≥ 0) → 0 ≥ a.tmod b | .negSucc m, .ofNat n, _, _ => Int.neg_le_neg (Int.tmod_nonneg (↑n) (Int.ofNat_le.mpr (Nat.zero_le (m + 1)))) diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 97c9bf45a8f6..81be3e7b1e58 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -90,6 +90,13 @@ Subtracts one `UnitVal` value from another of the same ratio. def sub (u1 : UnitVal α) (u2 : UnitVal α) : UnitVal α := ⟨u1.val - u2.val⟩ +/-- +Returns the absolute value of a `UnitVal`. +-/ +@[inline] +def abs (u : UnitVal α) : UnitVal α := + ⟨u.val.natAbs⟩ + /-- Converts an `Offset` to another unit type. -/ diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 4125c2d776ce..790083d62d4b 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -25,7 +25,7 @@ def now : IO PlainDateTime := do let rules ← Database.defaultGetLocalZoneRules let ltt := rules.findLocalTimeTypeForTimestamp tm - return PlainDateTime.ofTimestamp tm |>.addSeconds ltt.getTimeZone.toSeconds + return PlainDateTime.ofTimestampAssumingUTC tm |>.addSeconds ltt.getTimeZone.toSeconds end PlainDateTime diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index 0ae44837832a..40c286ebaf88 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -20,7 +20,7 @@ namespace Windows /-- Fetches the next timezone transition for a given timezone identifier and timestamp. -/ -@[extern "lean_get_windows_next_transition"] +@[extern "lean_windows_get_next_transition"] opaque getNextTransition : @&String -> @&Int64 -> IO (Option (Int64 × TimeZone)) /-- diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index e196a3f84dbb..512c668524af 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -147,7 +147,7 @@ If the timestamp falls between two transitions, it returns the most recent trans -/ def findTransitionIndexForTimestamp (transitions : Array Transition) (timestamp : Timestamp) : Option Nat := let value := timestamp.toSecondsSinceUnixEpoch - if let some idx := transitions.findIdx? (fun t => t.time.val ≥ value.val) + if let some idx := transitions.findIdx? (fun t => t.time.val > value.val) then some idx else none diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 420e944c03bb..5afd25dac864 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -58,7 +58,28 @@ Creates a new `ZonedDateTime` out of a `PlainDateTime` and a `ZoneRules`. @[inline] def ofPlainDateTime (pdt : PlainDateTime) (zr : TimeZone.ZoneRules) : ZonedDateTime := let tm := pdt.toTimestampAssumingUTC - let tz := zr.findLocalTimeTypeForTimestamp tm |>.getTimeZone + + let transition := + let value := tm.toSecondsSinceUnixEpoch + if let some idx := zr.transitions.findIdx? (fun t => t.time.val ≥ value.val) + then do + let last ← zr.transitions.get? (idx - 1) + let next ← zr.transitions.get? idx <|> zr.transitions.back? + + let utcNext := next.time.sub last.localTimeType.gmtOffset.second.abs + + if utcNext.val > tm.toSecondsSinceUnixEpoch.val + then pure last + else pure next + + else zr.transitions.back? + + let tz := + transition + |>.map (·.localTimeType) + |>.getD zr.initialLocalTimeType + |>.getTimeZone + let tm := tm.subSeconds tz.toSeconds ZonedDateTime.mk (Thunk.mk fun _ => tm.toPlainDateTimeAssumingUTC.addSeconds tz.toSeconds) tm zr tz diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 6a823204ab28..ac8c3b9bb1c2 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -648,8 +648,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } -/* Std.Time.Database.Windows.getNextTransition : @&String -> @&Int -> IO (Option (Int × TimeZone)) */ -extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(b_obj_arg timezone_str, b_obj_arg tm_obj, obj_arg /* w */) { +/* Std.Time.Database.Windows.getNextTransition : @&String -> @&Int64 -> IO (Option (Int64 × TimeZone)) */ +extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezone_str, uint64_t tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; const char* dst_name_id = lean_string_cstr(timezone_str); @@ -668,14 +668,14 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(b_obj_arg timezo return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); } - int64_t timestamp_secs = lean_scalar_to_int64(tm_obj); + int64_t timestamp_secs = (int64_t)tm_obj; ucal_setMillis(cal, timestamp_secs * 1000, &status); if (U_FAILURE(status)) { ucal_close(cal); return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); } - + int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); if (U_FAILURE(status)) { @@ -737,8 +737,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_next_transition(b_obj_arg timezo lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); lean_ctor_set(lean_tz, 0, lean_int_to_int(offset_seconds)); - lean_ctor_set(lean_tz, 1, lean_mk_ascii_string_unchecked(dst_name)); - lean_ctor_set(lean_tz, 2, lean_mk_ascii_string_unchecked(display_name_str)); + lean_ctor_set(lean_tz, 1, lean_mk_string_from_bytes_unchecked(dst_name)); + lean_ctor_set(lean_tz, 2, lean_mk_string_from_bytes_unchecked(display_name_str)); lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); lean_object *lean_pair = lean_alloc_ctor(0, 2, 0); diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean index de4ca9c399f3..39f1967df911 100644 --- a/tests/lean/run/timeClassOperations.lean +++ b/tests/lean/run/timeClassOperations.lean @@ -165,8 +165,8 @@ info: "3600s" #guard_msgs in #eval zoned("2000-12-20T00:00:00-03:00") - zoned("2000-12-20T00:00:00-02:00") -def now := PlainDateTime.ofTimestamp ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ -def after := PlainDateTime.ofTimestamp ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ +def now := PlainDateTime.ofTimestampAssumingUTC ⟨1724859638, ⟨907328169, by decide⟩, by decide⟩ +def after := PlainDateTime.ofTimestampAssumingUTC ⟨1724859639, ⟨907641224, by decide⟩, by decide⟩ def duration := after - now /-- diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index 6adbf9db23ae..f80d5294a33e 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -11,3 +11,60 @@ open Std.Time let res ← Database.defaultGetLocalZoneRules if res.transitions.size < 1 then throw <| IO.userError "invalid quantity for the local timezone" + +/- +Java: +2013-10-19T23:59:59-03:00[America/Sao_Paulo] 1382237999 +2013-10-20T01:00-02:00[America/Sao_Paulo] 1382238000 +2013-10-20T01:00:01-02:00[America/Sao_Paulo] 1382238001 +-/ + +/-- +info: 2013-10-19T23:59:59.000000000-03:00 +2013-10-20T00:00:00.000000000-02:00 +2013-10-20T00:00:01.000000000-02:00 +-/ +#guard_msgs in +#eval do + let zr ← Database.defaultGetZoneRules "America/Sao_Paulo" + println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-19T23:59:59") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-20T00:00:00") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-20T00:00:01") zr}" + +/- +Java: +2019-11-03T01:59:59-05:00[America/Chicago] 1572764399 +2019-11-03T02:00-06:00[America/Chicago] 1572768000 +2019-11-03T02:59:59-06:00[America/Chicago] 1572771599 +-/ + +/-- +info: 2019-11-03T01:59:59.000000000-05:00 +2019-11-03T02:00:00.000000000-06:00 +2019-11-03T02:59:59.000000000-06:00 +-/ +#guard_msgs in +#eval do + let zr ← Database.defaultGetZoneRules "America/Chicago" + println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T01:59:59") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T02:00:00") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T02:59:59") zr}" + +/- +Java: +2003-10-26T01:59:59-05:00[America/Monterrey] 1067151599 +2003-10-26T02:00-06:00[America/Monterrey] 1067155200 +2003-10-26T02:59:59-06:00[America/Monterrey] 1067158799 +-/ + +/-- +info: 2003-10-26T01:59:59.000000000-05:00 +2003-10-26T02:00:00.000000000-06:00 +2003-10-26T02:59:59.000000000-06:00 +-/ +#guard_msgs in +#eval do + let zr ← Database.defaultGetZoneRules "America/Monterrey" + println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T01:59:59") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T02:00:00") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T02:59:59") zr}" diff --git a/tests/lean/run/timeLocalDateTime.lean b/tests/lean/run/timeLocalDateTime.lean index 3d2123e641d8..4863a638729a 100644 --- a/tests/lean/run/timeLocalDateTime.lean +++ b/tests/lean/run/timeLocalDateTime.lean @@ -44,7 +44,7 @@ info: 736952399 info: "09/05/1993 12:59:59" -/ #guard_msgs in -#eval PlainDateTime.ofTimestamp 736952399 |> format +#eval PlainDateTime.ofTimestampAssumingUTC 736952399 |> format /-- info: 736952399 From 0ddca8cb42bc770cd3f15bb7db102f9e143a98a8 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 00:42:16 -0300 Subject: [PATCH 101/118] fix: docs --- src/Std/Time/Date/PlainDate.lean | 8 +++++--- src/Std/Time/Zoned/DateTime.lean | 2 +- 2 files changed, 6 insertions(+), 4 deletions(-) diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index c5a6898fc88f..0d92b36641f2 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -102,7 +102,7 @@ def quarter (date : PlainDate) : Bounded.LE 1 4 := date.month.sub 1 |>.ediv 3 (by decide) |>.add 1 /-- -Transforms a tuple of a `PlainDate` into a `Day.Ordinal.OfYear`. +Transforms a `PlainDate` into a `Day.Ordinal.OfYear`. -/ def dayOfYear (date : PlainDate) : Day.Ordinal.OfYear date.year.isLeap := ValidDate.dayOfYear ⟨(date.month, date.day), date.valid⟩ @@ -298,7 +298,9 @@ def alignedWeekOfMonth (date : PlainDate) : Week.Ordinal.OfMonth := days |>.ediv 7 (by decide) |>.add 1 /-- -Sets the date to the specified `desiredWeekday`. +Sets the date to the specified `desiredWeekday`. If the `desiredWeekday` is the same as the current weekday, +the original `date` is returned without modification. If the `desiredWeekday` is in the future, the +function adjusts the date forward to the next occurrence of that weekday. -/ def withWeekday (date : PlainDate) (desiredWeekday : Weekday) : PlainDate := let weekday := date |>.weekday |>.toOrdinal @@ -315,7 +317,7 @@ def withWeekday (date : PlainDate) (desiredWeekday : Weekday) : PlainDate := date.addDays (Day.Offset.ofInt offset.toInt) /-- -Calculates the starting Monday of the ISO week-based year for a given year. +Calculates the wekk of the year starting Monday for a given year. -/ def weekOfYear (date : PlainDate) : Week.Ordinal := let y := date.year diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 0d13c4bf88b9..b66a12833346 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -413,7 +413,7 @@ def weekOfMonth (date : DateTime tz) : Bounded.LE 1 5 := /-- Determines the week of the month for the given `DateTime`. The week of the month is calculated based -on the day of the month and the weekday. Each week starts on Sunday because the entire library is +on the day of the month and the weekday. Each week starts on Monday because the entire library is based on the Gregorian Calendar. -/ @[inline] From 14a193fd813ac36fa277ded7acb0e4fbf6eb6658 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 00:48:45 -0300 Subject: [PATCH 102/118] fix: windows time --- src/runtime/io.cpp | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index ac8c3b9bb1c2..fd10acbf67eb 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -751,8 +751,8 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo #endif } -/* Std.Time.Database.Windows.getLocalTimeZoneIdentifierAt : Int → IO String */ -extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(obj_arg tm_obj, obj_arg /* w */) { +/* Std.Time.Database.Windows.getLocalTimeZoneIdentifierAt : @&Int64 → IO String */ +extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(uint64_t tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; UCalendar* cal = ucal_open(NULL, -1, NULL, UCAL_GREGORIAN, &status); @@ -760,8 +760,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(obj_arg tm_ if (U_FAILURE(status)) { return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); } - - int64_t timestamp_secs = lean_scalar_to_int64(tm_obj); +0 + int64_t timestamp_secs = (int64_t)tm_obj; ucal_setMillis(cal, timestamp_secs * 1000, &status); if (U_FAILURE(status)) { From 12cc80119fbb5a8e145bf00c52de2eba5b82cc53 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 10:57:08 -0300 Subject: [PATCH 103/118] feat: all the datetime functions --- src/Std/Time/Date/PlainDate.lean | 2 +- src/Std/Time/Format/Basic.lean | 2 +- src/Std/Time/Zoned.lean | 11 ++++ src/Std/Time/Zoned/DateTime.lean | 90 +++++++++++++++++++------------- src/runtime/io.cpp | 12 ++--- 5 files changed, 74 insertions(+), 43 deletions(-) diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index 0d92b36641f2..ef48a544ceac 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -317,7 +317,7 @@ def withWeekday (date : PlainDate) (desiredWeekday : Weekday) : PlainDate := date.addDays (Day.Offset.ofInt offset.toInt) /-- -Calculates the wekk of the year starting Monday for a given year. +Calculates the week of the year starting Monday for a given year. -/ def weekOfYear (date : PlainDate) : Week.Ordinal := let y := date.year diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 8bdfcea96e4b..88b62af53dd4 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -1345,7 +1345,7 @@ private def build (builder : DateBuilder) (aw : Awareness) : Option aw.type := none match aw with - | .only newTz => (DateTime.ofLocalDateTime · newTz) <$> datetime + | .only newTz => (ofPlainDateTime · newTz) <$> datetime | .any => (ZonedDateTime.ofPlainDateTime · (ZoneRules.ofTimeZone tz)) <$> datetime end DateBuilder diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean index 790083d62d4b..570007ce5ecb 100644 --- a/src/Std/Time/Zoned.lean +++ b/src/Std/Time/Zoned.lean @@ -73,6 +73,17 @@ Converts a `DateTime` to a `PlainTime` def toPlainTime (dt : DateTime tz) : PlainTime := dt.date.get.time +end DateTime +namespace DateTime + +/-- +Gets the current `ZonedDateTime`. +-/ +@[inline] +def now : IO (DateTime tz) := do + let tm ← Timestamp.now + return DateTime.ofTimestamp tm tz + end DateTime namespace ZonedDateTime diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index b66a12833346..41c646ceb743 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -74,7 +74,7 @@ Creates a new `DateTime` from a `PlainDateTime`, assuming that the `PlainDateTim is relative to the given `TimeZone`. -/ @[inline] -def ofLocalDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := +def ofPlainDateTime (date : PlainDateTime) (tz : TimeZone) : DateTime tz := let tm := date.subSeconds tz.toSeconds DateTime.mk (Timestamp.ofPlainDateTimeAssumingUTC tm) (Thunk.mk fun _ => date) @@ -83,118 +83,118 @@ Add `Hour.Offset` to a `DateTime`. -/ @[inline] def addHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addHours hours) tz + ofPlainDateTime (dt.date.get.addHours hours) tz /-- Subtract `Hour.Offset` from a `DateTime`. -/ @[inline] def subHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subHours hours) tz + ofPlainDateTime (dt.date.get.subHours hours) tz /-- Add `Minute.Offset` to a `DateTime`. -/ @[inline] def addMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addMinutes minutes) tz + ofPlainDateTime (dt.date.get.addMinutes minutes) tz /-- Subtract `Minute.Offset` from a `DateTime`. -/ @[inline] def subMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subMinutes minutes) tz + ofPlainDateTime (dt.date.get.subMinutes minutes) tz /-- Add `Second.Offset` to a `DateTime`. -/ @[inline] def addSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addSeconds seconds) tz + ofPlainDateTime (dt.date.get.addSeconds seconds) tz /-- Subtract `Second.Offset` from a `DateTime`. -/ @[inline] def subSeconds (dt : DateTime tz) (seconds : Second.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subSeconds seconds) tz + ofPlainDateTime (dt.date.get.subSeconds seconds) tz /-- Add `Millisecond.Offset` to a `DateTime`. -/ @[inline] def addMilliseconds (dt : DateTime tz) (milliseconds : Millisecond.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addMilliseconds milliseconds) tz + ofPlainDateTime (dt.date.get.addMilliseconds milliseconds) tz /-- Subtract `Millisecond.Offset` from a `DateTime`. -/ @[inline] def subMilliseconds (dt : DateTime tz) (milliseconds : Millisecond.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subMilliseconds milliseconds) tz + ofPlainDateTime (dt.date.get.subMilliseconds milliseconds) tz /-- Add `Nanosecond.Offset` to a `DateTime`. -/ @[inline] def addNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addNanoseconds nanoseconds) tz + ofPlainDateTime (dt.date.get.addNanoseconds nanoseconds) tz /-- Subtract `Nanosecond.Offset` from a `DateTime`. -/ @[inline] def subNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subNanoseconds nanoseconds) tz + ofPlainDateTime (dt.date.get.subNanoseconds nanoseconds) tz /-- Add `Day.Offset` to a `DateTime`. -/ @[inline] def addDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addDays days) tz + ofPlainDateTime (dt.date.get.addDays days) tz /-- Subtracts `Day.Offset` to a `DateTime`. -/ @[inline] def subDays (dt : DateTime tz) (days : Day.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subDays days) tz + ofPlainDateTime (dt.date.get.subDays days) tz /-- Add `Week.Offset` to a `DateTime`. -/ @[inline] def addWeeks (dt : DateTime tz) (weeks : Week.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addWeeks weeks) tz + ofPlainDateTime (dt.date.get.addWeeks weeks) tz /-- Subtracts `Week.Offset` to a `DateTime`. -/ @[inline] def subWeeks (dt : DateTime tz) (weeks : Week.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subWeeks weeks) tz + ofPlainDateTime (dt.date.get.subWeeks weeks) tz /-- Add `Month.Offset` to a `DateTime`, it clips the day to the last valid day of that month. -/ def addMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addMonthsClip months) tz + ofPlainDateTime (dt.date.get.addMonthsClip months) tz /-- Subtracts `Month.Offset` from a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def subMonthsClip (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subMonthsClip months) tz + ofPlainDateTime (dt.date.get.subMonthsClip months) tz /-- Add `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following month. -/ def addMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addMonthsRollOver months) tz + ofPlainDateTime (dt.date.get.addMonthsRollOver months) tz /-- Subtract `Month.Offset` from a `DateTime`, this function rolls over any excess days into the following @@ -202,7 +202,7 @@ month. -/ @[inline] def subMonthsRollOver (dt : DateTime tz) (months : Month.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subMonthsRollOver months) tz + ofPlainDateTime (dt.date.get.subMonthsRollOver months) tz /-- Add `Year.Offset` to a `DateTime`, this function rolls over any excess days into the following @@ -210,14 +210,14 @@ month. -/ @[inline] def addYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addYearsRollOver years) tz + ofPlainDateTime (dt.date.get.addYearsRollOver years) tz /-- Add `Year.Offset` to a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def addYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.addYearsClip years) tz + ofPlainDateTime (dt.date.get.addYearsClip years) tz /-- Subtract `Year.Offset` from a `DateTime`, this function rolls over any excess days into the following @@ -225,14 +225,14 @@ month. -/ @[inline] def subYearsRollOver (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subYearsRollOver years) tz + ofPlainDateTime (dt.date.get.subYearsRollOver years) tz /-- Subtract `Year.Offset` from to a `DateTime`, it clips the day to the last valid day of that month. -/ @[inline] def subYearsClip (dt : DateTime tz) (years : Year.Offset) : DateTime tz := - ofLocalDateTime (dt.date.get.subYearsClip years) tz + ofPlainDateTime (dt.date.get.subYearsClip years) tz /-- Creates a new `DateTime tz` by adjusting the day of the month to the given `days` value, with any @@ -240,7 +240,7 @@ out-of-range days clipped to the nearest valid date. -/ @[inline] def withDaysClip (dt : DateTime tz) (days : Day.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withDaysClip days) tz + ofPlainDateTime (dt.date.get.withDaysClip days) tz /-- Creates a new `DateTime tz` by adjusting the day of the month to the given `days` value, with any @@ -248,7 +248,7 @@ out-of-range days rolled over to the next month or year as needed. -/ @[inline] def withDaysRollOver (dt : DateTime tz) (days : Day.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withDaysRollOver days) tz + ofPlainDateTime (dt.date.get.withDaysRollOver days) tz /-- Creates a new `DateTime tz` by adjusting the month to the given `month` value. @@ -256,7 +256,7 @@ The day remains unchanged, and any invalid days for the new month will be handle -/ @[inline] def withMonthClip (dt : DateTime tz) (month : Month.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withMonthClip month) tz + ofPlainDateTime (dt.date.get.withMonthClip month) tz /-- Creates a new `DateTime tz` by adjusting the month to the given `month` value. @@ -264,7 +264,7 @@ The day is rolled over to the next valid month if necessary. -/ @[inline] def withMonthRollOver (dt : DateTime tz) (month : Month.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withMonthRollOver month) tz + ofPlainDateTime (dt.date.get.withMonthRollOver month) tz /-- Creates a new `DateTime tz` by adjusting the year to the given `year` value. The month and day remain unchanged, @@ -272,7 +272,7 @@ and any invalid days for the new year will be handled according to the `clip` be -/ @[inline] def withYearClip (dt : DateTime tz) (year : Year.Offset) : DateTime tz := - .ofLocalDateTime (dt.date.get.withYearClip year) tz + ofPlainDateTime (dt.date.get.withYearClip year) tz /-- Creates a new `DateTime tz` by adjusting the year to the given `year` value. The month and day are rolled @@ -280,35 +280,42 @@ over to the next valid month and day if necessary. -/ @[inline] def withYearRollOver (dt : DateTime tz) (year : Year.Offset) : DateTime tz := - .ofLocalDateTime (dt.date.get.withYearRollOver year) tz + ofPlainDateTime (dt.date.get.withYearRollOver year) tz /-- Creates a new `DateTime tz` by adjusting the `hour` component. -/ @[inline] def withHours (dt : DateTime tz) (hour : Hour.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withHours hour) tz + ofPlainDateTime (dt.date.get.withHours hour) tz /-- Creates a new `DateTime tz` by adjusting the `minute` component. -/ @[inline] def withMinutes (dt : DateTime tz) (minute : Minute.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withMinutes minute) tz + ofPlainDateTime (dt.date.get.withMinutes minute) tz /-- Creates a new `DateTime tz` by adjusting the `second` component. -/ @[inline] def withSeconds (dt : DateTime tz) (second : Sigma Second.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withSeconds second) tz + ofPlainDateTime (dt.date.get.withSeconds second) tz /-- Creates a new `DateTime tz` by adjusting the `nano` component. -/ @[inline] def withNanoseconds (dt : DateTime tz) (nano : Nanosecond.Ordinal) : DateTime tz := - .ofLocalDateTime (dt.date.get.withNanoseconds nano) tz + ofPlainDateTime (dt.date.get.withNanoseconds nano) tz + +/-- +Creates a new `DateTime tz` by adjusting the `millisecond` component. +-/ +@[inline] +def withMilliseconds (dt : DateTime tz) (milli : Millisecond.Ordinal) : DateTime tz := + ofPlainDateTime (dt.date.get.withMilliseconds milli) tz /-- Converts a `Timestamp` to a `PlainDateTime` @@ -363,14 +370,14 @@ def second (dt : DateTime tz) : Second.Ordinal dt.date.get.time.second.fst := Getter for the `Milliseconds` inside of a `DateTime` -/ @[inline] -def milliseconds (dt : DateTime tz) : Millisecond.Ordinal := +def millisecond (dt : DateTime tz) : Millisecond.Ordinal := dt.date.get.time.nanosecond.emod 1000 (by decide) /-- Getter for the `Nanosecond` inside of a `DateTime` -/ @[inline] -def nanoseconds (dt : DateTime tz) : Nanosecond.Ordinal := +def nanosecond (dt : DateTime tz) : Nanosecond.Ordinal := dt.date.get.time.nanosecond /-- @@ -386,6 +393,12 @@ Determines the era of the given `DateTime` based on its year. def era (date : DateTime tz) : Year.Era := date.year.era +/-- +Sets the `DateTime` to the specified `desiredWeekday`. +-/ +def withWeekday (dt : DateTime tz) (desiredWeekday : Weekday) : DateTime tz := + ofPlainDateTime (dt.date.get.withWeekday desiredWeekday) tz + /-- Checks if the `DateTime` is in a leap year. -/ @@ -427,6 +440,13 @@ Determines the quarter of the year for the given `DateTime`. def quarter (date : DateTime tz) : Bounded.LE 1 4 := date.date.get.quarter +/-- +Getter for the `PlainTime` inside of a `DateTime` +-/ +@[inline] +def time (zdt : DateTime tz) : PlainTime := + zdt.date.get.time + instance : HAdd (DateTime tz) (Day.Offset) (DateTime tz) where hAdd := addDays diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index fd10acbf67eb..acca7525f2ca 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -648,7 +648,7 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } -/* Std.Time.Database.Windows.getNextTransition : @&String -> @&Int64 -> IO (Option (Int64 × TimeZone)) */ +/* Std.Time.Database.Windows.getNextTransition : @&String -> Int64 -> IO (Option (Int64 × TimeZone)) */ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezone_str, uint64_t tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; @@ -697,7 +697,7 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo } int32_t tzIDLength = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_DST : UCAL_STANDARD, "en_US", tzID, 32, &status); - + if (U_FAILURE(status)) { ucal_close(cal); return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to timezone identifier"))); @@ -742,7 +742,7 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); lean_object *lean_pair = lean_alloc_ctor(0, 2, 0); - lean_ctor_set(lean_pair, 0, lean_int64_to_int(tm)); + lean_ctor_set(lean_pair, 0, (uint64_t)tm); lean_ctor_set(lean_pair, 1, lean_tz); return lean_io_result_mk_ok(mk_option_some(lean_pair)); @@ -751,7 +751,7 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo #endif } -/* Std.Time.Database.Windows.getLocalTimeZoneIdentifierAt : @&Int64 → IO String */ +/* Std.Time.Database.Windows.getLocalTimeZoneIdentifierAt : Int64 → IO String */ extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(uint64_t tm_obj, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; @@ -760,10 +760,10 @@ extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(uint64_t tm if (U_FAILURE(status)) { return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); } -0 + int64_t timestamp_secs = (int64_t)tm_obj; ucal_setMillis(cal, timestamp_secs * 1000, &status); - + if (U_FAILURE(status)) { ucal_close(cal); return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); From 9b3421277b4a9787907244ee2e698f1fc648ad2f Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 11:05:57 -0300 Subject: [PATCH 104/118] fix: docs --- src/Std/Time/Format/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 88b62af53dd4..f6c16e349d11 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -914,9 +914,9 @@ private def dateFromModifier (date : DateTime tz) : TypeFormat modifier := | .H _ => date.hour | .m _ => date.minute | .s _ => date.date.get.time.second - | .S _ => date.nanoseconds + | .S _ => date.nanosecond | .A _ => date.date.get.time.toMilliseconds - | .n _ => date.nanoseconds + | .n _ => date.nanosecond | .N _ => date.date.get.time.toNanoseconds | .V => tz.name | .z .short => tz.abbreviation From a41a8c400ac4b23cdd6083e527dc7b6c63cd4421 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 11:41:44 -0300 Subject: [PATCH 105/118] fix: two tests --- tests/lean/run/timeFormats.lean | 6 +++--- tests/lean/run/timeParse.lean | 6 +++--- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index d71ea8b70a19..ee9cfd59eed8 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -134,9 +134,9 @@ This PlainDate is relative to the local time. -/ def PlainDate : PlainDateTime := Timestamp.toPlainDateTimeAssumingUTC (Timestamp.ofSecondsSinceUnixEpoch localTm) -def dateBR₁ := DateTime.ofLocalDateTime PlainDate brTZ -def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ -def dateUTC₁ := DateTime.ofLocalDateTime PlainDate .UTC +def dateBR₁ := DateTime.ofPlainDateTime PlainDate brTZ +def dateJP₁ := DateTime.ofPlainDateTime PlainDate jpTZ +def dateUTC₁ := DateTime.ofPlainDateTime PlainDate .UTC /-- info: "Thursday, August 15, 2024 14:03:47 -0300" diff --git a/tests/lean/run/timeParse.lean b/tests/lean/run/timeParse.lean index a16c3df1cbfc..fcd7b6be0b77 100644 --- a/tests/lean/run/timeParse.lean +++ b/tests/lean/run/timeParse.lean @@ -143,9 +143,9 @@ This PlainDate is relative to the local time. -/ def PlainDate : PlainDateTime := Timestamp.toPlainDateTimeAssumingUTC (Timestamp.ofSecondsSinceUnixEpoch localTm) -def dateBR₁ := DateTime.ofLocalDateTime PlainDate brTZ -def dateJP₁ := DateTime.ofLocalDateTime PlainDate jpTZ -def dateUTC₁ := DateTime.ofLocalDateTime PlainDate .UTC +def dateBR₁ := DateTime.ofPlainDateTime PlainDate brTZ +def dateJP₁ := DateTime.ofPlainDateTime PlainDate jpTZ +def dateUTC₁ := DateTime.ofPlainDateTime PlainDate .UTC /-- info: "2024-08-15T14:03:47.000000000-03:00" From 68a12ac0bf5cf30a14b0fbd7f4faacc8ebaaae11 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 11:42:57 -0300 Subject: [PATCH 106/118] feat: add better empty transitions error --- src/Std/Time/Zoned/Database/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 349e9fadb926..66aa46f91c07 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -92,7 +92,7 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := let first ← if let some res := transitions.get? 0 then .ok res - else .error "empty transitions" + else .error s!"empty transitions for {id}" .ok { transitions, initialLocalTimeType := first.localTimeType } From 74e63c02248f4af50e45f05300f783757fb614a8 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 12:10:16 -0300 Subject: [PATCH 107/118] fix: timeIO test --- tests/lean/run/timeIO.lean | 5 ----- 1 file changed, 5 deletions(-) diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index f80d5294a33e..5198c94e400f 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -7,11 +7,6 @@ open Std.Time if res.transitions.size < 1 then throw <| IO.userError "invalid quantity for America/Sao_Paulo" -#eval do - let res ← Database.defaultGetLocalZoneRules - if res.transitions.size < 1 then - throw <| IO.userError "invalid quantity for the local timezone" - /- Java: 2013-10-19T23:59:59-03:00[America/Sao_Paulo] 1382237999 From 8158cf95aeaf3966f1655110ad73d43a49d97908 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Fri, 8 Nov 2024 19:32:53 -0300 Subject: [PATCH 108/118] fix: tz on windows --- src/Std/Time/Zoned/Database/Windows.lean | 10 ++++++---- src/runtime/io.cpp | 14 ++++++++------ 2 files changed, 14 insertions(+), 10 deletions(-) diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index 40c286ebaf88..2f704baa2b92 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -21,13 +21,13 @@ namespace Windows Fetches the next timezone transition for a given timezone identifier and timestamp. -/ @[extern "lean_windows_get_next_transition"] -opaque getNextTransition : @&String -> @&Int64 -> IO (Option (Int64 × TimeZone)) +opaque getNextTransition : @&String -> Int64 -> IO (Option (Int64 × TimeZone)) /-- Fetches the timezone at a timestamp. -/ @[extern "lean_get_windows_local_timezone_id_at"] -opaque getLocalTimeZoneIdentifierAt : @&Int64 → IO String +opaque getLocalTimeZoneIdentifierAt : Int64 → IO String /-- Retrieves the timezone rules, including all transitions, for a given timezone identifier. @@ -38,8 +38,9 @@ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do while true do let result ← Windows.getNextTransition id start + if let some res := result then - transitions := transitions.push { time := Second.Offset.ofInt res.fst.toInt, localTimeType := { + transitions := transitions.push { time := Second.Offset.ofInt start.toInt, localTimeType := { gmtOffset := res.snd.offset, abbreviation := res.snd.abbreviation, identifier := res.snd.name, @@ -48,7 +49,8 @@ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do utLocal := .local }} - if res.fst ≤ start then + -- Avoid zone rules for more than year 3000 + if res.fst ≤ start ∨ res.fst >= 32503690800 then break start := res.fst diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index acca7525f2ca..83c530b985da 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -704,7 +704,8 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo } char dst_name[256]; - u_strToUTF8(dst_name, sizeof(dst_name), NULL, tzID, tzIDLength, &status); + int32_t dst_name_len; + u_strToUTF8(dst_name, sizeof(dst_name), &dst_name_len, tzID, tzIDLength, &status); UChar display_name[32]; int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_SHORT_DST : UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); @@ -715,7 +716,8 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo } char display_name_str[256]; - u_strToUTF8(display_name_str, sizeof(display_name_str), NULL, display_name, display_name_len, &status); + int32_t display_name_str_len; + u_strToUTF8(display_name_str, sizeof(display_name_str), &display_name_str_len, display_name, display_name_len, &status); if (U_FAILURE(status)) { ucal_close(cal); @@ -737,12 +739,12 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); lean_ctor_set(lean_tz, 0, lean_int_to_int(offset_seconds)); - lean_ctor_set(lean_tz, 1, lean_mk_string_from_bytes_unchecked(dst_name)); - lean_ctor_set(lean_tz, 2, lean_mk_string_from_bytes_unchecked(display_name_str)); + lean_ctor_set(lean_tz, 1, lean_mk_string_from_bytes_unchecked(dst_name, dst_name_len)); + lean_ctor_set(lean_tz, 2, lean_mk_string_from_bytes_unchecked(display_name_str, display_name_str_len)); lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); - + lean_object *lean_pair = lean_alloc_ctor(0, 2, 0); - lean_ctor_set(lean_pair, 0, (uint64_t)tm); + lean_ctor_set(lean_pair, 0, lean_box_uint64((uint64_t)tm)); lean_ctor_set(lean_pair, 1, lean_tz); return lean_io_result_mk_ok(mk_option_some(lean_pair)); From 70c1f0f6ab66221079971029a0ca2f4b1dd05a95 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 11 Nov 2024 09:58:10 -0300 Subject: [PATCH 109/118] fix: RFC 8536 states that it theres no transition the index 0 of the local times is the default time --- src/Std/Time/Zoned/Database/Basic.lean | 10 ++++++---- tests/lean/run/timeIO.lean | 17 +++++++++++++++++ 2 files changed, 23 insertions(+), 4 deletions(-) diff --git a/src/Std/Time/Zoned/Database/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean index 66aa46f91c07..7b046d2ce75c 100644 --- a/src/Std/Time/Zoned/Database/Basic.lean +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -89,13 +89,15 @@ def convertTZifV1 (tz : TZif.TZifV1) (id : String) : Except String ZoneRules := then transitions := transitions.push result else .error s!"cannot convert transition {i} of the file" - let first ← - if let some res := transitions.get? 0 + -- Local time for timestamps before the first transition is specified by the first time + -- type (time type 0). + + let initialLocalTimeType ← + if let some res := convertLocalTimeType 0 tz id then .ok res else .error s!"empty transitions for {id}" - - .ok { transitions, initialLocalTimeType := first.localTimeType } + .ok { transitions, initialLocalTimeType } /-- Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index 5198c94e400f..20da7418b8b6 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -2,11 +2,28 @@ import Std.Time import Init open Std.Time +/- +Test for quantity +-/ + #eval do let res ← Database.defaultGetZoneRules "America/Sao_Paulo" if res.transitions.size < 1 then throw <| IO.userError "invalid quantity for America/Sao_Paulo" +/-- +info: { gmtOffset := { second := 0 }, + isDst := false, + abbreviation := "UTC", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "Etc/UTC" } +-/ +#guard_msgs in +#eval do + let res ← Database.defaultGetZoneRules "Etc/UTC" + println! repr res.initialLocalTimeType + /- Java: 2013-10-19T23:59:59-03:00[America/Sao_Paulo] 1382237999 From 4f6b2cd614f1d653e1916118b8ede8bcb1d295f6 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 11 Nov 2024 10:31:17 -0300 Subject: [PATCH 110/118] fix: timezone on windows --- src/Std/Time/Zoned/Database/Windows.lean | 34 ++++++++++-------- src/runtime/io.cpp | 45 +++++++++++++----------- tests/lean/run/timeIO.lean | 9 ++--- 3 files changed, 46 insertions(+), 42 deletions(-) diff --git a/src/Std/Time/Zoned/Database/Windows.lean b/src/Std/Time/Zoned/Database/Windows.lean index 2f704baa2b92..2a9a0bea23f3 100644 --- a/src/Std/Time/Zoned/Database/Windows.lean +++ b/src/Std/Time/Zoned/Database/Windows.lean @@ -21,7 +21,7 @@ namespace Windows Fetches the next timezone transition for a given timezone identifier and timestamp. -/ @[extern "lean_windows_get_next_transition"] -opaque getNextTransition : @&String -> Int64 -> IO (Option (Int64 × TimeZone)) +opaque getNextTransition : @&String → Int64 → Bool → IO (Option (Int64 × TimeZone)) /-- Fetches the timezone at a timestamp. @@ -36,18 +36,16 @@ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do let mut start := -2147483648 let mut transitions : Array TimeZone.Transition := #[] + let mut initialLocalTimeType ← + if let some res := ← Windows.getNextTransition id start true + then pure (toLocalTime res.snd) + else throw (IO.userError "cannot find first transition in zone rules") + while true do - let result ← Windows.getNextTransition id start + let result ← Windows.getNextTransition id start false if let some res := result then - transitions := transitions.push { time := Second.Offset.ofInt start.toInt, localTimeType := { - gmtOffset := res.snd.offset, - abbreviation := res.snd.abbreviation, - identifier := res.snd.name, - isDst := res.snd.isDST, - wall := .wall, - utLocal := .local - }} + transitions := transitions.push { time := Second.Offset.ofInt start.toInt, localTimeType := toLocalTime res.snd } -- Avoid zone rules for more than year 3000 if res.fst ≤ start ∨ res.fst >= 32503690800 then @@ -57,13 +55,19 @@ def getZoneRules (id : String) : IO TimeZone.ZoneRules := do else break - let initialLocalTimeType ← do - if let some res := transitions.get? 0 - then pure res.localTimeType - else throw (IO.userError "cannot find first transition in zone rules") - return { transitions, initialLocalTimeType } + where + toLocalTime (res : TimeZone) : TimeZone.LocalTimeType := + { + gmtOffset := res.offset, + abbreviation := res.abbreviation, + identifier := res.name, + isDst := res.isDST, + wall := .wall, + utLocal := .local + } + end Windows /-- diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 83c530b985da..1a043e3868e0 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -648,8 +648,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { return lean_io_result_mk_ok(lean_ts); } -/* Std.Time.Database.Windows.getNextTransition : @&String -> Int64 -> IO (Option (Int64 × TimeZone)) */ -extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezone_str, uint64_t tm_obj, obj_arg /* w */) { +/* Std.Time.Database.Windows.getNextTransition : @&String -> Int64 -> Bool -> IO (Option (Int64 × TimeZone)) */ +extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezone_str, uint64_t tm_obj, uint8 default_time, obj_arg /* w */) { #if defined(LEAN_WINDOWS) UErrorCode status = U_ZERO_ERROR; const char* dst_name_id = lean_string_cstr(timezone_str); @@ -668,34 +668,40 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); } - int64_t timestamp_secs = (int64_t)tm_obj; + int64_t tm = 0; - ucal_setMillis(cal, timestamp_secs * 1000, &status); - if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); - } + if (!default_time) { + int64_t timestamp_secs = (int64_t)tm_obj; - int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); + ucal_setMillis(cal, timestamp_secs * 1000, &status); + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); + } - if (U_FAILURE(status)) { - ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); - } + UDate nextTransition; + if (!ucal_getTimeZoneTransitionDate(cal, UCAL_TZ_TRANSITION_NEXT, &nextTransition, &status)) { + ucal_close(cal); + return io_result_mk_ok(mk_option_none()); + } - int is_dst = dst_offset != 0; + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get next transation"))); + } - UDate nextTransition; - if (!ucal_getTimeZoneTransitionDate(cal, UCAL_TZ_TRANSITION_NEXT, &nextTransition, &status)) { - ucal_close(cal); - return io_result_mk_ok(mk_option_none()); + tm = (int64_t)(nextTransition / 1000.0); } + + int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); if (U_FAILURE(status)) { ucal_close(cal); - return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get next transation"))); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); } + int is_dst = dst_offset != 0; + int32_t tzIDLength = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_DST : UCAL_STANDARD, "en_US", tzID, 32, &status); if (U_FAILURE(status)) { @@ -732,7 +738,6 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); } - int64_t tm = (int64_t)(nextTransition / 1000.0); ucal_close(cal); int offset_seconds = zone_offset / 1000; diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index 20da7418b8b6..de1dda2eacda 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -12,17 +12,12 @@ Test for quantity throw <| IO.userError "invalid quantity for America/Sao_Paulo" /-- -info: { gmtOffset := { second := 0 }, - isDst := false, - abbreviation := "UTC", - wall := Std.Time.TimeZone.StdWall.standard, - utLocal := Std.Time.TimeZone.UTLocal.ut, - identifier := "Etc/UTC" } +info: { second := 0 } -/ #guard_msgs in #eval do let res ← Database.defaultGetZoneRules "Etc/UTC" - println! repr res.initialLocalTimeType + println! repr res.initialLocalTimeType.gmtOffset /- Java: From 0f6a93025f0092c0c3aa7312cc36afc994a03311 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Mon, 11 Nov 2024 11:34:54 -0300 Subject: [PATCH 111/118] test: add test --- tests/lean/run/timeIO.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index de1dda2eacda..3d9a805ae8e7 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -19,6 +19,12 @@ info: { second := 0 } let res ← Database.defaultGetZoneRules "Etc/UTC" println! repr res.initialLocalTimeType.gmtOffset +/- +The idea is just to check if there's no errors while computing the local zone rules. +-/ +#eval do + discard <| Database.defaultGetLocalZoneRules + /- Java: 2013-10-19T23:59:59-03:00[America/Sao_Paulo] 1382237999 From 08310f0d3a2ef097e02b8aeec2f2942ba864ae3d Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 12 Nov 2024 09:49:42 -0300 Subject: [PATCH 112/118] fix: test for tzif --- tests/lean/run/timeTzifParse.lean | 41 ++++++++++++++++++++++++++++++- 1 file changed, 40 insertions(+), 1 deletion(-) diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean index e0f5f2e492be..a82603691cba 100644 --- a/tests/lean/run/timeTzifParse.lean +++ b/tests/lean/run/timeTzifParse.lean @@ -1,7 +1,46 @@ import Std.Time open Std.Time -def file := ByteArray.mk #[84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0, 0, 0, 3, 0, 0, 0, 12, 150, 170, 114, 180, 184, 15, 73, 224, 184, 253, 64, 160, 185, 241, 52, 48, 186, 222, 116, 32, 218, 56, 174, 48, 218, 235, 250, 48, 220, 25, 225, 176, 220, 185, 89, 32, 221, 251, 21, 48, 222, 155, 222, 32, 223, 221, 154, 48, 224, 84, 51, 32, 244, 90, 9, 48, 245, 5, 94, 32, 246, 192, 100, 48, 247, 14, 30, 160, 248, 81, 44, 48, 248, 199, 197, 32, 250, 10, 210, 176, 250, 168, 248, 160, 251, 236, 6, 48, 252, 139, 125, 160, 29, 201, 142, 48, 30, 120, 215, 160, 31, 160, 53, 176, 32, 51, 207, 160, 33, 129, 105, 48, 34, 11, 200, 160, 35, 88, 16, 176, 35, 226, 112, 32, 37, 55, 242, 176, 37, 212, 199, 32, 39, 33, 15, 48, 39, 189, 227, 160, 41, 0, 241, 48, 41, 148, 139, 32, 42, 234, 13, 176, 43, 107, 50, 160, 44, 192, 181, 48, 45, 102, 196, 32, 46, 160, 151, 48, 47, 70, 166, 32, 48, 128, 121, 48, 49, 29, 77, 160, 50, 87, 32, 176, 51, 6, 106, 32, 52, 56, 84, 48, 52, 248, 193, 32, 54, 32, 31, 48, 54, 207, 104, 160, 55, 246, 198, 176, 56, 184, 133, 32, 57, 223, 227, 48, 58, 143, 44, 160, 59, 200, 255, 176, 60, 111, 14, 160, 61, 196, 145, 48, 62, 78, 240, 160, 63, 145, 254, 48, 64, 46, 210, 160, 65, 134, 248, 48, 66, 23, 239, 32, 67, 81, 194, 48, 67, 247, 209, 32, 69, 77, 83, 176, 69, 224, 237, 160, 71, 17, 134, 48, 71, 183, 149, 32, 72, 250, 162, 176, 73, 151, 119, 32, 74, 218, 132, 176, 75, 128, 147, 160, 76, 186, 102, 176, 77, 96, 117, 160, 78, 154, 72, 176, 79, 73, 146, 32, 80, 131, 101, 48, 81, 32, 57, 160, 82, 99, 71, 48, 83, 0, 27, 160, 84, 67, 41, 48, 84, 233, 56, 32, 86, 35, 11, 48, 86, 201, 26, 32, 88, 2, 237, 48, 88, 168, 252, 32, 89, 226, 207, 48, 90, 136, 222, 32, 91, 222, 96, 176, 92, 104, 192, 32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255, 213, 208, 0, 8, 76, 77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0, 0, 0, 3, 0, 0, 0, 12, 255, 255, 255, 255, 150, 170, 114, 180, 255, 255, 255, 255, 184, 15, 73, 224, 255, 255, 255, 255, 184, 253, 64, 160, 255, 255, 255, 255, 185, 241, 52, 48, 255, 255, 255, 255, 186, 222, 116, 32, 255, 255, 255, 255, 218, 56, 174, 48, 255, 255, 255, 255, 218, 235, 250, 48, 255, 255, 255, 255, 220, 25, 225, 176, 255, 255, 255, 255, 220, 185, 89, 32, 255, 255, 255, 255, 221, 251, 21, 48, 255, 255, 255, 255, 222, 155, 222, 32, 255, 255, 255, 255, 223, 221, 154, 48, 255, 255, 255, 255, 224, 84, 51, 32, 255, 255, 255, 255, 244, 90, 9, 48, 255, 255, 255, 255, 245, 5, 94, 32, 255, 255, 255, 255, 246, 192, 100, 48, 255, 255, 255, 255, 247, 14, 30, 160, 255, 255, 255, 255, 248, 81, 44, 48, 255, 255, 255, 255, 248, 199, 197, 32, 255, 255, 255, 255, 250, 10, 210, 176, 255, 255, 255, 255, 250, 168, 248, 160, 255, 255, 255, 255, 251, 236, 6, 48, 255, 255, 255, 255, 252, 139, 125, 160, 0, 0, 0, 0, 29, 201, 142, 48, 0, 0, 0, 0, 30, 120, 215, 160, 0, 0, 0, 0, 31, 160, 53, 176, 0, 0, 0, 0, 32, 51, 207, 160, 0, 0, 0, 0, 33, 129, 105, 48, 0, 0, 0, 0, 34, 11, 200, 160, 0, 0, 0, 0, 35, 88, 16, 176, 0, 0, 0, 0, 35, 226, 112, 32, 0, 0, 0, 0, 37, 55, 242, 176, 0, 0, 0, 0, 37, 212, 199, 32, 0, 0, 0, 0, 39, 33, 15, 48, 0, 0, 0, 0, 39, 189, 227, 160, 0, 0, 0, 0, 41, 0, 241, 48, 0, 0, 0, 0, 41, 148, 139, 32, 0, 0, 0, 0, 42, 234, 13, 176, 0, 0, 0, 0, 43, 107, 50, 160, 0, 0, 0, 0, 44, 192, 181, 48, 0, 0, 0, 0, 45, 102, 196, 32, 0, 0, 0, 0, 46, 160, 151, 48, 0, 0, 0, 0, 47, 70, 166, 32, 0, 0, 0, 0, 48, 128, 121, 48, 0, 0, 0, 0, 49, 29, 77, 160, 0, 0, 0, 0, 50, 87, 32, 176, 0, 0, 0, 0, 51, 6, 106, 32, 0, 0, 0, 0, 52, 56, 84, 48, 0, 0, 0, 0, 52, 248, 193, 32, 0, 0, 0, 0, 54, 32, 31, 48, 0, 0, 0, 0, 54, 207, 104, 160, 0, 0, 0, 0, 55, 246, 198, 176, 0, 0, 0, 0, 56, 184, 133, 32, 0, 0, 0, 0, 57, 223, 227, 48, 0, 0, 0, 0, 58, 143, 44, 160, 0, 0, 0, 0, 59, 200, 255, 176, 0, 0, 0, 0, 60, 111, 14, 160, 0, 0, 0, 0, 61, 196, 145, 48, 0, 0, 0, 0, 62, 78, 240, 160, 0, 0, 0, 0, 63, 145, 254, 48, 0, 0, 0, 0, 64, 46, 210, 160, 0, 0, 0, 0, 65, 134, 248, 48, 0, 0, 0, 0, 66, 23, 239, 32, 0, 0, 0, 0, 67, 81, 194, 48, 0, 0, 0, 0, 67, 247, 209, 32, 0, 0, 0, 0, 69, 77, 83, 176, 0, 0, 0, 0, 69, 224, 237, 160, 0, 0, 0, 0, 71, 17, 134, 48, 0, 0, 0, 0, 71, 183, 149, 32, 0, 0, 0, 0, 72, 250, 162, 176, 0, 0, 0, 0, 73, 151, 119, 32, 0, 0, 0, 0, 74, 218, 132, 176, 0, 0, 0, 0, 75, 128, 147, 160, 0, 0, 0, 0, 76, 186, 102, 176, 0, 0, 0, 0, 77, 96, 117, 160, 0, 0, 0, 0, 78, 154, 72, 176, 0, 0, 0, 0, 79, 73, 146, 32, 0, 0, 0, 0, 80, 131, 101, 48, 0, 0, 0, 0, 81, 32, 57, 160, 0, 0, 0, 0, 82, 99, 71, 48, 0, 0, 0, 0, 83, 0, 27, 160, 0, 0, 0, 0, 84, 67, 41, 48, 0, 0, 0, 0, 84, 233, 56, 32, 0, 0, 0, 0, 86, 35, 11, 48, 0, 0, 0, 0, 86, 201, 26, 32, 0, 0, 0, 0, 88, 2, 237, 48, 0, 0, 0, 0, 88, 168, 252, 32, 0, 0, 0, 0, 89, 226, 207, 48, 0, 0, 0, 0, 90, 136, 222, 32, 0, 0, 0, 0, 91, 222, 96, 176, 0, 0, 0, 0, 92, 104, 192, 32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255, 213, 208, 0, 8, 76, 77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 10, 60, 45, 48, 51, 62, 51, 10] +def file := ByteArray.mk <| +#[84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0] +++ #[0, 0, 3, 0, 0, 0, 12, 150, 170, 114, 180, 184, 15, 73, 224, 184, 253, 64, 160, 185, 241, 52, 48, 186, 222, 116, 32, 218, 56, 174, 48, 218, 235, 250, 48, 220, 25] +++ #[225, 176, 220, 185, 89, 32, 221, 251, 21, 48, 222, 155, 222, 32, 223, 221, 154, 48, 224, 84, 51, 32, 244, 90, 9, 48, 245, 5, 94, 32, 246, 192, 100, 48, 247, 14, 30] +++ #[160, 248, 81, 44, 48, 248, 199, 197, 32, 250, 10, 210, 176, 250, 168, 248, 160, 251, 236, 6, 48, 252, 139, 125, 160, 29, 201, 142, 48, 30, 120, 215, 160, 31, 160, 53, 176] +++ #[32, 51, 207, 160, 33, 129, 105, 48, 34, 11, 200, 160, 35, 88, 16, 176, 35, 226, 112, 32, 37, 55, 242, 176, 37, 212, 199, 32, 39, 33, 15, 48, 39, 189, 227, 160, 41] +++ #[0, 241, 48, 41, 148, 139, 32, 42, 234, 13, 176, 43, 107, 50, 160, 44, 192, 181, 48, 45, 102, 196, 32, 46, 160, 151, 48, 47, 70, 166, 32, 48, 128, 121, 48, 49, 29] +++ #[77, 160, 50, 87, 32, 176, 51, 6, 106, 32, 52, 56, 84, 48, 52, 248, 193, 32, 54, 32, 31, 48, 54, 207, 104, 160, 55, 246, 198, 176, 56, 184, 133, 32, 57, 223, 227] +++ #[48, 58, 143, 44, 160, 59, 200, 255, 176, 60, 111, 14, 160, 61, 196, 145, 48, 62, 78, 240, 160, 63, 145, 254, 48, 64, 46, 210, 160, 65, 134, 248, 48, 66, 23, 239, 32] +++ #[67, 81, 194, 48, 67, 247, 209, 32, 69, 77, 83, 176, 69, 224, 237, 160, 71, 17, 134, 48, 71, 183, 149, 32, 72, 250, 162, 176, 73, 151, 119, 32, 74, 218, 132, 176, 75] +++ #[128, 147, 160, 76, 186, 102, 176, 77, 96, 117, 160, 78, 154, 72, 176, 79, 73, 146, 32, 80, 131, 101, 48, 81, 32, 57, 160, 82, 99, 71, 48, 83, 0, 27, 160, 84, 67] +++ #[41, 48, 84, 233, 56, 32, 86, 35, 11, 48, 86, 201, 26, 32, 88, 2, 237, 48, 88, 168, 252, 32, 89, 226, 207, 48, 90, 136, 222, 32, 91, 222, 96, 176, 92, 104, 192] +++ #[32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1] +++ #[2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2] +++ #[1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255, 213, 208, 0, 8, 76] +++ #[77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 84, 90, 105, 102, 50, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] +++ #[0, 0, 0, 0, 0, 0, 0, 0, 0, 91, 0, 0, 0, 3, 0, 0, 0, 12, 255, 255, 255, 255, 150, 170, 114, 180, 255, 255, 255, 255, 184, 15, 73, 224, 255, 255, 255] +++ #[255, 184, 253, 64, 160, 255, 255, 255, 255, 185, 241, 52, 48, 255, 255, 255, 255, 186, 222, 116, 32, 255, 255, 255, 255, 218, 56, 174, 48, 255, 255, 255, 255, 218, 235, 250, 48] +++ #[255, 255, 255, 255, 220, 25, 225, 176, 255, 255, 255, 255, 220, 185, 89, 32, 255, 255, 255, 255, 221, 251, 21, 48, 255, 255, 255, 255, 222, 155, 222, 32, 255, 255, 255, 255, 223] +++ #[221, 154, 48, 255, 255, 255, 255, 224, 84, 51, 32, 255, 255, 255, 255, 244, 90, 9, 48, 255, 255, 255, 255, 245, 5, 94, 32, 255, 255, 255, 255, 246, 192, 100, 48, 255, 255] +++ #[255, 255, 247, 14, 30, 160, 255, 255, 255, 255, 248, 81, 44, 48, 255, 255, 255, 255, 248, 199, 197, 32, 255, 255, 255, 255, 250, 10, 210, 176, 255, 255, 255, 255, 250, 168, 248] +++ #[160, 255, 255, 255, 255, 251, 236, 6, 48, 255, 255, 255, 255, 252, 139, 125, 160, 0, 0, 0, 0, 29, 201, 142, 48, 0, 0, 0, 0, 30, 120, 215, 160, 0, 0, 0, 0] +++ #[31, 160, 53, 176, 0, 0, 0, 0, 32, 51, 207, 160, 0, 0, 0, 0, 33, 129, 105, 48, 0, 0, 0, 0, 34, 11, 200, 160, 0, 0, 0, 0, 35, 88, 16, 176, 0] +++ #[0, 0, 0, 35, 226, 112, 32, 0, 0, 0, 0, 37, 55, 242, 176, 0, 0, 0, 0, 37, 212, 199, 32, 0, 0, 0, 0, 39, 33, 15, 48, 0, 0, 0, 0, 39, 189] +++ #[227, 160, 0, 0, 0, 0, 41, 0, 241, 48, 0, 0, 0, 0, 41, 148, 139, 32, 0, 0, 0, 0, 42, 234, 13, 176, 0, 0, 0, 0, 43, 107, 50, 160, 0, 0, 0] +++ #[0, 44, 192, 181, 48, 0, 0, 0, 0, 45, 102, 196, 32, 0, 0, 0, 0, 46, 160, 151, 48, 0, 0, 0, 0, 47, 70, 166, 32, 0, 0, 0, 0, 48, 128, 121, 48] +++ #[0, 0, 0, 0, 49, 29, 77, 160, 0, 0, 0, 0, 50, 87, 32, 176, 0, 0, 0, 0, 51, 6, 106, 32, 0, 0, 0, 0, 52, 56, 84, 48, 0, 0, 0, 0, 52] +++ #[248, 193, 32, 0, 0, 0, 0, 54, 32, 31, 48, 0, 0, 0, 0, 54, 207, 104, 160, 0, 0, 0, 0, 55, 246, 198, 176, 0, 0, 0, 0, 56, 184, 133, 32, 0, 0] +++ #[0, 0, 57, 223, 227, 48, 0, 0, 0, 0, 58, 143, 44, 160, 0, 0, 0, 0, 59, 200, 255, 176, 0, 0, 0, 0, 60, 111, 14, 160, 0, 0, 0, 0, 61, 196, 145] +++ #[48, 0, 0, 0, 0, 62, 78, 240, 160, 0, 0, 0, 0, 63, 145, 254, 48, 0, 0, 0, 0, 64, 46, 210, 160, 0, 0, 0, 0, 65, 134, 248, 48, 0, 0, 0, 0] +++ #[66, 23, 239, 32, 0, 0, 0, 0, 67, 81, 194, 48, 0, 0, 0, 0, 67, 247, 209, 32, 0, 0, 0, 0, 69, 77, 83, 176, 0, 0, 0, 0, 69, 224, 237, 160, 0] +++ #[0, 0, 0, 71, 17, 134, 48, 0, 0, 0, 0, 71, 183, 149, 32, 0, 0, 0, 0, 72, 250, 162, 176, 0, 0, 0, 0, 73, 151, 119, 32, 0, 0, 0, 0, 74, 218] +++ #[132, 176, 0, 0, 0, 0, 75, 128, 147, 160, 0, 0, 0, 0, 76, 186, 102, 176, 0, 0, 0, 0, 77, 96, 117, 160, 0, 0, 0, 0, 78, 154, 72, 176, 0, 0, 0] +++ #[0, 79, 73, 146, 32, 0, 0, 0, 0, 80, 131, 101, 48, 0, 0, 0, 0, 81, 32, 57, 160, 0, 0, 0, 0, 82, 99, 71, 48, 0, 0, 0, 0, 83, 0, 27, 160] +++ #[0, 0, 0, 0, 84, 67, 41, 48, 0, 0, 0, 0, 84, 233, 56, 32, 0, 0, 0, 0, 86, 35, 11, 48, 0, 0, 0, 0, 86, 201, 26, 32, 0, 0, 0, 0, 88] +++ #[2, 237, 48, 0, 0, 0, 0, 88, 168, 252, 32, 0, 0, 0, 0, 89, 226, 207, 48, 0, 0, 0, 0, 90, 136, 222, 32, 0, 0, 0, 0, 91, 222, 96, 176, 0, 0] +++ #[0, 0, 92, 104, 192, 32, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2] +++ #[1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1] +++ #[2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 1, 2, 255, 255, 212, 76, 0, 0, 255, 255, 227, 224, 1, 4, 255, 255] +++ #[213, 208, 0, 8, 76, 77, 84, 0, 45, 48, 50, 0, 45, 48, 51, 0, 10, 60, 45, 48, 51, 62, 51, 10] def code := Std.Time.TimeZone.TZif.parse.run file |>.toOption |>.get! From 7c60f5075762e7694dee15454b637b6d744524f0 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 12 Nov 2024 19:45:20 -0300 Subject: [PATCH 113/118] fix: a bunch of small fixes and a huge test --- src/Std/Time/Date/PlainDate.lean | 2 +- src/Std/Time/Date/Unit/Day.lean | 19 +- src/Std/Time/Date/Unit/Month.lean | 11 +- src/Std/Time/Date/Unit/Week.lean | 12 +- src/Std/Time/Date/Unit/Weekday.lean | 6 +- src/Std/Time/Date/Unit/Year.lean | 6 + src/Std/Time/DateTime/PlainDateTime.lean | 16 +- src/Std/Time/DateTime/Timestamp.lean | 18 +- src/Std/Time/Format.lean | 16 +- src/Std/Time/Format/Basic.lean | 6 +- src/Std/Time/Internal/Bounded.lean | 2 +- src/Std/Time/Internal/UnitVal.lean | 15 +- src/Std/Time/Time/Unit/Hour.lean | 11 +- src/Std/Time/Time/Unit/Millisecond.lean | 4 +- src/Std/Time/Time/Unit/Minute.lean | 4 +- src/Std/Time/Time/Unit/Nanosecond.lean | 6 +- src/Std/Time/Time/Unit/Second.lean | 4 +- src/Std/Time/Zoned/DateTime.lean | 13 + src/Std/Time/Zoned/Offset.lean | 15 +- src/Std/Time/Zoned/ZoneRules.lean | 23 + src/Std/Time/Zoned/ZonedDateTime.lean | 22 + tests/lean/run/timeAPI.lean | 654 +++++++++++++++++++++++ tests/lean/run/timeIO.lean | 18 +- 23 files changed, 844 insertions(+), 59 deletions(-) create mode 100644 tests/lean/run/timeAPI.lean diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean index ef48a544ceac..a85f3d0c119a 100644 --- a/src/Std/Time/Date/PlainDate.lean +++ b/src/Std/Time/Date/PlainDate.lean @@ -289,7 +289,7 @@ def weekday (date : PlainDate) : Weekday := /-- Determines the week of the month for the given `PlainDate`. The week of the month is calculated based -on the day of the month and the weekday. Each week starts on Sunday because the entire library is +on the day of the month and the weekday. Each week starts on Monday because the entire library is based on the Gregorian Calendar. -/ def alignedWeekOfMonth (date : PlainDate) : Week.Ordinal.OfMonth := diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index e3d9c3fa786a..ea82499c6945 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -47,12 +47,25 @@ instance {x y : Offset} : Decidable (x < y) := namespace Ordinal +/-- +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 1 ≤ data ∧ data ≤ 31) : Ordinal := + Bounded.LE.mk data h + /-- `OfYear` represents the day ordinal within a year, which can be bounded between 1 and 365 or 366, depending on whether it's a leap year. -/ def OfYear (leap : Bool) := Bounded.LE 1 (.ofNat (if leap then 366 else 365)) +instance : Repr (OfYear leap) where + reprPrec r p := reprPrec r.val p + +instance : ToString (OfYear leap) where + toString r := toString r.val + namespace OfYear /-- @@ -104,7 +117,7 @@ namespace OfYear Converts an `OfYear` ordinal to a `Offset`. -/ def toOffset (ofYear : OfYear leap) : Offset := - UnitVal.mk ofYear.val + UnitVal.ofInt ofYear.val end OfYear end Ordinal @@ -123,14 +136,14 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Day.Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Creates an `Offset` from an integer. -/ @[inline] def ofInt (data : Int) : Day.Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Convert `Day.Offset` into `Nanosecond.Offset`. diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index b607ccee96e2..96601b9f10f7 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -147,6 +147,13 @@ Converts a `Ordinal` into a `Offset`. def toOffset (month : Ordinal) : Offset := month.val +/-- +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 1 ≤ data ∧ data ≤ 12) : Ordinal := + Bounded.LE.mk data h + /-- Creates an `Ordinal` from a `Nat`, ensuring the value is within bounds. -/ @@ -227,7 +234,7 @@ def days (leap : Bool) (month : Ordinal) : Day.Ordinal := if leap then 29 else 28 else let ⟨months, p⟩ := monthSizesNonLeap - let index : Fin 12 := (month.sub 1).toFin (by decide) (by decide) + let index : Fin 12 := (month.sub 1).toFin (by decide) months.get (index.cast (by rw [p])) theorem days_gt_27 (leap : Bool) (i : Month.Ordinal) : days leap i > 27 := by @@ -245,7 +252,7 @@ Returns the number of days until the `month`. -/ def cumulativeDays (leap : Bool) (month : Ordinal) : Day.Offset := by let ⟨months, p⟩ := cumulativeSizes - let index : Fin 12 := (month.sub 1).toFin (by decide) (by decide) + let index : Fin 12 := (month.sub 1).toFin (by decide) rw [← p] at index let res := months.get index exact res + (if leap ∧ month.val > 2 then 1 else 0) diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 1f6311feb97f..58c236de2c80 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -43,11 +43,19 @@ instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ namespace Ordinal +/-- +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 1 ≤ data ∧ data ≤ 53) : Ordinal := + Bounded.LE.mk data h + /-- `OfMonth` represents the number of weeks within a month. It ensures that the week is within the correct bounds—either 1 to 6, representing the possible weeks in a month. -/ def OfMonth := Bounded.LE 1 6 + deriving Repr /-- Creates an `Ordinal` from a natural number, ensuring the value is within bounds. @@ -78,14 +86,14 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Week.Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Creates an `Offset` from an integer. -/ @[inline] def ofInt (data : Int) : Week.Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Convert `Week.Offset` into `Millisecond.Offset`. diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean index 2df96a8d5001..14c90848a869 100644 --- a/src/Std/Time/Date/Unit/Weekday.lean +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -18,9 +18,6 @@ set_option linter.all true Defines the enumeration for days of the week. Each variant corresponds to a day of the week. -/ inductive Weekday - /-- Sunday. -/ - | sunday - /-- Monday. -/ | monday @@ -38,6 +35,9 @@ inductive Weekday /-- Saturday. -/ | saturday + + /-- Sunday. -/ + | sunday deriving Repr, Inhabited, BEq namespace Weekday diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index b512e339ad58..1b7d38cececa 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -26,6 +26,12 @@ inductive Era /-- The Common Era (CE), represents dates from year 0 onwards. -/ | ce + deriving Repr, Inhabited + +instance : ToString Era where + toString + | .bce => "BCE" + | .ce => "CE" /-- `Offset` represents a year offset, defined as an `Int`. diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 83715f8dad53..1ff146abe3e3 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -126,6 +126,20 @@ def ofTimestampAssumingUTC (stamp : Timestamp) : PlainDateTime := Id.run do time := PlainTime.ofHourMinuteSecondsNano (leap := false) (hour.expandTop (by decide)) minute second nano } +/-- +Converts a `PlainDateTime` to the number of days since the UNIX epoch. +-/ +@[inline] +def toDaysSinceUNIXEpoch (pdt : PlainDateTime) : Day.Offset := + pdt.date.toDaysSinceUNIXEpoch + +/-- +Converts a `PlainDateTime` to the number of days since the UNIX epoch. +-/ +@[inline] +def ofDaysSinceUNIXEpoch (days : Day.Offset) (time : PlainTime) : PlainDateTime := + PlainDateTime.mk (PlainDate.ofDaysSinceUNIXEpoch days) time + /-- Sets the `PlainDateTime` to the specified `desiredWeekday`. -/ @@ -482,7 +496,7 @@ def weekOfMonth (date : PlainDateTime) : Bounded.LE 1 5 := /-- Determines the week of the month for the given `PlainDateTime`. The week of the month is calculated based -on the day of the month and the weekday. Each week starts on Sunday because the entire library is +on the day of the month and the weekday. Each week starts on Monday because the entire library is based on the Gregorian Calendar. -/ @[inline] diff --git a/src/Std/Time/DateTime/Timestamp.lean b/src/Std/Time/DateTime/Timestamp.lean index 7a32102323aa..f2c38ddea7b2 100644 --- a/src/Std/Time/DateTime/Timestamp.lean +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -75,8 +75,15 @@ def ofSecondsSinceUnixEpoch (secs : Second.Offset) : Timestamp := Creates a `Timestamp` from a `Nanosecond.Offset` since the Unix epoch. -/ @[inline] -def ofNanosecondsSinceUnixEpoch (secs : Nanosecond.Offset) : Timestamp := - ⟨Duration.ofNanoseconds secs⟩ +def ofNanosecondsSinceUnixEpoch (nanos : Nanosecond.Offset) : Timestamp := + ⟨Duration.ofNanoseconds nanos⟩ + +/-- +Creates a `Timestamp` from a `Millisecond.Offset` since the Unix epoch. +-/ +@[inline] +def ofMillisecondsSinceUnixEpoch (milli : Millisecond.Offset) : Timestamp := + ⟨Duration.ofNanoseconds milli.toNanoseconds⟩ /-- Converts a `Timestamp` to seconds as `Second.Offset`. @@ -94,6 +101,13 @@ def toNanosecondsSinceUnixEpoch (tm : Timestamp) : Nanosecond.Offset := let nanos := nanos + (.ofInt tm.val.nano.val) nanos +/-- +Converts a `Timestamp` to nanoseconds as `Millisecond.Offset`. +-/ +@[inline] +def toMillisecondsSinceUnixEpoch (tm : Timestamp) : Millisecond.Offset := + tm.toNanosecondsSinceUnixEpoch.toMilliseconds + /-- Calculates the duration from the given `Timestamp` to the current time. -/ diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index c374540e09de..72de792abc39 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -102,14 +102,14 @@ The leanDateTimeWithIdentifier format, which follows the pattern `uuuu-MM-dd'T'H for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithIdentifier : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss'['z']'") +def leanDateTimeWithIdentifier : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss'['zzzz']'") /-- The leanDateTimeWithIdentifierAndNanos format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSS'[z]'` for representing date, time, and time zone. It uses the default value that can be parsed with the notation of dates. -/ -def leanDateTimeWithIdentifierAndNanos : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSS'['z']'") +def leanDateTimeWithIdentifierAndNanos : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSS'['zzzz']'") /-- The Lean Date format, which follows the pattern `uuuu-MM-dd`. It uses the default value that can be parsed with the @@ -406,10 +406,15 @@ def fromLeanDateTimeWithIdentifierString (input : String) : Except String ZonedD <|> Formats.leanDateTimeWithIdentifierAndNanos.parse input /-- -Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notationg. +Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notation. -/ def toLeanDateTimeWithZoneString (zdt : ZonedDateTime) : String := Formats.leanDateTimeWithZone.formatBuilder zdt.year zdt.month zdt.day zdt.hour zdt.minute zdt.date.get.time.second zdt.nanosecond zdt.offset +/-- +Formats a `DateTime` value into a simple date time with timezone string that can be parsed by the date% notation with the timezone identifier. +-/ +def toLeanDateTimeWithIdentifierString (zdt : ZonedDateTime) : String := + Formats.leanDateTimeWithIdentifierAndNanos.formatBuilder zdt.year zdt.month zdt.day zdt.hour zdt.minute zdt.date.get.time.second zdt.nanosecond zdt.timezone.name /-- Parses a `String` in the `ISO8601`, `RFC822` or `RFC850` format and returns a `ZonedDateTime`. @@ -420,7 +425,7 @@ def parse (input : String) : Except String ZonedDateTime := <|> fromRFC850String input instance : ToString ZonedDateTime where - toString := toLeanDateTimeWithZoneString + toString := toLeanDateTimeWithIdentifierString instance : Repr ZonedDateTime where reprPrec data := Repr.addAppParen ("zoned(\"" ++ toLeanDateTimeWithZoneString data ++ "\")") @@ -609,4 +614,7 @@ def parse (date : String) : Except String (DateTime .GMT) := instance : Repr (DateTime tz) where reprPrec data := Repr.addAppParen (toLeanDateTimeWithZoneString data) +instance : ToString (DateTime tz) where + toString := toLeanDateTimeWithZoneString + end DateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index f6c16e349d11..be78e24f364a 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -1090,15 +1090,15 @@ private inductive Reason private def parseOffset (withMinutes : Reason) (withSeconds : Reason) (withColon : Bool) : Parser Offset := do let sign : Second.Offset ← (pchar '+' *> pure 1) <|> (pchar '-' *> pure (-1)) - let hours : Hour.Offset ← UnitVal.mk <$> parseNum 2 + let hours : Hour.Offset ← UnitVal.ofInt <$> parseNum 2 let colon := if withColon then pchar ':' else pure ':' let parseUnit {n} (reason : Reason) : Parser (Option (UnitVal n)) := match reason with - | .yes => some <$> (colon *> UnitVal.mk <$> parseNum 2) + | .yes => some <$> (colon *> UnitVal.ofInt <$> parseNum 2) | .no => pure none - | .optional => optional (colon *> UnitVal.mk <$> parseNum 2) + | .optional => optional (colon *> UnitVal.ofInt <$> parseNum 2) let minutes : Option Minute.Offset ← parseUnit withMinutes let seconds : Option Second.Offset ← parseUnit withSeconds diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean index 7b4bc514c0b2..897636c9a077 100644 --- a/src/Std/Time/Internal/Bounded.lean +++ b/src/Std/Time/Internal/Bounded.lean @@ -193,7 +193,7 @@ def toInt (n : Bounded.LE lo hi) : Int := Convert a `Bounded.LE` to a `Fin`. -/ @[inline, simp] -def toFin (n : Bounded.LE lo hi) (h₀ : 0 ≤ lo) (h₁ : lo < hi) : Fin (hi + 1).toNat := by +def toFin (n : Bounded.LE lo hi) (h₀ : 0 ≤ lo) : Fin (hi + 1).toNat := by let h := n.property.right let h₁ := Int.le_trans h₀ n.property.left refine ⟨n.val.toNat, (Int.toNat_lt h₁).mpr ?_⟩ diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 81be3e7b1e58..9d11737e1c7c 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -19,6 +19,10 @@ set_option linter.all true A structure representing a unit of a given ratio type `α`. -/ structure UnitVal (α : Rat) where + /-- + Creates a `UnitVal` from an `Int`. + -/ + ofInt :: /-- Value inside the UnitVal Value. @@ -34,13 +38,6 @@ instance { x y : UnitVal z }: Decidable (x ≤ y) := namespace UnitVal -/-- -Creates a `UnitVal` from an `Int`. --/ -@[inline] -def ofInt (value : Int) : UnitVal α := - ⟨value⟩ - /-- Creates a `UnitVal` from a `Nat`. -/ @@ -113,9 +110,9 @@ instance : LE (UnitVal α) where le x y := x.val ≤ y.val instance : LT (UnitVal α) where lt x y := x.val < y.val -instance : Add (UnitVal α) where add x y := ⟨x.val + y.val⟩ +instance : Add (UnitVal α) where add := UnitVal.add -instance : Sub (UnitVal α) where sub x y := ⟨x.val - y.val⟩ +instance : Sub (UnitVal α) where sub := UnitVal.sub instance : Mul (UnitVal α) where mul x y := ⟨x.val * y.val⟩ diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 056fd26c619b..3e9aa93b599d 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -47,6 +47,13 @@ instance : OfNat Offset n := namespace Ordinal +/-- +Creates an `Ordinal` from an integer, ensuring the value is within bounds. +-/ +@[inline] +def ofInt (data : Int) (h : 0 ≤ data ∧ data ≤ 23) : Ordinal := + Bounded.LE.mk data h + /-- Converts an `Ordinal` into a relative hour in the range of 1 to 12. -/ @@ -89,14 +96,14 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Creates an `Offset` from an integer. -/ @[inline] def ofInt (data : Int) : Offset := - UnitVal.mk data + UnitVal.ofInt data end Offset end Hour diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index b01fc897384f..0c11c3b351a0 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -50,14 +50,14 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Creates an `Offset` from an integer. -/ @[inline] def ofInt (data : Int) : Offset := - UnitVal.mk data + UnitVal.ofInt data end Offset namespace Ordinal diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index 822b14e172a9..22eacab4c2a0 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -81,14 +81,14 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Creates an `Offset` from an integer. -/ @[inline] def ofInt (data : Int) : Offset := - UnitVal.mk data + UnitVal.ofInt data end Offset end Minute diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index d009a48e15b7..ed4d541f9827 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -52,14 +52,14 @@ Creates an `Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Creates an `Offset` from an integer. -/ @[inline] def ofInt (data : Int) : Offset := - UnitVal.mk data + UnitVal.ofInt data end Offset @@ -78,7 +78,7 @@ namespace Span Creates a new `Offset` out of a `Span`. -/ def toOffset (span : Span) : Offset := - UnitVal.mk span.val + UnitVal.ofInt span.val end Span diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 1dc112a8cd41..6b6a520f40b1 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -61,14 +61,14 @@ Creates an `Second.Offset` from a natural number. -/ @[inline] def ofNat (data : Nat) : Second.Offset := - UnitVal.mk data + UnitVal.ofInt data /-- Creates an `Second.Offset` from an integer. -/ @[inline] def ofInt (data : Int) : Second.Offset := - UnitVal.mk data + UnitVal.ofInt data end Offset diff --git a/src/Std/Time/Zoned/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean index 41c646ceb743..8082a6da89cf 100644 --- a/src/Std/Time/Zoned/DateTime.lean +++ b/src/Std/Time/Zoned/DateTime.lean @@ -46,6 +46,12 @@ Creates a new `DateTime` out of a `Timestamp` that is in a `TimeZone`. def ofTimestamp (tm : Timestamp) (tz : TimeZone) : DateTime tz := DateTime.mk tm (Thunk.mk fun _ => tm.toPlainDateTimeAssumingUTC |>.addSeconds tz.toSeconds) +/-- +Converts a `DateTime` to the number of days since the UNIX epoch. +-/ +def toDaysSinceUNIXEpoch (date : DateTime tz) : Day.Offset := + date.date.get.toDaysSinceUNIXEpoch + /-- Creates a `Timestamp` out of a `DateTime`. -/ @@ -447,6 +453,13 @@ Getter for the `PlainTime` inside of a `DateTime` def time (zdt : DateTime tz) : PlainTime := zdt.date.get.time +/-- +Converts a `DateTime` to the number of days since the UNIX epoch. +-/ +@[inline] +def ofDaysSinceUNIXEpoch (days : Day.Offset) (time : PlainTime) (tz : TimeZone) : DateTime tz := + DateTime.ofPlainDateTime (PlainDateTime.ofDaysSinceUNIXEpoch days time) tz + instance : HAdd (DateTime tz) (Day.Offset) (DateTime tz) where hAdd := addDays diff --git a/src/Std/Time/Zoned/Offset.lean b/src/Std/Time/Zoned/Offset.lean index 5e954b55e62c..c6efccf550a7 100644 --- a/src/Std/Time/Zoned/Offset.lean +++ b/src/Std/Time/Zoned/Offset.lean @@ -18,6 +18,11 @@ Represents a timezone offset with an hour and second component. -/ structure Offset where + /-- + Creates an `Offset` from a given number of seconds. + -/ + ofSeconds :: + /-- The same timezone offset in seconds. -/ @@ -55,19 +60,13 @@ def zero : Offset := Creates an `Offset` from a given number of hour. -/ def ofHours (n : Hour.Offset) : Offset := - mk n.toSeconds + ofSeconds n.toSeconds /-- Creates an `Offset` from a given number of hours and minutes. -/ def ofHoursAndMinutes (n : Hour.Offset) (m : Minute.Offset) : Offset := - mk (n.toSeconds + m.toSeconds) - -/-- -Creates an `Offset` from a given number of seconds. --/ -def ofSeconds (n : Second.Offset) : Offset := - ⟨n⟩ + ofSeconds (n.toSeconds + m.toSeconds) end Offset end TimeZone diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 512c668524af..896910ffb293 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -171,6 +171,29 @@ def timezoneAt (transitions : Array Transition) (tm : Timestamp) : Except String end Transition namespace ZoneRules +/-- +Creates ZoneRules with a fixed GMT offset. +-/ +def fixedOffsetZone (second : Second.Offset) : ZoneRules := + let offset : Offset := { second } + { + transitions := #[], + initialLocalTimeType := { + gmtOffset := offset, + isDst := false, abbreviation := offset.toIsoString true, + wall := .standard, + utLocal := .ut, + identifier := offset.toIsoString true + } + } + +/-- +Creates ZoneRules with a fixed offset of UTC (GMT+0). +-/ +@[inline] +def UTC : ZoneRules := + fixedOffsetZone 0 + /-- Finds the `LocalTimeType` corresponding to a given `Timestamp` in `ZoneRules`. If the timestamp falls between two transitions, it returns the most recent transition before the timestamp. diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index 5afd25dac864..e5121c5da6fc 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -229,6 +229,15 @@ Returns the unaligned week of the month for a `ZonedDateTime` (day divided by 7, def weekOfMonth (date : ZonedDateTime) : Internal.Bounded.LE 1 5 := date.date.get.weekOfMonth +/-- +Determines the week of the month for the given `ZonedDateTime`. The week of the month is calculated based +on the day of the month and the weekday. Each week starts on Monday because the entire library is +based on the Gregorian Calendar. +-/ +@[inline] +def alignedWeekOfMonth (date : ZonedDateTime) : Week.Ordinal.OfMonth := + date.date.get.alignedWeekOfMonth + /-- Determines the quarter of the year for the given `ZonedDateTime`. -/ @@ -508,6 +517,19 @@ Checks if the `ZonedDateTime` is in a leap year. def inLeapYear (date : ZonedDateTime) : Bool := date.year.isLeap +/-- +Converts a `ZonedDateTime` to the number of days since the UNIX epoch. +-/ +def toDaysSinceUNIXEpoch (date : ZonedDateTime) : Day.Offset := + date.date.get.toDaysSinceUNIXEpoch + +/-- +Converts a `ZonedDateTime` to the number of days since the UNIX epoch. +-/ +@[inline] +def ofDaysSinceUNIXEpoch (days : Day.Offset) (time : PlainTime) (zt : TimeZone.ZoneRules) : ZonedDateTime := + ZonedDateTime.ofPlainDateTime (PlainDateTime.ofDaysSinceUNIXEpoch days time) zt + instance : HAdd ZonedDateTime Day.Offset ZonedDateTime where hAdd := addDays diff --git a/tests/lean/run/timeAPI.lean b/tests/lean/run/timeAPI.lean new file mode 100644 index 000000000000..50ceecb6eab6 --- /dev/null +++ b/tests/lean/run/timeAPI.lean @@ -0,0 +1,654 @@ +import Std.Time +import Init +open Std.Time + +/- +Unit conversion tests. +-/ + +/-- +info: --- nanosecond +nanoseconds: 1209600000000000 +milliseconds: 1209600000 +seconds: 1209600 +minutes: 20160 +hours: 336 +days: 14 +weeks: 2 +--- millisecond +nanoseconds: 1209600000000000 +seconds: 1209600 +milliseconds: 1209600000 +minutes: 20160 +hours: 336 +days: 14 +weeks: 2 +--- second +nanoseconds: 1209600000000000 +milliseconds: 1209600000 +seconds: 1209600 +minutes: 20160 +hours: 336 +days: 14 +weeks: 2 +--- minute +nanoseconds: 1209600000000000 +milliseconds: 1209600000 +seconds: 1209600 +minutes: 20160 +hours: 336 +days: 14 +weeks: 2 +--- hour +nanoseconds: 1209600000000000 +milliseconds: 1209600000 +seconds: 1209600 +minutes: 20160 +hours: 336 +days: 14 +weeks: 2 +--- day +nanoseconds: 1209600000000000 +milliseconds: 1209600000 +seconds: 1209600 +minutes: 20160 +hours: 336 +days: 14 +weeks: 2 +--- week +nanoseconds: 1209600000000000 +milliseconds: 1209600000 +seconds: 1209600 +minutes: 20160 +hours: 336 +days: 14 +weeks: 2 + +-/ +#guard_msgs in +#eval do + let nanosecond: Nanosecond.Offset := 1209600000000000 + println! s!"--- nanosecond" + println! s!"nanoseconds: {nanosecond}" + println! s!"milliseconds: {nanosecond.toMilliseconds}" + println! s!"seconds: {nanosecond.toSeconds}" + println! s!"minutes: {nanosecond.toMinutes}" + println! s!"hours: {nanosecond.toHours}" + println! s!"days: {nanosecond.toDays}" + println! s!"weeks: {nanosecond.toWeeks}" + -- Cannot do this for months or years because sizes are variable. + + let millisecond: Millisecond.Offset := 1209600000 + println! s!"--- millisecond" + println! s!"nanoseconds: {millisecond.toNanoseconds}" + println! s!"seconds: {millisecond.toSeconds}" + println! s!"milliseconds: {millisecond}" + println! s!"minutes: {millisecond.toMinutes}" + println! s!"hours: {millisecond.toHours}" + println! s!"days: {millisecond.toDays}" + println! s!"weeks: {millisecond.toWeeks}" + -- Cannot do this for months or years because sizes are variable. + + let second : Second.Offset := 1209600 + println! s!"--- second" + println! s!"nanoseconds: {second.toNanoseconds}" + println! s!"milliseconds: {second.toMilliseconds}" + println! s!"seconds: {second}" + println! s!"minutes: {second.toMinutes}" + println! s!"hours: {second.toHours}" + println! s!"days: {second.toDays}" + println! s!"weeks: {second.toWeeks}" + -- Cannot do this for months or years because sizes are variable. + + let minute: Minute.Offset := 20160 + println! s!"--- minute" + println! s!"nanoseconds: {minute.toNanoseconds}" + println! s!"milliseconds: {minute.toMilliseconds}" + println! s!"seconds: {minute.toSeconds}" + println! s!"minutes: {minute}" + println! s!"hours: {minute.toHours}" + println! s!"days: {minute.toDays}" + println! s!"weeks: {minute.toWeeks}" + -- Cannot do this for months or years because sizes are variable. + + let hour: Hour.Offset := 336 + println! s!"--- hour" + println! s!"nanoseconds: {hour.toNanoseconds}" + println! s!"milliseconds: {hour.toMilliseconds}" + println! s!"seconds: {hour.toSeconds}" + println! s!"minutes: {hour.toMinutes}" + println! s!"hours: {hour}" + println! s!"days: {hour.toDays}" + println! s!"weeks: {hour.toWeeks}" + -- Cannot do this for months or years because sizes are variable. + + let day: Day.Offset := 14 + println! s!"--- day" + println! s!"nanoseconds: {day.toNanoseconds}" + println! s!"milliseconds: {day.toMilliseconds}" + println! s!"seconds: {day.toSeconds}" + println! s!"minutes: {day.toMinutes}" + println! s!"hours: {day.toHours}" + println! s!"days: {day}" + println! s!"weeks: {day.toWeeks}" + -- Cannot do this for months or years because sizes are variable. + + let week: Week.Offset := 2 + println! s!"--- week" + println! s!"nanoseconds: {week.toNanoseconds}" + println! s!"milliseconds: {week.toMilliseconds}" + println! s!"seconds: {week.toSeconds}" + println! s!"minutes: {week.toMinutes}" + println! s!"hours: {week.toHours}" + println! s!"days: {week.toDays}" + println! s!"weeks: {week}" + -- Cannot do this for months or years because sizes are variable. + +theorem nano_nano_nano : (1 : Nanosecond.Offset) + (1 : Nanosecond.Offset) = (2 : Nanosecond.Offset) := rfl +theorem nano_milli_nano : (1 : Nanosecond.Offset) + (1 : Millisecond.Offset) = (1000001 : Nanosecond.Offset) := rfl +theorem nano_sec_nano : (1 : Nanosecond.Offset) + (1 : Second.Offset) = (1000000001 : Nanosecond.Offset) := rfl +theorem nano_min_nano : (1 : Nanosecond.Offset) + (1 : Minute.Offset) = (60000000001 : Nanosecond.Offset) := rfl +theorem nano_hour_nano : (1 : Nanosecond.Offset) + (1 : Hour.Offset) = (3600000000001 : Nanosecond.Offset) := rfl +theorem nano_day_nano : (1 : Nanosecond.Offset) + (1 : Day.Offset) = (86400000000001 : Nanosecond.Offset) := rfl +theorem nano_week_nano : (1 : Nanosecond.Offset) + (1 : Week.Offset) = (604800000000001 : Nanosecond.Offset) := rfl + +theorem milli_nano_nano : (1 : Millisecond.Offset) + (1 : Nanosecond.Offset) = (1000001 : Nanosecond.Offset) := rfl +theorem milli_milli_milli : (1 : Millisecond.Offset) + (1 : Millisecond.Offset) = (2 : Millisecond.Offset) := rfl +theorem milli_sec_milli : (1 : Millisecond.Offset) + (1 : Second.Offset) = (1001 : Millisecond.Offset) := rfl +theorem milli_min_milli : (1 : Millisecond.Offset) + (1 : Minute.Offset) = (60001 : Millisecond.Offset) := rfl +theorem milli_hour_milli : (1 : Millisecond.Offset) + (1 : Hour.Offset) = (3600001 : Millisecond.Offset) := rfl +theorem milli_day_milli : (1 : Millisecond.Offset) + (1 : Day.Offset) = (86400001 : Millisecond.Offset) := rfl +theorem milli_week_milli : (1 : Millisecond.Offset) + (1 : Week.Offset) = (604800001 : Millisecond.Offset) := rfl + +theorem sec_nano_nano : (1 : Second.Offset) + (1 : Nanosecond.Offset) = (1000000001 : Nanosecond.Offset) := rfl +theorem sec_milli_milli : (1 : Second.Offset) + (1 : Millisecond.Offset) = (1001 : Millisecond.Offset) := rfl +theorem sec_sec_sec : (1 : Second.Offset) + (1 : Second.Offset) = (2 : Second.Offset) := rfl +theorem sec_min_sec : (1 : Second.Offset) + (1 : Minute.Offset) = (61 : Second.Offset) := rfl +theorem sec_hour_sec : (1 : Second.Offset) + (1 : Hour.Offset) = (3601 : Second.Offset) := rfl +theorem sec_day_sec : (1 : Second.Offset) + (1 : Day.Offset) = (86401 : Second.Offset) := rfl +theorem sec_week_sec : (1 : Second.Offset) + (1 : Week.Offset) = (604801 : Second.Offset) := rfl + +theorem min_nano_nano : (1 : Minute.Offset) + (1 : Nanosecond.Offset) = (60000000001 : Nanosecond.Offset) := rfl +theorem min_milli_milli : (1 : Minute.Offset) + (1 : Millisecond.Offset) = (60001 : Millisecond.Offset) := rfl +theorem min_sec_sec : (1 : Minute.Offset) + (1 : Second.Offset) = (61 : Second.Offset) := rfl +theorem min_min_min : (1 : Minute.Offset) + (1 : Minute.Offset) = (2 : Minute.Offset) := rfl +theorem min_hour_min : (1 : Minute.Offset) + (1 : Hour.Offset) = (61 : Minute.Offset) := rfl +theorem min_day_min : (1 : Minute.Offset) + (1 : Day.Offset) = (1441 : Minute.Offset) := rfl +theorem min_week_min : (1 : Minute.Offset) + (1 : Week.Offset) = (10081 : Minute.Offset) := rfl + +theorem hour_nano_nano : (1 : Hour.Offset) + (1 : Nanosecond.Offset) = (3600000000001 : Nanosecond.Offset) := rfl +theorem hour_milli_milli : (1 : Hour.Offset) + (1 : Millisecond.Offset) = (3600001 : Millisecond.Offset) := rfl +theorem hour_sec_sec : (1 : Hour.Offset) + (1 : Second.Offset) = (3601 : Second.Offset) := rfl +theorem hour_min_min : (1 : Hour.Offset) + (1 : Minute.Offset) = (61 : Minute.Offset) := rfl +theorem hour_hour_hour : (1 : Hour.Offset) + (1 : Hour.Offset) = (2 : Hour.Offset) := rfl +theorem hour_day_hour : (1 : Hour.Offset) + (1 : Day.Offset) = (25 : Hour.Offset) := rfl +theorem hour_week_hour : (1 : Hour.Offset) + (1 : Week.Offset) = (169 : Hour.Offset) := rfl + +theorem day_nano_nano : (1 : Day.Offset) + (1 : Nanosecond.Offset) = (86400000000001 : Nanosecond.Offset) := rfl +theorem day_milli_milli : (1 : Day.Offset) + (1 : Millisecond.Offset) = (86400001 : Millisecond.Offset) := rfl +theorem day_sec_sec : (1 : Day.Offset) + (1 : Second.Offset) = (86401 : Second.Offset) := rfl +theorem day_min_min : (1 : Day.Offset) + (1 : Minute.Offset) = (1441 : Minute.Offset) := rfl +theorem day_hour_hour : (1 : Day.Offset) + (1 : Hour.Offset) = (25 : Hour.Offset) := rfl +theorem day_day_day : (1 : Day.Offset) + (1 : Day.Offset) = (2 : Day.Offset) := rfl +theorem day_week_day : (1 : Day.Offset) + (1 : Week.Offset) = (8 : Day.Offset) := rfl + +theorem week_nano_nano : (1 : Week.Offset) + (1 : Nanosecond.Offset) = (604800000000001 : Nanosecond.Offset) := rfl +theorem week_milli_milli : (1 : Week.Offset) + (1 : Millisecond.Offset) = (604800001 : Millisecond.Offset) := rfl +theorem week_sec_sec : (1 : Week.Offset) + (1 : Second.Offset) = (604801 : Second.Offset) := rfl +theorem week_min_min : (1 : Week.Offset) + (1 : Minute.Offset) = (10081 : Minute.Offset) := rfl +theorem week_hour_hour : (1 : Week.Offset) + (1 : Hour.Offset) = (169 : Hour.Offset) := rfl +theorem week_day_day : (1 : Week.Offset) + (1 : Day.Offset) = (8 : Day.Offset) := rfl +theorem week_week_week : (1 : Week.Offset) + (1 : Week.Offset) = (2 : Week.Offset) := rfl + +theorem nano_nano_nano_sub : (1 : Nanosecond.Offset) - (1 : Nanosecond.Offset) = (0 : Nanosecond.Offset) := rfl +theorem nano_milli_nano_sub : (1 : Nanosecond.Offset) - (1 : Millisecond.Offset) = (-999999 : Nanosecond.Offset) := rfl +theorem nano_sec_nano_sub : (1 : Nanosecond.Offset) - (1 : Second.Offset) = (-999999999 : Nanosecond.Offset) := rfl +theorem nano_min_nano_sub : (1 : Nanosecond.Offset) - (1 : Minute.Offset) = (-59999999999 : Nanosecond.Offset) := rfl +theorem nano_hour_nano_sub : (1 : Nanosecond.Offset) - (1 : Hour.Offset) = (-3599999999999 : Nanosecond.Offset) := rfl +theorem nano_day_nano_sub : (1 : Nanosecond.Offset) - (1 : Day.Offset) = (-86399999999999 : Nanosecond.Offset) := rfl +theorem nano_week_nano_sub : (1 : Nanosecond.Offset) - (1 : Week.Offset) = (-604799999999999 : Nanosecond.Offset) := rfl + +theorem milli_nano_nano_sub : (1 : Millisecond.Offset) - (1 : Nanosecond.Offset) = (999999 : Nanosecond.Offset) := rfl +theorem milli_milli_milli_sub : (1 : Millisecond.Offset) - (1 : Millisecond.Offset) = (0 : Millisecond.Offset) := rfl +theorem milli_sec_milli_sub : (1 : Millisecond.Offset) - (1 : Second.Offset) = (-999 : Millisecond.Offset) := rfl +theorem milli_min_milli_sub : (1 : Millisecond.Offset) - (1 : Minute.Offset) = (-59999 : Millisecond.Offset) := rfl +theorem milli_hour_milli_sub : (1 : Millisecond.Offset) - (1 : Hour.Offset) = (-3599999 : Millisecond.Offset) := rfl +theorem milli_day_milli_sub : (1 : Millisecond.Offset) - (1 : Day.Offset) = (-86399999 : Millisecond.Offset) := rfl +theorem milli_week_milli_sub : (1 : Millisecond.Offset) - (1 : Week.Offset) = (-604799999 : Millisecond.Offset) := rfl + +theorem sec_nano_nano_sub : (1 : Second.Offset) - (1 : Nanosecond.Offset) = (999999999 : Nanosecond.Offset) := rfl +theorem sec_milli_milli_sub : (1 : Second.Offset) - (1 : Millisecond.Offset) = (999 : Millisecond.Offset) := rfl +theorem sec_sec_sec_sub : (1 : Second.Offset) - (1 : Second.Offset) = (0 : Second.Offset) := rfl +theorem sec_min_sec_sub : (1 : Second.Offset) - (1 : Minute.Offset) = (-59 : Second.Offset) := rfl +theorem sec_hour_sec_sub : (1 : Second.Offset) - (1 : Hour.Offset) = (-3599 : Second.Offset) := rfl +theorem sec_day_sec_sub : (1 : Second.Offset) - (1 : Day.Offset) = (-86399 : Second.Offset) := rfl +theorem sec_week_sec_sub : (1 : Second.Offset) - (1 : Week.Offset) = (-604799 : Second.Offset) := rfl + +theorem min_nano_nano_sub : (1 : Minute.Offset) - (1 : Nanosecond.Offset) = (59999999999 : Nanosecond.Offset) := rfl +theorem min_milli_milli_sub : (1 : Minute.Offset) - (1 : Millisecond.Offset) = (59999 : Millisecond.Offset) := rfl +theorem min_sec_sec_sub : (1 : Minute.Offset) - (1 : Second.Offset) = (59 : Second.Offset) := rfl +theorem min_min_min_sub : (1 : Minute.Offset) - (1 : Minute.Offset) = (0 : Minute.Offset) := rfl +theorem min_hour_min_sub : (1 : Minute.Offset) - (1 : Hour.Offset) = (-59 : Minute.Offset) := rfl +theorem min_day_min_sub : (1 : Minute.Offset) - (1 : Day.Offset) = (-1439 : Minute.Offset) := rfl +theorem min_week_min_sub : (1 : Minute.Offset) - (1 : Week.Offset) = (-10079 : Minute.Offset) := rfl + +theorem hour_nano_nano_sub : (1 : Hour.Offset) - (1 : Nanosecond.Offset) = (3599999999999 : Nanosecond.Offset) := rfl +theorem hour_milli_milli_sub : (1 : Hour.Offset) - (1 : Millisecond.Offset) = (3599999 : Millisecond.Offset) := rfl +theorem hour_sec_sec_sub : (1 : Hour.Offset) - (1 : Second.Offset) = (3599 : Second.Offset) := rfl +theorem hour_min_min_sub : (1 : Hour.Offset) - (1 : Minute.Offset) = (59 : Minute.Offset) := rfl +theorem hour_hour_hour_sub : (1 : Hour.Offset) - (1 : Hour.Offset) = (0 : Hour.Offset) := rfl +theorem hour_day_hour_sub : (1 : Hour.Offset) - (1 : Day.Offset) = (-23 : Hour.Offset) := rfl +theorem hour_week_hour_sub : (1 : Hour.Offset) - (1 : Week.Offset) = (-167 : Hour.Offset) := rfl + +theorem day_nano_nano_sub : (1 : Day.Offset) - (1 : Nanosecond.Offset) = (86399999999999 : Nanosecond.Offset) := rfl +theorem day_milli_milli_sub : (1 : Day.Offset) - (1 : Millisecond.Offset) = (86399999 : Millisecond.Offset) := rfl +theorem day_sec_sec_sub : (1 : Day.Offset) - (1 : Second.Offset) = (86399 : Second.Offset) := rfl +theorem day_min_min_sub : (1 : Day.Offset) - (1 : Minute.Offset) = (1439 : Minute.Offset) := rfl +theorem day_hour_hour_sub : (1 : Day.Offset) - (1 : Hour.Offset) = (23 : Hour.Offset) := rfl +theorem day_day_day_sub : (1 : Day.Offset) - (1 : Day.Offset) = (0 : Day.Offset) := rfl +theorem day_week_day_sub : (1 : Day.Offset) - (1 : Week.Offset) = (-6 : Day.Offset) := rfl + +theorem week_nano_nano_sub : (1 : Week.Offset) - (1 : Nanosecond.Offset) = (604799999999999 : Nanosecond.Offset) := rfl +theorem week_milli_milli_sub : (1 : Week.Offset) - (1 : Millisecond.Offset) = (604799999 : Millisecond.Offset) := rfl +theorem week_sec_sec_sub : (1 : Week.Offset) - (1 : Second.Offset) = (604799 : Second.Offset) := rfl +theorem week_min_min_sub : (1 : Week.Offset) - (1 : Minute.Offset) = (10079 : Minute.Offset) := rfl +theorem week_hour_hour_sub : (1 : Week.Offset) - (1 : Hour.Offset) = (167 : Hour.Offset) := rfl +theorem week_day_day_sub : (1 : Week.Offset) - (1 : Day.Offset) = (6 : Day.Offset) := rfl +theorem week_week_week_sub : (1 : Week.Offset) - (1 : Week.Offset) = (0 : Week.Offset) := rfl + +/- +Of and To basic units +-/ + +example : (1 : Nanosecond.Offset).toInt = (1 : Int) := rfl +example : (1 : Millisecond.Offset).toInt = (1 : Int) := rfl +example : (1 : Second.Offset).toInt = (1 : Int) := rfl +example : (1 : Minute.Offset).toInt = (1 : Int) := rfl +example : (1 : Hour.Offset).toInt = (1 : Int) := rfl +example : (1 : Day.Offset).toInt = (1 : Int) := rfl +example : (1 : Week.Offset).toInt = (1 : Int) := rfl + +example : Nanosecond.Offset.ofInt 1 = (1 : Nanosecond.Offset) := rfl +example : Millisecond.Offset.ofInt 1 = (1 : Millisecond.Offset) := rfl +example : Second.Offset.ofInt 1 = (1 : Second.Offset) := rfl +example : Minute.Offset.ofInt 1 = (1 : Minute.Offset) := rfl +example : Hour.Offset.ofInt 1 = (1 : Hour.Offset) := rfl +example : Day.Offset.ofInt 1 = (1 : Day.Offset) := rfl +example : Week.Offset.ofInt 1 = (1 : Week.Offset) := rfl + +example : Nanosecond.Offset.ofNat 1 = (1 : Nanosecond.Offset) := rfl +example : Millisecond.Offset.ofNat 1 = (1 : Millisecond.Offset) := rfl +example : Second.Offset.ofNat 1 = (1 : Second.Offset) := rfl +example : Minute.Offset.ofNat 1 = (1 : Minute.Offset) := rfl +example : Hour.Offset.ofNat 1 = (1 : Hour.Offset) := rfl +example : Day.Offset.ofNat 1 = (1 : Day.Offset) := rfl +example : Week.Offset.ofNat 1 = (1 : Week.Offset) := rfl + +example := Nanosecond.Ordinal.ofInt 1 (by decide) +example := Millisecond.Ordinal.ofInt 1 (by decide) +example := Second.Ordinal.ofInt (leap := false) 59 (by decide) +example := Second.Ordinal.ofInt (leap := true) 60 (by decide) +example := Minute.Ordinal.ofInt 1 (by decide) +example := Hour.Ordinal.ofInt 1 (by decide) +example := Day.Ordinal.ofInt 1 (by decide) +example := Week.Ordinal.ofInt 1 (by decide) + +example := Nanosecond.Ordinal.ofFin 1 +example := Millisecond.Ordinal.ofFin 1 +example := Second.Ordinal.ofFin (leap := false) (Fin.ofNat 1) +example := Second.Ordinal.ofFin (leap := true) (Fin.ofNat 1) +example := Minute.Ordinal.ofFin 1 +example := Hour.Ordinal.ofFin 1 +example := Day.Ordinal.ofFin 1 +example := Week.Ordinal.ofFin 1 + +example := Nanosecond.Ordinal.ofNat 1 +example := Millisecond.Ordinal.ofNat 1 +example := Second.Ordinal.ofNat (leap := false) 1 +example := Second.Ordinal.ofNat (leap := true) 1 +example := Minute.Ordinal.ofNat 1 +example := Hour.Ordinal.ofNat 1 +example := Day.Ordinal.ofNat 1 +example := Week.Ordinal.ofNat 1 + +example := Nanosecond.Ordinal.toOffset 1 +example := Millisecond.Ordinal.toOffset 1 +example := Second.Ordinal.toOffset (leap := false) 1 +example := Second.Ordinal.toOffset (leap := true) 1 +example := Minute.Ordinal.toOffset 1 +example := Hour.Ordinal.toOffset 1 +example := Day.Ordinal.toOffset 1 +example := Week.Ordinal.toOffset 1 + +example : (1 : Nanosecond.Ordinal).toInt = (1 : Int) := rfl +example : (1 : Millisecond.Ordinal).toInt = (1 : Int) := rfl +example : (1 : Second.Ordinal false).toInt = (1 : Int) := rfl +example : (1 : Second.Ordinal true).toInt = (1 : Int) := rfl +example : (1 : Minute.Ordinal).toInt = (1 : Int) := rfl +example : (1 : Hour.Ordinal).toInt = (1 : Int) := rfl +example : (1 : Day.Ordinal).toInt = (1 : Int) := rfl +example : (1 : Week.Ordinal).toInt = (1 : Int) := rfl + +example : ((1 : Nanosecond.Ordinal).toFin (by decide) |>.val) = 1 := rfl +example : ((1 : Millisecond.Ordinal).toFin (by decide) |>.val) = 1 := rfl +example : ((1 : Second.Ordinal false).toFin (by decide) |>.val) = 1 := rfl +example : ((1 : Second.Ordinal true).toFin (by decide) |>.val) = 1 := rfl +example : ((1 : Minute.Ordinal).toFin (by decide) |>.val) = 1 := rfl +example : ((1 : Hour.Ordinal).toFin (by decide) |>.val) = 1 := rfl +example : ((1 : Day.Ordinal).toFin (by decide) |>.val) = 1 := rfl +example : ((1 : Week.Ordinal).toFin (by decide) |>.val) = 1 := rfl + +example : (1 : Nanosecond.Ordinal).toNat = 1 := rfl +example : (1 : Millisecond.Ordinal).toNat = 1 := rfl +example : (1 : Second.Ordinal false).toNat = 1 := rfl +example : (1 : Second.Ordinal true).toNat = 1 := rfl +example : (1 : Minute.Ordinal).toNat = 1 := rfl +example : (1 : Hour.Ordinal).toNat = 1 := rfl +example : (1 : Day.Ordinal).toNat = 1 := rfl +example : (1 : Week.Ordinal).toNat = 1 := rfl + +/-- +info: 9 +2023-06-10 +2023-06-16 +2023-07-09 +2023-07-09 +2024-06-09 +2024-06-09 +2023-06-08 +2023-06-02 +2023-05-09 +2023-05-09 +2022-06-09 +2022-06-09 +2023-06-01 +2023-06-01 +2023-06-11 +2023-01-09 +2023-01-09 +0001-06-09 +0001-06-09 +23 2023 2023 2023 23 2023 2023 23 J +false +160 +CE +2 +2 +Std.Time.Weekday.friday +23 +2 +06-09-2023 +2023-06-09 +2023-06-09 +19517 +1686268800000 +1970-01-02 + +-/ +#guard_msgs in +#eval do + -- :> + let date := date("2023-06-09") + + println! repr date.day + println! date.addDays 1 + println! date.addWeeks 1 + println! date.addMonthsClip 1 + println! date.addMonthsRollOver 1 + println! date.addYearsClip 1 + println! date.addYearsRollOver 1 + + println! date.subDays 1 + println! date.subWeeks 1 + println! date.subMonthsClip 1 + println! date.subMonthsRollOver 1 + println! date.subYearsClip 1 + println! date.subYearsRollOver 1 + + println! date.withDaysClip 1 + println! date.withDaysRollOver 1 + println! date.withWeekday .sunday + println! date.withMonthClip 1 + println! date.withMonthRollOver 1 + println! date.withYearClip 1 + println! date.withYearRollOver 1 + + println! date.format "yy yyyy yyy yyy yy uuuu uuu uu MMMMM" + println! date.inLeapYear + println! date.dayOfYear + println! date.era + println! repr date.quarter + println! repr date.alignedWeekOfMonth + println! repr date.weekday + println! repr date.weekOfYear + println! repr date.weekOfMonth + + println! date.toAmericanDateString + println! date.toLeanDateString + println! date.toSQLDateString + + println! date.toDaysSinceUNIXEpoch + println! date.toTimestampAssumingUTC + println! PlainDate.ofDaysSinceUNIXEpoch 1 + +/-- +info: 1997-03-19T02:03:04.000000000 +1997-03-25T02:03:04.000000000 +1997-04-18T02:03:04.000000000 +1997-04-18T02:03:04.000000000 +1998-03-18T02:03:04.000000000 +1998-03-18T02:03:04.000000000 +1997-03-17T02:03:04.000000000 +1997-03-11T02:03:04.000000000 +1997-02-18T02:03:04.000000000 +1997-02-18T02:03:04.000000000 +1996-03-18T02:03:04.000000000 +1996-03-18T02:03:04.000000000 +1997-03-01T02:03:04.000000000 +1997-03-01T02:03:04.000000000 +1997-03-23T02:03:04.000000000 +1997-01-18T02:03:04.000000000 +1997-01-18T02:03:04.000000000 +0001-03-18T02:03:04.000000000 +0001-03-18T02:03:04.000000000 +97 1997 1997 1997 97 1997 1997 97 M +false +77 +CE +1 +4 +Std.Time.Weekday.tuesday +12 +3 +9938 +858650584000 +1970-01-02T00:00:00.000000000 + +-/ +#guard_msgs in +#eval do + let plaindatetime := datetime("1997-03-18T02:03:04") + + println! plaindatetime.addDays 1 + println! plaindatetime.addWeeks 1 + println! plaindatetime.addMonthsClip 1 + println! plaindatetime.addMonthsRollOver 1 + println! plaindatetime.addYearsClip 1 + println! plaindatetime.addYearsRollOver 1 + + println! plaindatetime.subDays 1 + println! plaindatetime.subWeeks 1 + println! plaindatetime.subMonthsClip 1 + println! plaindatetime.subMonthsRollOver 1 + println! plaindatetime.subYearsClip 1 + println! plaindatetime.subYearsRollOver 1 + + println! plaindatetime.withDaysClip 1 + println! plaindatetime.withDaysRollOver 1 + println! plaindatetime.withWeekday .sunday + println! plaindatetime.withMonthClip 1 + println! plaindatetime.withMonthRollOver 1 + println! plaindatetime.withYearClip 1 + println! plaindatetime.withYearRollOver 1 + + println! plaindatetime.format "yy yyyy yyy yyy yy uuuu uuu uu MMMMM" + println! plaindatetime.inLeapYear + println! plaindatetime.dayOfYear + println! plaindatetime.era + println! repr plaindatetime.quarter + println! repr plaindatetime.alignedWeekOfMonth + println! repr plaindatetime.weekday + println! repr plaindatetime.weekOfYear + println! repr plaindatetime.weekOfMonth + + println! plaindatetime.toDaysSinceUNIXEpoch + println! plaindatetime.toTimestampAssumingUTC + println! PlainDateTime.ofDaysSinceUNIXEpoch 1 PlainTime.midnight + +/-- +info: 2024-09-13T02:01:02.000000000-03:00 +2024-09-19T02:01:02.000000000-03:00 +2024-10-12T02:01:02.000000000-03:00 +2024-10-12T02:01:02.000000000-03:00 +2025-09-12T02:01:02.000000000-03:00 +2025-09-12T02:01:02.000000000-03:00 +2024-09-11T02:01:02.000000000-03:00 +2024-09-05T02:01:02.000000000-03:00 +2024-08-12T02:01:02.000000000-03:00 +2024-08-12T02:01:02.000000000-03:00 +2023-09-12T02:01:02.000000000-03:00 +2023-09-12T02:01:02.000000000-03:00 +2024-09-01T02:01:02.000000000-03:00 +2024-09-01T02:01:02.000000000-03:00 +2024-09-15T02:01:02.000000000-03:00 +2024-01-12T02:01:02.000000000-03:00 +2024-01-12T02:01:02.000000000-03:00 +0001-09-12T02:01:02.000000000-03:00 +0001-09-12T02:01:02.000000000-03:00 +24 2024 2024 2024 24 2024 2024 24 S +true +256 +CE +3 +3 +Std.Time.Weekday.thursday +37 +2 +19978 +1726117262000 +1970-01-02T00:00:00.000000000Z + +-/ +#guard_msgs in +#eval do + let zoned := DateTime.ofPlainDateTimeAssumingUTC datetime("2024-09-12T05:01:02") timezone("America/Sao_Paulo -03:00") + + println! zoned.addDays 1 + println! zoned.addWeeks 1 + println! zoned.addMonthsClip 1 + println! zoned.addMonthsRollOver 1 + println! zoned.addYearsClip 1 + println! zoned.addYearsRollOver 1 + + println! zoned.subDays 1 + println! zoned.subWeeks 1 + println! zoned.subMonthsClip 1 + println! zoned.subMonthsRollOver 1 + println! zoned.subYearsClip 1 + println! zoned.subYearsRollOver 1 + + println! zoned.withDaysClip 1 + println! zoned.withDaysRollOver 1 + println! zoned.withWeekday .sunday + println! zoned.withMonthClip 1 + println! zoned.withMonthRollOver 1 + println! zoned.withYearClip 1 + println! zoned.withYearRollOver 1 + + println! zoned.format "yy yyyy yyy yyy yy uuuu uuu uu MMMMM" + println! zoned.inLeapYear + println! zoned.dayOfYear + println! zoned.era + println! repr zoned.quarter + println! repr zoned.alignedWeekOfMonth + println! repr zoned.weekday + println! repr zoned.weekOfYear + println! repr zoned.weekOfMonth + + println! zoned.toDaysSinceUNIXEpoch + println! zoned.toTimestamp + println! DateTime.ofDaysSinceUNIXEpoch 1 PlainTime.midnight .UTC + +/-- +info: 1997-03-19T02:03:04.000000000[America/Sao_Paulo] +1997-03-25T02:03:04.000000000[America/Sao_Paulo] +1997-04-18T02:03:04.000000000[America/Sao_Paulo] +1997-04-18T02:03:04.000000000[America/Sao_Paulo] +1998-03-18T02:03:04.000000000[America/Sao_Paulo] +1998-03-18T02:03:04.000000000[America/Sao_Paulo] +1997-03-17T02:03:04.000000000[America/Sao_Paulo] +1997-03-11T02:03:04.000000000[America/Sao_Paulo] +1997-02-18T02:03:04.000000000[America/Sao_Paulo] +1997-02-18T02:03:04.000000000[America/Sao_Paulo] +1996-03-18T02:03:04.000000000[America/Sao_Paulo] +1996-03-18T02:03:04.000000000[America/Sao_Paulo] +1997-03-01T02:03:04.000000000[America/Sao_Paulo] +1997-03-01T02:03:04.000000000[America/Sao_Paulo] +1997-03-23T02:03:04.000000000[America/Sao_Paulo] +1997-01-18T02:03:04.000000000[America/Sao_Paulo] +1997-01-18T02:03:04.000000000[America/Sao_Paulo] +0001-03-17T02:03:04.000000000[America/Sao_Paulo] +0001-03-17T02:03:04.000000000[America/Sao_Paulo] +97 1997 1997 1997 97 1997 1997 97 M +false +77 +CE +1 +4 +Std.Time.Weekday.tuesday +12 +3 +9938 +858661384000 +1970-01-02T00:00:00.000000000[+00:00] + +-/ +#guard_msgs in +#eval do + let zoned ← zoned("1997-03-18T02:03:04[America/Sao_Paulo]") + + println! zoned.addDays 1 + println! zoned.addWeeks 1 + println! zoned.addMonthsClip 1 + println! zoned.addMonthsRollOver 1 + println! zoned.addYearsClip 1 + println! zoned.addYearsRollOver 1 + + println! zoned.subDays 1 + println! zoned.subWeeks 1 + println! zoned.subMonthsClip 1 + println! zoned.subMonthsRollOver 1 + println! zoned.subYearsClip 1 + println! zoned.subYearsRollOver 1 + + println! zoned.withDaysClip 1 + println! zoned.withDaysRollOver 1 + println! zoned.withWeekday .sunday + println! zoned.withMonthClip 1 + println! zoned.withMonthRollOver 1 + println! zoned.withYearClip 1 + println! zoned.withYearRollOver 1 + + println! zoned.format "yy yyyy yyy yyy yy uuuu uuu uu MMMMM" + println! zoned.inLeapYear + println! zoned.dayOfYear + println! zoned.era + println! repr zoned.quarter + println! repr zoned.alignedWeekOfMonth + println! repr zoned.weekday + println! repr zoned.weekOfYear + println! repr zoned.weekOfMonth + + println! zoned.toDaysSinceUNIXEpoch + println! zoned.toTimestamp + println! ZonedDateTime.ofDaysSinceUNIXEpoch 1 PlainTime.midnight .UTC diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index 3d9a805ae8e7..fc157e2a7045 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -40,9 +40,9 @@ info: 2013-10-19T23:59:59.000000000-03:00 #guard_msgs in #eval do let zr ← Database.defaultGetZoneRules "America/Sao_Paulo" - println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-19T23:59:59") zr}" - println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-20T00:00:00") zr}" - println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-20T00:00:01") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-19T23:59:59") zr |>.toLeanDateTimeWithZoneString}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-20T00:00:00") zr |>.toLeanDateTimeWithZoneString}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2013-10-20T00:00:01") zr |>.toLeanDateTimeWithZoneString}" /- Java: @@ -59,9 +59,9 @@ info: 2019-11-03T01:59:59.000000000-05:00 #guard_msgs in #eval do let zr ← Database.defaultGetZoneRules "America/Chicago" - println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T01:59:59") zr}" - println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T02:00:00") zr}" - println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T02:59:59") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T01:59:59") zr |>.toLeanDateTimeWithZoneString}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T02:00:00") zr |>.toLeanDateTimeWithZoneString}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2019-11-03T02:59:59") zr |>.toLeanDateTimeWithZoneString}" /- Java: @@ -78,6 +78,6 @@ info: 2003-10-26T01:59:59.000000000-05:00 #guard_msgs in #eval do let zr ← Database.defaultGetZoneRules "America/Monterrey" - println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T01:59:59") zr}" - println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T02:00:00") zr}" - println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T02:59:59") zr}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T01:59:59") zr |>.toLeanDateTimeWithZoneString}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T02:00:00") zr |>.toLeanDateTimeWithZoneString}" + println! "{ZonedDateTime.ofPlainDateTime datetime("2003-10-26T02:59:59") zr |>.toLeanDateTimeWithZoneString}" From 50a17560d80bc12a0afff1bb698694cb1ceaa01e Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 12 Nov 2024 20:20:51 -0300 Subject: [PATCH 114/118] feat: more tests --- src/Std/Time/Zoned/TimeZone.lean | 2 +- src/Std/Time/Zoned/ZoneRules.lean | 8 ++--- src/Std/Time/Zoned/ZonedDateTime.lean | 3 +- tests/lean/run/timeAPI.lean | 52 ++++++++++++++++++++++++++- tests/lean/run/timeIO.lean | 2 +- 5 files changed, 59 insertions(+), 8 deletions(-) diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean index c4b63456d038..1dd0d6c3d887 100644 --- a/src/Std/Time/Zoned/TimeZone.lean +++ b/src/Std/Time/Zoned/TimeZone.lean @@ -44,7 +44,7 @@ namespace TimeZone A zeroed `Timezone` representing UTC (no offset). -/ def UTC : TimeZone := - TimeZone.mk (Offset.zero) "Coordinated Universal Time" "UTC" false + TimeZone.mk (Offset.zero) "UTC" "UTC" false /-- A zeroed `Timezone` representing GMT (no offset). diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean index 896910ffb293..8bfdcb48bc27 100644 --- a/src/Std/Time/Zoned/ZoneRules.lean +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -174,16 +174,16 @@ namespace ZoneRules /-- Creates ZoneRules with a fixed GMT offset. -/ -def fixedOffsetZone (second : Second.Offset) : ZoneRules := +def fixedOffsetZone (second : Second.Offset) (identifier : Option String := none) (abbreviation : Option String := none) : ZoneRules := let offset : Offset := { second } { transitions := #[], initialLocalTimeType := { gmtOffset := offset, - isDst := false, abbreviation := offset.toIsoString true, + isDst := false, abbreviation := abbreviation.getD (offset.toIsoString true), wall := .standard, utLocal := .ut, - identifier := offset.toIsoString true + identifier := identifier.getD (offset.toIsoString true) } } @@ -192,7 +192,7 @@ Creates ZoneRules with a fixed offset of UTC (GMT+0). -/ @[inline] def UTC : ZoneRules := - fixedOffsetZone 0 + fixedOffsetZone 0 "UTC" "UTC" /-- Finds the `LocalTimeType` corresponding to a given `Timestamp` in `ZoneRules`. diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean index e5121c5da6fc..bd6686b75df2 100644 --- a/src/Std/Time/Zoned/ZonedDateTime.lean +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -112,7 +112,8 @@ def convertZoneRules (date : ZonedDateTime) (tz₁ : TimeZone.ZoneRules) : Zoned ofTimestamp date.toTimestamp tz₁ /-- -Creates a new `ZonedDateTime` out of a `PlainDateTime` +Creates a new `ZonedDateTime` out of a `PlainDateTime`. It assumes that the `PlainDateTime` is relative +to UTC. -/ @[inline] def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone.ZoneRules) : ZonedDateTime := diff --git a/tests/lean/run/timeAPI.lean b/tests/lean/run/timeAPI.lean index 50ceecb6eab6..b0c7dcf548f9 100644 --- a/tests/lean/run/timeAPI.lean +++ b/tests/lean/run/timeAPI.lean @@ -587,6 +587,7 @@ info: 1997-03-19T02:03:04.000000000[America/Sao_Paulo] 1998-03-18T02:03:04.000000000[America/Sao_Paulo] 1998-03-18T02:03:04.000000000[America/Sao_Paulo] 1997-03-17T02:03:04.000000000[America/Sao_Paulo] +1997-03-17T02:03:04.000000000[America/Sao_Paulo] 1997-03-11T02:03:04.000000000[America/Sao_Paulo] 1997-02-18T02:03:04.000000000[America/Sao_Paulo] 1997-02-18T02:03:04.000000000[America/Sao_Paulo] @@ -610,7 +611,6 @@ Std.Time.Weekday.tuesday 3 9938 858661384000 -1970-01-02T00:00:00.000000000[+00:00] -/ #guard_msgs in @@ -624,6 +624,7 @@ Std.Time.Weekday.tuesday println! zoned.addYearsClip 1 println! zoned.addYearsRollOver 1 + println! zoned.subDays 1 println! zoned.subDays 1 println! zoned.subWeeks 1 println! zoned.subMonthsClip 1 @@ -651,4 +652,53 @@ Std.Time.Weekday.tuesday println! zoned.toDaysSinceUNIXEpoch println! zoned.toTimestamp + +/-- +info: 2023-06-09T00:00:00.000000000 +0001-01-01T12:32:43.000000000 +2033-11-18T12:32:43.000000000 +-/ +#guard_msgs in +#eval do + println! PlainDateTime.ofPlainDate date("2023-06-09") + println! PlainDateTime.ofPlainTime time("12:32:43") + println! PlainDateTime.ofDaysSinceUNIXEpoch 23332 time("12:32:43") + +/-- +info: 1970-01-02T00:00:00.000000000Z +1997-03-18T00:00:00.000000000Z +1997-03-18T00:01:02.000000000Z +1997-03-18T00:01:02.000000000Z +2024-02-16T22:07:14.000000000Z + +-/ +#guard_msgs in +#eval do + println! DateTime.ofDaysSinceUNIXEpoch 1 PlainTime.midnight .UTC + println! DateTime.ofPlainDate date("1997-03-18") .UTC + println! DateTime.ofPlainDateTime datetime("1997-03-18T00:01:02") .UTC + println! DateTime.ofPlainDateTimeAssumingUTC datetime("1997-03-18T00:01:02") .UTC + println! DateTime.ofTimestamp 1708121234 .UTC + +/-- +info: 1970-01-02T00:00:00.000000000[UTC] +1997-03-18T00:00:00.000000000[UTC] +1997-03-18T00:00:00.000000000[UTC] +1997-03-18T00:01:02.000000000[UTC] +1997-03-18T00:01:02.000000000[UTC] +1997-03-18T00:01:02.000000000[UTC] +1997-03-18T00:01:02.000000000[UTC] +2024-02-16T22:07:14.000000000[UTC] +2024-02-16T22:07:14.000000000[UTC] +-/ +#guard_msgs in +#eval do println! ZonedDateTime.ofDaysSinceUNIXEpoch 1 PlainTime.midnight .UTC + println! ZonedDateTime.ofPlainDate date("1997-03-18") .UTC + println! ZonedDateTime.ofPlainDateWithZone date("1997-03-18") .UTC + println! ZonedDateTime.ofPlainDateTime datetime("1997-03-18T00:01:02") .UTC + println! ZonedDateTime.ofPlainDateTimeAssumingUTC datetime("1997-03-18T00:01:02") .UTC + println! ZonedDateTime.ofPlainDateTimeWithZone datetime("1997-03-18T00:01:02") .UTC + println! ← ZonedDateTime.of datetime("1997-03-18T00:01:02") "UTC" + println! ZonedDateTime.ofTimestamp 1708121234 .UTC + println! ZonedDateTime.ofTimestampWithZone 1708121234 .UTC diff --git a/tests/lean/run/timeIO.lean b/tests/lean/run/timeIO.lean index fc157e2a7045..4bb46e57265d 100644 --- a/tests/lean/run/timeIO.lean +++ b/tests/lean/run/timeIO.lean @@ -16,7 +16,7 @@ info: { second := 0 } -/ #guard_msgs in #eval do - let res ← Database.defaultGetZoneRules "Etc/UTC" + let res ← Database.defaultGetZoneRules "UTC" println! repr res.initialLocalTimeType.gmtOffset /- From 97e312d6abfb458c07e29c279734dc36633cbae4 Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Tue, 12 Nov 2024 20:32:52 -0300 Subject: [PATCH 115/118] fix: nix ci --- tests/lean/run/timeAPI.lean | 210 +++++++++++++++++++++++++++++++++++- 1 file changed, 207 insertions(+), 3 deletions(-) diff --git a/tests/lean/run/timeAPI.lean b/tests/lean/run/timeAPI.lean index b0c7dcf548f9..881a38bad547 100644 --- a/tests/lean/run/timeAPI.lean +++ b/tests/lean/run/timeAPI.lean @@ -2,6 +2,212 @@ import Std.Time import Init open Std.Time +def sao_paulo : TimeZone.ZoneRules := + { + initialLocalTimeType := + { + gmtOffset := { second := -11188 }, + isDst := false, + abbreviation := "LMT", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + }, + transitions := #[ + { + time := 782276400, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 793159200, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 813726000, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 824004000, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 844570800, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 856058400, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 876106800, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 888717600, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 908074800, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 919562400, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 938919600, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 951616800, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 970974000, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 982461600, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 1003028400, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 1013911200, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 1036292400, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 1045360800, + localTimeType := { + gmtOffset := { second := -10800 }, + isDst := false, + abbreviation := "-03", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + }, + { time := 1066532400, + localTimeType := { + gmtOffset := { second := -7200 }, + isDst := true, + abbreviation := "-02", + wall := Std.Time.TimeZone.StdWall.standard, + utLocal := Std.Time.TimeZone.UTLocal.ut, + identifier := "America/Sao_Paulo" + } + } + ] + } + /- Unit conversion tests. -/ @@ -615,7 +821,7 @@ Std.Time.Weekday.tuesday -/ #guard_msgs in #eval do - let zoned ← zoned("1997-03-18T02:03:04[America/Sao_Paulo]") + let zoned := zoned("1997-03-18T02:03:04", sao_paulo) println! zoned.addDays 1 println! zoned.addWeeks 1 @@ -687,7 +893,6 @@ info: 1970-01-02T00:00:00.000000000[UTC] 1997-03-18T00:01:02.000000000[UTC] 1997-03-18T00:01:02.000000000[UTC] 1997-03-18T00:01:02.000000000[UTC] -1997-03-18T00:01:02.000000000[UTC] 2024-02-16T22:07:14.000000000[UTC] 2024-02-16T22:07:14.000000000[UTC] -/ @@ -699,6 +904,5 @@ info: 1970-01-02T00:00:00.000000000[UTC] println! ZonedDateTime.ofPlainDateTime datetime("1997-03-18T00:01:02") .UTC println! ZonedDateTime.ofPlainDateTimeAssumingUTC datetime("1997-03-18T00:01:02") .UTC println! ZonedDateTime.ofPlainDateTimeWithZone datetime("1997-03-18T00:01:02") .UTC - println! ← ZonedDateTime.of datetime("1997-03-18T00:01:02") "UTC" println! ZonedDateTime.ofTimestamp 1708121234 .UTC println! ZonedDateTime.ofTimestampWithZone 1708121234 .UTC From 2a30e1e4383a31c3f7d494264d9ec9400e086afd Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 13 Nov 2024 08:20:23 -0300 Subject: [PATCH 116/118] fix: remove mul/div instances for unitval and fix 2038 bug --- src/Std/Time/Date/Unit/Day.lean | 2 +- src/Std/Time/Date/Unit/Month.lean | 3 ++- src/Std/Time/Date/Unit/Week.lean | 2 +- src/Std/Time/Date/Unit/Year.lean | 2 +- src/Std/Time/DateTime/PlainDateTime.lean | 8 ++++---- src/Std/Time/Format/Basic.lean | 4 ++-- src/Std/Time/Internal/UnitVal.lean | 4 ---- src/Std/Time/Time/Unit/Hour.lean | 2 +- src/Std/Time/Time/Unit/Millisecond.lean | 2 +- src/Std/Time/Time/Unit/Minute.lean | 2 +- src/Std/Time/Time/Unit/Nanosecond.lean | 2 +- src/Std/Time/Time/Unit/Second.lean | 2 +- src/runtime/io.cpp | 4 ++-- 13 files changed, 18 insertions(+), 21 deletions(-) diff --git a/src/Std/Time/Date/Unit/Day.lean b/src/Std/Time/Date/Unit/Day.lean index ea82499c6945..9020ea068b47 100644 --- a/src/Std/Time/Date/Unit/Day.lean +++ b/src/Std/Time/Date/Unit/Day.lean @@ -35,7 +35,7 @@ instance : Inhabited Ordinal where default := 1 (the number of seconds in a day). -/ def Offset : Type := UnitVal 86400 - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean index 96601b9f10f7..c54e836e46d9 100644 --- a/src/Std/Time/Date/Unit/Month.lean +++ b/src/Std/Time/Date/Unit/Month.lean @@ -183,7 +183,8 @@ Transforms `Month.Ordinal` into `Second.Offset`. -/ def toSeconds (leap : Bool) (month : Ordinal) : Second.Offset := let daysAcc := #[0, 31, 59, 90, 120, 151, 181, 212, 243, 273, 304, 334] - let time := daysAcc[month.toNat]! * 86400 + let days : Day.Offset := daysAcc[month.toNat]! + let time := days.toSeconds if leap && month.toNat ≥ 2 then time + 86400 else time diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean index 58c236de2c80..84fe108728a2 100644 --- a/src/Std/Time/Date/Unit/Week.lean +++ b/src/Std/Time/Date/Unit/Week.lean @@ -37,7 +37,7 @@ instance : Inhabited Ordinal where `Offset` represents an offset in weeks. -/ def Offset : Type := UnitVal (86400 * 7) - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Date/Unit/Year.lean b/src/Std/Time/Date/Unit/Year.lean index 1b7d38cececa..94ac4d67ea49 100644 --- a/src/Std/Time/Date/Unit/Year.lean +++ b/src/Std/Time/Date/Unit/Year.lean @@ -37,7 +37,7 @@ instance : ToString Era where `Offset` represents a year offset, defined as an `Int`. -/ def Offset : Type := Int - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString instance {x y : Offset} : Decidable (x ≤ y) := let x : Int := x diff --git a/src/Std/Time/DateTime/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean index 1ff146abe3e3..ab9ce145fdfd 100644 --- a/src/Std/Time/DateTime/PlainDateTime.lean +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -70,11 +70,11 @@ def ofTimestampAssumingUTC (stamp : Timestamp) : PlainDateTime := Id.run do let h := rem.truncateBottom (Int.not_le.mp h) (h, rawDays) - let mut quadracentennialCycles := days / daysPer400Y; - let mut remDays := days.val % daysPer400Y.val; + let mut quadracentennialCycles := days.val / daysPer400Y; + let mut remDays := days.val % daysPer400Y; if remDays < 0 then - remDays := remDays + daysPer400Y.val + remDays := remDays + daysPer400Y quadracentennialCycles := quadracentennialCycles - 1 let mut centenialCycles := remDays / daysPer100Y; @@ -96,7 +96,7 @@ def ofTimestampAssumingUTC (stamp : Timestamp) : PlainDateTime := Id.run do remDays := remDays - remYears * 365 - let mut year := 2000 + remYears + 4 * quadrennialCycles + 100 * centenialCycles + 400 * quadracentennialCycles.val + let mut year := 2000 + remYears + 4 * quadrennialCycles + 100 * centenialCycles + 400 * quadracentennialCycles let months := [31, 30, 31, 30, 31, 31, 30, 31, 30, 31, 31, 29]; let mut mon : Fin 13 := 0; diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index be78e24f364a..391d694f0308 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -1089,7 +1089,7 @@ private inductive Reason | optional private def parseOffset (withMinutes : Reason) (withSeconds : Reason) (withColon : Bool) : Parser Offset := do - let sign : Second.Offset ← (pchar '+' *> pure 1) <|> (pchar '-' *> pure (-1)) + let sign ← (pchar '+' *> pure 1) <|> (pchar '-' *> pure (-1)) let hours : Hour.Offset ← UnitVal.ofInt <$> parseNum 2 let colon := if withColon then pchar ':' else pure ':' @@ -1105,7 +1105,7 @@ private def parseOffset (withMinutes : Reason) (withSeconds : Reason) (withColon let hours := hours.toSeconds + (minutes.getD 0).toSeconds + (seconds.getD 0) - return Offset.ofSeconds (sign * hours) + return Offset.ofSeconds ⟨hours.val * sign⟩ private def parseWith : (mod : Modifier) → Parser (TypeFormat mod) | .G format => diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean index 9d11737e1c7c..cc97c0b6481c 100644 --- a/src/Std/Time/Internal/UnitVal.lean +++ b/src/Std/Time/Internal/UnitVal.lean @@ -114,10 +114,6 @@ instance : Add (UnitVal α) where add := UnitVal.add instance : Sub (UnitVal α) where sub := UnitVal.sub -instance : Mul (UnitVal α) where mul x y := ⟨x.val * y.val⟩ - -instance : Div (UnitVal α) where div x y := ⟨x.val / y.val⟩ - instance : Neg (UnitVal α) where neg x := ⟨-x.val⟩ instance : ToString (UnitVal n) where toString n := toString n.val diff --git a/src/Std/Time/Time/Unit/Hour.lean b/src/Std/Time/Time/Unit/Hour.lean index 3e9aa93b599d..7eceb2c7648b 100644 --- a/src/Std/Time/Time/Unit/Hour.lean +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -40,7 +40,7 @@ instance {x y : Ordinal} : Decidable (x < y) := or differences in hours. -/ def Offset : Type := UnitVal 3600 - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, ToString instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Time/Unit/Millisecond.lean b/src/Std/Time/Time/Unit/Millisecond.lean index 0c11c3b351a0..754ddf22c477 100644 --- a/src/Std/Time/Time/Unit/Millisecond.lean +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -38,7 +38,7 @@ instance {x y : Ordinal} : Decidable (x < y) := `Offset` represents a duration offset in milliseconds. -/ def Offset : Type := UnitVal (1 / 1000) - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/Std/Time/Time/Unit/Minute.lean b/src/Std/Time/Time/Unit/Minute.lean index 22eacab4c2a0..0accf7f45dab 100644 --- a/src/Std/Time/Time/Unit/Minute.lean +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -38,7 +38,7 @@ instance {x y : Ordinal} : Decidable (x < y) := `Offset` represents a duration offset in minutes. -/ def Offset : Type := UnitVal 60 - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, ToString instance : OfNat Offset n := ⟨UnitVal.ofInt <| Int.ofNat n⟩ diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean index ed4d541f9827..1f092cd4bd80 100644 --- a/src/Std/Time/Time/Unit/Nanosecond.lean +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -37,7 +37,7 @@ instance {x y : Ordinal} : Decidable (x < y) := `Offset` represents a time offset in nanoseconds. -/ def Offset : Type := UnitVal (1 / 1000000000) - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString instance { x y : Offset } : Decidable (x ≤ y) := inferInstanceAs (Decidable (x.val ≤ y.val)) diff --git a/src/Std/Time/Time/Unit/Second.lean b/src/Std/Time/Time/Unit/Second.lean index 6b6a520f40b1..39c409f56056 100644 --- a/src/Std/Time/Time/Unit/Second.lean +++ b/src/Std/Time/Time/Unit/Second.lean @@ -49,7 +49,7 @@ instance {x y : Ordinal l} : Decidable (x < y) := `Offset` represents an offset in seconds. It is defined as an `Int`. -/ def Offset : Type := UnitVal 1 - deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, LE, LT, ToString + deriving Repr, BEq, Inhabited, Add, Sub, Neg, LE, LT, ToString instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 1a043e3868e0..b95cc311f036 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -642,8 +642,8 @@ extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { long long nano = timestamp % 1000000000; lean_object *lean_ts = lean_alloc_ctor(0, 2, 0); - lean_ctor_set(lean_ts, 0, lean_int_to_int(static_cast(secs))); - lean_ctor_set(lean_ts, 1, lean_int_to_int(static_cast(nano))); + lean_ctor_set(lean_ts, 0, lean_int64_to_int(secs)); + lean_ctor_set(lean_ts, 1, lean_int64_to_int(nano)); return lean_io_result_mk_ok(lean_ts); } From d5093c295008bfe3a0581a00917457cbcb52e02d Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 13 Nov 2024 09:08:09 -0300 Subject: [PATCH 117/118] fix: u_failure check --- src/runtime/io.cpp | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index b95cc311f036..44f18cc58733 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -713,6 +713,10 @@ extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezo int32_t dst_name_len; u_strToUTF8(dst_name, sizeof(dst_name), &dst_name_len, tzID, tzIDLength, &status); + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to convert DST name to UTF-8"))); + } + UChar display_name[32]; int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_SHORT_DST : UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); From d4882293599662992bf3fdd81ba8fa854b69c9fc Mon Sep 17 00:00:00 2001 From: Sofia Rodrigues Date: Wed, 13 Nov 2024 14:25:28 -0300 Subject: [PATCH 118/118] fix: some formats and parser that was not parsing more than the numbers needed with padding --- src/Std/Time/Format.lean | 7 ++- src/Std/Time/Format/Basic.lean | 98 +++++++++++++++++---------------- tests/lean/run/timeFormats.lean | 16 ++++++ 3 files changed, 73 insertions(+), 48 deletions(-) diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean index 72de792abc39..6d4b31c3af37 100644 --- a/src/Std/Time/Format.lean +++ b/src/Std/Time/Format.lean @@ -50,10 +50,10 @@ representing date, time, and time zone. def dateTime24Hour : GenericFormat (.only .GMT) := datespec("uuuu-MM-dd:HH:mm:ss.SSSSSSSSS") /-- -The DateTimeWithZone format, which follows the pattern `uuuu-MM-dd:HH:mm:ss.SSSSSSSSSZZZ` +The DateTimeWithZone format, which follows the pattern `uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZ` for representing date, time, and time zone. -/ -def dateTimeWithZone : GenericFormat .any := datespec("uuuu-MM-dd:HH:mm:ss.SSSSSSSSSZZZ") +def dateTimeWithZone : GenericFormat .any := datespec("uuuu-MM-dd'T'HH:mm:ss.SSSSSSSSSZZZ") /-- The leanTime24Hour format, which follows the pattern `HH:mm:ss.SSSSSSSSS` for representing time @@ -423,6 +423,8 @@ def parse (input : String) : Except String ZonedDateTime := fromISO8601String input <|> fromRFC822String input <|> fromRFC850String input + <|> fromDateTimeWithZoneString input + <|> fromLeanDateTimeWithIdentifierString input instance : ToString ZonedDateTime where toString := toLeanDateTimeWithIdentifierString @@ -530,6 +532,7 @@ def parse (date : String) : Except String PlainDateTime := fromAscTimeString date <|> fromLongDateFormatString date <|> fromDateTimeString date + <|> fromLeanDateTimeString date instance : ToString PlainDateTime where toString := toLeanDateTimeString diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean index 391d694f0308..0ecde2581c55 100644 --- a/src/Std/Time/Format/Basic.lean +++ b/src/Std/Time/Format/Basic.lean @@ -599,7 +599,7 @@ private def leftPad (n : Nat) (a : Char) (s : String) : String := private def rightPad (n : Nat) (a : Char) (s : String) : String := s ++ "".pushn a (n - s.length) -private def truncate (size : Nat) (n : Int) (cut : Bool := false) : String := +private def pad (size : Nat) (n : Int) (cut : Bool := false) : String := let (sign, n) := if n < 0 then ("-", -n) else ("", n) let numStr := toString n @@ -787,34 +787,34 @@ private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : Strin let info := data.toInt let info := if info ≤ 0 then -info + 1 else info match format with - | .twoDigit => truncate 2 (info % 100) - | .fourDigit => truncate 4 info - | .extended n => truncate n info + | .twoDigit => pad 2 (info % 100) + | .fourDigit => pad 4 info + | .extended n => pad n info | .u format => match format with - | .twoDigit => truncate 2 (data.toInt % 100) - | .fourDigit => truncate 4 data.toInt - | .extended n => truncate n data.toInt + | .twoDigit => pad 2 (data.toInt % 100) + | .fourDigit => pad 4 data.toInt + | .extended n => pad n data.toInt | .D format => - truncate format.padding data.snd.val + pad format.padding data.snd.val | .MorL format => match format with - | .inl format => truncate format.padding data.val + | .inl format => pad format.padding data.val | .inr .short => formatMonthShort data | .inr .full => formatMonthLong data | .inr .narrow => formatMonthNarrow data | .d format => - truncate format.padding data.val + pad format.padding data.val | .Qorq format => match format with - | .inl format => truncate format.padding data.val + | .inl format => pad format.padding data.val | .inr .short => formatQuarterShort data | .inr .full => formatQuarterLong data | .inr .narrow => formatQuarterNumber data | .w format => - truncate format.padding data.val + pad format.padding data.val | .W format => - truncate format.padding data.val + pad format.padding data.val | .E format => match format with | .short => formatWeekdayShort data @@ -822,33 +822,33 @@ private def formatWith (modifier : Modifier) (data: TypeFormat modifier) : Strin | .narrow => formatWeekdayNarrow data | .eorc format => match format with - | .inl format => truncate format.padding data.toOrdinal.val + | .inl format => pad format.padding data.toOrdinal.val | .inr .short => formatWeekdayShort data | .inr .full => formatWeekdayLong data | .inr .narrow => formatWeekdayNarrow data | .F format => - truncate format.padding data.val + pad format.padding data.val | .a format => match format with | .short => formatMarkerShort data | .full => formatMarkerLong data | .narrow => formatMarkerNarrow data - | .h format => truncate format.padding (data.val % 12) - | .K format => truncate format.padding (data.val % 12) - | .k format => truncate format.padding (data.val) - | .H format => truncate format.padding (data.val) - | .m format => truncate format.padding (data.val) - | .s format => truncate format.padding (data.snd.val) + | .h format => pad format.padding (data.val % 12) + | .K format => pad format.padding (data.val % 12) + | .k format => pad format.padding (data.val) + | .H format => pad format.padding (data.val) + | .m format => pad format.padding (data.val) + | .s format => pad format.padding (data.snd.val) | .S format => match format with - | .nano => truncate 9 data.val + | .nano => pad 9 data.val | .truncated n => rightTruncate n data.val (cut := true) | .A format => - truncate format.padding data.val + pad format.padding data.val | .n format => - truncate format.padding data.val + pad format.padding data.val | .N format => - truncate format.padding data.val + pad format.padding data.val | .V => data | .z format => match format with @@ -1070,6 +1070,12 @@ private def parseSigned (parser : Parser Nat) : Parser Int := do private def parseNum (size : Nat) : Parser Nat := String.toNat! <$> exactlyChars (satisfy Char.isDigit) size +private def parseAtLeastNum (size : Nat) : Parser Nat := + String.toNat! <$> do + let start ← exactlyChars (satisfy Char.isDigit) size + let end_ ← manyChars (satisfy Char.isDigit) + pure (start ++ end_) + private def parseFractionNum (size : Nat) (pad : Nat) : Parser Nat := String.toNat! <$> rightPad pad '0' <$> exactlyChars (satisfy Char.isDigit) size @@ -1117,28 +1123,28 @@ private def parseWith : (mod : Modifier) → Parser (TypeFormat mod) match format with | .twoDigit => (2000 + ·) <$> (Int.ofNat <$> parseNum 2) | .fourDigit => Int.ofNat <$> parseNum 4 - | .extended n => Int.ofNat <$> parseNum n + | .extended n => Int.ofNat <$> parseAtLeastNum n | .u format => match format with | .twoDigit => (2000 + ·) <$> (parseSigned <| parseNum 2) - | .fourDigit => parseSigned <| parseNum 4 - | .extended n => parseSigned <| parseNum n - | .D format => Sigma.mk true <$> parseNatToBounded (parseNum format.padding) + | .fourDigit => parseSigned <| parseAtLeastNum 4 + | .extended n => parseSigned <| parseAtLeastNum n + | .D format => Sigma.mk true <$> parseNatToBounded (parseAtLeastNum format.padding) | .MorL format => match format with - | .inl format => parseNatToBounded (parseNum format.padding) + | .inl format => parseNatToBounded (parseAtLeastNum format.padding) | .inr .short => parseMonthShort | .inr .full => parseMonthLong | .inr .narrow => parseMonthNarrow - | .d format => parseNatToBounded (parseNum format.padding) + | .d format => parseNatToBounded (parseAtLeastNum format.padding) | .Qorq format => match format with - | .inl format => parseNatToBounded (parseNum format.padding) + | .inl format => parseNatToBounded (parseAtLeastNum format.padding) | .inr .short => parseQuarterShort | .inr .full => parseQuarterLong | .inr .narrow => parseQuarterNumber - | .w format => parseNatToBounded (parseNum format.padding) - | .W format => parseNatToBounded (parseNum format.padding) + | .w format => parseNatToBounded (parseAtLeastNum format.padding) + | .W format => parseNatToBounded (parseAtLeastNum format.padding) | .E format => match format with | .short => parseWeekdayShort @@ -1146,29 +1152,29 @@ private def parseWith : (mod : Modifier) → Parser (TypeFormat mod) | .narrow => parseWeekdayNarrow | .eorc format => match format with - | .inl format => Weekday.ofOrdinal <$> parseNatToBounded (parseNum format.padding) + | .inl format => Weekday.ofOrdinal <$> parseNatToBounded (parseAtLeastNum format.padding) | .inr .short => parseWeekdayShort | .inr .full => parseWeekdayLong | .inr .narrow => parseWeekdayNarrow - | .F format => parseNatToBounded (parseNum format.padding) + | .F format => parseNatToBounded (parseAtLeastNum format.padding) | .a format => match format with | .short => parseMarkerShort | .full => parseMarkerLong | .narrow => parseMarkerNarrow - | .h format => parseNatToBounded (parseNum format.padding) - | .K format => parseNatToBounded (parseNum format.padding) - | .k format => parseNatToBounded (parseNum format.padding) - | .H format => parseNatToBounded (parseNum format.padding) - | .m format => parseNatToBounded (parseNum format.padding) - | .s format => Sigma.mk true <$> (parseNatToBounded (parseNum format.padding)) + | .h format => parseNatToBounded (parseAtLeastNum format.padding) + | .K format => parseNatToBounded (parseAtLeastNum format.padding) + | .k format => parseNatToBounded (parseAtLeastNum format.padding) + | .H format => parseNatToBounded (parseAtLeastNum format.padding) + | .m format => parseNatToBounded (parseAtLeastNum format.padding) + | .s format => Sigma.mk true <$> (parseNatToBounded (parseAtLeastNum format.padding)) | .S format => match format with - | .nano => parseNatToBounded (parseNum 9) + | .nano => parseNatToBounded (parseAtLeastNum 9) | .truncated n => parseNatToBounded (parseFractionNum n 9) - | .A format => Millisecond.Offset.ofNat <$> (parseNum format.padding) - | .n format => parseNatToBounded (parseNum format.padding) - | .N format => Nanosecond.Offset.ofNat <$> (parseNum format.padding) + | .A format => Millisecond.Offset.ofNat <$> (parseAtLeastNum format.padding) + | .n format => parseNatToBounded (parseAtLeastNum format.padding) + | .N format => Nanosecond.Offset.ofNat <$> (parseAtLeastNum format.padding) | .V => parseIdentifier | .z format => match format with diff --git a/tests/lean/run/timeFormats.lean b/tests/lean/run/timeFormats.lean index ee9cfd59eed8..31be8b6faf55 100644 --- a/tests/lean/run/timeFormats.lean +++ b/tests/lean/run/timeFormats.lean @@ -794,3 +794,19 @@ info: 1 #eval let t : ZonedDateTime := .ofPlainDateTime datetime("2018-12-31T12:00:00") (TimeZone.ZoneRules.ofTimeZone TimeZone.UTC) IO.println s!"{t.format "w"}" + +/- +Truncation Test +-/ + +/-- +info: ("19343232432-01-04T01:04:03.000000000", + Except.ok (datetime("19343232432-01-04T01:04:03.000000000")), + datetime("1932-01-02T05:04:03.000000000")) +-/ +#guard_msgs in +#eval + let r := (PlainDateTime.mk (PlainDate.ofYearMonthDayClip 19343232432 1 4) (PlainTime.mk 25 64 ⟨true, 3⟩ 0)) + let s := r.toLeanDateTimeString + let r := PlainDateTime.parse s + (s, r, datetime("1932-01-02T05:04:03.000000000"))