luau/prototyping/Luau/Type.agda

164 lines
5.4 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module Luau.Type where
open import FFI.Data.Maybe using (Maybe; just; nothing; just-inv)
open import Agda.Builtin.Equality using (_≡_; refl)
open import Properties.Dec using (Dec; yes; no)
open import Properties.Equality using (cong)
open import FFI.Data.Maybe using (Maybe; just; nothing)
data Type : Set where
nil : Type
_⇒_ : Type Type Type
bot : Type
top : Type
boolean : Type
number : Type
__ : Type Type Type
_∩_ : Type Type Type
lhs : Type Type
lhs (T _) = T
lhs (T _) = T
lhs (T _) = T
lhs nil = nil
lhs bot = bot
lhs top = top
lhs number = number
lhs boolean = boolean
rhs : Type Type
rhs (_ T) = T
rhs (_ T) = T
rhs (_ T) = T
rhs nil = nil
rhs bot = bot
rhs top = top
rhs number = number
rhs boolean = boolean
_≡ᵀ_ : (T U : Type) Dec(T U)
nil ≡ᵀ nil = yes refl
nil ≡ᵀ (S T) = no (λ ())
nil ≡ᵀ bot = no (λ ())
nil ≡ᵀ top = no (λ ())
nil ≡ᵀ number = no (λ ())
nil ≡ᵀ boolean = no (λ ())
nil ≡ᵀ (S T) = no (λ ())
nil ≡ᵀ (S T) = no (λ ())
(S T) ≡ᵀ nil = no (λ ())
(S T) ≡ᵀ (U V) with (S ≡ᵀ U) | (T ≡ᵀ V)
(S T) ≡ᵀ (S T) | yes refl | yes refl = yes refl
(S T) ≡ᵀ (U V) | _ | no p = no (λ q p (cong rhs q))
(S T) ≡ᵀ (U V) | no p | _ = no (λ q p (cong lhs q))
(S T) ≡ᵀ bot = no (λ ())
(S T) ≡ᵀ top = no (λ ())
(S T) ≡ᵀ number = no (λ ())
(S T) ≡ᵀ boolean = no (λ ())
(S T) ≡ᵀ (U V) = no (λ ())
(S T) ≡ᵀ (U V) = no (λ ())
bot ≡ᵀ nil = no (λ ())
bot ≡ᵀ (U V) = no (λ ())
bot ≡ᵀ bot = yes refl
bot ≡ᵀ top = no (λ ())
bot ≡ᵀ number = no (λ ())
bot ≡ᵀ boolean = no (λ ())
bot ≡ᵀ (U V) = no (λ ())
bot ≡ᵀ (U V) = no (λ ())
top ≡ᵀ nil = no (λ ())
top ≡ᵀ (U V) = no (λ ())
top ≡ᵀ bot = no (λ ())
top ≡ᵀ top = yes refl
top ≡ᵀ number = no (λ ())
top ≡ᵀ boolean = no (λ ())
top ≡ᵀ (U V) = no (λ ())
top ≡ᵀ (U V) = no (λ ())
number ≡ᵀ nil = no (λ ())
number ≡ᵀ (T U) = no (λ ())
number ≡ᵀ bot = no (λ ())
number ≡ᵀ top = no (λ ())
number ≡ᵀ number = yes refl
number ≡ᵀ boolean = no (λ ())
number ≡ᵀ (T U) = no (λ ())
number ≡ᵀ (T U) = no (λ ())
boolean ≡ᵀ nil = no (λ ())
boolean ≡ᵀ (T U) = no (λ ())
boolean ≡ᵀ bot = no (λ ())
boolean ≡ᵀ top = no (λ ())
boolean ≡ᵀ boolean = yes refl
boolean ≡ᵀ number = no (λ ())
boolean ≡ᵀ (T U) = no (λ ())
boolean ≡ᵀ (T U) = no (λ ())
(S T) ≡ᵀ nil = no (λ ())
(S T) ≡ᵀ (U V) = no (λ ())
(S T) ≡ᵀ bot = no (λ ())
(S T) ≡ᵀ top = no (λ ())
(S T) ≡ᵀ number = no (λ ())
(S T) ≡ᵀ boolean = no (λ ())
(S T) ≡ᵀ (U V) with (S ≡ᵀ U) | (T ≡ᵀ V)
(S T) ≡ᵀ (S T) | yes refl | yes refl = yes refl
(S T) ≡ᵀ (U V) | _ | no p = no (λ q p (cong rhs q))
(S T) ≡ᵀ (U V) | no p | _ = no (λ q p (cong lhs q))
(S T) ≡ᵀ (U V) = no (λ ())
(S T) ≡ᵀ nil = no (λ ())
(S T) ≡ᵀ (U V) = no (λ ())
(S T) ≡ᵀ bot = no (λ ())
(S T) ≡ᵀ top = no (λ ())
(S T) ≡ᵀ number = no (λ ())
(S T) ≡ᵀ boolean = no (λ ())
(S T) ≡ᵀ (U V) = no (λ ())
(S T) ≡ᵀ (U V) with (S ≡ᵀ U) | (T ≡ᵀ V)
(S T) ≡ᵀ (U V) | yes refl | yes refl = yes refl
(S T) ≡ᵀ (U V) | _ | no p = no (λ q p (cong rhs q))
(S T) ≡ᵀ (U V) | no p | _ = no (λ q p (cong lhs q))
_≡ᴹᵀ_ : (T U : Maybe Type) Dec(T U)
nothing ≡ᴹᵀ nothing = yes refl
nothing ≡ᴹᵀ just U = no (λ ())
just T ≡ᴹᵀ nothing = no (λ ())
just T ≡ᴹᵀ just U with T ≡ᵀ U
(just T ≡ᴹᵀ just T) | yes refl = yes refl
(just T ≡ᴹᵀ just U) | no p = no (λ q p (just-inv q))
data Mode : Set where
strict : Mode
nonstrict : Mode
src : Mode Type Type
src m nil = bot
src m number = bot
src m boolean = bot
src m (S T) = S
-- In nonstrict mode, functions are covaraiant, in strict mode they're contravariant
src strict (S T) = (src strict S) (src strict T)
src nonstrict (S T) = (src nonstrict S) (src nonstrict T)
src strict (S T) = (src strict S) (src strict T)
src nonstrict (S T) = (src nonstrict S) (src nonstrict T)
src strict bot = top
src nonstrict bot = bot
src strict top = bot
src nonstrict top = top
tgt : Type Type
tgt nil = bot
tgt (S T) = T
tgt bot = bot
tgt top = top
tgt number = bot
tgt boolean = bot
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