Start working on the improved type-safe interpreter.
This commit is contained in:
parent
12aedfce92
commit
a3c299b057
99
code/typesafe-interpreter/TypesafeIntrV2.idr
Normal file
99
code/typesafe-interpreter/TypesafeIntrV2.idr
Normal file
|
@ -0,0 +1,99 @@
|
|||
data ExprType
|
||||
= IntType
|
||||
| BoolType
|
||||
| StringType
|
||||
|
||||
repr : ExprType -> Type
|
||||
repr IntType = Int
|
||||
repr BoolType = Bool
|
||||
repr StringType = String
|
||||
|
||||
intBoolImpossible : IntType = BoolType -> Void
|
||||
intBoolImpossible Refl impossible
|
||||
|
||||
intStringImpossible : IntType = StringType -> Void
|
||||
intStringImpossible Refl impossible
|
||||
|
||||
boolStringImpossible : BoolType = StringType -> Void
|
||||
boolStringImpossible Refl impossible
|
||||
|
||||
decEq : (a : ExprType) -> (b : ExprType) -> Dec (a = b)
|
||||
decEq {a = IntType} {b = IntType} = Yes Refl
|
||||
decEq {a = BoolType} {b = BoolType} = Yes Refl
|
||||
decEq {a = StringType} {b = StringType} = Yes Refl
|
||||
decEq {a = IntType} {b = BoolType} = No intBoolImpossible
|
||||
decEq {a = BoolType} {b = IntType} = No $ intBoolImpossible . sym
|
||||
decEq {a = IntType} {b = StringType} = No intStringImpossible
|
||||
decEq {a = StringType} {b = IntType} = No $ intStringImpossible . sym
|
||||
decEq {a = BoolType} {b = StringType} = No boolStringImpossible
|
||||
decEq {a = StringType} {b = BoolType} = No $ boolStringImpossible . sym
|
||||
|
||||
data Op
|
||||
= Add
|
||||
| Subtract
|
||||
| Multiply
|
||||
| Divide
|
||||
|
||||
data Expr
|
||||
= IntLit Int
|
||||
| BoolLit Bool
|
||||
| StringLit String
|
||||
| BinOp Op Expr Expr
|
||||
| IfElse Expr Expr Expr
|
||||
|
||||
data SafeExpr : ExprType -> Type where
|
||||
IntLiteral : Int -> SafeExpr IntType
|
||||
BoolLiteral : Bool -> SafeExpr BoolType
|
||||
StringLiteral : String -> SafeExpr StringType
|
||||
BinOperation : (repr a -> repr b -> repr c) -> SafeExpr a -> SafeExpr b -> SafeExpr c
|
||||
IfThenElse : { t : ExprType } -> SafeExpr BoolType -> SafeExpr t -> SafeExpr t -> SafeExpr t
|
||||
|
||||
typecheckOp : Op -> (a : ExprType) -> (b : ExprType) -> Either String (c : ExprType ** repr a -> repr b -> repr c)
|
||||
typecheckOp Add IntType IntType = Right (IntType ** (+))
|
||||
typecheckOp Subtract IntType IntType = Right (IntType ** (-))
|
||||
typecheckOp Multiply IntType IntType = Right (IntType ** (*))
|
||||
typecheckOp Divide IntType IntType = Right (IntType ** div)
|
||||
typecheckOp _ _ _ = Left "Invalid binary operator application"
|
||||
|
||||
requireBool : (n : ExprType ** SafeExpr n) -> Either String (SafeExpr BoolType)
|
||||
requireBool (BoolType ** e) = Right e
|
||||
requireBool _ = Left "Not a boolean."
|
||||
|
||||
typecheck : Expr -> Either String (n : ExprType ** SafeExpr n)
|
||||
typecheck (IntLit i) = Right (_ ** IntLiteral i)
|
||||
typecheck (BoolLit b) = Right (_ ** BoolLiteral b)
|
||||
typecheck (StringLit s) = Right (_ ** StringLiteral s)
|
||||
typecheck (BinOp o l r) = do
|
||||
(lt ** le) <- typecheck l
|
||||
(rt ** re) <- typecheck r
|
||||
(ot ** f) <- typecheckOp o lt rt
|
||||
pure (_ ** BinOperation f le re)
|
||||
typecheck (IfElse c t e) =
|
||||
do
|
||||
ce <- typecheck c >>= requireBool
|
||||
(tt ** te) <- typecheck t
|
||||
(et ** ee) <- typecheck e
|
||||
case decEq tt et of
|
||||
Yes p => pure (_ ** IfThenElse ce (replace p te) ee)
|
||||
No _ => Left "Incompatible branch types."
|
||||
|
||||
eval : SafeExpr t -> repr t
|
||||
eval (IntLiteral i) = i
|
||||
eval (BoolLiteral b) = b
|
||||
eval (StringLiteral s) = s
|
||||
eval (BinOperation f l r) = f (eval l) (eval r)
|
||||
eval (IfThenElse c t e) = if (eval c) then (eval t) else (eval e)
|
||||
|
||||
resultStr : {t : ExprType} -> repr t -> String
|
||||
resultStr {t=IntType} i = show i
|
||||
resultStr {t=BoolType} b = show b
|
||||
resultStr {t=StringType} s = show s
|
||||
|
||||
tryEval : Expr -> String
|
||||
tryEval ex =
|
||||
case typecheck ex of
|
||||
Left err => "Type error: " ++ err
|
||||
Right (t ** e) => resultStr $ eval {t} e
|
||||
|
||||
main : IO ()
|
||||
main = putStrLn $ tryEval $ BinOp Add (IfElse (BoolLit True) (IntLit 6) (IntLit 7)) (BinOp Multiply (IntLit 160) (IntLit 2))
|
Loading…
Reference in New Issue
Block a user