diff --git a/prototyping/Examples/Run.agda b/prototyping/Examples/Run.agda index 70863cc9..daf627a5 100644 --- a/prototyping/Examples/Run.agda +++ b/prototyping/Examples/Run.agda @@ -3,8 +3,9 @@ module Examples.Run where open import Agda.Builtin.Equality using (_≡_; refl) -open import Luau.Syntax using (nil; var; _$_; function_is_end; return; _∙_; done; _⟨_⟩; number; binexp; +) -open import Luau.Value using (nil; number) +open import Agda.Builtin.Bool using (true; false) +open import Luau.Syntax using (nil; var; _$_; function_is_end; return; _∙_; done; _⟨_⟩; number; binexp; +; <; true; false) +open import Luau.Value using (nil; number; bool) open import Luau.Run using (run; return) ex1 : (run (function "id" ⟨ var "x" ⟩ is return (var "x") ∙ done end ∙ return (var "id" $ nil) ∙ done) ≡ return nil _) @@ -15,3 +16,6 @@ ex2 = refl ex3 : (run (function "fn" ⟨ var "x" ⟩ is return (binexp (number 1.0) + (number 2.0)) ∙ done end ∙ return (var "fn" $ nil) ∙ done) ≡ return (number 3.0) _) ex3 = refl + +ex4 : (run (function "fn" ⟨ var "x" ⟩ is return (binexp (number 1.0) < (number 2.0)) ∙ done end ∙ return (var "fn" $ nil) ∙ done) ≡ return (bool true) _) +ex4 = refl diff --git a/prototyping/FFI/Data/Bool.agda b/prototyping/FFI/Data/Bool.agda deleted file mode 100644 index 4b04b033..00000000 --- a/prototyping/FFI/Data/Bool.agda +++ /dev/null @@ -1,8 +0,0 @@ -module FFI.Data.Bool where - -{-# FOREIGN GHC import qualified Data.Bool #-} - -data Bool : Set where - false : Bool - true : Bool -{-# COMPILE GHC Bool = data Data.Bool.Bool (Data.Bool.False|Data.Bool.True) #-} diff --git a/prototyping/FFI/Data/Vector.agda b/prototyping/FFI/Data/Vector.agda index 30a51d69..08761edf 100644 --- a/prototyping/FFI/Data/Vector.agda +++ b/prototyping/FFI/Data/Vector.agda @@ -6,7 +6,7 @@ open import Agda.Builtin.Equality using (_≡_) open import Agda.Builtin.Equality.Rewrite using () open import Agda.Builtin.Int using (Int; pos; negsuc) open import Agda.Builtin.Nat using (Nat) -open import FFI.Data.Bool using (Bool; false; true) +open import Agda.Builtin.Bool using (Bool; false; true) open import FFI.Data.HaskellInt using (HaskellInt; haskellIntToInt; intToHaskellInt) open import FFI.Data.Maybe using (Maybe; just; nothing) open import Properties.Equality using (_≢_) diff --git a/prototyping/Luau/OpSem.agda b/prototyping/Luau/OpSem.agda index c97cb973..e3f3aa6f 100644 --- a/prototyping/Luau/OpSem.agda +++ b/prototyping/Luau/OpSem.agda @@ -3,18 +3,50 @@ module Luau.OpSem where open import Agda.Builtin.Equality using (_≡_) -open import Agda.Builtin.Float using (Float; primFloatPlus; primFloatMinus; primFloatTimes; primFloatDiv) +open import Agda.Builtin.Float using (Float; primFloatPlus; primFloatMinus; primFloatTimes; primFloatDiv; primFloatEquality; primFloatLess; primFloatInequality) +open import Agda.Builtin.Bool using (Bool; true; false) +open import Utility.Bool using (not; _or_; _and_) +open import Agda.Builtin.Nat using () renaming (_==_ to _==ᴬ_) open import FFI.Data.Maybe using (just) open import Luau.Heap using (Heap; _≡_⊕_↦_; _[_]; function_is_end) open import Luau.Substitution using (_[_/_]ᴮ) -open import Luau.Syntax using (Expr; Stat; Block; nil; addr; var; function_is_end; _$_; block_is_end; local_←_; _∙_; done; return; name; fun; arg; binexp; BinaryOperator; +; -; *; /; number) -open import Luau.Value using (addr; val; number) +open import Luau.Syntax using (Expr; Stat; Block; nil; addr; var; function_is_end; _$_; block_is_end; local_←_; _∙_; done; return; name; fun; arg; binexp; BinaryOperator; +; -; *; /; <; >; ==; ~=; <=; >=; number) +open import Luau.Value using (addr; val; number; Value; bool) +open import Luau.RuntimeType using (RuntimeType; valueType) -evalBinOp : Float → BinaryOperator → Float → Float -evalBinOp x + y = primFloatPlus x y -evalBinOp x - y = primFloatMinus x y -evalBinOp x * y = primFloatTimes x y -evalBinOp x / y = primFloatDiv x y +evalNumOp : Float → BinaryOperator → Float → Value +evalNumOp x + y = number (primFloatPlus x y) +evalNumOp x - y = number (primFloatMinus x y) +evalNumOp x * y = number (primFloatTimes x y) +evalNumOp x / y = number (primFloatDiv x y) +evalNumOp x < y = bool (primFloatLess x y) +evalNumOp x > y = bool (primFloatLess y x) +evalNumOp x == y = bool (primFloatEquality x y) +evalNumOp x ~= y = bool (primFloatInequality x y) +evalNumOp x <= y = bool ((primFloatLess x y) or (primFloatEquality x y)) +evalNumOp x >= y = bool ((primFloatLess y x) or (primFloatEquality x y)) + +evalEqOp : Value → Value → Value +evalEqOp Value.nil Value.nil = bool true +evalEqOp (addr x) (addr y) = bool (x ==ᴬ y) +evalEqOp (number x) (number y) = bool (primFloatEquality x y) +evalEqOp (bool true) (bool y) = bool y +evalEqOp (bool false) (bool y) = bool (not y) +evalEqOp _ _ = bool false + +evalNeqOp : Value → Value → Value +evalNeqOp Value.nil Value.nil = bool false +evalNeqOp (addr x) (addr y) = bool (not (x ==ᴬ y)) +evalNeqOp (number x) (number y) = bool (primFloatInequality x y) +evalNeqOp (bool true) (bool y) = bool (not y) +evalNeqOp (bool false) (bool y) = bool y +evalNeqOp _ _ = bool true + +coerceToBool : Value → Bool +coerceToBool Value.nil = false +coerceToBool (addr x) = true +coerceToBool (number x) = true +coerceToBool (bool x) = x data _⊢_⟶ᴮ_⊣_ {a} : Heap a → Block a → Block a → Heap a → Set data _⊢_⟶ᴱ_⊣_ {a} : Heap a → Expr a → Expr a → Heap a → Set @@ -62,10 +94,20 @@ data _⊢_⟶ᴱ_⊣_ where --------------------------------- H ⊢ (block b is done end) ⟶ᴱ nil ⊣ H - binOpEval : ∀ {H op} m n → - - -------------------------------------------------------------------------- - H ⊢ (binexp (number m) op (number n)) ⟶ᴱ (number (evalBinOp m op n)) ⊣ H + binOpEquality : + ∀ {H x y} → + --------------------------------------------------------------------------- + H ⊢ (binexp (val x) == (val y)) ⟶ᴱ (val (evalEqOp x y)) ⊣ H + + binOpInequality : + ∀ {H x y} → + ---------------------------------------------------------------------------- + H ⊢ (binexp (val x) ~= (val y)) ⟶ᴱ (val (evalNeqOp x y)) ⊣ H + + binOpNumbers : + ∀ {H x op y} → + ----------------------------------------------------------------------- + H ⊢ (binexp (number x) op (number y)) ⟶ᴱ (val (evalNumOp x op y)) ⊣ H binOp₁ : ∀ {H H′ x x′ op y} → diff --git a/prototyping/Luau/RuntimeType.agda b/prototyping/Luau/RuntimeType.agda index 25c0283a..375fc540 100644 --- a/prototyping/Luau/RuntimeType.agda +++ b/prototyping/Luau/RuntimeType.agda @@ -1,13 +1,15 @@ module Luau.RuntimeType where -open import Luau.Value using (Value; nil; addr; number) +open import Luau.Value using (Value; nil; addr; number; bool) data RuntimeType : Set where function : RuntimeType number : RuntimeType nil : RuntimeType + boolean : RuntimeType valueType : Value → RuntimeType valueType nil = nil valueType (addr x) = function valueType (number x) = number +valueType (bool _) = boolean diff --git a/prototyping/Luau/RuntimeType/ToString.agda b/prototyping/Luau/RuntimeType/ToString.agda index be67ee0c..dfb7c955 100644 --- a/prototyping/Luau/RuntimeType/ToString.agda +++ b/prototyping/Luau/RuntimeType/ToString.agda @@ -1,9 +1,10 @@ module Luau.RuntimeType.ToString where open import FFI.Data.String using (String) -open import Luau.RuntimeType using (RuntimeType; function; number; nil) +open import Luau.RuntimeType using (RuntimeType; function; number; nil; boolean) runtimeTypeToString : RuntimeType → String runtimeTypeToString function = "function" runtimeTypeToString number = "number" runtimeTypeToString nil = "nil" +runtimeTypeToString boolean = "boolean" diff --git a/prototyping/Luau/StrictMode/ToString.agda b/prototyping/Luau/StrictMode/ToString.agda new file mode 100644 index 00000000..0630aeec --- /dev/null +++ b/prototyping/Luau/StrictMode/ToString.agda @@ -0,0 +1,39 @@ +{-# OPTIONS --rewriting #-} + +module Luau.StrictMode.ToString where + +open import FFI.Data.String using (String; _++_) +open import Luau.StrictMode using (Warningᴱ; Warningᴮ; UnallocatedAddress; UnboundVariable; FunctionCallMismatch; FunctionDefnMismatch; BlockMismatch; app₁; app₂; BinopMismatch₁; BinopMismatch₂; bin₁; bin₂; block₁; return; LocalVarMismatch; local₁; local₂; function₁; function₂; heap; expr; block; addr) +open import Luau.Syntax using (Expr; yes; var; var_∈_; _⟨_⟩∈_; _$_; addr; number; binexp; nil; function_is_end; block_is_end; done; return; local_←_; _∙_; fun; arg; name) +open import Luau.Type using (strict) +open import Luau.TypeCheck(strict) using (_⊢ᴮ_∈_; _⊢ᴱ_∈_) +open import Luau.Addr.ToString using (addrToString) +open import Luau.Var.ToString using (varToString) +open import Luau.Type.ToString using (typeToString) +open import Luau.Syntax.ToString using (binOpToString) + +warningToStringᴱ : ∀ {H Γ T} M → {D : Γ ⊢ᴱ M ∈ T} → Warningᴱ H D → String +warningToStringᴮ : ∀ {H Γ T} B → {D : Γ ⊢ᴮ B ∈ T} → Warningᴮ H D → String + +warningToStringᴱ (var x) (UnboundVariable p) = "Unbound variable " ++ varToString x +warningToStringᴱ (addr a) (UnallocatedAddress p) = "Unallocated adress " ++ addrToString a +warningToStringᴱ (M $ N) (FunctionCallMismatch {T = T} {U = U} p) = "Function has type " ++ typeToString T ++ " but argument has type " ++ typeToString U +warningToStringᴱ (M $ N) (app₁ W) = warningToStringᴱ M W +warningToStringᴱ (M $ N) (app₂ W) = warningToStringᴱ N W +warningToStringᴱ (function f ⟨ var x ∈ T ⟩∈ U is B end) (FunctionDefnMismatch {V = V} p) = "Function expresion " ++ varToString f ++ " has return type " ++ typeToString U ++ " but body returns " ++ typeToString V +warningToStringᴱ (function f ⟨ var x ∈ T ⟩∈ U is B end) (function₁ W) = warningToStringᴮ B W ++ "\n in function expression " ++ varToString f +warningToStringᴱ block var b ∈ T is B end (BlockMismatch {U = U} p) = "Block " ++ varToString b ++ " has type " ++ typeToString T ++ " but body returns " ++ typeToString U +warningToStringᴱ block var b ∈ T is B end (block₁ W) = warningToStringᴮ B W ++ "\n in block " ++ varToString b +warningToStringᴱ (binexp M op N) (BinopMismatch₁ {T = T} p) = "Binary operator " ++ binOpToString op ++ " lhs has type " ++ typeToString T ++ " not number" +warningToStringᴱ (binexp M op N) (BinopMismatch₂ {U = U} p) = "Binary operator " ++ binOpToString op ++ " rhs has type " ++ typeToString U ++ " not number" +warningToStringᴱ (binexp M op N) (bin₁ W) = warningToStringᴱ M W +warningToStringᴱ (binexp M op N) (bin₂ W) = warningToStringᴱ N W + +warningToStringᴮ (function f ⟨ var x ∈ T ⟩∈ U is C end ∙ B) (FunctionDefnMismatch {V = V} p) = "Function declaration " ++ varToString f ++ " has return type " ++ typeToString U ++ " but body returns " ++ typeToString V +warningToStringᴮ (function f ⟨ var x ∈ T ⟩∈ U is C end ∙ B) (function₁ W) = warningToStringᴮ C W ++ "\n in function declaration " ++ varToString f +warningToStringᴮ (function f ⟨ var x ∈ T ⟩∈ U is C end ∙ B) (function₂ W) = warningToStringᴮ B W +warningToStringᴮ (local var x ∈ T ← M ∙ B) (LocalVarMismatch {U = U} p) = "Local variable " ++ varToString x ++ " has type " ++ typeToString T ++ " but expression has type " ++ typeToString U +warningToStringᴮ (local var x ∈ T ← M ∙ B) (local₁ W) = warningToStringᴱ M W ++ "\n in local variable declaration " ++ varToString x +warningToStringᴮ (local var x ∈ T ← M ∙ B) (local₂ W) = warningToStringᴮ B W +warningToStringᴮ (return M ∙ B) (return W) = warningToStringᴱ M W ++ "\n in return statement" + diff --git a/prototyping/Luau/Substitution.agda b/prototyping/Luau/Substitution.agda index 0ba59745..abf2ee7f 100644 --- a/prototyping/Luau/Substitution.agda +++ b/prototyping/Luau/Substitution.agda @@ -1,6 +1,6 @@ module Luau.Substitution where -open import Luau.Syntax using (Expr; Stat; Block; nil; addr; var; function_is_end; _$_; block_is_end; local_←_; _∙_; done; return; _⟨_⟩ ; name; fun; arg; number; binexp) +open import Luau.Syntax using (Expr; Stat; Block; nil; true; false; addr; var; function_is_end; _$_; block_is_end; local_←_; _∙_; done; return; _⟨_⟩ ; name; fun; arg; number; binexp) open import Luau.Value using (Value; val) open import Luau.Var using (Var; _≡ⱽ_) open import Properties.Dec using (Dec; yes; no) @@ -11,6 +11,8 @@ var_[_/_]ᴱwhenever_ : ∀ {a P} → Var → Value → Var → (Dec P) → Expr _[_/_]ᴮunless_ : ∀ {a P} → Block a → Value → Var → (Dec P) → Block a nil [ v / x ]ᴱ = nil +true [ v / x ]ᴱ = true +false [ v / x ]ᴱ = false var y [ v / x ]ᴱ = var y [ v / x ]ᴱwhenever (x ≡ⱽ y) addr a [ v / x ]ᴱ = addr a (number y) [ v / x ]ᴱ = number y diff --git a/prototyping/Luau/Syntax.agda b/prototyping/Luau/Syntax.agda index 0278a91e..b1320efb 100644 --- a/prototyping/Luau/Syntax.agda +++ b/prototyping/Luau/Syntax.agda @@ -38,6 +38,12 @@ data BinaryOperator : Set where - : BinaryOperator * : BinaryOperator / : BinaryOperator + < : BinaryOperator + > : BinaryOperator + == : BinaryOperator + ~= : BinaryOperator + <= : BinaryOperator + >= : BinaryOperator data Block (a : Annotated) : Set data Stat (a : Annotated) : Set @@ -54,6 +60,8 @@ data Stat a where data Expr a where nil : Expr a + true : Expr a + false : Expr a var : Var → Expr a addr : Addr → Expr a _$_ : Expr a → Expr a → Expr a @@ -68,6 +76,8 @@ isAnnotatedᴮ : ∀ {a} → Block a → Maybe (Block yes) isAnnotatedᴱ nil = just nil isAnnotatedᴱ (var x) = just (var x) isAnnotatedᴱ (addr a) = just (addr a) +isAnnotatedᴱ true = just true +isAnnotatedᴱ false = just false isAnnotatedᴱ (M $ N) with isAnnotatedᴱ M | isAnnotatedᴱ N isAnnotatedᴱ (M $ N) | just M′ | just N′ = just (M′ $ N′) isAnnotatedᴱ (M $ N) | _ | _ = nothing diff --git a/prototyping/Luau/Syntax/FromJSON.agda b/prototyping/Luau/Syntax/FromJSON.agda index 1cbc0064..41c0c1c9 100644 --- a/prototyping/Luau/Syntax/FromJSON.agda +++ b/prototyping/Luau/Syntax/FromJSON.agda @@ -2,13 +2,13 @@ module Luau.Syntax.FromJSON where -open import Luau.Syntax using (Block; Stat ; Expr; nil; _$_; var; var_∈_; function_is_end; _⟨_⟩; _⟨_⟩∈_; local_←_; return; done; _∙_; maybe; VarDec; number; binexp; BinaryOperator; +; -; *; /) +open import Luau.Syntax using (Block; Stat ; Expr; nil; _$_; var; var_∈_; function_is_end; _⟨_⟩; _⟨_⟩∈_; local_←_; return; done; _∙_; maybe; VarDec; number; binexp; BinaryOperator; +; -; *; /; ==; ~=; <; >; <=; >=) open import Luau.Type.FromJSON using (typeFromJSON) open import Agda.Builtin.List using (List; _∷_; []) +open import Agda.Builtin.Bool using (true; false) open import FFI.Data.Aeson using (Value; Array; Object; object; array; string; fromString; lookup) -open import FFI.Data.Bool using (true; false) open import FFI.Data.Either using (Either; Left; Right) open import FFI.Data.Maybe using (Maybe; nothing; just) open import FFI.Data.Scientific using (toFloat) @@ -59,6 +59,12 @@ binOpFromString "Add" = Right + binOpFromString "Sub" = Right - binOpFromString "Mul" = Right * binOpFromString "Div" = Right / +binOpFromString "CompareEq" = Right == +binOpFromString "CompareNe" = Right ~= +binOpFromString "CompareLt" = Right < +binOpFromString "CompareLe" = Right <= +binOpFromString "CompareGt" = Right > +binOpFromString "CompareGe" = Right >= binOpFromString s = Left ("'" ++ s ++ "' is not a valid operator") varDecFromJSON (object arg) = varDecFromObject arg @@ -116,6 +122,11 @@ exprFromObject obj | just (string "AstExprConstantNumber") with lookup value obj exprFromObject obj | just (string "AstExprConstantNumber") | just (FFI.Data.Aeson.Value.number x) = Right (number (toFloat x)) exprFromObject obj | just (string "AstExprConstantNumber") | just _ = Left "AstExprConstantNumber value is not a number" exprFromObject obj | just (string "AstExprConstantNumber") | nothing = Left "AstExprConstantNumber missing value" +exprFromObject obj | just (string "AstExprConstantBool") with lookup value obj +exprFromObject obj | just (string "AstExprConstantBool") | just (FFI.Data.Aeson.Value.bool true) = Right Expr.true +exprFromObject obj | just (string "AstExprConstantBool") | just (FFI.Data.Aeson.Value.bool false) = Right Expr.false +exprFromObject obj | just (string "AstExprConstantBool") | just _ = Left "AstExprConstantBool value is not a bool" +exprFromObject obj | just (string "AstExprConstantBool") | nothing = Left "AstExprConstantBool missing value" exprFromObject obj | just (string "AstExprBinary") with lookup op obj | lookup left obj | lookup right obj exprFromObject obj | just (string "AstExprBinary") | just o | just l | just r with binOpFromJSON o | exprFromJSON l | exprFromJSON r exprFromObject obj | just (string "AstExprBinary") | just o | just l | just r | Right o′ | Right l′ | Right r′ = Right (binexp l′ o′ r′) diff --git a/prototyping/Luau/Syntax/ToString.agda b/prototyping/Luau/Syntax/ToString.agda index a16af0b4..70ef672d 100644 --- a/prototyping/Luau/Syntax/ToString.agda +++ b/prototyping/Luau/Syntax/ToString.agda @@ -1,7 +1,7 @@ module Luau.Syntax.ToString where open import Agda.Builtin.Float using (primShowFloat) -open import Luau.Syntax using (Block; Stat; Expr; VarDec; FunDec; nil; var; var_∈_; addr; _$_; function_is_end; return; local_←_; _∙_; done; block_is_end; _⟨_⟩; _⟨_⟩∈_; number; BinaryOperator; +; -; *; /; binexp) +open import Luau.Syntax using (Block; Stat; Expr; VarDec; FunDec; nil; var; var_∈_; addr; _$_; function_is_end; return; local_←_; _∙_; done; block_is_end; _⟨_⟩; _⟨_⟩∈_; number; BinaryOperator; +; -; *; /; <; >; ≡; ≅; ≤; ≥; binexp; true; false) open import FFI.Data.String using (String; _++_) open import Luau.Addr.ToString using (addrToString) open import Luau.Type.ToString using (typeToString) @@ -22,6 +22,12 @@ binOpToString + = "+" binOpToString - = "-" binOpToString * = "*" binOpToString / = "/" +binOpToString < = "<" +binOpToString > = ">" +binOpToString ≡ = "==" +binOpToString ≅ = "~=" +binOpToString ≤ = "<=" +binOpToString ≥ = ">=" exprToString′ : ∀ {a} → String → Expr a → String statToString′ : ∀ {a} → String → Stat a → String @@ -45,6 +51,8 @@ exprToString′ lb (block b is B end) = "end)()" exprToString′ lb (number x) = primShowFloat x exprToString′ lb (binexp x op y) = exprToString′ lb x ++ " " ++ binOpToString op ++ " " ++ exprToString′ lb y +exprToString′ lb true = "true" +exprToString′ lb false = "false" statToString′ lb (function F is B end) = "local " ++ funDecToString F ++ lb ++ diff --git a/prototyping/Luau/Type.agda b/prototyping/Luau/Type.agda index ce066233..59c4db66 100644 --- a/prototyping/Luau/Type.agda +++ b/prototyping/Luau/Type.agda @@ -11,6 +11,7 @@ data Type : Set where _⇒_ : Type → Type → Type bot : Type top : Type + boolean : Type number : Type _∪_ : Type → Type → Type _∩_ : Type → Type → Type @@ -23,6 +24,7 @@ lhs nil = nil lhs bot = bot lhs top = top lhs number = number +lhs boolean = boolean rhs : Type → Type rhs (_ ⇒ T) = T @@ -32,6 +34,7 @@ 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 @@ -39,6 +42,7 @@ 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 (λ ()) @@ -49,6 +53,7 @@ nil ≡ᵀ (S ∩ T) = 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) = no (λ ()) bot ≡ᵀ nil = no (λ ()) @@ -56,6 +61,7 @@ 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 (λ ()) @@ -63,20 +69,31 @@ 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 ≡ᵀ (U ⇒ U₁) = no (λ ()) +number ≡ᵀ (T ⇒ U) = no (λ ()) number ≡ᵀ bot = no (λ ()) number ≡ᵀ top = no (λ ()) number ≡ᵀ number = yes refl -number ≡ᵀ (U ∪ U₁) = no (λ ()) -number ≡ᵀ (U ∩ U₁) = no (λ ()) +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)) @@ -87,6 +104,7 @@ number ≡ᵀ (U ∩ U₁) = 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 @@ -108,6 +126,7 @@ data Mode : Set where 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) @@ -125,6 +144,7 @@ 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) diff --git a/prototyping/Luau/Type/FromJSON.agda b/prototyping/Luau/Type/FromJSON.agda index 99089f51..805460c1 100644 --- a/prototyping/Luau/Type/FromJSON.agda +++ b/prototyping/Luau/Type/FromJSON.agda @@ -5,9 +5,9 @@ module Luau.Type.FromJSON where open import Luau.Type using (Type; nil; _⇒_; _∪_; _∩_; top; number) open import Agda.Builtin.List using (List; _∷_; []) +open import Agda.Builtin.Bool using (true; false) open import FFI.Data.Aeson using (Value; Array; Object; object; array; string; fromString; lookup) -open import FFI.Data.Bool using (true; false) open import FFI.Data.Either using (Either; Left; Right) open import FFI.Data.Maybe using (Maybe; nothing; just) open import FFI.Data.String using (String; _++_) diff --git a/prototyping/Luau/Type/ToString.agda b/prototyping/Luau/Type/ToString.agda index a864b84c..958840a7 100644 --- a/prototyping/Luau/Type/ToString.agda +++ b/prototyping/Luau/Type/ToString.agda @@ -1,7 +1,7 @@ module Luau.Type.ToString where open import FFI.Data.String using (String; _++_) -open import Luau.Type using (Type; nil; _⇒_; bot; top; number; _∪_; _∩_; normalizeOptional) +open import Luau.Type using (Type; nil; _⇒_; bot; top; number; boolean; _∪_; _∩_; normalizeOptional) {-# TERMINATING #-} typeToString : Type → String @@ -13,6 +13,7 @@ typeToString (S ⇒ T) = "(" ++ (typeToString S) ++ ") -> " ++ (typeToString T) typeToString bot = "bot" typeToString top = "top" typeToString number = "number" +typeToString boolean = "boolean" typeToString (S ∪ T) with normalizeOptional(S ∪ T) typeToString (S ∪ T) | ((S′ ⇒ T′) ∪ nil) = "(" ++ typeToString (S′ ⇒ T′) ++ ")?" typeToString (S ∪ T) | (S′ ∪ nil) = typeToString S′ ++ "?" diff --git a/prototyping/Luau/TypeCheck.agda b/prototyping/Luau/TypeCheck.agda index 217abce2..39d0420f 100644 --- a/prototyping/Luau/TypeCheck.agda +++ b/prototyping/Luau/TypeCheck.agda @@ -6,12 +6,12 @@ module Luau.TypeCheck (m : Mode) where open import Agda.Builtin.Equality using (_≡_) open import FFI.Data.Maybe using (Maybe; just) -open import Luau.Syntax using (Expr; Stat; Block; yes; nil; addr; number; var; var_∈_; _⟨_⟩∈_; function_is_end; _$_; block_is_end; binexp; local_←_; _∙_; done; return; name) +open import Luau.Syntax using (Expr; Stat; Block; yes; nil; addr; number; true; false; var; var_∈_; _⟨_⟩∈_; function_is_end; _$_; block_is_end; binexp; local_←_; _∙_; done; return; name) open import Luau.Var using (Var) open import Luau.Addr using (Addr) open import Luau.Heap using (Heap; Object; function_is_end) renaming (_[_] to _[_]ᴴ) open import Luau.Value using (addr; val) -open import Luau.Type using (Type; Mode; nil; bot; top; number; _⇒_; tgt) +open import Luau.Type using (Type; Mode; nil; bot; top; number; boolean; _⇒_; tgt) open import Luau.VarCtxt using (VarCtxt; ∅; _⋒_; _↦_; _⊕_↦_; _⊝_) renaming (_[_] to _[_]ⱽ) open import FFI.Data.Vector using (Vector) open import FFI.Data.Maybe using (Maybe; just; nothing) @@ -78,6 +78,16 @@ data _⊢ᴱ_∈_ where ------------------------ Γ ⊢ᴱ (number n) ∈ number + true : ∀ {Γ} → + + ------------------- + Γ ⊢ᴱ true ∈ boolean + + false : ∀ {Γ} → + + ------------------- + Γ ⊢ᴱ false ∈ boolean + app : ∀ {M N T U Γ} → Γ ⊢ᴱ M ∈ T → diff --git a/prototyping/Luau/Value.agda b/prototyping/Luau/Value.agda index 1086d39c..ed51abd7 100644 --- a/prototyping/Luau/Value.agda +++ b/prototyping/Luau/Value.agda @@ -1,16 +1,20 @@ module Luau.Value where +open import Agda.Builtin.Bool using (Bool; true; false) open import Agda.Builtin.Float using (Float) open import Luau.Addr using (Addr) -open import Luau.Syntax using (Block; Expr; nil; addr; number) +open import Luau.Syntax using (Block; Expr; nil; addr; number; true; false) open import Luau.Var using (Var) data Value : Set where nil : Value addr : Addr → Value number : Float → Value + bool : Bool → Value val : ∀ {a} → Value → Expr a val nil = nil val (addr a) = addr a val (number x) = number x +val (bool false) = false +val (bool true) = true diff --git a/prototyping/Luau/Value/ToString.agda b/prototyping/Luau/Value/ToString.agda index 51c3fa78..3421d4cd 100644 --- a/prototyping/Luau/Value/ToString.agda +++ b/prototyping/Luau/Value/ToString.agda @@ -2,10 +2,13 @@ module Luau.Value.ToString where open import Agda.Builtin.String using (String) open import Agda.Builtin.Float using (primShowFloat) -open import Luau.Value using (Value; nil; addr; number) +open import Agda.Builtin.Bool using (true; false) +open import Luau.Value using (Value; nil; addr; number; bool) open import Luau.Addr.ToString using (addrToString) valueToString : Value → String valueToString nil = "nil" valueToString (addr a) = addrToString a valueToString (number x) = primShowFloat x +valueToString (bool false) = "false" +valueToString (bool true) = "true" diff --git a/prototyping/Properties/Step.agda b/prototyping/Properties/Step.agda index 348e6802..545e84c7 100644 --- a/prototyping/Properties/Step.agda +++ b/prototyping/Properties/Step.agda @@ -4,14 +4,15 @@ module Properties.Step where open import Agda.Builtin.Equality using (_≡_; refl) open import Agda.Builtin.Float using (primFloatPlus; primFloatMinus; primFloatTimes; primFloatDiv) +open import Agda.Builtin.Bool using (true; false) open import FFI.Data.Maybe using (just; nothing) open import Luau.Heap using (Heap; _[_]; alloc; ok; function_is_end) -open import Luau.Syntax using (Block; Expr; nil; var; addr; function_is_end; block_is_end; _$_; local_←_; return; done; _∙_; name; fun; arg; number; binexp; +) -open import Luau.OpSem using (_⊢_⟶ᴱ_⊣_; _⊢_⟶ᴮ_⊣_; app₁ ; app₂ ; beta; function; block; return; done; local; subst; binOpEval; evalBinOp; binOp₁; binOp₂) +open import Luau.Syntax using (Block; Expr; nil; var; addr; true; false; function_is_end; block_is_end; _$_; local_←_; return; done; _∙_; name; fun; arg; number; binexp; +; ==; ~=) +open import Luau.OpSem using (_⊢_⟶ᴱ_⊣_; _⊢_⟶ᴮ_⊣_; app₁ ; app₂ ; beta; function; block; return; done; local; subst; binOpNumbers; evalNumOp; binOp₁; binOp₂; evalEqOp; evalNeqOp; binOpEquality; binOpInequality) open import Luau.RuntimeError using (RuntimeErrorᴱ; RuntimeErrorᴮ; FunctionMismatch; BinopMismatch₁; BinopMismatch₂; UnboundVariable; SEGV; app₁; app₂; block; local; return; bin₁; bin₂) open import Luau.RuntimeType using (function; number) open import Luau.Substitution using (_[_/_]ᴮ) -open import Luau.Value using (nil; addr; val; number) +open import Luau.Value using (nil; addr; val; number; bool) open import Properties.Remember using (remember; _,_) data StepResultᴮ {a} (H : Heap a) (B : Block a) : Set @@ -35,6 +36,8 @@ stepᴱ H nil = value nil refl stepᴱ H (var x) = error UnboundVariable stepᴱ H (addr a) = value (addr a) refl stepᴱ H (number x) = value (number x) refl +stepᴱ H (true) = value (bool true) refl +stepᴱ H (false) = value (bool false) refl stepᴱ H (M $ N) with stepᴱ H M stepᴱ H (M $ N) | step H′ M′ D = step H′ (M′ $ N) (app₁ D) stepᴱ H (_ $ N) | value v refl with stepᴱ H N @@ -44,6 +47,7 @@ stepᴱ H (_ $ _) | value (addr a) refl | value w refl | (nothing , p) = error stepᴱ H (_ $ _) | value (addr a) refl | value w refl | (just(function F is B end) , p) = step H (block (fun F) is B [ w / name (arg F) ]ᴮ end) (beta function F is B end w refl p) stepᴱ H (_ $ _) | value nil refl | value w refl = error (FunctionMismatch nil w (λ ())) stepᴱ H (_ $ _) | value (number x) refl | value w refl = error (FunctionMismatch (number x) w (λ ())) +stepᴱ H (_ $ _) | value (bool b) refl | value w refl = error (FunctionMismatch (bool b) w (λ ())) stepᴱ H (M $ N) | value V p | error E = error (app₂ E) stepᴱ H (M $ N) | error E = error (app₁ E) stepᴱ H (block b is B end) with stepᴮ H B @@ -57,11 +61,15 @@ stepᴱ H (binexp M op N) with stepᴱ H M stepᴱ H (binexp M op N) | value v refl with stepᴱ H N stepᴱ H (binexp M op N) | value v refl | step H′ N′ s = step H′ (binexp (val v) op N′) (binOp₂ s) stepᴱ H (binexp M op N) | value v refl | error E = error (bin₂ E) -stepᴱ H (binexp M op N) | value (number m) refl | value (number n) refl = step H (number (evalBinOp m op n)) (binOpEval m n) +stepᴱ H (binexp M == N) | value v refl | value w refl = step H (val (evalEqOp v w)) binOpEquality +stepᴱ H (binexp M ~= N) | value v refl | value w refl = step H (val (evalNeqOp v w)) binOpInequality +stepᴱ H (binexp M op N) | value (number m) refl | value (number n) refl = step H (val (evalNumOp m op n)) binOpNumbers stepᴱ H (binexp M op N) | value nil refl | value w refl = error (BinopMismatch₁ nil w λ ()) stepᴱ H (binexp M op N) | value (addr a) refl | value w refl = error (BinopMismatch₁ (addr a) w λ ()) +stepᴱ H (binexp M op N) | value (bool b) refl | value w refl = error (BinopMismatch₁ (bool b) w λ ()) stepᴱ H (binexp M op N) | value v refl | value nil refl = error (BinopMismatch₂ v nil (λ ())) stepᴱ H (binexp M op N) | value v refl | value (addr a) refl = error (BinopMismatch₂ v (addr a) (λ ())) +stepᴱ H (binexp M op N) | value v refl | value (bool b) refl = error (BinopMismatch₂ v (bool b) (λ ())) stepᴱ H (binexp M op N) | step H′ M′ s = step H′ (binexp M′ op N) (binOp₁ s) stepᴱ H (binexp M op N) | error E = error (bin₁ E) diff --git a/prototyping/Properties/TypeCheck.agda b/prototyping/Properties/TypeCheck.agda index 04f58e67..7b7ee63d 100644 --- a/prototyping/Properties/TypeCheck.agda +++ b/prototyping/Properties/TypeCheck.agda @@ -5,16 +5,17 @@ open import Luau.Type using (Mode) module Properties.TypeCheck (m : Mode) where open import Agda.Builtin.Equality using (_≡_; refl) +open import Agda.Builtin.Bool using (Bool; true; false) open import FFI.Data.Maybe using (Maybe; just; nothing) open import FFI.Data.Either using (Either) -open import Luau.TypeCheck(m) using (_⊢ᴱ_∈_; _⊢ᴮ_∈_; ⊢ᴼ_; ⊢ᴴ_; _⊢ᴴᴱ_▷_∈_; _⊢ᴴᴮ_▷_∈_; nil; var; addr; number; app; function; block; binexp; done; return; local; nothing; orBot) -open import Luau.Syntax using (Block; Expr; yes; nil; var; addr; number; binexp; _$_; function_is_end; block_is_end; _∙_; return; done; local_←_; _⟨_⟩; _⟨_⟩∈_; var_∈_; name; fun; arg) -open import Luau.Type using (Type; nil; top; bot; number; _⇒_; tgt) +open import Luau.TypeCheck(m) using (_⊢ᴱ_∈_; _⊢ᴮ_∈_; ⊢ᴼ_; ⊢ᴴ_; _⊢ᴴᴱ_▷_∈_; _⊢ᴴᴮ_▷_∈_; nil; var; addr; number; true; false; app; function; block; binexp; done; return; local; nothing; orBot) +open import Luau.Syntax using (Block; Expr; yes; nil; var; addr; number; true; false; binexp; _$_; function_is_end; block_is_end; _∙_; return; done; local_←_; _⟨_⟩; _⟨_⟩∈_; var_∈_; name; fun; arg) +open import Luau.Type using (Type; nil; top; bot; number; boolean; _⇒_; tgt) open import Luau.RuntimeType using (RuntimeType; nil; number; function; valueType) open import Luau.VarCtxt using (VarCtxt; ∅; _↦_; _⊕_↦_; _⋒_; _⊝_) renaming (_[_] to _[_]ⱽ) open import Luau.Addr using (Addr) open import Luau.Var using (Var; _≡ⱽ_) -open import Luau.Value using (Value; nil; addr; number; val) +open import Luau.Value using (Value; nil; addr; number; bool; val) open import Luau.Heap using (Heap; Object; function_is_end) renaming (_[_] to _[_]ᴴ) open import Properties.Contradiction using (CONTRADICTION) open import Properties.Dec using (yes; no) @@ -34,6 +35,7 @@ typeOfᴹᴼ (just O) = just (typeOfᴼ O) typeOfⱽ : Heap yes → Value → Maybe Type typeOfⱽ H nil = just nil +typeOfⱽ H (bool b) = just boolean typeOfⱽ H (addr a) = typeOfᴹᴼ (H [ a ]ᴴ) typeOfⱽ H (number n) = just number @@ -44,6 +46,8 @@ typeOfᴱ H Γ nil = nil typeOfᴱ H Γ (var x) = orBot(Γ [ x ]ⱽ) typeOfᴱ H Γ (addr a) = orBot(typeOfᴹᴼ (H [ a ]ᴴ)) typeOfᴱ H Γ (number n) = number +typeOfᴱ H Γ true = boolean +typeOfᴱ H Γ false = boolean typeOfᴱ H Γ (M $ N) = tgt(typeOfᴱ H Γ M) typeOfᴱ H Γ (function f ⟨ var x ∈ S ⟩∈ T is B end) = S ⇒ T typeOfᴱ H Γ (block var b ∈ T is B end) = T @@ -56,6 +60,8 @@ typeOfᴮ H Γ done = nil typeOfᴱⱽ : ∀ {H Γ} v → (typeOfᴱ H Γ (val v) ≡ orBot(typeOfⱽ H v)) typeOfᴱⱽ nil = refl +typeOfᴱⱽ (bool true) = refl +typeOfᴱⱽ (bool false) = refl typeOfᴱⱽ (addr a) = refl typeOfᴱⱽ (number n) = refl @@ -63,6 +69,8 @@ mustBeFunction : ∀ H Γ v → (bot ≢ src (typeOfᴱ H Γ (val v))) → (func mustBeFunction H Γ nil p = CONTRADICTION (p refl) mustBeFunction H Γ (addr a) p = refl mustBeFunction H Γ (number n) p = CONTRADICTION (p refl) +mustBeFunction H Γ (bool true) p = CONTRADICTION (p refl) +mustBeFunction H Γ (bool false) p = CONTRADICTION (p refl) mustBeNumber : ∀ H Γ v → (number ≡ typeOfᴱ H Γ (val v)) → (number ≡ valueType(v)) mustBeNumber H Γ nil () @@ -72,6 +80,8 @@ mustBeNumber H Γ (addr a) p | (just function f ⟨ var x ∈ T ⟩∈ U is B en mustBeNumber H Γ (addr a) p | (nothing , q) with trans p (cong orBot (cong typeOfᴹᴼ q)) mustBeNumber H Γ (addr a) p | nothing , q | () mustBeNumber H Γ (number n) p = refl +mustBeNumber H Γ (bool true) () +mustBeNumber H Γ (bool false) () typeCheckᴱ : ∀ H Γ M → (Γ ⊢ᴱ M ∈ (typeOfᴱ H Γ M)) typeCheckᴮ : ∀ H Γ B → (Γ ⊢ᴮ B ∈ (typeOfᴮ H Γ B)) @@ -80,6 +90,8 @@ typeCheckᴱ H Γ nil = nil typeCheckᴱ H Γ (var x) = var refl typeCheckᴱ H Γ (addr a) = addr (orBot (typeOfᴹᴼ (H [ a ]ᴴ))) typeCheckᴱ H Γ (number n) = number +typeCheckᴱ H Γ true = true +typeCheckᴱ H Γ false = false typeCheckᴱ H Γ (M $ N) = app (typeCheckᴱ H Γ M) (typeCheckᴱ H Γ N) typeCheckᴱ H Γ (function f ⟨ var x ∈ T ⟩∈ U is B end) = function (typeCheckᴮ H (Γ ⊕ x ↦ T) B) typeCheckᴱ H Γ (block var b ∈ T is B end) = block (typeCheckᴮ H Γ B) diff --git a/prototyping/Tests/Interpreter/binary_equality_bools/in.lua b/prototyping/Tests/Interpreter/binary_equality_bools/in.lua new file mode 100644 index 00000000..c9c72dc6 --- /dev/null +++ b/prototyping/Tests/Interpreter/binary_equality_bools/in.lua @@ -0,0 +1 @@ +return true == false diff --git a/prototyping/Tests/Interpreter/binary_equality_bools/out.txt b/prototyping/Tests/Interpreter/binary_equality_bools/out.txt new file mode 100644 index 00000000..c508d536 --- /dev/null +++ b/prototyping/Tests/Interpreter/binary_equality_bools/out.txt @@ -0,0 +1 @@ +false diff --git a/prototyping/Tests/Interpreter/binary_equality_numbers/in.lua b/prototyping/Tests/Interpreter/binary_equality_numbers/in.lua new file mode 100644 index 00000000..4efc68a7 --- /dev/null +++ b/prototyping/Tests/Interpreter/binary_equality_numbers/in.lua @@ -0,0 +1 @@ +return 1 == 1 diff --git a/prototyping/Tests/Interpreter/binary_equality_numbers/out.txt b/prototyping/Tests/Interpreter/binary_equality_numbers/out.txt new file mode 100644 index 00000000..27ba77dd --- /dev/null +++ b/prototyping/Tests/Interpreter/binary_equality_numbers/out.txt @@ -0,0 +1 @@ +true diff --git a/prototyping/Utility/Bool.agda b/prototyping/Utility/Bool.agda new file mode 100644 index 00000000..1afffb0d --- /dev/null +++ b/prototyping/Utility/Bool.agda @@ -0,0 +1,16 @@ +module Utility.Bool where + +open import Agda.Builtin.Bool using (Bool; true; false) + +not : Bool → Bool +not false = true +not true = false + +_or_ : Bool → Bool → Bool +true or _ = true +_ or true = true +_ or _ = false + +_and_ : Bool → Bool → Bool +true and true = true +_ and _ = false