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/Std.lean b/src/Std.lean index aeee8c6e31f5..e48e80154693 100644 --- a/src/Std.lean +++ b/src/Std.lean @@ -6,5 +6,6 @@ Authors: Sebastian Ullrich prelude import Std.Data import Std.Sat +import Std.Time import Std.Tactic import Std.Internal 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 127d9eb707d7..eddd2890a36e 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.lean b/src/Std/Time.lean new file mode 100644 index 000000000000..62eeefc49da8 --- /dev/null +++ b/src/Std/Time.lean @@ -0,0 +1,209 @@ +/- +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.Zoned +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 + +/-! +# Time + +The Lean4 API for date, time, and duration functionalities. + +# 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 an 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. + +## Date +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. + +- **`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`. +- **`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. +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`. + +## Duration +Represents spans of time and the difference between two points in time. + +- **`Duration`**: Represents the time span or difference between two `Timestamp`s values with nanosecond precision. + +# 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. + +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. +- **`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.lean b/src/Std/Time/Date.lean new file mode 100644 index 000000000000..8caabbe8e108 --- /dev/null +++ b/src/Std/Time/Date.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 Std.Time.Date.Basic +import Std.Time.Date.PlainDate diff --git a/src/Std/Time/Date/Basic.lean b/src/Std/Time/Date/Basic.lean new file mode 100644 index 000000000000..e50fbc17d028 --- /dev/null +++ b/src/Std/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 Std.Time.Date.Unit.Basic diff --git a/src/Std/Time/Date/PlainDate.lean b/src/Std/Time/Date/PlainDate.lean new file mode 100644 index 000000000000..f0b191405d4f --- /dev/null +++ b/src/Std/Time/Date/PlainDate.lean @@ -0,0 +1,304 @@ +/- +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 +import Std.Internal.Rat + +namespace Std +namespace Time +open Std.Internal +open Internal +open Lean + +set_option linter.all true + +/-- +`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 PlainDate 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 + + +instance : BEq PlainDate where + beq x y := x.day == y.day && x.month == y.month && x.year == y.year + +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 + +instance : Inhabited PlainDate where + default := mk 0 1 1 (by decide) + +/-- +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) + else none + +/-- +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⟩ + +/-- +Creates a `PlainDate` from the number of days since the UNIX epoch (January 1st, 1970). +-/ +def ofDaysSinceUNIXEpoch (day : Day.Offset) : PlainDate := + let z := day.toInt + 719468 + let era := (if z ≥ 0 then z else z - 146096).tdiv 146097 + let doe := z - era * 146097 + 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.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)) + +/-- +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 6).emod 7 (by decide) + + .ofOrdinal (d.add 1) + +/-- +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. +-/ +@[inline] +def inLeapYear (date : PlainDate) : Bool := + date.year.isLeap + +/-- +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).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).tdiv 5 + d - 1 + let doe := yoe * 365 + yoe.tdiv 4 - yoe.tdiv 100 + doy + + .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`. +-/ +@[inline] +def addDays (date : PlainDate) (days : Day.Offset) : PlainDate := + let dateDays := date.toDaysSinceUNIXEpoch + ofDaysSinceUNIXEpoch (dateDays + days) + +/-- +Subtracts a given number of days from a `PlainDate`. +-/ +@[inline] +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. +-/ +def addMonthsClip (date : PlainDate) (months : Month.Offset) : PlainDate := + 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. +-/ +@[inline] +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 := + 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. +-/ +@[inline] +def subMonthsRollOver (date : PlainDate) (months : Month.Offset) : PlainDate := + addMonthsRollOver date (-months) + +/-- +Adds `Year.Offset` to a `PlainDate`, rolling over excess days to the next month, or next year. +-/ +@[inline] +def addYearsRollOver (date : PlainDate) (years : Year.Offset) : PlainDate := + addMonthsRollOver date (years.mul 12) + +/-- +Subtracts `Year.Offset` from a `PlainDate`, rolling over excess days to the next month. +-/ +@[inline] +def subYearsRollOver (date : PlainDate) (years : Year.Offset) : PlainDate := + addMonthsRollOver date (- years.mul 12) + +/-- +Adds `Year.Offset` to a `PlainDate`, clipping the day to the last valid day of the month. +-/ +@[inline] +def addYearsClip (date : PlainDate) (years : Year.Offset) : PlainDate := + addMonthsClip date (years.mul 12) + +/-- +Subtracts `Year.Offset` from a `PlainDate`, clipping the day to the last valid day of the month. +-/ +@[inline] +def subYearsClip (date : PlainDate) (years : Year.Offset) : PlainDate := + addMonthsClip date (- years.mul 12) + +instance : HAdd PlainDate Day.Offset PlainDate where + hAdd := addDays + +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 new file mode 100644 index 000000000000..b16bb3219082 --- /dev/null +++ b/src/Std/Time/Date/Unit/Basic.lean @@ -0,0 +1,59 @@ +/- +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.Day +import Std.Time.Date.Unit.Month +import Std.Time.Date.Unit.Year +import Std.Time.Date.Unit.Weekday +import Std.Time.Date.Unit.Week + +/-! +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. +-/ + +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 new file mode 100644 index 000000000000..751ff3c05676 --- /dev/null +++ b/src/Std/Time/Date/Unit/Day.lean @@ -0,0 +1,251 @@ +/- +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 + +namespace Std +namespace Time +namespace Day +open Lean Internal + +set_option linter.all true + +/-- +`Ordinal` represents a bounded value for days, which ranges between 1 and 31. +-/ +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) + +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 + +/-- +`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 + +instance : OfNat Offset n := ⟨UnitVal.ofNat n⟩ + +instance {x y : Offset} : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +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. -/ + | 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.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) + +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` 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 := + Bounded.LE.ofFin' data (by decide) + +/-- +Converts an `Ordinal` to an `Offset`. +-/ +@[inline] +def toOffset (ordinal : Ordinal) : Offset := + UnitVal.ofInt ordinal.val + +namespace OfYear + +/-- +Converts an `OfYear` ordinal to a `Offset`. +-/ +def toOffset (ofYear : OfYear leap) : Offset := + UnitVal.mk ofYear.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 + +/-- +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 `Nanosecond.Offset`. +-/ +@[inline] +def toNanoseconds (days : Offset) : Nanosecond.Offset := + days.mul 86400000000000 + +/-- +Convert `Nanosecond.Offset` into `Day.Offset`. +-/ +@[inline] +def ofNanoseconds (ns : Nanosecond.Offset) : Offset := + ns.ediv 86400000000000 + +/-- +Convert `Day.Offset` into `Millisecond.Offset`. +-/ +@[inline] +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`. +-/ +@[inline] +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`. +-/ +@[inline] +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`. +-/ +@[inline] +def ofHours (hours : Hour.Offset) : Offset := + hours.ediv 24 + +end Offset +end Day +end Time +end Std diff --git a/src/Std/Time/Date/Unit/Month.lean b/src/Std/Time/Date/Unit/Month.lean new file mode 100644 index 000000000000..6a60a5488d65 --- /dev/null +++ b/src/Std/Time/Date/Unit/Month.lean @@ -0,0 +1,359 @@ +/- +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 + +namespace Std +namespace Time +namespace Month +open Std.Internal +open Internal + +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, LT + +instance : ToString Ordinal where + toString x := toString x.val + +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 1 (1 + (11 : Nat))) n) + +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, LT, LE, DecidableEq + +instance : OfNat Offset n := + ⟨Int.ofNat n⟩ + + +/-- +`Quarter` represents a value between 1 and 4, inclusive, corresponding to the four quarters of a year. +-/ +def Quarter := Bounded.LE 1 4 + +namespace Quarter + +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. +-/ +@[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 + +/-- +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 + +/-- +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. +-/ +@[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 := 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 +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 toMinutes (leap : Bool) (month : Ordinal) : Minute.Offset := + toSeconds leap month + |>.ediv 60 + +/-- +Transforms `Month.Ordinal` into `Hour.Offset`. +-/ +@[inline] +def toHours (leap : Bool) (month : Ordinal) : Hour.Offset := + toMinutes leap month + |>.ediv 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 a leap year. +-/ +@[inline] +private def monthSizesNonLeap : { val : Array Day.Ordinal // val.size = 12 } := + ⟨#[31, 28, 31, 30, 31, 30, 31, 31, 30, 31, 30, 31], by decide⟩ + +/-- +Returns the cumulative size in days of each month for a non-leap year. +-/ +@[inline] +private 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 index : Fin 12 := (month.sub 1).toFin (by decide) (by decide) + rw [← p] at index + exact months.get index + +theorem days_gt_27 (leap : Bool) (i : Month.Ordinal) : days leap i > 27 := by + match i with + | ⟨2, _⟩ => + simp [days] + split <;> decide + | ⟨1, _⟩ | ⟨3, _⟩ | ⟨4, _⟩ | ⟨5, _⟩ | ⟨6, _⟩ | ⟨7, _⟩ + | ⟨8, _⟩ | ⟨9, _⟩ | ⟨10, _⟩ | ⟨11, _⟩ | ⟨12, _⟩ => + simp [days, monthSizesNonLeap, Bounded.LE.sub, Bounded.LE.add, Bounded.LE.toFin] + decide + +/-- +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) + 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 ≤ (days leap month).val + +/-- +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)⟩ + +/-- +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 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.Ordinal := + let max : Day.Ordinal := month.days leap + if day.val > max.val + then max + else day + +/-- +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 + simp [Valid, clipDay] + split + exact Int.le_refl (days leap month).val + next h => exact Int.not_lt.mp h + +end Ordinal +end Month +end Time +end Std diff --git a/src/Std/Time/Date/Unit/Week.lean b/src/Std/Time/Date/Unit/Week.lean new file mode 100644 index 000000000000..84e0e1475757 --- /dev/null +++ b/src/Std/Time/Date/Unit/Week.lean @@ -0,0 +1,158 @@ +/- +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 + +namespace Std +namespace Time +namespace Week +open Std.Internal +open Internal + +set_option linter.all true + +/-- +`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 : ToString Ordinal where + toString x := toString x.val + +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 1 (1 + (52 : Nat))) n) + +instance : Inhabited Ordinal where + default := 1 + +/-- +`Offset` represents an offset in weeks. +-/ +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 + +/-- +`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. +-/ +@[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 + +/-- +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`. +-/ +@[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 Week +end Time +end Std diff --git a/src/Std/Time/Date/Unit/Weekday.lean b/src/Std/Time/Date/Unit/Weekday.lean new file mode 100644 index 000000000000..2a6437859572 --- /dev/null +++ b/src/Std/Time/Date/Unit/Weekday.lean @@ -0,0 +1,134 @@ +/- +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 + +namespace Std +namespace Time +open Std.Internal +open Internal + +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 + + /-- Tuesday. -/ + | tuesday + + /-- Wednesday. -/ + | wednesday + + /-- Thursday. -/ + | thursday + + /-- Friday. -/ + | friday + + /-- Saturday. -/ + | saturday + deriving Repr, Inhabited, BEq + +namespace Weekday + +/-- +`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 ofOrdinal : Ordinal → Weekday + | 1 => .sunday + | 2 => .monday + | 3 => .tuesday + | 4 => .wednesday + | 5 => .thursday + | 6 => .friday + | 7 => .saturday + +/-- +Converts a `Weekday` to a `Ordinal`. +-/ +def toOrdinal : Weekday → Ordinal + | .sunday => 1 + | .monday => 2 + | .tuesday => 3 + | .wednesday => 4 + | .thursday => 5 + | .friday => 6 + | .saturday => 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 + +/-- +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 + | _ => 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 + | .monday => .tuesday + | .tuesday => .wednesday + | .wednesday => .thursday + | .thursday => .friday + | .friday => .saturday + | .saturday => .sunday + | .sunday => .monday + +/-- +Check if it's a Weekend. +-/ +def weekend : Weekday → Bool + | .saturday => true + | .sunday => 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 new file mode 100644 index 000000000000..b594ba752761 --- /dev/null +++ b/src/Std/Time/Date/Unit/Year.lean @@ -0,0 +1,96 @@ +/- +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.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 + +/-- +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`. +-/ +def Offset : Type := Int + deriving Repr, BEq, Inhabited, Add, Sub, Mul, Div, Neg, ToString, LT, LE + +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`. +-/ +@[inline] +def toInt (offset : Offset) : Int := + offset + +/-- +Converts the `Year` offset to a `Month` offset. +-/ +@[inline] +def toMonths (val : Offset) : Month.Offset := + val.mul 12 + +/-- +Determines if a year is a leap year in the proleptic Gregorian calendar. +-/ +@[inline] +def isLeap (y : Offset) : Bool := + y.toInt.tmod 4 = 0 ∧ (y.toInt.tmod 100 ≠ 0 ∨ y.toInt.tmod 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 : 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 +end Std diff --git a/src/Std/Time/DateTime.lean b/src/Std/Time/DateTime.lean new file mode 100644 index 000000000000..93b5a1a4d7ab --- /dev/null +++ b/src/Std/Time/DateTime.lean @@ -0,0 +1,154 @@ +/- +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.Timestamp +import Std.Time.DateTime.PlainDateTime + +namespace Std +namespace Time + +namespace Timestamp + +set_option linter.all true + +/-- +Converts a `PlainDateTime` to a `Timestamp` +-/ +@[inline] +def ofPlainDateTimeAssumingUTC (pdt : PlainDateTime) : Timestamp := + pdt.toTimestampAssumingUTC + +/-- +Converts a `Timestamp` to a `PlainDateTime` +-/ +@[inline] +def toPlainDateTimeAssumingUTC (timestamp : Timestamp) : PlainDateTime := + PlainDateTime.ofUTCTimestamp timestamp + +/-- +Converts a `PlainDate` to a `Timestamp` +-/ +@[inline] +def ofPlainDateAssumingUTC (pd : PlainDate) : Timestamp := + let days := pd.toDaysSinceUNIXEpoch + let secs := days.toSeconds + Timestamp.ofSecondsSinceUnixEpoch secs + +/-- +Converts a `Timestamp` to a `PlainDate` +-/ +@[inline] +def toPlainDateAssumingUTC (timestamp : Timestamp) : PlainDate := + let secs := timestamp.toSecondsSinceUnixEpoch + 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 := + let nanos := timestamp.toNanosecondsSinceUnixEpoch + PlainTime.ofNanoseconds nanos + +end Timestamp +namespace PlainDate + +/-- +Converts a `PlainDate` to a `Timestamp` +-/ +@[inline] +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 + +end PlainDate +namespace PlainDateTime + +/-- +Converts a `PlainDate` to a `Timestamp` +-/ +@[inline] +def ofPlainDate (date : PlainDate) : PlainDateTime := + { date, time := PlainTime.midnight } + +/-- +Converts a `PlainDateTime` to a `PlainDate` +-/ +@[inline] +def toPlainDate (pdt : PlainDateTime) : PlainDate := + pdt.date + +/-- +Converts a `PlainTime` to a `PlainDateTime` +-/ +@[inline] +def ofPlainTime (time : PlainTime) : PlainDateTime := + { date := ⟨1, 1, 1, by decide⟩, time } + +/-- +Converts a `PlainDateTime` to a `PlainTime` +-/ +@[inline] +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 := 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 + 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/PlainDateTime.lean b/src/Std/Time/DateTime/PlainDateTime.lean new file mode 100644 index 000000000000..2a1eb3d4f019 --- /dev/null +++ b/src/Std/Time/DateTime/PlainDateTime.lean @@ -0,0 +1,469 @@ +/- +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.Internal +import Std.Time.DateTime.Timestamp + +namespace Std +namespace Time +open Internal + +set_option linter.all true + +/-- +Represents a date and time with components for Year, Month, Day, Hour, Minute, Second, and Nanosecond. +-/ +structure PlainDateTime where + + /-- + The `Date` component of a `PlainDate` + -/ + date : PlainDate + + /-- + The `Time` component of a `PlainTime` + -/ + time : PlainTime + + deriving Inhabited, BEq + +namespace PlainDateTime + +/-- +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 + Timestamp.ofNanosecondsSinceUnixEpoch (Nanosecond.Offset.ofInt nanos) + +/-- +Converts a UNIX `Timestamp` to a `PlainDateTime`. +-/ +def ofUTCTimestamp (stamp : Timestamp) : PlainDateTime := Id.run do + let leapYearEpoch := 11017 + let daysPer400Y := 365 * 400 + 97 + let daysPer100Y := 365 * 100 + 24 + let daysPer4Y := 365 * 4 + 1 + + let nanos := stamp.toNanosecondsSinceUnixEpoch + 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.val 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.val % daysPer400Y.val; + + if remDays < 0 then + remDays := remDays + daysPer400Y.val + 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.val + 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.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 := PlainDate.clip year hmon (Day.Ordinal.ofFin (Fin.succ mday)) + 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`. +-/ +@[inline] +def addDays (dt : PlainDateTime) (days : Day.Offset) : PlainDateTime := + { dt with date := dt.date.addDays days } + +/-- +Subtracts a `Day.Offset` from a `PlainDateTime`. +-/ +@[inline] +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. +-/ +def addMonthsClip (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := + { dt with date := dt.date.addMonthsClip months } + +/-- +Subtracts `Month.Offset` from a `PlainDateTime`, it clips the day to the last valid day of that month. +-/ +@[inline] +def subMonthsClip (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := + { dt with date := dt.date.subMonthsClip months } + +/-- +Adds a `Month.Offset` to a `PlainDateTime`, rolling over excess days to the following month if needed. +-/ +def addMonthsRollOver (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := + { dt with date := dt.date.addMonthsRollOver months } + +/-- +Subtracts a `Month.Offset` from a `PlainDateTime`, adjusting the day to the last valid day of the +resulting month. +-/ +@[inline] +def subMonthsRollOver (dt : PlainDateTime) (months : Month.Offset) : PlainDateTime := + { dt with date := dt.date.subMonthsRollOver months } + +/-- +Adds a `Month.Offset` to a `PlainDateTime`, rolling over excess days to the following month if needed. +-/ +@[inline] +def addYearsRollOver (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := + { dt with date := dt.date.addYearsRollOver years } + +/-- +Subtracts a `Month.Offset` from a `PlainDateTime`, rolling over excess days to the following month if +needed. +-/ +@[inline] +def addYearsClip (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := + { dt with date := dt.date.addYearsClip years } + +/-- +Subtracts a `Year.Offset` from a `PlainDateTime`, this function rolls over any excess days into the +following month. +-/ +@[inline] +def subYearsRollOver (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := + { dt with date := dt.date.subYearsRollOver years } + +/-- +Subtracts a `Year.Offset` from a `PlainDateTime`, adjusting the day to the last valid day of the +resulting month. +-/ +@[inline] +def subYearsClip (dt : PlainDateTime) (years : Year.Offset) : PlainDateTime := + { dt with date := dt.date.subYearsClip years } + + +/-- +Adds an `Hour.Offset` to a `PlainDateTime`, adjusting the date if the hour overflows. +-/ +@[inline] +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 `PlainDateTime`, adjusting the date if the hour underflows. +-/ +@[inline] +def subHours (dt : PlainDateTime) (hours : Hour.Offset) : PlainDateTime := + addHours dt (-hours) + +/-- +Adds a `Minute.Offset` to a `PlainDateTime`, adjusting the hour and date if the minutes overflow. +-/ +@[inline] +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 `PlainDateTime`, adjusting the hour and date if the minutes underflow. +-/ +@[inline] +def subMinutes (dt : PlainDateTime) (minutes : Minute.Offset) : PlainDateTime := + addMinutes dt (-minutes) + +/-- +Adds a `Second.Offset` to a `PlainDateTime`, adjusting the minute, hour, and date if the seconds overflow. +-/ +@[inline] +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 `PlainDateTime`, adjusting the minute, hour, and date if the seconds underflow. +-/ +@[inline] +def subSeconds (dt : PlainDateTime) (seconds : Second.Offset) : PlainDateTime := + addSeconds dt (-seconds) + +/-- +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 totalNanos := nano + nanos + 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 `PlainDateTime`, adjusting the seconds, minutes, hours, and date if the nanoseconds underflow. +-/ +@[inline] +def subNanoseconds (dt : PlainDateTime) (nanos : Nanosecond.Offset) : PlainDateTime := + addNanoseconds dt (-nanos) + +/-- +Getter for the `Year` inside of a `PlainDateTime`. +-/ +@[inline] +def year (dt : PlainDateTime) : Year.Offset := + dt.date.year + +/-- +Getter for the `Month` inside of a `PlainDateTime`. +-/ +@[inline] +def month (dt : PlainDateTime) : Month.Ordinal := + dt.date.month + +/-- +Getter for the `Day` inside of a `PlainDateTime`. +-/ +@[inline] +def day (dt : PlainDateTime) : Day.Ordinal := + dt.date.day + +/-- +Getter for the `Hour` inside of a `PlainDateTime`. +-/ +@[inline] +def hour (dt : PlainDateTime) : Hour.Ordinal := + dt.time.hour + +/-- +Getter for the `Minute` inside of a `PlainDateTime`. +-/ +@[inline] +def minute (dt : PlainDateTime) : Minute.Ordinal := + dt.time.minute + +/-- +Getter for the `Second` inside of a `PlainDateTime`. +-/ +@[inline] +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. +-/ +def era (date : PlainDateTime) : Year.Era := + if date.year.toInt ≥ 0 then + .ce + else + .bce + +/-- +Checks if the `PlainDateTime` is in a leap year. +-/ +@[inline] +def inLeapYear (date : PlainDateTime) : Bool := + date.year.isLeap + +instance : HAdd PlainDateTime Day.Offset PlainDateTime where + hAdd := addDays + +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 + +instance : HSub PlainDateTime Hour.Offset PlainDateTime where + hSub := subHours + +instance : HAdd PlainDateTime Minute.Offset PlainDateTime where + hAdd := addMinutes + +instance : HSub PlainDateTime Minute.Offset PlainDateTime where + hSub := subMinutes + +instance : HAdd PlainDateTime Second.Offset PlainDateTime where + hAdd := addSeconds + +instance : HSub PlainDateTime Second.Offset PlainDateTime where + hSub := subSeconds + +instance : HAdd PlainDateTime Nanosecond.Offset PlainDateTime where + hAdd := addNanoseconds + +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 new file mode 100644 index 000000000000..1cd9434a678d --- /dev/null +++ b/src/Std/Time/DateTime/Timestamp.lean @@ -0,0 +1,253 @@ +/- +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 Init.Data.Int +import Std.Time.Time +import Std.Time.Date +import Std.Time.Duration + +namespace Std +namespace Time +open Internal + +set_option linter.all true + +/-- +`Timestamp` represents a specific point in time since the UNIX Epoch. +-/ +structure Timestamp where + + /-- + Duration since the unix epoch. + -/ + 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⟩ + +instance : ToString Timestamp where + toString s := toString s.val + +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 + +/-- +Fetches the current duration from the system. +-/ +@[extern "lean_get_current_time"] +opaque now : IO Timestamp + +/-- +Converts a `Timestamp` to a `Minute.Offset` +-/ +def toMinutes (tm : Timestamp) : Minute.Offset := + tm.val.second.ediv 60 + +/-- +Converts a `Timestamp` to a `Day.Offset` +-/ +def toDays (tm : Timestamp) : Day.Offset := + tm.val.second.ediv 86400 + +/-- +Creates a new `Timestamp` out of `Second.Offset`. +-/ +def ofSecondsSinceUnixEpoch (secs : Second.Offset) : Timestamp := + ⟨Duration.ofSeconds secs⟩ + +/-- +Creates a new `Timestamp` out of `Nanosecond.Offset`. +-/ +def ofNanosecondsSinceUnixEpoch (secs : Nanosecond.Offset) : Timestamp := + ⟨Duration.ofNanoseconds secs⟩ + +/-- +Converts a `Timestamp` into its equivalent `Second.Offset`. +-/ +@[inline] +def toSecondsSinceUnixEpoch (t : Timestamp) : Second.Offset := + t.val.second + +/-- +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.val.nano.val) + nanos + +/-- +Calculates a `Timestamp` out of two `Timestamp`s. +-/ +@[inline] +def since (f : Timestamp) : IO Duration := do + let cur ← Timestamp.now + return Std.Time.Duration.sub cur.val f.val + +/-- +Creates a duration out of the `Timestamp` since the unix epoch. +-/ +@[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.val + s⟩ + +/-- +Adds a `Nanosecond.Offset` to a `Timestamp` +-/ +@[inline] +def subNanoseconds (t : Timestamp) (s : Nanosecond.Offset) : Timestamp := + ⟨t.val - s⟩ + +/-- +Adds a `Second.Offset` to a `Timestamp` +-/ +@[inline] +def addSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := + ⟨t.val + s⟩ + +/-- +Subtracts a `Second.Offset` from a `Timestamp` +-/ +@[inline] +def subSeconds (t : Timestamp) (s : Second.Offset) : Timestamp := + ⟨t.val - s⟩ + +/-- +Adds a `Minute.Offset` to a `Timestamp` +-/ +@[inline] +def addMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := + ⟨t.val + m⟩ + +/-- +Subtracts a `Minute.Offset` from a `Timestamp` +-/ +@[inline] +def subMinutes (t : Timestamp) (m : Minute.Offset) : Timestamp := + ⟨t.val - m⟩ + +/-- +Adds an `Hour.Offset` to a `Timestamp` +-/ +@[inline] +def addHours (t : Timestamp) (h : Hour.Offset) : Timestamp := + ⟨t.val + h⟩ + +/-- +Subtracts an `Hour.Offset` from a `Timestamp` +-/ +@[inline] +def subHours (t : Timestamp) (h : Hour.Offset) : Timestamp := + ⟨t.val - h⟩ + +/-- +Adds a `Day.Offset` to a `Timestamp` +-/ +@[inline] +def addDays (t : Timestamp) (d : Day.Offset) : Timestamp := + ⟨t.val + d⟩ + +/-- +Subtracts a `Day.Offset` from a `Timestamp` +-/ +@[inline] +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 Duration Timestamp where + hAdd x y := ⟨x.val + y⟩ + +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 + +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 + +instance : HAdd Timestamp Nanosecond.Offset Timestamp where + hAdd := addNanoseconds + +instance : HSub Timestamp Nanosecond.Offset Timestamp where + hSub := subNanoseconds + +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 new file mode 100644 index 000000000000..1b867bfbde91 --- /dev/null +++ b/src/Std/Time/Duration.lean @@ -0,0 +1,323 @@ +/- +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 + +namespace Std +namespace Time +open Internal + +set_option linter.all true + +/-- +Represents a point in time relative to the UNIX Epoch, with nanoseconds precision. +-/ +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, 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 ++ (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) + +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 + +/-- +Negates a `Duration`, flipping its second and nanosecond values. +-/ +@[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) + +/-- +Creates a new `Duration` out of `Second.Offset`. +-/ +@[inline] +def ofSeconds (s : Second.Offset) : Duration := by + refine ⟨s, ⟨0, by decide⟩, ?_⟩ + simp <;> exact Int.le_total s.val 0 |>.symm + +/-- +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)) + 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)))) + | 0, n, _, _ => Int.eq_iff_le_and_ge.mp (Int.zero_tmod 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 + +/-- +Converts a `Duration` from a `Nanosecond.Offset` +-/ +@[inline] +def toSeconds (duration : Duration) : Second.Offset := + duration.second + +/-- +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 + +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` +-/ +@[inline] +def toMinutes (tm : Duration) : Minute.Offset := + tm.second.ediv 60 + +/-- +Converts a `Duration` to a `Day.Offset` +-/ +@[inline] +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` +-/ +@[inline] +def addNanoseconds (t : Duration) (s : Nanosecond.Offset) : Duration := + t.add (ofNanoseconds s) + +/-- +Adds a `Nanosecond.Offset` to a `Duration` +-/ +@[inline] +def subNanoseconds (t : Duration) (s : Nanosecond.Offset) : Duration := + t.sub (ofNanoseconds s) + +/-- +Adds a `Second.Offset` to a `Duration` +-/ +@[inline] +def addSeconds (t : Duration) (s : Second.Offset) : Duration := + t.add (ofSeconds s) + +/-- +Subtracts a `Second.Offset` from a `Duration` +-/ +@[inline] +def subSeconds (t : Duration) (s : Second.Offset) : Duration := + t.sub (ofSeconds s) + +/-- +Adds a `Minute.Offset` to a `Duration` +-/ +@[inline] +def addMinutes (t : Duration) (m : Minute.Offset) : Duration := + let seconds := m.mul 60 + t.addSeconds seconds + +/-- +Subtracts a `Minute.Offset` from a `Duration` +-/ +@[inline] +def subMinutes (t : Duration) (m : Minute.Offset) : Duration := + let seconds := m.mul 60 + t.subSeconds seconds + +/-- +Adds an `Hour.Offset` to a `Duration` +-/ +@[inline] +def addHours (t : Duration) (h : Hour.Offset) : Duration := + let seconds := h.mul 3600 + t.addSeconds seconds + +/-- +Subtracts an `Hour.Offset` from a `Duration` +-/ +@[inline] +def subHours (t : Duration) (h : Hour.Offset) : Duration := + let seconds := h.mul 3600 + t.subSeconds seconds + +/-- +Adds a `Day.Offset` to a `Duration` +-/ +@[inline] +def addDays (t : Duration) (d : Day.Offset) : Duration := + let seconds := d.mul 86400 + t.addSeconds seconds + +/-- +Subtracts a `Day.Offset` from a `Duration` +-/ +@[inline] +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 + +instance : HSub Duration Hour.Offset Duration where + hSub := subHours + +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 + +instance : HSub Duration Duration Duration where + hSub := sub + +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 i d := Duration.ofNanoseconds <| Nanosecond.Offset.ofInt (d.toNanoseconds.val * i) + +instance : HMul Duration Int Duration where + 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 +end Std diff --git a/src/Std/Time/Format.lean b/src/Std/Time/Format.lean new file mode 100644 index 000000000000..ef6210d83456 --- /dev/null +++ b/src/Std/Time/Format.lean @@ -0,0 +1,562 @@ +/- +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.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: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`. +-/ +def americanDate : Format .any := datespec("MM-dd-yyyy") + +/-- +The europeanDate format, which follows the pattern `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 +in a 12-hour clock format with an upper case AM/PM marker. +-/ +def time12Hour : Format .any := datespec("hh:mm:ss aa") + +/-- +The Time24Hour format, which follows the pattern `HH:mm:ss:SSSSSSSSS` for representing time +in a 24-hour clock format. +-/ +def time24Hour : Format .any := datespec("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) := 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") + +/-- +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 := datespec("HH:mm:ss.SSSSSSSSS") + +/-- +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 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 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 leanDateTimeWithZoneNoNanos : Format .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 +notation of dates. +-/ +def leanDate : Format .any := datespec("yyyy-MM-dd") + +/-- +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") + +/-- +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") + +/-- +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") + +/-- +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") + +/-- +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") + +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 + +/-- +Formats a `PlainDate` using a specific format. +-/ +def format (date : PlainDate) (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 + | .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`. +-/ +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`. +-/ +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 PlainDate := do + Formats.sqlDate.parseBuilder (PlainDate.ofYearMonthDay) input + +/-- +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.sqlDate.formatBuilder input.year input.month input.day + +/-- +Parses a `String` in the `AmericanDate` or `SQLDate` format and returns a `PlainDate`. +-/ +def parse (input : String) : Except String PlainDate := + fromAmericanDateString input + <|> fromSQLDateString input + +instance : ToString PlainDate where + toString := toLeanDateString + +instance : Repr PlainDate where + reprPrec data := Repr.addAppParen ("date(\"" ++ toLeanDateString data ++ "\")") + +end PlainDate + +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 + match format with + | .error err => s!"error: {err}" + | .ok res => + let res := res.formatGeneric fun + | .H _ => some time.hour + | .m _ => some time.minute + | .n _ => some time.nano + | .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.sssssssss`) 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 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 (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`). +-/ +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`. +-/ +def fromTime12Hour (input : String) : Except String PlainTime := do + let builder h m s a : Option PlainTime := do + let value ← Internal.Bounded.ofInt? h.val + some <| PlainTime.ofHourMinuteSeconds (HourMarker.toAbsolute a value) m s.snd + + Formats.time12Hour.parseBuilder builder input + +/-- +Formats a `PlainTime` value into a 12-hour format string (`hh:mm:ss aa`). +-/ +def toTime12Hour (input : PlainTime) : String := + 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`. +-/ +def parse (input : String) : Except String PlainTime := + fromTime12Hour input + <|> fromTime24Hour input + +instance : ToString PlainTime where + toString := toLeanTime24Hour + +instance : Repr PlainTime where + reprPrec data := Repr.addAppParen ("time(\"" ++ toLeanTime24Hour data ++ "\")") + +end PlainTime + +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`. +-/ +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 `ZonedDateTime`. +-/ +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 `ZonedDateTime`. +-/ +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 + +/-- +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 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.leanDateTimeWithZoneNoNanos.parse input + +/-- +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 + +/-- +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 + +instance : ToString ZonedDateTime where + toString := toLeanDateTimeWithZoneString + +instance : Repr ZonedDateTime where + reprPrec data := Repr.addAppParen ("zoned(\"" ++ toLeanDateTimeWithZoneString data ++ "\")") + +end ZonedDateTime + +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 + match format with + | .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 + | .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 `PlainDateTime` object in the GMT time zone. +-/ +def fromAscTimeString (input : String) : Except String PlainDateTime := + Formats.ascTime.parse input + |>.map DateTime.toPlainDateTime + +/-- +Formats a `PlainDateTime` value into an AscTime format string. +-/ +def toAscTimeString (pdt : PlainDateTime) : String := + Formats.ascTime.format (DateTime.ofPlainDateTimeAssumingUTC pdt .UTC) + +/-- +Parses a `String` in the `LongDateFormat` and returns a `PlainDateTime` object in the GMT time zone. +-/ +def fromLongDateFormatString (input : String) : Except String PlainDateTime := + Formats.longDateFormat.parse input + |>.map DateTime.toPlainDateTime + +/-- +Formats a `PlainDateTime` value into a LongDateFormat string. +-/ +def toLongDateFormatString (pdt : PlainDateTime) : String := + Formats.longDateFormat.format (DateTime.ofPlainDateTimeAssumingUTC pdt .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 `PlainDateTime` value into a `DateTime` format string. +-/ +def toDateTimeString (pdt : PlainDateTime) : String := + 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`. +-/ +def fromLeanDateTimeString (input : String) : Except String PlainDateTime := + (Formats.leanDateTime24Hour.parse input <|> Formats.leanDateTime24HourNoNanos.parse input) + |>.map DateTime.toPlainDateTime + +/-- +Formats a `PlainDateTime` value into a `DateTime` format string. +-/ +def toLeanDateTimeString (pdt : PlainDateTime) : String := + 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`. +-/ +def parse (date : String) : Except String PlainDateTime := + fromAscTimeString date + <|> fromLongDateFormatString date + <|> fromDateTimeString date + +instance : ToString PlainDateTime where + toString := toLeanDateTimeString + +instance : Repr PlainDateTime where + reprPrec data := Repr.addAppParen ("datetime(\"" ++ toLeanDateTimeString data ++ "\")") + +end PlainDateTime + +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. +-/ +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 `DateTime` value into an ISO8601 string. +-/ +def toISO8601String (date : DateTime tz) : String := + Formats.iso8601.format date + +/-- +Formats a `DateTime` value into an RFC822 format string. +-/ +def toRFC822String (date : DateTime tz) : String := + Formats.rfc822.format date + +/-- +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 (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`. +-/ +def parse (date : String) : Except String (DateTime .GMT) := + fromAscTimeString date + <|> fromLongDateFormatString date + +instance : Repr (DateTime tz) where + reprPrec data := Repr.addAppParen (toLeanDateTimeWithZoneString data) + +end DateTime diff --git a/src/Std/Time/Format/Basic.lean b/src/Std/Time/Format/Basic.lean new file mode 100644 index 000000000000..1c41bd8a9b83 --- /dev/null +++ b/src/Std/Time/Format/Basic.lean @@ -0,0 +1,1492 @@ +/- +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.Parsec +import Std.Time.Date +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 +open Std.Internal.Parsec.String +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 + +/-- +`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 + /-- + `G`: Era (e.g., AD, Anno Domini, A). + -/ + | G (presentation : Text) + + /-- + `y`: Year of era (e.g., 2004, 04). + -/ + | y (presentation : Year) + + /-- + `D`: Day of year (e.g., 189). + -/ + | D (presentation : Number) + + /-- + `M`: Month of year as number or text (e.g., 7, 07, Jul, July, J). + -/ + | MorL (presentation : Number ⊕ Text) + + /-- + `d`: Day of month (e.g., 10). + -/ + | d (presentation : Number) + + /-- + `Q`: Quarter of year as number or text (e.g., 3, 03, Q3, 3rd quarter). + -/ + | Qorq (presentation : Number ⊕ Text) + + /-- + `w`: Week of week-based year (e.g., 27). + -/ + | w (presentation : Number) + + /-- + `W`: Week of month (e.g., 4). + -/ + | W (presentation : Number) + + /-- + `E`: Day of week as text (e.g., Tue, Tuesday, T). + -/ + | E (presentation : Text) + + /-- + `e`: Localized day of week as number or text (e.g., 2, 02, Tue, Tuesday, T). + -/ + | eorc (presentation : Number ⊕ Text) + + /-- + `F`: Aligned week of month (e.g., 3). + -/ + | F (presentation : Number) + + /-- + `a`: AM/PM of day (e.g., PM). + -/ + | 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). + -/ + | h (presentation : Number) + + /-- + `K`: Hour of AM/PM (0-11) (e.g., 0). + -/ + | K (presentation : Number) + + /-- + `k`: Clock hour of day (1-24) (e.g., 24). + -/ + | k (presentation : Number) + + /-- + `H`: Hour of day (0-23) (e.g., 0). + -/ + | H (presentation : Number) + + /-- + `m`: Minute of hour (e.g., 30). + -/ + | m (presentation : Number) + + /-- + `s`: Second of minute (e.g., 55). + -/ + | s (presentation : Number) + + /-- + `S`: Fraction of second (e.g., 978). + -/ + | S (presentation : Fraction) + + /-- + `A`: Millisecond of day (e.g., 1234). + -/ + | A (presentation : Number) + + /-- + `n`: Nanosecond of second (e.g., 987654321). + -/ + | n (presentation : Number) + + /-- + `N`: Nanosecond of day (e.g., 1234000000). + -/ + | N (presentation : Number) + + /-- + `V`: Time zone ID (e.g., America/Los_Angeles, Z, -08:30). + -/ + | V + + /-- + `z`: Time zone name (e.g., Pacific Standard Time, PST). + -/ + | z (presentation : ZoneName) + + /-- + `O`: Localized zone offset (e.g., GMT+8, GMT+08:00, UTC-08:00). + -/ + | O (presentation : OffsetO) + + /-- + `X`: Zone offset with 'Z' for zero (e.g., Z, -08, -0830, -08:30). + -/ + | X (presentation : OffsetX) + + /-- + `x`: Zone offset without 'Z' (e.g., +0000, -08, -0830, -08:30). + -/ + | x (presentation : OffsetX) + + /-- + `Z`: Zone offset with 'Z' for UTC (e.g., +0000, -0800, -08:00). + -/ + | Z (presentation : OffsetZ) + deriving Repr, Inhabited + +/-- +`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. + -/ + | string (val : String) + + /-- + A modifier that renders some data into text. + -/ + | 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 + +/-- +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 + +instance : Coe TimeZone Awareness where + coe := .only + +@[simp] +private 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 + +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 + +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 => ¬Char.isAlpha x ∧ x ≠ '\'' ∧ x ≠ '\"')) <&> .string + +private def specParser : Parser FormatString := + (Array.toList <$> many parseFormatPart) <* eof + +private def specParse (s : String) : Except String FormatString := + specParser.run s + +-- Pretty printer + +private def leftPad (n : Nat) (a : Char) (s : String) : String := + "".pushn a (n - s.length) ++ s + +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" + | .wednesday => "Wednesday" + | .thursday => "Thursday" + | .friday => "Friday" + | .saturday => "Saturday" + +private def formatWeekdayShort : Weekday → String + | .sunday => "Sun" + | .monday => "Mon" + | .tuesday => "Tue" + | .wednesday => "Wed" + | .thursday => "Thu" + | .friday => "Fri" + | .saturday => "Sat" + +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 : Month.Quarter → String + |⟨1, _⟩ => "1" + |⟨2, _⟩ => "2" + |⟨3, _⟩ => "3" + |⟨4, _⟩ => "4" + +private def formatQuarterShort : Month.Quarter → String + | ⟨1, _⟩ => "Q1" + | ⟨2, _⟩ => "Q2" + | ⟨3, _⟩ => "Q3" + | ⟨4, _⟩ => "Q4" + +private def formatQuarterLong : Month.Quarter → 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.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 + | .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.weekOfYear + | .W _ => date.weekOfMonth + | .E _ => date.weekday + | .eorc _ => date.weekday + | .F _ => date.weekOfMonth + | .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 + | .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⟩ + +/-- +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⟩ + <|> 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.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.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.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 + return if signed.isSome then -res else 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 => (TypeFormat entry) → (FormatType result xs) + | .string _ :: xs => (FormatType result xs) + | [] => result + +namespace Format + +private structure DateBuilder where + 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.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 .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 } + | .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 + 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 + + match aw with + | .only newTz => (DateTime.ofLocalDateTime · newTz) <$> datetime + | .any => (ZonedDateTime.ofLocalDateTime · tz) <$> datetime + +end DateBuilder + +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 +custom formats, such as "YYYY, MMMM, D". +-/ +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 mapper (part : FormatPart) := + match aw with + | .any => formatPartWithDate date part + | .only tz => formatPartWithDate (date.convertTimeZone tz) part + + format.string.map mapper + |> String.join + +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 => parseWithDate builder x >>= (go · xs) + | [] => + match builder.build aw with + | some res => pure res + | none => fail "could not parse the date" + go {} format + +/-- +Parser for a format with a builder. +-/ +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 ← parseWith x + go xs (func res) + | .string s :: xs => skipString s *> (go xs func) + | [] => + match func with + | some res => eof *> 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 <* eof).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 (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 (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. +-/ +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.lean b/src/Std/Time/Internal.lean new file mode 100644 index 000000000000..596a2a0489c7 --- /dev/null +++ b/src/Std/Time/Internal.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 Std.Time.Internal.Bounded +import Std.Time.Internal.UnitVal diff --git a/src/Std/Time/Internal/Bounded.lean b/src/Std/Time/Internal/Bounded.lean new file mode 100644 index 000000000000..b1e9df28aaae --- /dev/null +++ b/src/Std/Time/Internal/Bounded.lean @@ -0,0 +1,446 @@ +/- +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 + +namespace Std +namespace Time +namespace Internal + +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 } + +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) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +/-- +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 + +/-- +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⟩ + +/-- +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⟩ + +/-- +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 + +/-- +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. +-/ +@[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 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₁⟩ + else ⟨hi, And.intro h (Int.le_refl hi)⟩ + else ⟨lo, And.intro (Int.le_refl lo) h⟩ + +/-- +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 => by contradiction + +/-- +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, 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 + 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.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) + 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.tmod_lt_of_pos (m + 1) hi + exact Int.le_sub_one_of_lt h + · 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). +-/ +@[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, 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 + · 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, 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. +-/ +@[inline] +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 h + 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 : 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 h + simp at h + exact h + +/-- +Adds two `Bounded` and adjust the boundaries. +-/ +@[inline] +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 + +/-- +Adjust the bounds of a `Bounded` by subtracting a constant value to both the lower and upper bounds. +-/ +@[inline, simp] +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 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] +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.ediv num, And.intro ?_ ?_⟩ + apply Int.ediv_le_ediv + · exact h + · exact left + · apply Int.ediv_le_ediv + · 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 range of a bounded value. +-/ +@[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. +-/ +@[inline] +def expandTop (bounded : Bounded.LE lo hi) (h : hi ≤ nhi) : Bounded.LE lo nhi := + 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 := + 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. +-/ +@[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)⟩ + +/-- +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`. +-/ +@[inline] +def abs (bo : Bounded.LE (-i) i) : Bounded.LE 0 i := + 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 +end Time +end Std diff --git a/src/Std/Time/Internal/UnitVal.lean b/src/Std/Time/Internal/UnitVal.lean new file mode 100644 index 000000000000..97c9bf45a8f6 --- /dev/null +++ b/src/Std/Time/Internal/UnitVal.lean @@ -0,0 +1,125 @@ +/- +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 +import Std.Internal.Rat + +namespace Std +namespace Time +namespace Internal +open Std.Internal +open Lean + +set_option linter.all true + +/-- +A structure representing a unit of a given ratio type `α`. +-/ +structure UnitVal (α : Rat) where + + /-- + Value inside the UnitVal Value. + -/ + 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 + +/-- +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 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.tdiv 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. +-/ +@[inline] +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 + + +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..60ee6df8b127 --- /dev/null +++ b/src/Std/Time/Notation.lean @@ -0,0 +1,178 @@ +/- +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 + +namespace Std +namespace Time +open Lean Parser Command Std + +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)) + +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 + let info ← MonadRef.mkInfoFromRefPos + pure { raw := Syntax.node1 info `str (Lean.Syntax.atom info (toString n)) } + +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 + `(Std.Time.Internal.Bounded.LE.ofNatWrapping $(← syntaxInt n) (by decide)) + +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)) + +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 + `(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 + `(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 + `(Std.Time.PlainDateTime.mk $(← convertPlainDate d.date) $(← convertPlainTime d.time)) + +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))) + +syntax "zoned(" str ")" : term + +syntax "datetime(" str ")" : term + +syntax "date(" str ")" : term + +syntax "time(" str ")" : term + +syntax "offset(" str ")" : term + +syntax "timezone(" str ")" : term + +macro_rules + | `(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..72d44dca9748 --- /dev/null +++ b/src/Std/Time/Notation/Spec.lean @@ -0,0 +1,113 @@ +/- +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 + +namespace Std +namespace Time +open Lean Parser Command Std + +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.lean b/src/Std/Time/Time.lean new file mode 100644 index 000000000000..43590cf600a0 --- /dev/null +++ b/src/Std/Time/Time.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 Std.Time.Time.Basic +import Std.Time.Time.HourMarker +import Std.Time.Time.PlainTime diff --git a/src/Std/Time/Time/Basic.lean b/src/Std/Time/Time/Basic.lean new file mode 100644 index 000000000000..754d0bc6bea7 --- /dev/null +++ b/src/Std/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 Std.Time.Time.Unit.Basic diff --git a/src/Std/Time/Time/HourMarker.lean b/src/Std/Time/Time/HourMarker.lean new file mode 100644 index 000000000000..eaaa2f6577d1 --- /dev/null +++ b/src/Std/Time/Time/HourMarker.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 Std.Time.Time.Basic + +namespace Std +namespace Time +open Internal + +set_option linter.all true + +/-- +`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 + /-- Ante meridiem. -/ + | am + /-- Post meridiem. -/ + | 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 toAbsolute (marker : HourMarker) (time : Bounded.LE 1 12) : Hour.Ordinal := + match marker with + | .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 +end Std diff --git a/src/Std/Time/Time/PlainTime.lean b/src/Std/Time/Time/PlainTime.lean new file mode 100644 index 000000000000..4658c7d6227a --- /dev/null +++ b/src/Std/Time/Time/PlainTime.lean @@ -0,0 +1,267 @@ +/- +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 +open Internal + +set_option linter.all true + +/-- +Represents a specific point in a day, including hours, minutes, seconds, and nanoseconds. +-/ +structure PlainTime where + + /-- + `Hour` component of the `PlainTime` + -/ + hour : Hour.Ordinal + + /-- + `Minute` component of the `PlainTime` + -/ + minute : Minute.Ordinal + + /-- + `Second` component of the `PlainTime` + -/ + second : Sigma Second.Ordinal + + /-- + `Nanoseconds` component of the `PlainTime` + -/ + nano : Nanosecond.Ordinal + deriving Repr + +instance : Inhabited PlainTime where + default := ⟨0, 0, Sigma.mk false 0, 0, by decide⟩ + +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 + +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. +-/ +@[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. +-/ +@[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 := + 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 := + 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.toOffset.toSeconds + + time.minute.toOffset.toSeconds + + time.second.snd.toOffset + +/-- +Converts a `PlainTime` value to the total number of minutes. +-/ +def toMinutes (time : PlainTime) : Minute.Offset := + time.hour.toOffset.toMinutes + + time.minute.toOffset + + time.second.snd.toOffset.toMinutes + +/-- +Converts a `PlainTime` value to the total number of hours. +-/ +def toHours (time : PlainTime) : Hour.Offset := + time.hour.toOffset + +/-- +Creates a `PlainTime` value from a total number of nanoseconds. +-/ +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) + 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) + 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 := + ofNanoseconds secs.toNanoseconds + +/-- +Adds seconds to a `PlainTime`. +-/ +@[inline] +def addSeconds (time : PlainTime) (secondsToAdd : Second.Offset) : PlainTime := + let totalSeconds := time.toSeconds + secondsToAdd + ofSeconds totalSeconds + +/-- +Subtracts seconds from a `PlainTime`. +-/ +@[inline] +def subSeconds (time : PlainTime) (secondsToSub : Second.Offset) : PlainTime := + addSeconds time (-secondsToSub) + +/-- +Adds minutes to a `PlainTime`. +-/ +@[inline] +def addMinutes (time : PlainTime) (minutesToAdd : Minute.Offset) : PlainTime := + let total := time.toSeconds + minutesToAdd.toSeconds + ofSeconds total + +/-- +Subtracts minutes from a `PlainTime`. +-/ +@[inline] +def subMinutes (time : PlainTime) (minutesToSub : Minute.Offset) : PlainTime := + addMinutes time (-minutesToSub) + +/-- +Adds hours to a `PlainTime`. +-/ +@[inline] +def addHours (time : PlainTime) (hoursToAdd : Hour.Offset) : PlainTime := + let total := time.toSeconds + hoursToAdd.toSeconds + ofSeconds total + +/-- +Subtracts hours from a `PlainTime`. +-/ +@[inline] +def subHours (time : PlainTime) (hoursToSub : Hour.Offset) : PlainTime := + addHours time (-hoursToSub) + +/-- +Adds nanoseconds to a `PlainTime`. +-/ +def addNanoseconds (time : PlainTime) (nanosToAdd : Nanosecond.Offset) : PlainTime := + let total := time.toNanoseconds + nanosToAdd + ofNanoseconds total + +/-- +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) + +/-- +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 + +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 + +instance : HSub PlainTime Second.Offset PlainTime where + hSub := subSeconds + +instance : HAdd PlainTime Minute.Offset PlainTime where + hAdd := addMinutes + +instance : HSub PlainTime Minute.Offset PlainTime where + hSub := subMinutes + +instance : HAdd PlainTime Hour.Offset PlainTime where + hAdd := addHours + +instance : HSub PlainTime Hour.Offset PlainTime where + hSub := subHours + +end PlainTime +end Time +end Std diff --git a/src/Std/Time/Time/Unit/Basic.lean b/src/Std/Time/Time/Unit/Basic.lean new file mode 100644 index 000000000000..ab5f9d282d0c --- /dev/null +++ b/src/Std/Time/Time/Unit/Basic.lean @@ -0,0 +1,329 @@ +/- +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.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 + +/-! +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 +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.div 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.div 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.div 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.div 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.div 1000000 + +/-- +Converts a `Millisecond.Offset` to a `Second.Offset`. +-/ +@[inline] +def toSeconds (offset : Millisecond.Offset) : Second.Offset := + offset.div 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.div 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.div 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.div 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.div 1000 + +/-- +Converts a `Second.Offset` to a `Minute.Offset`. +-/ +@[inline] +def toMinutes (offset : Second.Offset) : Minute.Offset := + offset.div 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`. +-/ +@[inline] +def toHours (offset : Second.Offset) : Hour.Offset := + offset.div 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.div 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.div 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.div 60 + +/-- +Converts a `Minute.Offset` to an `Hour.Offset`. +-/ +@[inline] +def toHours (offset : Minute.Offset) : Hour.Offset := + offset.div 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.div 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.div 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.div 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.div 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 new file mode 100644 index 000000000000..a09c5c4d94cf --- /dev/null +++ b/src/Std/Time/Time/Unit/Hour.lean @@ -0,0 +1,83 @@ +/- +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.Internal +import Std.Time.Time.Unit.Minute +import Std.Time.Time.Unit.Second + +namespace Std +namespace Time +namespace Hour +open Std.Internal +open Internal + +set_option linter.all true + +/-- +`Ordinal` represents a bounded value for hours, ranging from 0 to 23. +-/ +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) + +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 +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⟩ + +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 ≤ 23) : Ordinal := + Bounded.LE.ofNat data h + +/-- +Creates an `Ordinal` from a `Fin` value. +-/ +@[inline] +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) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal +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..4ffe993743c5 --- /dev/null +++ b/src/Std/Time/Time/Unit/Millisecond.lean @@ -0,0 +1,100 @@ +/- +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.Internal +import Std.Time.Time.Unit.Nanosecond + +namespace Std +namespace Time +namespace Millisecond +open Std.Internal +open Internal + +set_option linter.all true + +/-- +`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 : ToString Ordinal where + toString x := toString x.val + +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 0 (0 + (999 : Nat))) 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. +-/ +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⟩ + +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 + +/-- +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 new file mode 100644 index 000000000000..59c8d31be401 --- /dev/null +++ b/src/Std/Time/Time/Unit/Minute.lean @@ -0,0 +1,82 @@ +/- +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.Internal +import Std.Time.Time.Unit.Second + +namespace Std +namespace Time +namespace Minute +open Std.Internal +open Internal + +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, LT + +instance : ToString Ordinal where + toString x := toString x.val + +instance : OfNat Ordinal n := + inferInstanceAs (OfNat (Bounded.LE 0 (0 + (59 : Nat))) 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. +-/ +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 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) : 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 +end Minute +end Time +end Std diff --git a/src/Std/Time/Time/Unit/Nanosecond.lean b/src/Std/Time/Time/Unit/Nanosecond.lean new file mode 100644 index 000000000000..8d635c5a8f50 --- /dev/null +++ b/src/Std/Time/Time/Unit/Nanosecond.lean @@ -0,0 +1,127 @@ +/- +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.Internal + +namespace Std +namespace Time +namespace Nanosecond +open Std.Internal +open Internal + +set_option linter.all true + +/-- +`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 : ToString Ordinal where + toString x := toString x.val + +instance : OfNat Ordinal n where + ofNat := Bounded.LE.ofFin (Fin.ofNat 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 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 + +instance { x y : Offset } : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.val ≤ y.val)) + +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. +-/ +def Span := Bounded.LE (-999999999) 999999999 + deriving Repr, BEq, LE, LT + +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 + +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 new file mode 100644 index 000000000000..cf86e6f76c99 --- /dev/null +++ b/src/Std/Time/Time/Unit/Second.lean @@ -0,0 +1,113 @@ +/- +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.Time.Unit.Nanosecond + +namespace Std +namespace Time +namespace Second +open Std.Internal +open Internal + +set_option linter.all true + +/-- +`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)) + +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 + +instance : Repr (Ordinal l) where + reprPrec r := reprPrec r.val + +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 {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`. +-/ +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 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 + +/-- +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. +-/ +@[inline] +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 (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 leap) : Offset := + UnitVal.ofInt ordinal.val + +end Ordinal +end Second +end Time +end Std diff --git a/src/Std/Time/Zoned.lean b/src/Std/Time/Zoned.lean new file mode 100644 index 000000000000..73a0eddc68c3 --- /dev/null +++ b/src/Std/Time/Zoned.lean @@ -0,0 +1,149 @@ +/- +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.DateTime +import Std.Time.Zoned.ZoneRules +import Std.Time.Zoned.ZonedDateTime +import Std.Time.Zoned.ZonedDateTime + +namespace Std +namespace Time + +namespace PlainDateTime + +/-- +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 monotonic time. +-/ +@[inline] +def now : IO PlainDateTime := do + let tm ← Timestamp.now + let tz ← TimeZone.getSystemTimezone + return PlainDateTime.ofTimestamp tm tz + +end PlainDateTime + +namespace DateTime + +/-- +Converts a `PlainDate` to a `DateTime` +-/ +@[inline] +def ofPlainDate (pd : PlainDate) (tz : TimeZone) : DateTime tz := + DateTime.ofTimestamp (Timestamp.ofPlainDateAssumingUTC pd) tz + +/-- +Converts a `DateTime` to a `PlainDate` +-/ +@[inline] +def toPlainDate (dt : DateTime tz) : PlainDate := + Timestamp.toPlainDateAssumingUTC dt.toTimestamp + +/-- +Converts a `PlainTime` to a `DateTime` +-/ +@[inline] +def ofPlainTime (pt : PlainTime) (tz : TimeZone) : DateTime tz := + DateTime.ofTimestamp (Timestamp.ofPlainTime pt) tz + +/-- +Converts a `DateTime` to a `PlainTime` +-/ +@[inline] +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.toTimestamp + let since := ToTimestamp.toTimestamp since + Std.Time.Duration.sub date.toDurationSinceUnixEpoch since.toDurationSinceUnixEpoch + +end DateTime +namespace ZonedDateTime + +/-- +Gets the current `DateTime`. +-/ +@[inline] +def now : IO ZonedDateTime := do + let date ← Timestamp.now + let tz ← TimeZone.getSystemTimezone + return ofTimestamp date tz + +/-- +Converts a `PlainDate` to a `ZonedDateTime` +-/ +@[inline] +def ofPlainDate (pd : PlainDate) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofTimestamp (Timestamp.ofPlainDateAssumingUTC pd) tz⟩ + +/-- +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⟩ + +/-- +Converts a `ZonedDateTime` to a `PlainTime` +-/ +@[inline] +def toPlainTime (dt : ZonedDateTime) : PlainTime := + DateTime.toPlainTime dt.snd + +/-- +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 := + 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/Database.lean b/src/Std/Time/Zoned/Database.lean new file mode 100644 index 000000000000..68400253e900 --- /dev/null +++ b/src/Std/Time/Zoned/Database.lean @@ -0,0 +1,43 @@ +/- +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.ZonedDateTime +import Std.Time.Zoned.Database.Basic +import Std.Time.Zoned.Database.TZdb + +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. +-/ +-- def applyLeapSecondsOnUTCTimestamp [Database α] (db : α) (tm : Timestamp) : IO Timestamp := do +-- (applyLeapSeconds tm ·) <$> Database.localRules db + +/-- +Gets the TimeZone at the local timezone. +-/ +def getPlainTimeZoneAt [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 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/Basic.lean b/src/Std/Time/Zoned/Database/Basic.lean new file mode 100644 index 000000000000..a1a843ee98e6 --- /dev/null +++ b/src/Std/Time/Zoned/Database/Basic.lean @@ -0,0 +1,111 @@ +/- +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.ZoneRules +import Std.Time.Zoned.Database.Leap +import Std.Time.Zoned.Database.TzIf + +namespace Std +namespace Time + +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 + + /-- + Loads a `ZoneRules` by its id. + -/ + load : α → String → IO TimeZone.ZoneRules + + /-- + 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 + +/-- +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 given time index into a `LocalTimeType` by using a time zone (`tz`) and its identifier. +-/ +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) + let utLocal := convertUt (tz.utLocalIndicators.getD index true) + + return { + gmtOffset := Offset.ofSeconds <| .ofInt localType.gmtOffset + isDst := localType.isDst + abbreviation + wall := wallflag + utLocal + identifier + } + +/-- +Converts a transition. +-/ +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, 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 LocalTimeType := #[] + + for i in [0:tz.header.typecnt.toNat] do + 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 := #[] + + 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 { transitions, localTimes := times } + +/-- +Converts a `TZif.TZifV2` structure to a `ZoneRules` structure. +-/ +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) (id : String) : Except String ZoneRules := do + if let some v2 := tz.v2 + then convertTZifV2 v2 id + else convertTZifV1 tz.v1 id 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 new file mode 100644 index 000000000000..2a85dbcf76e5 --- /dev/null +++ b/src/Std/Time/Zoned/Database/TZdb.lean @@ -0,0 +1,99 @@ +/- +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 + +/-- +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 +open TimeZone + +/-- +Returns a default `TZdb` instance with common timezone data paths for most Linux distributions and macOS. +-/ +@[inline] +def default : TZdb := {} + +/-- +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 + +/-- +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) + +/-- +Extracts a timezone ID from a file path. +-/ +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 + +/-- +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] } + + if let some id := idFromPath PlainTimePath + then parseTZIfFromDisk path id + else throw (IO.userError "cannot read the id of the path.") + +/-- +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/Database/TzIf.lean b/src/Std/Time/Zoned/Database/TzIf.lean new file mode 100644 index 000000000000..0b36e1f1e22c --- /dev/null +++ b/src/Std/Time/Zoned/Database/TzIf.lean @@ -0,0 +1,328 @@ +/- +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 + +-- Based on: https://www.rfc-editor.org/rfc/rfc8536.html + +namespace Std +namespace Time +namespace TimeZone +namespace TZif +open Std.Internal.Parsec Std.Internal.Parsec.ByteArray + +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 + +/-- +Represents the local time type information, including offset and daylight saving details. +-/ +structure PlainTimeType 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 + +/-- +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 + +/-- +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 PlainTimeType + + /-- + 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 + +/-- +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 + +/-- +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 + +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 parsePlainTimeType : Parser PlainTimeType := + PlainTimeType.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 parseLocalTimeType (n : UInt32) : Parser (Array PlainTimeType) := + manyN (n.toNat) parsePlainTimeType + +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 + + 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 ← 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 + + 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 (· ≠ 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 ← 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 + + 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/DateTime.lean b/src/Std/Time/Zoned/DateTime.lean new file mode 100644 index 000000000000..e082cac78471 --- /dev/null +++ b/src/Std/Time/Zoned/DateTime.lean @@ -0,0 +1,463 @@ +/- +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.Internal + +namespace Std +namespace Time +open Internal + +set_option linter.all true + +/-- +It stores a `Timestamp`, a `PlainDateTime` and a `TimeZone` +-/ +structure DateTime (tz : TimeZone) where + private mk :: + + /-- + `Timestamp` represents the exact moment in time. It's a UTC related `Timestamp`. + -/ + timestamp : Timestamp + + /-- + `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 PlainDateTime + +instance : BEq (DateTime tz) where + beq x y := x.timestamp == y.timestamp + +instance : Inhabited (DateTime tz) where + default := ⟨Inhabited.default, Thunk.mk fun _ => Inhabited.default⟩ + +namespace DateTime + +/-- +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) + +/-- +Creates a new zone aware `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.timestamp tz₁ + +/-- +Creates a new `DateTime` out of a `PlainDateTime`. It assumes that the `PlainDateTime` is relative +to UTC. +-/ +@[inline] +def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : DateTime tz := + let tm := Timestamp.ofPlainDateTimeAssumingUTC date + DateTime.mk tm (Thunk.mk fun _ => date.addSeconds tz.toSeconds) + +/-- +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 := + let tm := Timestamp.ofPlainDateTimeAssumingUTC date + DateTime.mk (tm.subSeconds tz.toSeconds) (Thunk.mk fun _ => date) + +/-- +Add `Hour.Offset` to a `DateTime`. +-/ +@[inline] +def addHours (dt : DateTime tz) (hours : Hour.Offset) : DateTime 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 := + ofLocalDateTime (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 + +/-- +Subtract `Minute.Offset` from a `DateTime`. +-/ +@[inline] +def subMinutes (dt : DateTime tz) (minutes : Minute.Offset) : DateTime 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 := + ofLocalDateTime (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 + +/-- +Add `Nanosecond.Offset` to a `DateTime`. +-/ +@[inline] +def addNanoseconds (dt : DateTime tz) (nanoseconds : Nanosecond.Offset) : DateTime 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 := + ofLocalDateTime (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 + +/-- +Subtracts `Day.Offset` to a `DateTime`. +-/ +@[inline] +def subDays (dt : DateTime tz) (days : Day.Offset) : DateTime 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 := + 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 := + 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 := + ofLocalDateTime (dt.date.get.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.date.get.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.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 + +/-- +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.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 + +/-- +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` +-/ +@[inline] +def toPlainDateTime (dt : DateTime tz) : PlainDateTime := + dt.date.get + +/-- +Getter for the `Year` inside of a `DateTime` +-/ +@[inline] +def year (dt : DateTime tz) : Year.Offset := + dt.date.get.year + +/-- +Getter for the `Month` inside of a `DateTime` +-/ +@[inline] +def month (dt : DateTime tz) : Month.Ordinal := + dt.date.get.month + +/-- +Getter for the `Day` inside of a `DateTime` +-/ +@[inline] +def day (dt : DateTime tz) : Day.Ordinal := + dt.date.get.day + +/-- +Getter for the `Hour` inside of a `DateTime` +-/ +@[inline] +def hour (dt : DateTime tz) : Hour.Ordinal := + dt.date.get.hour + +/-- +Getter for the `Minute` inside of a `DateTime` +-/ +@[inline] +def minute (dt : DateTime tz) : Minute.Ordinal := + dt.date.get.minute + +/-- +Getter for the `Second` inside of a `DateTime` +-/ +@[inline] +def second (dt : DateTime tz) : Second.Ordinal dt.date.get.time.second.fst := + dt.date.get.second + +/-- +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) + +/-- +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. +-/ +@[inline] +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.Ordinal.Period := + Day.Ordinal.Period.fromHour dt.hour + +/-- +Determines the era of the given `DateTime` based on its year. +-/ +def era (date : DateTime tz) : Year.Era := + date.year.era + +/-- +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 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 + +/-- +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 + +/-- +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 + +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) (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 + +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 + +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/Offset.lean b/src/Std/Time/Zoned/Offset.lean new file mode 100644 index 000000000000..70e60ff5e234 --- /dev/null +++ b/src/Std/Time/Zoned/Offset.lean @@ -0,0 +1,75 @@ +/- +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 + +namespace Std +namespace Time +namespace TimeZone +open Internal + +set_option linter.all true + +/-- +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 + deriving Repr, Inhabited, BEq + +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.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 + 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 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. +-/ +def ofSeconds (n : Second.Offset) : Offset := + mk n.toHours n + +end Offset +end TimeZone +end Time +end Std diff --git a/src/Std/Time/Zoned/TimeZone.lean b/src/Std/Time/Zoned/TimeZone.lean new file mode 100644 index 000000000000..5b9d62ab6eac --- /dev/null +++ b/src/Std/Time/Zoned/TimeZone.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 +-/ +prelude +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 + + /-- + The abbreviation of the time zone/ + -/ + abbreviation : String + + /-- + Day light saving flag. + -/ + isDST : Bool + deriving Inhabited, Repr, BEq + +namespace TimeZone + +/-- +Fetches the current timestamp from the system. +-/ +@[extern "lean_get_timezone_offset"] +opaque getSystemTimezone : IO TimeZone + +/-- +A zeroed `Timezone` representing UTC (no offset). +-/ +def UTC : TimeZone := + 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" "GMT" false + +/-- +Creates a `Timestamp` from a given number of hour. +-/ +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) (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. +-/ +def toSeconds (tz : TimeZone) : Second.Offset := + tz.offset.second diff --git a/src/Std/Time/Zoned/ZoneRules.lean b/src/Std/Time/Zoned/ZoneRules.lean new file mode 100644 index 000000000000..048cbef73480 --- /dev/null +++ b/src/Std/Time/Zoned/ZoneRules.lean @@ -0,0 +1,165 @@ +/- +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 + +namespace Std +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 + +/-- +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 + +/-- +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 + +/-- +Gets the `TimeZone` offset from a `LocalTimeType`. +-/ +def getTimeZone (time : LocalTimeType) : TimeZone := + ⟨time.gmtOffset, time.identifier, time.abbreviation, time.isDst⟩ + +end LocalTimeType + +/-- +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 + +/-- +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 + + deriving Repr, Inhabited + +namespace Transition + +/-- +Create a TimeZone from a Transition. +-/ +def createTimeZoneFromTransition (transition : Transition) : TimeZone := + let name := transition.localTimeType.identifier + let offset := transition.localTimeType.gmtOffset + let abbreviation := transition.localTimeType.abbreviation + TimeZone.mk offset name abbreviation transition.localTimeType.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 + timestamp.addSeconds offsetInSeconds + +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. +-/ +def findTransitionForTimestamp (zoneRules : ZoneRules) (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? + +/-- +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." diff --git a/src/Std/Time/Zoned/ZonedDateTime.lean b/src/Std/Time/Zoned/ZonedDateTime.lean new file mode 100644 index 000000000000..b1beb3b485c9 --- /dev/null +++ b/src/Std/Time/Zoned/ZonedDateTime.lean @@ -0,0 +1,431 @@ +/- +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.DateTime +import Std.Time.Zoned.ZoneRules + +namespace Std +namespace Time + +set_option linter.all true + +/-- +An existential version of `DateTime` that encapsulates a `DateTime` value without explicitly storing +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 + +instance : Inhabited ZonedDateTime where + default := ⟨Inhabited.default, Inhabited.default⟩ + +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` +-/ +@[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 + +/-- +Creates a new `ZonedDateTime` 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 + +/-- +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. +-/ +@[inline] +def convertTimeZone (date : ZonedDateTime) (tz₁ : TimeZone) : ZonedDateTime := + ofTimestamp date.toTimestamp tz₁ + +/-- +Creates a new `ZonedDateTime` out of a `PlainDateTime` +-/ +@[inline] +def ofPlainDateTimeAssumingUTC (date : PlainDateTime) (tz : TimeZone) : ZonedDateTime := + ⟨tz, DateTime.ofPlainDateTimeAssumingUTC date tz⟩ + +/-- +Converts a `ZonedDateTime` to a `PlainDateTime` +-/ +@[inline] +def toPlainDateTime (dt : ZonedDateTime) : PlainDateTime := + DateTime.toPlainDateTime dt.snd + +/-- +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.date.get.time.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.date.get.time.second.fst := + zdt.snd.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 + +/-- +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 + +/-- +Returns the weekday. +-/ +@[inline] +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 `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. +-/ +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) + +/-- +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. +-/ +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) + +/-- +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. +-/ +@[inline] +def era (date : ZonedDateTime) : Year.Era := + if date.year.toInt ≥ 0 then + .ce + 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. +-/ +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 + +instance : HSub ZonedDateTime Day.Offset ZonedDateTime where + hSub := subDays + +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 + 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 + +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 + +end ZonedDateTime +end Time +end Std diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index db628b207d7f..f2742eff00b9 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -521,6 +521,60 @@ extern "C" LEAN_EXPORT obj_res lean_io_prim_handle_put_str(b_obj_arg h, b_obj_ar } } +/* Std.Time.Timestamp.now : IO Timestamp */ +extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { + using namespace std::chrono; + + 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_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); +} + +/* 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(LEAN_WINDOWS) + 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 + struct tm *tm_ptr = localtime_r(&now_time_t, &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; + + 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, 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/hpow.lean b/tests/lean/hpow.lean index 6dd24d41c369..fe896373faed 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 @@ -14,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 3acd48538b77..a52dcda22f4c 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` @@ -6,7 +7,7 @@ Tests for `pp.numericTypes` and `pp.natLit` RFC #3021 -/ -open Lean (Rat) +open Lean section 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 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 diff --git a/tests/lean/run/timeClassOperations.lean b/tests/lean/run/timeClassOperations.lean new file mode 100644 index 000000000000..2c92b96813d4 --- /dev/null +++ b/tests/lean/run/timeClassOperations.lean @@ -0,0 +1,188 @@ +import Std.Time +open Std.Time + +def date := date("1970-01-20") + +/-- +info: date("1970-02-01") +-/ +#guard_msgs in +#eval date + (12 : Day.Offset) + +/-- +info: date("1970-01-08") +-/ +#guard_msgs in +#eval date - (12 : Day.Offset) + +def datetime := datetime("2000-01-20T03:02:01") + +/-- +info: datetime("2000-01-20T04:02:01.000000000") +-/ +#guard_msgs in +#eval datetime + (1 : Hour.Offset) + +/-- +info: datetime("2000-01-20T02:02:01.000000000") +-/ +#guard_msgs in +#eval datetime - (1 : Hour.Offset) + +/-- +info: datetime("2000-01-20T03:12:01.000000000") +-/ +#guard_msgs in +#eval datetime + (10 : Minute.Offset) + +/-- +info: datetime("2000-01-20T02:52:01.000000000") +-/ +#guard_msgs in +#eval datetime - (10 : Minute.Offset) + +/-- +info: datetime("2000-01-20T03:03:01.000000000") +-/ +#guard_msgs in +#eval datetime + (60 : Second.Offset) + +/-- +info: datetime("2000-01-20T03:01:01.000000000") +-/ +#guard_msgs in +#eval datetime - (60 : Second.Offset) + +/-- +info: datetime("2000-01-21T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime + (1 : Day.Offset) + +/-- +info: datetime("2000-01-19T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime - (1 : Day.Offset) + +def time := time("13:02:01") + +/-- +info: time("14:02:01.000000000") +-/ +#guard_msgs in +#eval time + (1 : Hour.Offset) + +/-- +info: time("12:02:01.000000000") +-/ +#guard_msgs in +#eval time - (1 : Hour.Offset) + +/-- +info: time("13:12:01.000000000") +-/ +#guard_msgs in +#eval time + (10 : Minute.Offset) + +/-- +info: time("12:52:01.000000000") +-/ +#guard_msgs in +#eval time - (10 : Minute.Offset) + +/-- +info: time("13:03:01.000000000") +-/ +#guard_msgs in +#eval time + (60 : Second.Offset) + +/-- +info: time("13:01:01.000000000") +-/ +#guard_msgs in +#eval time - (60 : Second.Offset) + +def datetimetz := zoned("2000-01-20T06:02:01-03:00") + +/-- +info: zoned("2000-01-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz + +/-- +info: zoned("2000-01-22T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz + (2 : Day.Offset) + +/-- +info: zoned("2000-01-19T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz - (1 : Day.Offset) + +/-- +info: zoned("2000-01-20T07:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz + (1 : Hour.Offset) + +/-- +info: zoned("2000-01-20T05:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz - (1 : Hour.Offset) + +/-- +info: zoned("2000-01-20T06:12:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz + (10 : Minute.Offset) + +/-- +info: zoned("2000-01-20T05:52:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz - (10 : Minute.Offset) + +/-- +info: zoned("2000-01-20T06:03:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz + (60 : Second.Offset) + +/-- +info: zoned("2000-01-20T06:01:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz - (60 : Second.Offset) + +/-- +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 duration := after.since now + +/-- +info: "15:40:38.907328169" +-/ +#guard_msgs in +#eval now.format "HH:mm:ss.SSSSSSSSS" + +/-- +info: "15.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 new file mode 100644 index 000000000000..9e638ab4e352 --- /dev/null +++ b/tests/lean/run/timeFormats.lean @@ -0,0 +1,271 @@ +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") + +-- 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: "Monday, June 16, 2014 03:03:03 -0300" +-/ +#guard_msgs in +#eval FullDayTimeZone.format date₁.snd + +def tm := date₁.snd.timestamp +def date₂ := DateTime.ofTimestamp 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.ofTimestamp 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.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" +-/ +#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 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 + +/-- +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 + +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" +-/ +#guard_msgs in +#eval l +/-- +info: "03:11:01 AM" +-/ +#guard_msgs in +#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" +-/ +#guard_msgs in +#eval ShortDate.formatBuilder date₁.month date₁.day date₁.year + +/-- +info: "0053-06-19" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-700000⟩) .UTC) + +/-- +info: "-0002-09-16" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-720000⟩) .UTC) + +/-- +info: "-0084-07-28" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-750000⟩) .UTC) + +/-- +info: "-0221-09-04" +-/ +#guard_msgs in +#eval Formats.sqlDate.format (DateTime.ofPlainDate (PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩) .UTC) + +/-- +info: date("-0221-09-04") +-/ +#guard_msgs in +#eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ + +/-- +info: date("-0221-09-04") +-/ +#guard_msgs in +#eval PlainDate.ofDaysSinceUNIXEpoch ⟨-800000⟩ + +/-- +info: date("2002-07-14") +-/ +#guard_msgs in +#eval date("2002-07-14") + +/-- +info: time("14:13:12.000000000") +-/ +#guard_msgs in +#eval time("14:13:12") + +/-- +info: zoned("2002-07-14T14:13:12.000000000Z") +-/ +#guard_msgs in +#eval zoned("2002-07-14T14:13:12Z") + +/-- +info: zoned("2002-07-14T14:13:12.000000000+09:00") +-/ +#guard_msgs in +#eval zoned("2002-07-14T14:13:12+09:00") + +/-- +info: "2002-07-14" +-/ +#guard_msgs in +#eval zoned("2002-07-14T14:13:12+09:00").format "yyyy-MM-dd" + +/-- +info: "14-13-12" +-/ +#guard_msgs in +#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 new file mode 100644 index 000000000000..e79d6e202617 --- /dev/null +++ b/tests/lean/run/timeLocalDateTime.lean @@ -0,0 +1,87 @@ +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 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₁ := 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 + +/-- +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₁.toTimestampAssumingUTC.toSecondsSinceUnixEpoch + +/-- +info: 736952399 +-/ +#guard_msgs in +#eval date₂.toTimestampAssumingUTC.toSecondsSinceUnixEpoch + +/-- +info: "09/05/1993 12:59:59" +-/ +#guard_msgs in +#eval PlainDateTime.ofUTCTimestamp 736952399 |> format + +/-- +info: 736952399 +-/ +#guard_msgs in +#eval PlainDateTime.toTimestampAssumingUTC date₂ |>.toSecondsSinceUnixEpoch + +/-- +info: "16/08/2024" +-/ +#guard_msgs in +#eval PlainDate.ofDaysSinceUNIXEpoch 19951 |> format₂ + +/-- +info: 19951 +-/ +#guard_msgs in +#eval PlainDate.toDaysSinceUNIXEpoch date₃ + +/-- +info: Std.Time.Weekday.friday +-/ +#guard_msgs in +#eval PlainDate.weekday date₃ + +/-- +info: #[] +-/ +#guard_msgs in +#eval Id.run do + let mut res := #[] + + for i in [0:10000] do + let i := Int.ofNat i - 999975 + let date := PlainDate.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/timeOperations.lean b/tests/lean/run/timeOperations.lean new file mode 100644 index 000000000000..9f4eba8a858c --- /dev/null +++ b/tests/lean/run/timeOperations.lean @@ -0,0 +1,307 @@ +import Std.Time +open Std.Time + +def date := date("1970-01-20") + +/-- +info: date("1970-02-01") +-/ +#guard_msgs in +#eval date.addDays 12 + +/-- +info: date("1970-02-20") +-/ +#guard_msgs in +#eval date.addMonthsClip 1 + +/-- +info: date("1971-01-20") +-/ +#guard_msgs in +#eval date.addYearsRollOver 1 + +/-- +info: date("1969-01-20") +-/ +#guard_msgs in +#eval date.subYearsClip 1 + +/-- +info: date("1969-12-20") +-/ +#guard_msgs in +#eval date.subMonthsClip 1 + +def datetime := datetime("2000-01-20T03:02:01") + +/-- +info: datetime("2000-01-20T04:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.addHours 1 + +/-- +info: datetime("2000-01-20T02:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.subHours 1 + +/-- +info: datetime("2000-01-20T03:12:01.000000000") +-/ +#guard_msgs in +#eval datetime.addMinutes 10 + +/-- +info: datetime("2000-01-20T02:52:01.000000000") +-/ +#guard_msgs in +#eval datetime.subMinutes 10 + +/-- +info: datetime("2000-01-20T03:03:01.000000000") +-/ +#guard_msgs in +#eval datetime.addSeconds 60 + +/-- +info: datetime("2000-01-20T03:01:01.000000000") +-/ +#guard_msgs in +#eval datetime.subSeconds 60 + +/-- +info: datetime("2000-01-21T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.addDays 1 + +/-- +info: datetime("2000-01-19T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.subDays 1 + +/-- +info: datetime("2000-02-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.addMonthsClip 1 + +/-- +info: datetime("1999-12-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.subMonthsClip 1 + +/-- +info: datetime("2000-02-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.addMonthsRollOver 1 + +/-- +info: datetime("1999-12-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.subMonthsRollOver 1 + +/-- +info: datetime("2001-01-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.addYearsClip 1 + +/-- +info: datetime("1999-01-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.subYearsClip 1 + +/-- +info: datetime("2001-01-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.addYearsRollOver 1 + +/-- +info: datetime("1999-01-20T03:02:01.000000000") +-/ +#guard_msgs in +#eval datetime.subYearsRollOver 1 + +def time := time("13:02:01") + +/-- +info: time("14:02:01.000000000") +-/ +#guard_msgs in +#eval time.addHours 1 + +/-- +info: time("12:02:01.000000000") +-/ +#guard_msgs in +#eval time.subHours 1 + +/-- +info: time("13:12:01.000000000") +-/ +#guard_msgs in +#eval time.addMinutes 10 + +/-- +info: time("12:52:01.000000000") +-/ +#guard_msgs in +#eval time.subMinutes 10 + +/-- +info: time("13:03:01.000000000") +-/ +#guard_msgs in +#eval time.addSeconds 60 + +/-- +info: time("13:01:01.000000000") +-/ +#guard_msgs in +#eval time.subSeconds 60 + +def datetimetz := zoned("2000-01-20T06:02:01-03:00") + +/-- +info: zoned("2000-01-20T06:02:01.000000000-03:00") + +-/ +#guard_msgs in +#eval datetimetz + +/-- +info: zoned("2000-01-22T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addDays 2 + +/-- +info: zoned("2000-01-19T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.subDays 1 + +/-- +info: zoned("2000-02-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addMonthsClip 1 + +/-- +info: zoned("1999-12-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.subMonthsClip 1 + +/-- +info: zoned("2000-02-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addMonthsRollOver 1 + +/-- +info: zoned("1999-12-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.subMonthsRollOver 1 + +/-- +info: zoned("2001-01-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addYearsClip 1 + +/-- +info: zoned("2001-01-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addYearsClip 1 + +/-- +info: zoned("2001-01-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addYearsRollOver 1 + +/-- +info: zoned("1999-01-20T06:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.subYearsRollOver 1 + +/-- +info: zoned("2000-01-20T07:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addHours 1 + +/-- +info: zoned("2000-01-20T05:02:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.subHours 1 + +/-- +info: zoned("2000-01-20T06:12:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addMinutes 10 + +/-- +info: zoned("2000-01-20T05:52:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.subMinutes 10 + +/-- +info: zoned("2000-01-20T06:03:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.addSeconds 60 + +/-- +info: zoned("2000-01-20T06:01:01.000000000-03:00") +-/ +#guard_msgs in +#eval datetimetz.subSeconds 60 + +/-- +info: "86400s" +-/ +#guard_msgs in +#eval (date("2000-01-20").since (date("2000-01-19"))) + +/-- +info: "86399s" +-/ +#guard_msgs in +#eval (datetime("2000-01-20T00:00:00").since (datetime("2000-01-19T00:00:01"))) + +/-- +info: "86399s" +-/ +#guard_msgs in +#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: zoned("2024-08-29T10:56:43.276801081+02:00") +-/ +#guard_msgs in +#eval now + +/-- +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 new file mode 100644 index 000000000000..f843a5a818c1 --- /dev/null +++ b/tests/lean/run/timeParse.lean @@ -0,0 +1,205 @@ +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 + +def tm := date₁.snd.timestamp +def date₂ := DateTime.ofTimestamp tm brTZ + +/-- +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 + +def tm₃ := date₁.toTimestamp +def date₃ := DateTime.ofTimestamp tm₃ brTZ + +/-- +info: "2014-06-16T00:00:00.000000000Z" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := ShortDate.parse! "06/16/2014" + ISO8601UTC.format t.snd + +-- 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.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.000000000-03:00" +-/ +#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:00.000000000Z" +-/ +#guard_msgs in +#eval + let t : ZonedDateTime := LongDateTime.parse! "August 16, 2024 01:28 AM" + ISO8601UTC.format t.snd + +/-- +info: "0000-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) + +/-- +info: "0000-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) + +/-- +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-16T13:28:00.000000000Z" +-/ +#guard_msgs in +#eval + let t5 : ZonedDateTime := CustomDayTime.parse! "Thu 16 Aug 2024 13:28" + ISO8601UTC.format t5.snd + +/-- +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) + +/-- +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) + +/-- +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 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 + +/-- +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 + +/-- +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 + +/-- +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 + +/-- +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 + +/-- +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 + +/-- +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 + +/-- +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 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 diff --git a/tests/lean/run/timeTzifParse.lean b/tests/lean/run/timeTzifParse.lean new file mode 100644 index 000000000000..10ac2ca911d6 --- /dev/null +++ b/tests/lean/run/timeTzifParse.lean @@ -0,0 +1,71 @@ +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 } +-/ +#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 + +/-- +info: some (zoned("1969-12-30T21:00:00.000000000-03:00")) +-/ +#guard_msgs in +#eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 0) rules + +/-- +info: some (zoned("2012-12-10T00:35:47.000000000-02:00")) +-/ +#guard_msgs in +#eval ZonedDateTime.ofZoneRules (Timestamp.ofSecondsSinceUnixEpoch 1355106947) rules