2022-02-09 23:14:29 +00:00
|
|
|
|
module Luau.Type where
|
|
|
|
|
|
2022-02-11 20:38:35 +00:00
|
|
|
|
open import FFI.Data.Maybe using (Maybe; just; nothing)
|
|
|
|
|
|
2022-02-09 23:14:29 +00:00
|
|
|
|
data Type : Set where
|
|
|
|
|
nil : Type
|
|
|
|
|
_⇒_ : Type → Type → Type
|
|
|
|
|
none : Type
|
|
|
|
|
any : Type
|
2022-02-18 19:09:00 +00:00
|
|
|
|
number : Type
|
2022-02-09 23:14:29 +00:00
|
|
|
|
_∪_ : Type → Type → Type
|
|
|
|
|
_∩_ : Type → Type → Type
|
|
|
|
|
|
2022-02-11 20:38:35 +00:00
|
|
|
|
src : Type → Type
|
|
|
|
|
src nil = none
|
|
|
|
|
src (S ⇒ T) = S
|
|
|
|
|
src none = none
|
|
|
|
|
src any = any
|
2022-02-18 19:09:00 +00:00
|
|
|
|
src number = none
|
2022-02-11 20:38:35 +00:00
|
|
|
|
src (S ∪ T) = (src S) ∪ (src T)
|
|
|
|
|
src (S ∩ T) = (src S) ∩ (src T)
|
|
|
|
|
|
|
|
|
|
tgt : Type → Type
|
|
|
|
|
tgt nil = none
|
|
|
|
|
tgt (S ⇒ T) = T
|
|
|
|
|
tgt none = none
|
|
|
|
|
tgt any = any
|
2022-02-18 19:09:00 +00:00
|
|
|
|
tgt number = none
|
2022-02-11 20:38:35 +00:00
|
|
|
|
tgt (S ∪ T) = (tgt S) ∪ (tgt T)
|
|
|
|
|
tgt (S ∩ T) = (tgt S) ∩ (tgt T)
|
|
|
|
|
|
|
|
|
|
optional : Type → Type
|
|
|
|
|
optional nil = nil
|
|
|
|
|
optional (T ∪ nil) = (T ∪ nil)
|
|
|
|
|
optional T = (T ∪ nil)
|
|
|
|
|
|
|
|
|
|
normalizeOptional : Type → Type
|
|
|
|
|
normalizeOptional (S ∪ T) with normalizeOptional S | normalizeOptional T
|
|
|
|
|
normalizeOptional (S ∪ T) | (S′ ∪ nil) | (T′ ∪ nil) = (S′ ∪ T′) ∪ nil
|
|
|
|
|
normalizeOptional (S ∪ T) | S′ | (T′ ∪ nil) = (S′ ∪ T′) ∪ nil
|
|
|
|
|
normalizeOptional (S ∪ T) | (S′ ∪ nil) | T′ = (S′ ∪ T′) ∪ nil
|
|
|
|
|
normalizeOptional (S ∪ T) | S′ | nil = optional S′
|
|
|
|
|
normalizeOptional (S ∪ T) | nil | T′ = optional T′
|
|
|
|
|
normalizeOptional (S ∪ T) | S′ | T′ = S′ ∪ T′
|
|
|
|
|
normalizeOptional T = T
|