mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 12:42:02 +03:00
test node006
This commit is contained in:
parent
f8f09858e8
commit
525011eaff
@ -239,6 +239,10 @@ jsPrim (NS _ (UN "prim__os")) [] =
|
||||
pure sysos
|
||||
jsPrim x args = throw (InternalError $ "prim not implemented: " ++ (show x))
|
||||
|
||||
tag2es : Either Int String -> String
|
||||
tag2es (Left x) = show x
|
||||
tag2es (Right x) = jsString x
|
||||
|
||||
mutual
|
||||
impExp2es : {auto c : Ref ESs ESSt} -> ImperativeExp -> Core String
|
||||
impExp2es (IEVar n) =
|
||||
@ -255,13 +259,13 @@ mutual
|
||||
jsPrim n !(traverse impExp2es args)
|
||||
impExp2es (IEConstructorHead e) =
|
||||
pure $ !(impExp2es e) ++ ".h"
|
||||
impExp2es (IEConstructorTag t) =
|
||||
pure $ show t
|
||||
impExp2es (IEConstructorTag x) =
|
||||
pure $ tag2es x
|
||||
impExp2es (IEConstructorArg i e) =
|
||||
pure $ !(impExp2es e) ++ ".a" ++ show i
|
||||
impExp2es (IEConstructor h args) =
|
||||
let argPairs = zipWith (\i,a => "a" ++ show i ++ ": " ++ a ) [1..length args] !(traverse impExp2es args)
|
||||
in pure $ "({" ++ showSep ", " (("h:" ++ show h)::argPairs) ++ "})"
|
||||
in pure $ "({" ++ showSep ", " (("h:" ++ tag2es h)::argPairs) ++ "})"
|
||||
impExp2es (IEDelay e) =
|
||||
pure $ "(()=>" ++ !(impExp2es e) ++ ")"
|
||||
impExp2es (IEForce e) =
|
||||
|
@ -22,9 +22,9 @@ mutual
|
||||
| IEPrimFn (PrimFn arity) (Vect arity ImperativeExp)
|
||||
| IEPrimFnExt Name (List ImperativeExp)
|
||||
| IEConstructorHead ImperativeExp
|
||||
| IEConstructorTag Int
|
||||
| IEConstructorTag (Either Int String)
|
||||
| IEConstructorArg Int ImperativeExp
|
||||
| IEConstructor Int (List ImperativeExp)
|
||||
| IEConstructor (Either Int String) (List ImperativeExp)
|
||||
| IEDelay ImperativeExp
|
||||
| IEForce ImperativeExp
|
||||
| IENull
|
||||
@ -218,12 +218,10 @@ mutual
|
||||
do
|
||||
(s, a) <- impListExp args
|
||||
pure (s, IEPrimFnExt p a)
|
||||
impExp (NmCon fc x Nothing args) =
|
||||
throw (InternalError "MknConAlt without tag")
|
||||
impExp (NmCon fc x (Just tag) args) =
|
||||
impExp (NmCon fc x tag args) =
|
||||
do
|
||||
(s, a) <- impListExp args
|
||||
pure (s, IEConstructor tag a)
|
||||
pure (s, IEConstructor (impTag x tag) a)
|
||||
impExp (NmDelay fc t) =
|
||||
do
|
||||
(s, x) <- impExp t
|
||||
@ -242,15 +240,17 @@ mutual
|
||||
impExp (NmCrash fc msg) =
|
||||
pure (ErrorStatement msg, IENull)
|
||||
|
||||
impTag : Name -> Maybe Int -> Either Int String
|
||||
impTag n Nothing = Right $ show n
|
||||
impTag n (Just i) = Left i
|
||||
|
||||
impConAlt : {auto c : Ref Imps ImpSt} -> Name -> ImperativeExp -> NamedConAlt -> Core (ImperativeExp, ImperativeStatement)
|
||||
impConAlt res target (MkNConAlt n (Just tag) args exp) =
|
||||
impConAlt res target (MkNConAlt n tag args exp) =
|
||||
do
|
||||
(s, r) <- impExp exp
|
||||
let nargs = length args
|
||||
let reps = zipWith (\i, n => (n, IEConstructorArg (cast i) target)) [1..nargs] args
|
||||
pure (IEConstructorTag tag, replaceNamesExpS reps $ s <+> MutateStatement res r)
|
||||
impConAlt res target (MkNConAlt n Nothing args exp) =
|
||||
throw (InternalError "MknConAlt without tag")
|
||||
pure (IEConstructorTag (impTag n tag), replaceNamesExpS reps $ s <+> MutateStatement res r)
|
||||
|
||||
impConstAlt : {auto c : Ref Imps ImpSt} -> Name -> NamedConstAlt -> Core (ImperativeExp, ImperativeStatement)
|
||||
impConstAlt res (MkNConstAlt c exp) =
|
||||
|
@ -16,6 +16,12 @@ strangeId' _
|
||||
Main> Bye for now!
|
||||
1/1: Building TypeCase2 (TypeCase2.idr)
|
||||
TypeCase2.idr:5:14--5:17:While processing left hand side of strangeId at TypeCase2.idr:5:1--6:1:
|
||||
Can't match on Nat (Erased argument)
|
||||
TypeCase2.idr:9:5--9:9:While processing left hand side of foo at TypeCase2.idr:9:1--10:1:
|
||||
Can't match on Nat (Erased argument)
|
||||
Can't match on Nat (Erased argument) at
|
||||
5 strangeId {a=Nat} x = x+1
|
||||
^^^
|
||||
|
||||
TypeCase2.idr:9:5--9:8:While processing left hand side of foo at TypeCase2.idr:9:1--10:1:
|
||||
Can't match on Nat (Erased argument) at
|
||||
9 foo Nat = "Nat"
|
||||
^^^
|
||||
|
||||
|
Loading…
Reference in New Issue
Block a user