add file locations to Parser.Syntax
they're immediately thrown away currently. but one step at a time
This commit is contained in:
parent
97f51b4436
commit
7e079a9668
7 changed files with 822 additions and 667 deletions
63
lib/Quox/Loc.idr
Normal file
63
lib/Quox/Loc.idr
Normal file
|
@ -0,0 +1,63 @@
|
||||||
|
||| file locations
|
||||||
|
module Quox.Loc
|
||||||
|
|
||||||
|
import Text.Bounded
|
||||||
|
import Derive.Prelude
|
||||||
|
|
||||||
|
%default total
|
||||||
|
%language ElabReflection
|
||||||
|
|
||||||
|
public export
|
||||||
|
FileName : Type
|
||||||
|
FileName = String
|
||||||
|
|
||||||
|
public export
|
||||||
|
record Loc' where
|
||||||
|
constructor L
|
||||||
|
fname : FileName
|
||||||
|
startLine, startCol, endLine, endCol : Int
|
||||||
|
%name Loc' loc
|
||||||
|
%runElab derive "Loc'" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
public export
|
||||||
|
Loc : Type
|
||||||
|
Loc = Maybe Loc'
|
||||||
|
|
||||||
|
export
|
||||||
|
makeLoc : FileName -> Bounds -> Loc
|
||||||
|
makeLoc fname (MkBounds {startLine, startCol, endLine, endCol}) =
|
||||||
|
Just $ L {fname, startLine, startCol, endLine, endCol}
|
||||||
|
|
||||||
|
export
|
||||||
|
onlyStart : Loc -> Loc
|
||||||
|
onlyStart Nothing = Nothing
|
||||||
|
onlyStart (Just (L fname sl sc _ _)) = Just $ L fname sl sc sl sc
|
||||||
|
|
||||||
|
export
|
||||||
|
onlyEnd : Loc -> Loc
|
||||||
|
onlyEnd Nothing = Nothing
|
||||||
|
onlyEnd (Just (L fname _ _ el ec)) = Just $ L fname el ec el ec
|
||||||
|
|
||||||
|
|
||||||
|
export
|
||||||
|
extend : Loc -> Bounds -> Loc
|
||||||
|
extend Nothing _ = Nothing
|
||||||
|
extend (Just (L fname sl1 sc1 el1 ec1)) (MkBounds sl2 sc2 el2 ec2) =
|
||||||
|
let (sl, sc) = (sl1, sc1) `min` (sl2, sc2)
|
||||||
|
(el, ec) = (el1, ec1) `max` (el2, ec2)
|
||||||
|
in
|
||||||
|
Just $ L fname sl sc el ec
|
||||||
|
|
||||||
|
export
|
||||||
|
extend' : Loc -> Maybe Bounds -> Loc
|
||||||
|
extend' l b = maybe l (extend l) b
|
||||||
|
|
||||||
|
|
||||||
|
public export
|
||||||
|
interface Located a where (.loc) : a -> Loc
|
||||||
|
|
||||||
|
export
|
||||||
|
(.bounds) : Loc -> Maybe Bounds
|
||||||
|
(Just (L {fname, startLine, startCol, endLine, endCol})).bounds =
|
||||||
|
Just $ MkBounds {startLine, startCol, endLine, endCol}
|
||||||
|
(Nothing).bounds = Nothing
|
|
@ -23,15 +23,15 @@ import public Quox.Parser.FromParser.Error as Quox.Parser.FromParser
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 NDefinition : Type
|
NDefinition : Type
|
||||||
NDefinition = (Name, Definition)
|
NDefinition = (Name, Definition)
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 IncludePath : Type
|
IncludePath : Type
|
||||||
IncludePath = List String
|
IncludePath = List String
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 SeenFiles : Type
|
SeenFiles : Type
|
||||||
SeenFiles = SortedSet String
|
SeenFiles = SortedSet String
|
||||||
|
|
||||||
|
|
||||||
|
@ -39,26 +39,29 @@ public export
|
||||||
data StateTag = NS | SEEN
|
data StateTag = NS | SEEN
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 FromParserPure : List (Type -> Type)
|
FromParserPure : List (Type -> Type)
|
||||||
FromParserPure =
|
FromParserPure =
|
||||||
[Except Error, StateL DEFS Definitions, StateL NS Mods]
|
[Except Error, StateL DEFS Definitions, StateL NS Mods]
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 FromParserEff : List (Type -> Type)
|
FromParserEff : List (Type -> Type)
|
||||||
FromParserEff =
|
FromParserEff =
|
||||||
[Except Error, StateL DEFS Definitions, StateL NS Mods,
|
[Except Error, StateL DEFS Definitions, StateL NS Mods,
|
||||||
Reader IncludePath, StateL SEEN SeenFiles, IO]
|
Reader IncludePath, StateL SEEN SeenFiles, IO]
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 FromParser : Type -> Type
|
FromParser : Type -> Type
|
||||||
FromParser = Eff FromParserEff
|
FromParser = Eff FromParserEff
|
||||||
|
|
||||||
|
-- [todo] put the locs in the core ast, obv
|
||||||
|
|
||||||
|
|
||||||
parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
|
parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
|
||||||
(xs : Context' BName n)
|
(xs : Context' PatVar n)
|
||||||
private
|
private
|
||||||
fromBaseName : PBaseName -> m a
|
fromBaseName : PBaseName -> m a
|
||||||
fromBaseName x = maybe (f $ MakePName [<] x) b $ Context.find (== Just x) xs
|
fromBaseName x = maybe (f $ MakePName [<] x) b $
|
||||||
|
Context.find (\y => y.name == Just x) xs
|
||||||
|
|
||||||
private
|
private
|
||||||
fromName : PName -> m a
|
fromName : PName -> m a
|
||||||
|
@ -66,14 +69,14 @@ parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
|
||||||
|
|
||||||
export
|
export
|
||||||
fromPDimWith : Has (Except Error) fs =>
|
fromPDimWith : Has (Except Error) fs =>
|
||||||
Context' BName d -> PDim -> Eff fs (Dim d)
|
Context' PatVar d -> PDim -> Eff fs (Dim d)
|
||||||
fromPDimWith ds (K e) = pure $ K e
|
fromPDimWith ds (K e _) = pure $ K e
|
||||||
fromPDimWith ds (V i) =
|
fromPDimWith ds (V i _) =
|
||||||
fromBaseName (pure . B) (const $ throw $ DimNotInScope i) ds i
|
fromBaseName (pure . B) (const $ throw $ DimNotInScope i) ds i
|
||||||
|
|
||||||
private
|
private
|
||||||
avoidDim : Has (Except Error) fs =>
|
avoidDim : Has (Except Error) fs =>
|
||||||
Context' BName d -> PName -> Eff fs Name
|
Context' PatVar d -> PName -> Eff fs Name
|
||||||
avoidDim ds x =
|
avoidDim ds x =
|
||||||
fromName (const $ throw $ DimNameInTerm x.base) (pure . fromPName) ds x
|
fromName (const $ throw $ DimNameInTerm x.base) (pure . fromPName) ds x
|
||||||
|
|
||||||
|
@ -88,97 +91,112 @@ resolveName ns x =
|
||||||
| _ => throw $ TermNotInScope x
|
| _ => throw $ TermNotInScope x
|
||||||
resolveName ns x
|
resolveName ns x
|
||||||
|
|
||||||
|
export
|
||||||
|
fromPatVar : PatVar -> BaseName
|
||||||
|
fromPatVar (Unused _) = Unused
|
||||||
|
fromPatVar (PV x _) = UN x
|
||||||
|
|
||||||
|
export
|
||||||
|
fromPQty : PQty -> Qty
|
||||||
|
fromPQty (PQ q _) = q
|
||||||
|
|
||||||
|
export
|
||||||
|
fromPTagVal : PTagVal -> TagVal
|
||||||
|
fromPTagVal (PT t _) = t
|
||||||
|
|
||||||
|
|
||||||
mutual
|
mutual
|
||||||
export
|
export
|
||||||
fromPTermWith : Context' BName d -> Context' BName n ->
|
fromPTermWith : Context' PatVar d -> Context' PatVar n ->
|
||||||
PTerm -> Eff FromParserPure (Term d n)
|
PTerm -> Eff FromParserPure (Term d n)
|
||||||
fromPTermWith ds ns t0 = case t0 of
|
fromPTermWith ds ns t0 = case t0 of
|
||||||
TYPE k =>
|
TYPE k _ =>
|
||||||
pure $ TYPE $ k
|
pure $ TYPE k
|
||||||
|
|
||||||
Pi pi x s t =>
|
Pi pi x s t _ =>
|
||||||
Pi pi <$> fromPTermWith ds ns s
|
Pi (fromPQty pi)
|
||||||
<*> fromPTermTScope ds ns [< x] t
|
<$> fromPTermWith ds ns s
|
||||||
|
<*> fromPTermTScope ds ns [< x] t
|
||||||
|
|
||||||
Lam x s =>
|
Lam x s _ =>
|
||||||
Lam <$> fromPTermTScope ds ns [< x] s
|
Lam <$> fromPTermTScope ds ns [< x] s
|
||||||
|
|
||||||
s :@ t =>
|
App s t _ =>
|
||||||
map E $ (:@) <$> fromPTermElim ds ns s <*> fromPTermWith ds ns t
|
map E $ (:@) <$> fromPTermElim ds ns s <*> fromPTermWith ds ns t
|
||||||
|
|
||||||
Sig x s t =>
|
Sig x s t _ =>
|
||||||
Sig <$> fromPTermWith ds ns s <*> fromPTermTScope ds ns [< x] t
|
Sig <$> fromPTermWith ds ns s <*> fromPTermTScope ds ns [< x] t
|
||||||
|
|
||||||
Pair s t =>
|
Pair s t _ =>
|
||||||
Pair <$> fromPTermWith ds ns s <*> fromPTermWith ds ns t
|
Pair <$> fromPTermWith ds ns s <*> fromPTermWith ds ns t
|
||||||
|
|
||||||
Case pi pair (r, ret) (CasePair (x, y) body) =>
|
Case pi pair (r, ret) (CasePair (x, y) body _) _ =>
|
||||||
map E $ CasePair pi
|
map E $ CasePair (fromPQty pi)
|
||||||
<$> fromPTermElim ds ns pair
|
<$> fromPTermElim ds ns pair
|
||||||
<*> fromPTermTScope ds ns [< r] ret
|
<*> fromPTermTScope ds ns [< r] ret
|
||||||
<*> fromPTermTScope ds ns [< x, y] body
|
<*> fromPTermTScope ds ns [< x, y] body
|
||||||
|
|
||||||
Case pi tag (r, ret) (CaseEnum arms) =>
|
Case pi tag (r, ret) (CaseEnum arms _) _ =>
|
||||||
map E $ CaseEnum pi
|
map E $ CaseEnum (fromPQty pi)
|
||||||
<$> fromPTermElim ds ns tag
|
<$> fromPTermElim ds ns tag
|
||||||
<*> fromPTermTScope ds ns [< r] ret
|
<*> fromPTermTScope ds ns [< r] ret
|
||||||
<*> assert_total fromPTermEnumArms ds ns arms
|
<*> assert_total fromPTermEnumArms ds ns arms
|
||||||
|
|
||||||
Nat => pure Nat
|
Nat _ => pure Nat
|
||||||
Zero => pure Zero
|
Zero _ => pure Zero
|
||||||
Succ n => [|Succ $ fromPTermWith ds ns n|]
|
Succ n _ => [|Succ $ fromPTermWith ds ns n|]
|
||||||
|
|
||||||
Case pi nat (r, ret) (CaseNat zer (s, pi', ih, suc)) =>
|
Case pi nat (r, ret) (CaseNat zer (s, pi', ih, suc) _) _ =>
|
||||||
map E $ CaseNat pi pi'
|
map E $ CaseNat (fromPQty pi) (fromPQty pi')
|
||||||
<$> fromPTermElim ds ns nat
|
<$> fromPTermElim ds ns nat
|
||||||
<*> fromPTermTScope ds ns [< r] ret
|
<*> fromPTermTScope ds ns [< r] ret
|
||||||
<*> fromPTermWith ds ns zer
|
<*> fromPTermWith ds ns zer
|
||||||
<*> fromPTermTScope ds ns [< s, ih] suc
|
<*> fromPTermTScope ds ns [< s, ih] suc
|
||||||
|
|
||||||
Enum strs =>
|
Enum strs _ =>
|
||||||
let set = SortedSet.fromList strs in
|
let set = SortedSet.fromList strs in
|
||||||
if length strs == length (SortedSet.toList set) then
|
if length strs == length (SortedSet.toList set) then
|
||||||
pure $ Enum set
|
pure $ Enum set
|
||||||
else
|
else
|
||||||
throw $ DuplicatesInEnum strs
|
throw $ DuplicatesInEnum strs
|
||||||
|
|
||||||
Tag str => pure $ Tag str
|
Tag str _ => pure $ Tag str
|
||||||
|
|
||||||
Eq (i, ty) s t =>
|
Eq (i, ty) s t _ =>
|
||||||
Eq <$> fromPTermDScope ds ns [< i] ty
|
Eq <$> fromPTermDScope ds ns [< i] ty
|
||||||
<*> fromPTermWith ds ns s
|
<*> fromPTermWith ds ns s
|
||||||
<*> fromPTermWith ds ns t
|
<*> fromPTermWith ds ns t
|
||||||
|
|
||||||
DLam i s =>
|
DLam i s _ =>
|
||||||
DLam <$> fromPTermDScope ds ns [< i] s
|
DLam <$> fromPTermDScope ds ns [< i] s
|
||||||
|
|
||||||
BOX q ty => BOX q <$> fromPTermWith ds ns ty
|
DApp s p _ =>
|
||||||
|
map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p
|
||||||
|
|
||||||
Box val => Box <$> fromPTermWith ds ns val
|
BOX q ty _ => BOX (fromPQty q) <$> fromPTermWith ds ns ty
|
||||||
|
|
||||||
Case pi box (r, ret) (CaseBox b body) =>
|
Box val _ => Box <$> fromPTermWith ds ns val
|
||||||
map E $ CaseBox pi
|
|
||||||
|
Case pi box (r, ret) (CaseBox b body _) _ =>
|
||||||
|
map E $ CaseBox (fromPQty pi)
|
||||||
<$> fromPTermElim ds ns box
|
<$> fromPTermElim ds ns box
|
||||||
<*> fromPTermTScope ds ns [< r] ret
|
<*> fromPTermTScope ds ns [< r] ret
|
||||||
<*> fromPTermTScope ds ns [< b] body
|
<*> fromPTermTScope ds ns [< b] body
|
||||||
|
|
||||||
s :% p =>
|
V x _ =>
|
||||||
map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p
|
|
||||||
|
|
||||||
V x =>
|
|
||||||
fromName (pure . E . B) (resolveName !(getAt NS) <=< avoidDim ds) ns x
|
fromName (pure . E . B) (resolveName !(getAt NS) <=< avoidDim ds) ns x
|
||||||
|
|
||||||
s :# a =>
|
Ann s a _ =>
|
||||||
map E $ (:#) <$> fromPTermWith ds ns s <*> fromPTermWith ds ns a
|
map E $ (:#) <$> fromPTermWith ds ns s <*> fromPTermWith ds ns a
|
||||||
|
|
||||||
Coe (i, ty) p q val =>
|
Coe (i, ty) p q val _ =>
|
||||||
map E $ Coe
|
map E $ Coe
|
||||||
<$> fromPTermDScope ds ns [< i] ty
|
<$> fromPTermDScope ds ns [< i] ty
|
||||||
<*> fromPDimWith ds p
|
<*> fromPDimWith ds p
|
||||||
<*> fromPDimWith ds q
|
<*> fromPDimWith ds q
|
||||||
<*> fromPTermWith ds ns val
|
<*> fromPTermWith ds ns val
|
||||||
|
|
||||||
Comp (i, ty) p q val r (j0, val0) (j1, val1) =>
|
Comp (i, ty) p q val r (j0, val0) (j1, val1) _ =>
|
||||||
map E $ CompH'
|
map E $ CompH'
|
||||||
<$> fromPTermDScope ds ns [< i] ty
|
<$> fromPTermDScope ds ns [< i] ty
|
||||||
<*> fromPDimWith ds p
|
<*> fromPDimWith ds p
|
||||||
|
@ -189,46 +207,42 @@ mutual
|
||||||
<*> fromPTermDScope ds ns [< j1] val1
|
<*> fromPTermDScope ds ns [< j1] val1
|
||||||
|
|
||||||
private
|
private
|
||||||
fromPTermEnumArms : Context' BName d -> Context' BName n ->
|
fromPTermEnumArms : Context' PatVar d -> Context' PatVar n ->
|
||||||
List (TagVal, PTerm) ->
|
List (PTagVal, PTerm) ->
|
||||||
Eff FromParserPure (CaseEnumArms d n)
|
Eff FromParserPure (CaseEnumArms d n)
|
||||||
fromPTermEnumArms ds ns =
|
fromPTermEnumArms ds ns =
|
||||||
map SortedMap.fromList . traverse (traverse $ fromPTermWith ds ns)
|
map SortedMap.fromList .
|
||||||
|
traverse (bitraverse (pure . fromPTagVal) (fromPTermWith ds ns))
|
||||||
|
|
||||||
private
|
private
|
||||||
fromPTermElim : Context' BName d -> Context' BName n ->
|
fromPTermElim : Context' PatVar d -> Context' PatVar n ->
|
||||||
PTerm -> Eff FromParserPure (Elim d n)
|
PTerm -> Eff FromParserPure (Elim d n)
|
||||||
fromPTermElim ds ns e =
|
fromPTermElim ds ns e =
|
||||||
case !(fromPTermWith ds ns e) of
|
case !(fromPTermWith ds ns e) of
|
||||||
E e => pure e
|
E e => pure e
|
||||||
t => let ctx = MkNameContexts (map name ds) (map name ns) in
|
t => let ctx = MkNameContexts (map fromPatVar ds) (map fromPatVar ns) in
|
||||||
throw $ AnnotationNeeded ctx t
|
throw $ AnnotationNeeded ctx t
|
||||||
where
|
|
||||||
name : BName -> BaseName
|
|
||||||
name = maybe Unused UN
|
|
||||||
|
|
||||||
-- [todo] use SN if the var is named but still unused
|
-- [todo] use SN if the var is named but still unused
|
||||||
private
|
private
|
||||||
fromPTermTScope : {s : Nat} -> Context' BName d -> Context' BName n ->
|
fromPTermTScope : {s : Nat} -> Context' PatVar d -> Context' PatVar n ->
|
||||||
SnocVect s BName -> PTerm ->
|
SnocVect s PatVar -> PTerm ->
|
||||||
Eff FromParserPure (ScopeTermN s d n)
|
Eff FromParserPure (ScopeTermN s d n)
|
||||||
fromPTermTScope ds ns xs t =
|
fromPTermTScope ds ns xs t =
|
||||||
if all isNothing xs then
|
if all isUnused xs then
|
||||||
SN <$> fromPTermWith ds ns t
|
SN <$> fromPTermWith ds ns t
|
||||||
else
|
else
|
||||||
ST (fromSnocVect $ map (maybe Unused UN) xs)
|
ST (fromSnocVect $ map fromPatVar xs) <$> fromPTermWith ds (ns ++ xs) t
|
||||||
<$> fromPTermWith ds (ns ++ xs) t
|
|
||||||
|
|
||||||
private
|
private
|
||||||
fromPTermDScope : {s : Nat} -> Context' BName d -> Context' BName n ->
|
fromPTermDScope : {s : Nat} -> Context' PatVar d -> Context' PatVar n ->
|
||||||
SnocVect s BName -> PTerm ->
|
SnocVect s PatVar -> PTerm ->
|
||||||
Eff FromParserPure (DScopeTermN s d n)
|
Eff FromParserPure (DScopeTermN s d n)
|
||||||
fromPTermDScope ds ns xs t =
|
fromPTermDScope ds ns xs t =
|
||||||
if all isNothing xs then
|
if all isUnused xs then
|
||||||
SN <$> fromPTermWith ds ns t
|
SN <$> fromPTermWith ds ns t
|
||||||
else
|
else
|
||||||
DST (fromSnocVect $ map (maybe Unused UN) xs)
|
DST (fromSnocVect $ map fromPatVar xs) <$> fromPTermWith (ds ++ xs) ns t
|
||||||
<$> fromPTermWith (ds ++ xs) ns t
|
|
||||||
|
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
|
@ -237,30 +251,26 @@ fromPTerm = fromPTermWith [<] [<]
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
globalPQty : Has (Except Error) fs =>
|
globalPQty : (q : Qty) -> Eff [Except Error] (So $ isGlobal q)
|
||||||
(q : Qty) -> Eff fs (So $ isGlobal q)
|
|
||||||
globalPQty pi = case choose $ isGlobal pi of
|
globalPQty pi = case choose $ isGlobal pi of
|
||||||
Left y => pure y
|
Left y => pure y
|
||||||
Right _ => throw $ QtyNotGlobal pi
|
Right _ => throw $ QtyNotGlobal pi
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
fromPBaseNameNS : Has (StateL NS Mods) fs => PBaseName -> Eff fs Name
|
fromPBaseNameNS : PBaseName -> Eff [StateL NS Mods] Name
|
||||||
fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name
|
fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name
|
||||||
|
|
||||||
private
|
private
|
||||||
injTC : (Has (StateL DEFS Definitions) fs, Has (Except Error) fs) =>
|
injTC : TC a -> Eff FromParserPure a
|
||||||
TC a -> Eff fs a
|
|
||||||
injTC act = rethrow $ mapFst WrapTypeError $ runTC !(getAt DEFS) act
|
injTC act = rethrow $ mapFst WrapTypeError $ runTC !(getAt DEFS) act
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
fromPDef : (Has (StateL DEFS Definitions) fs,
|
fromPDef : PDefinition -> Eff FromParserPure NDefinition
|
||||||
Has (StateL NS Mods) fs,
|
fromPDef (MkPDef qty pname ptype pterm _) = do
|
||||||
Has (Except Error) fs) =>
|
name <- lift $ fromPBaseNameNS pname
|
||||||
PDefinition -> Eff fs NDefinition
|
let qty = fromPQty qty
|
||||||
fromPDef (MkPDef qty pname ptype pterm) = do
|
qtyGlobal <- lift $ globalPQty qty
|
||||||
name <- fromPBaseNameNS pname
|
|
||||||
qtyGlobal <- globalPQty qty
|
|
||||||
let gqty = Element qty qtyGlobal
|
let gqty = Element qty qtyGlobal
|
||||||
let sqty = globalToSubj gqty
|
let sqty = globalToSubj gqty
|
||||||
type <- lift $ traverse fromPTerm ptype
|
type <- lift $ traverse fromPTerm ptype
|
||||||
|
@ -280,21 +290,18 @@ fromPDef (MkPDef qty pname ptype pterm) = do
|
||||||
pure (name, def)
|
pure (name, def)
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
fromPDecl : (Has (StateL DEFS Definitions) fs,
|
fromPDecl : PDecl -> Eff FromParserPure (List NDefinition)
|
||||||
Has (StateL NS Mods) fs,
|
|
||||||
Has (Except Error) fs) =>
|
|
||||||
PDecl -> Eff fs (List NDefinition)
|
|
||||||
fromPDecl (PDef def) = singleton <$> fromPDef def
|
fromPDecl (PDef def) = singleton <$> fromPDef def
|
||||||
fromPDecl (PNs ns) =
|
fromPDecl (PNs ns) =
|
||||||
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
|
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
|
||||||
|
|
||||||
|
|
||||||
|
public export
|
||||||
|
LoadFile : List (Type -> Type)
|
||||||
|
LoadFile = [IO, StateL SEEN SeenFiles, Reader IncludePath, Except Error]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
loadFile : (Has IO fs,
|
loadFile : String -> Eff LoadFile (Maybe String)
|
||||||
Has (StateL SEEN SeenFiles) fs,
|
|
||||||
Has (Reader IncludePath) fs,
|
|
||||||
Has (Except Error) fs) =>
|
|
||||||
String -> Eff fs (Maybe String)
|
|
||||||
loadFile file =
|
loadFile file =
|
||||||
if contains file !(getAt SEEN) then
|
if contains file !(getAt SEEN) then
|
||||||
pure Nothing
|
pure Nothing
|
||||||
|
@ -305,27 +312,21 @@ loadFile file =
|
||||||
Right res => modifyAt SEEN (insert file) $> Just res
|
Right res => modifyAt SEEN (insert file) $> Just res
|
||||||
Left err => throw $ LoadError ifile err
|
Left err => throw $ LoadError ifile err
|
||||||
|
|
||||||
parameters {auto _ : (Has IO fs,
|
mutual
|
||||||
Has (StateL SEEN SeenFiles) fs,
|
|
||||||
Has (Reader IncludePath) fs,
|
|
||||||
Has (Except Error) fs,
|
|
||||||
Has (StateL DEFS Definitions) fs,
|
|
||||||
Has (StateL NS Mods) fs)}
|
|
||||||
mutual
|
|
||||||
export covering
|
export covering
|
||||||
loadProcessFile : String -> Eff fs (List NDefinition)
|
loadProcessFile : String -> FromParser (List NDefinition)
|
||||||
loadProcessFile file =
|
loadProcessFile file =
|
||||||
case !(loadFile file) of
|
case !(lift $ loadFile file) of
|
||||||
Just inp => do
|
Just inp => do
|
||||||
tl <- either (throw . WrapParseError file) pure $ lexParseInput inp
|
tl <- either (throw . WrapParseError file) pure $ lexParseInput file inp
|
||||||
concat <$> traverse fromPTopLevel tl
|
concat <$> traverse fromPTopLevel tl
|
||||||
Nothing => pure []
|
Nothing => pure []
|
||||||
|
|
||||||
||| populates the `defs` field of the state
|
||| populates the `defs` field of the state
|
||||||
export covering
|
export covering
|
||||||
fromPTopLevel : PTopLevel -> Eff fs (List NDefinition)
|
fromPTopLevel : PTopLevel -> FromParser (List NDefinition)
|
||||||
fromPTopLevel (PD decl) = fromPDecl decl
|
fromPTopLevel (PD decl) = lift $ fromPDecl decl
|
||||||
fromPTopLevel (PLoad file) = loadProcessFile file
|
fromPTopLevel (PLoad file _) = loadProcessFile file
|
||||||
|
|
||||||
export
|
export
|
||||||
fromParserPure : Definitions ->
|
fromParserPure : Definitions ->
|
||||||
|
|
|
@ -26,12 +26,27 @@ data Error =
|
||||||
%hide Lexer.Error
|
%hide Lexer.Error
|
||||||
%runElab derive "Error" [Show]
|
%runElab derive "Error" [Show]
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
lexParseWith : {c : Bool} -> Grammar c a -> String -> Either Error a
|
lexParseWith : {c : Bool} -> Grammar c a -> String -> Either Error a
|
||||||
lexParseWith grm input = do
|
lexParseWith grm input = do
|
||||||
toks <- mapFst LexError $ lex input
|
toks <- mapFst LexError $ lex input
|
||||||
bimap ParseError fst $ parse (grm <* eof) toks
|
bimap ParseError fst $ parse (grm <* eof) toks
|
||||||
|
|
||||||
|
export
|
||||||
|
withLoc : {c : Bool} -> FileName -> (Grammar c (Loc -> a)) -> Grammar c a
|
||||||
|
withLoc fname act = bounds act <&> \res =>
|
||||||
|
if res.isIrrelevant then res.val Nothing
|
||||||
|
else res.val $ makeLoc fname res.bounds
|
||||||
|
|
||||||
|
export
|
||||||
|
defLoc : FileName -> (Loc -> a) -> Grammar False a
|
||||||
|
defLoc fname f = position <&> f . makeLoc fname
|
||||||
|
|
||||||
|
export
|
||||||
|
unused : FileName -> Grammar False PatVar
|
||||||
|
unused fname = defLoc fname Unused
|
||||||
|
|
||||||
|
|
||||||
||| reserved token, like punctuation or keywords etc
|
||| reserved token, like punctuation or keywords etc
|
||||||
export
|
export
|
||||||
|
@ -114,26 +129,33 @@ dimConst = terminalMatchN "dimension constant"
|
||||||
|
|
||||||
||| quantity (0, 1, or ω)
|
||| quantity (0, 1, or ω)
|
||||||
export
|
export
|
||||||
qty : Grammar True Qty
|
qtyVal : Grammar True Qty
|
||||||
qty = terminalMatchN "quantity"
|
qtyVal = terminalMatchN "quantity"
|
||||||
[(`(Nat 0), `(Zero)), (`(Nat 1), `(One)), (`(Reserved "ω"), `(Any))]
|
[(`(Nat 0), `(Zero)), (`(Nat 1), `(One)), (`(Reserved "ω"), `(Any))]
|
||||||
|
|
||||||
|
|
||||||
|
||| quantity (0, 1, or ω)
|
||||||
|
export
|
||||||
|
qty : FileName -> Grammar True PQty
|
||||||
|
qty fname = withLoc fname [|PQ qtyVal|]
|
||||||
|
|
||||||
|
|
||||||
||| pattern var (unqualified name or _)
|
||| pattern var (unqualified name or _)
|
||||||
export
|
export
|
||||||
patVar : Grammar True BName
|
patVar : FileName -> Grammar True PatVar
|
||||||
patVar = [|Just baseName|] <|> Nothing <$ res "_"
|
patVar fname = withLoc fname $
|
||||||
|
[|PV baseName|] <|> Unused <$ res "_"
|
||||||
|
|
||||||
|
|
||||||
||| dimension (without `@` prefix)
|
||| dimension (without `@` prefix)
|
||||||
export
|
export
|
||||||
dim : Grammar True PDim
|
dim : FileName -> Grammar True PDim
|
||||||
dim = [|K dimConst|] <|> [|V baseName|]
|
dim fname = withLoc fname $ [|K dimConst|] <|> [|V baseName|]
|
||||||
|
|
||||||
||| dimension argument (with @)
|
||| dimension argument (with @)
|
||||||
export
|
export
|
||||||
dimArg : Grammar True PDim
|
dimArg : FileName -> Grammar True PDim
|
||||||
dimArg = resC "@" *> mustWork dim
|
dimArg fname = do resC "@"; mustWork $ dim fname
|
||||||
|
|
||||||
|
|
||||||
delim : (o, c : String) -> (0 _ : IsReserved o) => (0 _ : IsReserved c) =>
|
delim : (o, c : String) -> (0 _ : IsReserved o) => (0 _ : IsReserved c) =>
|
||||||
|
@ -168,313 +190,369 @@ enumType = delimSep "{" "}" "," bareTag
|
||||||
|
|
||||||
||| e.g. `case` or `case 1.`
|
||| e.g. `case` or `case 1.`
|
||||||
export
|
export
|
||||||
caseIntro : Grammar True Qty
|
caseIntro : FileName -> Grammar True PQty
|
||||||
caseIntro =
|
caseIntro fname =
|
||||||
Zero <$ res "case0"
|
withLoc fname (PQ Zero <$ res "case0")
|
||||||
<|> One <$ res "case1"
|
<|> withLoc fname (PQ One <$ res "case1")
|
||||||
<|> Any <$ res "caseω"
|
<|> withLoc fname (PQ Any <$ res "caseω")
|
||||||
<|> delim "case" "." qty
|
<|> delim "case" "." (qty fname)
|
||||||
|
|
||||||
export
|
export
|
||||||
qtyPatVar : Grammar True (Qty, BName)
|
qtyPatVar : FileName -> Grammar True (PQty, PatVar)
|
||||||
qtyPatVar = [|(,) qty (needRes "." *> patVar)|]
|
qtyPatVar fname = [|(,) (qty fname) (needRes "." *> patVar fname)|]
|
||||||
|
|
||||||
|
|
||||||
|
export
|
||||||
|
ptag : FileName -> Grammar True PTagVal
|
||||||
|
ptag fname = withLoc fname $ [|PT tag|]
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PCasePat =
|
data PCasePat =
|
||||||
PPair BName BName
|
PPair PatVar PatVar Loc
|
||||||
| PTag TagVal
|
| PTag PTagVal Loc
|
||||||
| PZero
|
| PZero Loc
|
||||||
| PSucc BName Qty BName
|
| PSucc PatVar PQty PatVar Loc
|
||||||
| PBox BName
|
| PBox PatVar Loc
|
||||||
%runElab derive "PCasePat" [Eq, Ord, Show]
|
%runElab derive "PCasePat" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
export
|
||||||
|
Located PCasePat where
|
||||||
|
(PPair _ _ loc).loc = loc
|
||||||
|
(PTag _ loc).loc = loc
|
||||||
|
(PZero loc).loc = loc
|
||||||
|
(PSucc _ _ _ loc).loc = loc
|
||||||
|
(PBox _ loc).loc = loc
|
||||||
|
|
||||||
||| either `zero` or `0`
|
||| either `zero` or `0`
|
||||||
export
|
export
|
||||||
zeroPat : Grammar True ()
|
zeroPat : Grammar True ()
|
||||||
zeroPat = resC "zero" <|> terminal "expected '0'" (guard . (== Nat 0))
|
zeroPat = resC "zero" <|> terminal "expected '0'" (guard . (== Nat 0))
|
||||||
|
|
||||||
export
|
export
|
||||||
casePat : Grammar True PCasePat
|
casePat : FileName -> Grammar True PCasePat
|
||||||
casePat =
|
casePat fname = withLoc fname $
|
||||||
delim "(" ")" [|PPair patVar (needRes "," *> patVar)|]
|
delim "(" ")" [|PPair (patVar fname) (needRes "," *> patVar fname)|]
|
||||||
<|> [|PTag tag|]
|
<|> [|PTag (ptag fname)|]
|
||||||
<|> PZero <$ zeroPat
|
<|> PZero <$ zeroPat
|
||||||
<|> do p <- resC "succ" *> patVar
|
<|> do p <- resC "succ" *> patVar fname
|
||||||
ih <- option (Zero, Nothing) $ resC "," *> qtyPatVar
|
ih <- resC "," *> qtyPatVar fname
|
||||||
|
<|> [|(,) (defLoc fname $ PQ Zero) (unused fname)|]
|
||||||
pure $ PSucc p (fst ih) (snd ih)
|
pure $ PSucc p (fst ih) (snd ih)
|
||||||
<|> delim "[" "]" [|PBox patVar|]
|
<|> delim "[" "]" [|PBox (patVar fname)|]
|
||||||
<|> fatalError "invalid pattern"
|
<|> fatalError "invalid pattern"
|
||||||
|
|
||||||
|
export covering
|
||||||
|
term : FileName -> Grammar True PTerm
|
||||||
|
-- defined after all the subterm parsers
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
term : Grammar True PTerm
|
typeLine : FileName -> Grammar True (PatVar, PTerm)
|
||||||
|
typeLine fname = do
|
||||||
export covering
|
|
||||||
typeLine : Grammar True (BName, PTerm)
|
|
||||||
typeLine = do
|
|
||||||
resC "["
|
resC "["
|
||||||
mustWork $ do
|
mustWork $ do
|
||||||
i <- option Nothing $ patVar <* resC "⇒"
|
i <- patVar fname <* resC "⇒" <|> unused fname
|
||||||
t <- term <* needRes "]"
|
t <- term fname <* needRes "]"
|
||||||
pure (i, t)
|
pure (i, t)
|
||||||
|
|
||||||
||| box term `[t]` or type `[π.A]`
|
||| box term `[t]` or type `[π.A]`
|
||||||
export covering
|
export covering
|
||||||
boxTerm : Grammar True PTerm
|
boxTerm : FileName -> Grammar True PTerm
|
||||||
boxTerm = do
|
boxTerm fname = withLoc fname $ do
|
||||||
res "["; commit
|
res "["; commit
|
||||||
q <- optional $ qty <* res "."
|
q <- optional $ qty fname <* res "."
|
||||||
t <- mustWork $ term <* needRes "]"
|
t <- mustWork $ term fname <* needRes "]"
|
||||||
pure $ maybe (Box t) (\q => BOX q t) q
|
pure $ maybe (Box t) (\q => BOX q t) q
|
||||||
|
|
||||||
||| tuple term like `(a, b)`, or parenthesised single term.
|
||| tuple term like `(a, b)`, or parenthesised single term.
|
||||||
||| allows terminating comma. more than two elements are nested on the right:
|
||| allows terminating comma. more than two elements are nested on the right:
|
||||||
||| `(a, b, c, d) = (a, (b, (c, d)))`.
|
||| `(a, b, c, d) = (a, (b, (c, d)))`.
|
||||||
export covering
|
export covering
|
||||||
tupleTerm : Grammar True PTerm
|
tupleTerm : FileName -> Grammar True PTerm
|
||||||
tupleTerm = foldr1 Pair <$> delimSep1 "(" ")" "," term
|
tupleTerm fname = withLoc fname $ do
|
||||||
|
terms <- delimSep1 "(" ")" "," $ term fname
|
||||||
|
pure $ \loc => foldr1 (\s, t => Pair s t loc) terms
|
||||||
|
|
||||||
||| argument/atomic term: single-token terms, or those with delimiters e.g.
|
||| argument/atomic term: single-token terms, or those with delimiters e.g.
|
||||||
||| `[t]`
|
||| `[t]`
|
||||||
export covering
|
export covering
|
||||||
termArg : Grammar True PTerm
|
termArg : FileName -> Grammar True PTerm
|
||||||
termArg = [|TYPE universe1|]
|
termArg fname = withLoc fname $
|
||||||
<|> [|Enum enumType|]
|
[|TYPE universe1|]
|
||||||
<|> [|Tag tag|]
|
<|> [|Enum enumType|]
|
||||||
<|> boxTerm
|
<|> [|Tag tag|]
|
||||||
<|> Nat <$ res "ℕ"
|
<|> const <$> boxTerm fname
|
||||||
<|> Zero <$ res "zero"
|
<|> Nat <$ res "ℕ"
|
||||||
<|> [|fromNat nat|]
|
<|> Zero <$ res "zero"
|
||||||
<|> [|V qname|]
|
<|> [|fromNat nat|]
|
||||||
<|> tupleTerm
|
<|> [|V qname|]
|
||||||
|
<|> const <$> tupleTerm fname
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
coeTerm : Grammar True PTerm
|
coeTerm : FileName -> Grammar True PTerm
|
||||||
coeTerm = resC "coe" *> mustWork [|Coe typeLine dimArg dimArg termArg|]
|
coeTerm fname = withLoc fname $ do
|
||||||
|
resC "coe"
|
||||||
|
mustWork [|Coe (typeLine fname) (dimArg fname) (dimArg fname)
|
||||||
|
(termArg fname)|]
|
||||||
|
|
||||||
|
public export
|
||||||
|
CompBranch : Type
|
||||||
|
CompBranch = (DimConst, PatVar, PTerm)
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
compBranch : Grammar True (DimConst, BName, PTerm)
|
compBranch : FileName -> Grammar True CompBranch
|
||||||
compBranch = [|(,,) dimConst patVar (needRes "⇒" *> term)|]
|
compBranch fname = [|(,,) dimConst (patVar fname) (needRes "⇒" *> term fname)|]
|
||||||
|
|
||||||
|
private
|
||||||
|
checkCompTermBody : (PatVar, PTerm) -> PDim -> PDim -> PTerm -> PDim ->
|
||||||
|
CompBranch -> CompBranch -> Bounds ->
|
||||||
|
Grammar False (Loc -> PTerm)
|
||||||
|
checkCompTermBody a p q s r (e0, s0) (e1, s1) bounds =
|
||||||
|
case (e0, e1) of
|
||||||
|
(Zero, One) => pure $ Comp a p q s r s0 s1
|
||||||
|
(One, Zero) => pure $ Comp a p q s r s1 s0
|
||||||
|
(_, _) =>
|
||||||
|
fatalLoc bounds "body of 'comp' needs one 0 case and one 1 case"
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
compTerm : Grammar True PTerm
|
compTerm : FileName -> Grammar True PTerm
|
||||||
compTerm = do
|
compTerm fname = withLoc fname $ do
|
||||||
resC "comp"
|
resC "comp"
|
||||||
mustWork $ do
|
mustWork $ do
|
||||||
a <- typeLine
|
a <- typeLine fname
|
||||||
p <- dimArg; q <- dimArg
|
p <- dimArg fname; q <- dimArg fname
|
||||||
s <- termArg; r <- dimArg
|
s <- termArg fname; r <- dimArg fname
|
||||||
needRes "{"
|
bodyStart <- bounds $ needRes "{"
|
||||||
s0 <- compBranch; needRes ";"
|
s0 <- compBranch fname; needRes ";"
|
||||||
s1 <- compBranch; optRes ";"
|
s1 <- compBranch fname; optRes ";"
|
||||||
needRes "}"
|
bodyEnd <- bounds $ needRes "}"
|
||||||
case (fst s0, fst s1) of
|
let body = bounds $ mergeBounds bodyStart bodyEnd
|
||||||
(Zero, One) => pure $ Comp a p q s r (snd s0) (snd s1)
|
checkCompTermBody a p q s r s0 s1 body
|
||||||
(One, Zero) => pure $ Comp a p q s r (snd s1) (snd s0)
|
|
||||||
(_, _) =>
|
|
||||||
fatalError "body of 'comp' needs one 0 case and one 1 case"
|
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
splitUniverseTerm : Grammar True PTerm
|
splitUniverseTerm : FileName -> Grammar True PTerm
|
||||||
splitUniverseTerm = resC "★" *> mustWork [|TYPE nat|]
|
splitUniverseTerm fname = withLoc fname $ resC "★" *> mustWork [|TYPE nat|]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
eqTerm : Grammar True PTerm
|
eqTerm : FileName -> Grammar True PTerm
|
||||||
eqTerm = resC "Eq" *> mustWork [|Eq typeLine termArg termArg|]
|
eqTerm fname = withLoc fname $
|
||||||
|
resC "Eq" *> mustWork [|Eq (typeLine fname) (termArg fname) (termArg fname)|]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
succTerm : Grammar True PTerm
|
succTerm : FileName -> Grammar True PTerm
|
||||||
succTerm = resC "succ" *> mustWork [|Succ termArg|]
|
succTerm fname = withLoc fname $
|
||||||
|
resC "succ" *> mustWork [|Succ (termArg fname)|]
|
||||||
|
|
||||||
||| a dimension argument with an `@` prefix, or
|
||| a dimension argument with an `@` prefix, or
|
||||||
||| a term argument with no prefix
|
||| a term argument with no prefix
|
||||||
export covering
|
export covering
|
||||||
anyArg : Grammar True (Either PDim PTerm)
|
anyArg : FileName -> Grammar True (Either PDim PTerm)
|
||||||
anyArg = dimArg <||> termArg
|
anyArg fname = dimArg fname <||> termArg fname
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
normalAppTerm : Grammar True PTerm
|
normalAppTerm : FileName -> Grammar True PTerm
|
||||||
normalAppTerm = foldl ap <$> termArg <*> many anyArg
|
normalAppTerm fname = withLoc fname $ do
|
||||||
where ap : PTerm -> Either PDim PTerm -> PTerm
|
head <- termArg fname
|
||||||
ap f (Left p) = f :% p
|
args <- many $ anyArg fname
|
||||||
ap f (Right s) = f :@ s
|
pure $ \loc => foldl (ap loc) head args
|
||||||
|
where ap : Loc -> PTerm -> Either PDim PTerm -> PTerm
|
||||||
|
ap loc f (Left p) = DApp f p loc
|
||||||
|
ap loc f (Right s) = App f s loc
|
||||||
|
|
||||||
||| application term `f x @y z`, or other terms that look like application
|
||| application term `f x @y z`, or other terms that look like application
|
||||||
||| like `succ` or `coe`.
|
||| like `succ` or `coe`.
|
||||||
export covering
|
export covering
|
||||||
appTerm : Grammar True PTerm
|
appTerm : FileName -> Grammar True PTerm
|
||||||
appTerm = coeTerm
|
appTerm fname =
|
||||||
<|> compTerm
|
coeTerm fname
|
||||||
<|> splitUniverseTerm
|
<|> compTerm fname
|
||||||
<|> eqTerm
|
<|> splitUniverseTerm fname
|
||||||
<|> succTerm
|
<|> eqTerm fname
|
||||||
<|> normalAppTerm
|
<|> succTerm fname
|
||||||
|
<|> normalAppTerm fname
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
infixEqTerm : Grammar True PTerm
|
infixEqTerm : FileName -> Grammar True PTerm
|
||||||
infixEqTerm = do
|
infixEqTerm fname = withLoc fname $ do
|
||||||
l <- appTerm; commit
|
l <- appTerm fname; commit
|
||||||
rest <- optional $ res "≡" *> [|(,) term (needRes ":" *> appTerm)|]
|
rest <- optional $
|
||||||
pure $ maybe l (\rest => Eq (Nothing, snd rest) l (fst rest)) rest
|
res "≡" *> [|(,) (term fname) (needRes ":" *> appTerm fname)|]
|
||||||
|
let u = Unused $ onlyStart l.loc
|
||||||
|
pure $ \loc => maybe l (\rest => Eq (u, snd rest) l (fst rest) loc) rest
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
annTerm : Grammar True PTerm
|
annTerm : FileName -> Grammar True PTerm
|
||||||
annTerm = do
|
annTerm fname = withLoc fname $ do
|
||||||
tm <- infixEqTerm; commit
|
tm <- infixEqTerm fname; commit
|
||||||
ty <- optional $ res "∷" *> term
|
ty <- optional $ res "∷" *> term fname
|
||||||
pure $ maybe tm (tm :#) ty
|
pure $ \loc => maybe tm (\ty => Ann tm ty loc) ty
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
lamTerm : Grammar True PTerm
|
lamTerm : FileName -> Grammar True PTerm
|
||||||
lamTerm = do
|
lamTerm fname = withLoc fname $ do
|
||||||
k <- DLam <$ res "δ" <|> Lam <$ res "λ"
|
k <- DLam <$ res "δ" <|> Lam <$ res "λ"
|
||||||
mustWork $ do
|
mustWork $ do
|
||||||
xs <- some patVar; needRes "⇒"
|
xs <- some $ patVar fname; needRes "⇒"
|
||||||
body <- term; commit
|
body <- term fname; commit
|
||||||
pure $ foldr k body xs
|
pure $ \loc => foldr (\x, s => k x s loc) body xs
|
||||||
|
|
||||||
-- [todo] fix the backtracking in e.g. (F x y z × B)
|
-- [todo] fix the backtracking in e.g. (F x y z × B)
|
||||||
export covering
|
export covering
|
||||||
properBinders : Grammar True (List1 BName, PTerm)
|
properBinders : FileName -> Grammar True (List1 PatVar, PTerm)
|
||||||
properBinders = do
|
properBinders fname = do
|
||||||
res "("
|
res "("
|
||||||
xs <- some patVar; resC ":"
|
xs <- some $ patVar fname; resC ":"
|
||||||
t <- term; needRes ")"
|
t <- term fname; needRes ")"
|
||||||
pure (xs, t)
|
pure (xs, t)
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
piTerm : Grammar True PTerm
|
piTerm : FileName -> Grammar True PTerm
|
||||||
piTerm = do
|
piTerm fname = withLoc fname $ do
|
||||||
q <- qty; resC "."
|
q <- qty fname; resC "."
|
||||||
dom <- piBinder; needRes "→"
|
dom <- piBinder; needRes "→"
|
||||||
cod <- term; commit
|
cod <- term fname; commit
|
||||||
pure $ foldr (\x => Pi q x (snd dom)) cod (fst dom)
|
pure $ \loc => foldr (\x, t => Pi q x (snd dom) t loc) cod (fst dom)
|
||||||
where
|
where
|
||||||
piBinder : Grammar True (List1 BName, PTerm)
|
piBinder : Grammar True (List1 PatVar, PTerm)
|
||||||
piBinder = properBinders <|> [|(singleton Nothing,) termArg|]
|
piBinder = properBinders fname
|
||||||
|
<|> [|(,) [|singleton $ unused fname|] (termArg fname)|]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
sigmaTerm : Grammar True PTerm
|
sigmaTerm : FileName -> Grammar True PTerm
|
||||||
sigmaTerm =
|
sigmaTerm fname =
|
||||||
(properBinders >>= continueDep)
|
(properBinders fname >>= continueDep)
|
||||||
<|> (annTerm >>= continueNondep)
|
<|> (annTerm fname >>= continueNondep)
|
||||||
where
|
where
|
||||||
continueDep : (List1 BName, PTerm) -> Grammar True PTerm
|
continueDep : (List1 PatVar, PTerm) -> Grammar True PTerm
|
||||||
continueDep (names, dom) = do
|
continueDep (names, fst) = withLoc fname $ do
|
||||||
cod <- needRes "×" *> sigmaTerm
|
snd <- needRes "×" *> sigmaTerm fname
|
||||||
pure $ foldr (\x => Sig x dom) cod names
|
pure $ \loc => foldr (\x, snd => Sig x fst snd loc) snd names
|
||||||
continueNondep : PTerm -> Grammar False PTerm
|
|
||||||
continueNondep dom = do
|
|
||||||
rest <- optional $ resC "×" *> sepBy1 (res "×") annTerm
|
|
||||||
Core.pure $ foldr1 (Sig Nothing) $ dom ::: maybe [] forget rest
|
|
||||||
|
|
||||||
|
cross : PTerm -> PTerm -> PTerm
|
||||||
|
cross l r = let loc = extend' l.loc r.loc.bounds in
|
||||||
|
Sig (Unused $ onlyStart l.loc) l r loc
|
||||||
|
|
||||||
|
continueNondep : PTerm -> Grammar False PTerm
|
||||||
|
continueNondep fst = do
|
||||||
|
rest <- optional $ resC "×" *> sepBy1 (res "×") (annTerm fname)
|
||||||
|
pure $ foldr1 cross $ fst ::: maybe [] toList rest
|
||||||
|
|
||||||
public export
|
public export
|
||||||
PCaseArm : Type
|
PCaseArm : Type
|
||||||
PCaseArm = (PCasePat, PTerm)
|
PCaseArm = (PCasePat, PTerm)
|
||||||
|
|
||||||
|
export covering
|
||||||
|
caseArm : FileName -> Grammar True PCaseArm
|
||||||
|
caseArm fname = [|(,) (casePat fname) (needRes "⇒" *> term fname)|]
|
||||||
|
|
||||||
export
|
export
|
||||||
checkCaseArms : List PCaseArm -> Grammar False PCaseBody
|
checkCaseArms : Loc -> List PCaseArm -> Grammar False PCaseBody
|
||||||
checkCaseArms [] =
|
checkCaseArms loc [] = pure $ CaseEnum [] loc
|
||||||
pure $ CaseEnum []
|
checkCaseArms loc ((PPair x y _, rhs) :: rest) =
|
||||||
checkCaseArms ((PPair x y, rhs) :: rest) = do
|
if null rest then pure $ CasePair (x, y) rhs loc
|
||||||
let [] = rest | _ => fatalError "unexpected pattern after pair"
|
else fatalError "unexpected pattern after pair"
|
||||||
pure $ CasePair (x, y) rhs
|
checkCaseArms loc ((PTag tag _, rhs1) :: rest) = do
|
||||||
checkCaseArms ((PTag tag, rhs1) :: rest) = do
|
|
||||||
let rest = for rest $ \case
|
let rest = for rest $ \case
|
||||||
(PTag tag, rhs) => Just (tag, rhs)
|
(PTag tag _, rhs) => Just (tag, rhs)
|
||||||
_ => Nothing
|
_ => Nothing
|
||||||
maybe (fatalError "expected all patterns to be tags")
|
maybe (fatalError "expected all patterns to be tags")
|
||||||
(pure . CaseEnum . ((tag, rhs1) ::)) rest
|
(\rest => pure $ CaseEnum ((tag, rhs1) :: rest) loc) rest
|
||||||
checkCaseArms ((PZero, rhs1) :: rest) = do
|
checkCaseArms loc ((PZero _, rhs1) :: rest) = do
|
||||||
let [(PSucc p q ih, rhs2)] = rest
|
let [(PSucc p q ih _, rhs2)] = rest
|
||||||
| _ => fatalError "expected succ pattern after zero"
|
| _ => fatalError "expected succ pattern after zero"
|
||||||
pure $ CaseNat rhs1 (p, q, ih, rhs2)
|
pure $ CaseNat rhs1 (p, q, ih, rhs2) loc
|
||||||
checkCaseArms ((PSucc p q ih, rhs1) :: rest) = do
|
checkCaseArms loc ((PSucc p q ih _, rhs1) :: rest) = do
|
||||||
let [(PZero, rhs2)] = rest
|
let [(PZero _, rhs2)] = rest
|
||||||
| _ => fatalError "expected zero pattern after succ"
|
| _ => fatalError "expected zero pattern after succ"
|
||||||
pure $ CaseNat rhs2 (p, q, ih, rhs1)
|
pure $ CaseNat rhs2 (p, q, ih, rhs1) loc
|
||||||
checkCaseArms ((PBox x, rhs) :: rest) = do
|
checkCaseArms loc ((PBox x _, rhs) :: rest) =
|
||||||
let [] = rest | _ => fatalError "unexpected pattern after box"
|
if null rest then pure $ CaseBox x rhs loc
|
||||||
pure $ CaseBox x rhs
|
else fatalError "unexpected pattern after box"
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
caseArm : Grammar True PCaseArm
|
caseBody : FileName -> Grammar True PCaseBody
|
||||||
caseArm = [|(,) casePat (needRes "⇒" *> term)|]
|
caseBody fname = do
|
||||||
|
body <- bounds $ delimSep "{" "}" ";" $ caseArm fname
|
||||||
|
let loc = makeLoc fname body.bounds
|
||||||
|
checkCaseArms loc body.val
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
caseBody : Grammar True PCaseBody
|
caseReturn : FileName -> Grammar True (PatVar, PTerm)
|
||||||
caseBody = delimSep "{" "}" ";" caseArm >>= checkCaseArms
|
caseReturn fname = do
|
||||||
|
x <- patVar fname <* resC "⇒" <|> unused fname
|
||||||
|
ret <- term fname
|
||||||
|
pure (x, ret)
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
caseTerm : Grammar True PTerm
|
caseTerm : FileName -> Grammar True PTerm
|
||||||
caseTerm = do
|
caseTerm fname = withLoc fname $ do
|
||||||
qty <- caseIntro; commit
|
qty <- caseIntro fname; commit
|
||||||
head <- mustWork term; needRes "return"
|
head <- mustWork $ term fname; needRes "return"
|
||||||
ret <- mustWork [|(,) (option Nothing $ patVar <* resC "⇒") term|]
|
ret <- mustWork $ caseReturn fname; needRes "of"
|
||||||
needRes "of"
|
body <- mustWork $ caseBody fname
|
||||||
body <- mustWork caseBody
|
|
||||||
pure $ Case qty head ret body
|
pure $ Case qty head ret body
|
||||||
|
|
||||||
|
-- export covering
|
||||||
|
-- term : FileName -> Grammar True PTerm
|
||||||
|
term fname = lamTerm fname
|
||||||
|
<|> caseTerm fname
|
||||||
|
<|> piTerm fname
|
||||||
|
<|> sigmaTerm fname
|
||||||
|
|
||||||
term = lamTerm
|
|
||||||
<|> caseTerm
|
|
||||||
<|> piTerm
|
|
||||||
<|> sigmaTerm
|
|
||||||
|
|
||||||
|
export covering
|
||||||
|
decl : FileName -> Grammar True PDecl
|
||||||
|
|
||||||
||| `def` alone means `defω`
|
||| `def` alone means `defω`
|
||||||
export
|
export
|
||||||
defIntro : Grammar True Qty
|
defIntro : FileName -> Grammar True PQty
|
||||||
defIntro = Zero <$ resC "def0"
|
defIntro fname =
|
||||||
<|> Any <$ resC "defω"
|
withLoc fname (PQ Zero <$ resC "def0")
|
||||||
<|> resC "def" *> option Any (qty <* needRes ".")
|
<|> withLoc fname (PQ Any <$ resC "defω")
|
||||||
|
<|> do pos <- bounds $ resC "def"
|
||||||
|
let any = PQ Any $ makeLoc fname pos.bounds
|
||||||
|
option any $ qty fname <* needRes "."
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
decl : Grammar True PDecl
|
definition : FileName -> Grammar True PDefinition
|
||||||
|
definition fname = withLoc fname $ do
|
||||||
export covering
|
qty <- defIntro fname
|
||||||
definition : Grammar True PDefinition
|
|
||||||
definition = do
|
|
||||||
qty <- defIntro
|
|
||||||
name <- baseName
|
name <- baseName
|
||||||
type <- optional $ resC ":" *> mustWork term
|
type <- optional $ resC ":" *> mustWork (term fname)
|
||||||
term <- needRes "=" *> mustWork term
|
term <- needRes "=" *> mustWork (term fname)
|
||||||
optRes ";"
|
optRes ";"
|
||||||
pure $ MkPDef {qty, name, type, term}
|
pure $ MkPDef qty name type term
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
namespace_ : Grammar True PNamespace
|
namespace_ : FileName -> Grammar True PNamespace
|
||||||
namespace_ = do
|
namespace_ fname = withLoc fname $ do
|
||||||
ns <- resC "namespace" *> qname; needRes "{"
|
ns <- resC "namespace" *> qname; needRes "{"
|
||||||
decls <- nsInner; optRes ";"
|
decls <- nsInner; optRes ";"
|
||||||
pure $ MkPNamespace (ns.mods :< ns.base) decls
|
pure $ MkPNamespace (ns.mods :< ns.base) decls
|
||||||
where
|
where
|
||||||
nsInner : Grammar True (List PDecl)
|
nsInner : Grammar True (List PDecl)
|
||||||
nsInner = [] <$ resC "}"
|
nsInner = [] <$ resC "}"
|
||||||
<|> [|(decl <* commit) :: nsInner|]
|
<|> [|(decl fname <* commit) :: nsInner|]
|
||||||
|
|
||||||
decl = [|PDef definition|] <|> [|PNs namespace_|]
|
decl fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
load : Grammar True String
|
load : FileName -> Grammar True PTopLevel
|
||||||
load = resC "load" *> mustWork strLit <* optRes ";"
|
load fname = withLoc fname $
|
||||||
|
resC "load" *> mustWork [|PLoad strLit|] <* optRes ";"
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
topLevel : Grammar True PTopLevel
|
topLevel : FileName -> Grammar True PTopLevel
|
||||||
topLevel = [|PLoad load|] <|> [|PD decl|]
|
topLevel fname = load fname <|> [|PD $ decl fname|]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
input : Grammar False (List PTopLevel)
|
input : FileName -> Grammar False (List PTopLevel)
|
||||||
input = [] <$ eof
|
input fname = [] <$ eof
|
||||||
<|> [|(topLevel <* commit) :: input|]
|
<|> [|(topLevel fname <* commit) :: input fname|]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
lexParseTerm : String -> Either Error PTerm
|
lexParseTerm : FileName -> String -> Either Error PTerm
|
||||||
lexParseTerm = lexParseWith term
|
lexParseTerm = lexParseWith . term
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
lexParseInput : String -> Either Error (List PTopLevel)
|
lexParseInput : FileName -> String -> Either Error (List PTopLevel)
|
||||||
lexParseInput = lexParseWith input
|
lexParseInput = lexParseWith . input
|
||||||
|
|
|
@ -1,5 +1,6 @@
|
||||||
module Quox.Parser.Syntax
|
module Quox.Parser.Syntax
|
||||||
|
|
||||||
|
import public Quox.Loc
|
||||||
import public Quox.Syntax
|
import public Quox.Syntax
|
||||||
import public Quox.Definition
|
import public Quox.Definition
|
||||||
|
|
||||||
|
@ -11,81 +12,148 @@ import Derive.Prelude
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
BName : Type
|
data PatVar = Unused Loc | PV PBaseName Loc
|
||||||
BName = Maybe String
|
%name PatVar v
|
||||||
|
%runElab derive "PatVar" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
export
|
||||||
|
Located PatVar where
|
||||||
|
(Unused loc).loc = loc
|
||||||
|
(PV _ loc).loc = loc
|
||||||
|
|
||||||
|
export
|
||||||
|
(.name) : PatVar -> Maybe PBaseName
|
||||||
|
(Unused _).name = Nothing
|
||||||
|
(PV nm _).name = Just nm
|
||||||
|
|
||||||
|
export
|
||||||
|
isUnused : PatVar -> Bool
|
||||||
|
isUnused (Unused {}) = True
|
||||||
|
isUnused _ = False
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
PUniverse : Type
|
data PQty = PQ Qty Loc
|
||||||
PUniverse = Nat
|
%name PQty qty
|
||||||
|
%runElab derive "PQty" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
export Located PQty where (PQ _ loc).loc = loc
|
||||||
|
|
||||||
namespace PDim
|
namespace PDim
|
||||||
public export
|
public export
|
||||||
data PDim = K DimConst | V PBaseName
|
data PDim = K DimConst Loc | V PBaseName Loc
|
||||||
%name PDim p, q
|
%name PDim p, q
|
||||||
%runElab derive "PDim" [Eq, Ord, Show]
|
%runElab derive "PDim" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
export
|
||||||
|
Located PDim where
|
||||||
|
(K _ loc).loc = loc
|
||||||
|
(V _ loc).loc = loc
|
||||||
|
|
||||||
|
public export
|
||||||
|
data PTagVal = PT TagVal Loc
|
||||||
|
%name PTagVal tag
|
||||||
|
%runElab derive "PTagVal" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
namespace PTerm
|
namespace PTerm
|
||||||
mutual
|
mutual
|
||||||
||| terms out of the parser with BVs and bidirectionality still tangled up
|
||| terms out of the parser with BVs and bidirectionality still tangled up
|
||||||
public export
|
public export
|
||||||
data PTerm =
|
data PTerm =
|
||||||
TYPE Nat
|
TYPE Universe Loc
|
||||||
|
|
||||||
| Pi Qty BName PTerm PTerm
|
| Pi PQty PatVar PTerm PTerm Loc
|
||||||
| Lam BName PTerm
|
| Lam PatVar PTerm Loc
|
||||||
| (:@) PTerm PTerm
|
| App PTerm PTerm Loc
|
||||||
|
|
||||||
| Sig BName PTerm PTerm
|
| Sig PatVar PTerm PTerm Loc
|
||||||
| Pair PTerm PTerm
|
| Pair PTerm PTerm Loc
|
||||||
| Case Qty PTerm (BName, PTerm) PCaseBody
|
| Case PQty PTerm (PatVar, PTerm) PCaseBody Loc
|
||||||
|
|
||||||
| Enum (List TagVal)
|
| Enum (List TagVal) Loc
|
||||||
| Tag TagVal
|
| Tag TagVal Loc
|
||||||
|
|
||||||
| Eq (BName, PTerm) PTerm PTerm
|
| Eq (PatVar, PTerm) PTerm PTerm Loc
|
||||||
| DLam BName PTerm
|
| DLam PatVar PTerm Loc
|
||||||
| (:%) PTerm PDim
|
| DApp PTerm PDim Loc
|
||||||
|
|
||||||
| Nat
|
| Nat Loc
|
||||||
| Zero | Succ PTerm
|
| Zero Loc | Succ PTerm Loc
|
||||||
|
|
||||||
| BOX Qty PTerm
|
| BOX PQty PTerm Loc
|
||||||
| Box PTerm
|
| Box PTerm Loc
|
||||||
|
|
||||||
| V PName
|
| V PName Loc
|
||||||
| (:#) PTerm PTerm
|
| Ann PTerm PTerm Loc
|
||||||
|
|
||||||
| Coe (BName, PTerm) PDim PDim PTerm
|
| Coe (PatVar, PTerm) PDim PDim PTerm Loc
|
||||||
| Comp (BName, PTerm) PDim PDim PTerm PDim (BName, PTerm) (BName, PTerm)
|
| Comp (PatVar, PTerm) PDim PDim PTerm PDim
|
||||||
|
(PatVar, PTerm) (PatVar, PTerm) Loc
|
||||||
%name PTerm s, t
|
%name PTerm s, t
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PCaseBody =
|
data PCaseBody =
|
||||||
CasePair (BName, BName) PTerm
|
CasePair (PatVar, PatVar) PTerm Loc
|
||||||
| CaseEnum (List (TagVal, PTerm))
|
| CaseEnum (List (PTagVal, PTerm)) Loc
|
||||||
| CaseNat PTerm (BName, Qty, BName, PTerm)
|
| CaseNat PTerm (PatVar, PQty, PatVar, PTerm) Loc
|
||||||
| CaseBox BName PTerm
|
| CaseBox PatVar PTerm Loc
|
||||||
%name PCaseBody body
|
%name PCaseBody body
|
||||||
|
|
||||||
%runElab deriveMutual ["PTerm", "PCaseBody"] [Eq, Ord, Show]
|
%runElab deriveMutual ["PTerm", "PCaseBody"] [Eq, Ord, Show]
|
||||||
|
|
||||||
|
export
|
||||||
|
Located PTerm where
|
||||||
|
(TYPE _ loc).loc = loc
|
||||||
|
(Pi _ _ _ _ loc).loc = loc
|
||||||
|
(Lam _ _ loc).loc = loc
|
||||||
|
(App _ _ loc).loc = loc
|
||||||
|
(Sig _ _ _ loc).loc = loc
|
||||||
|
(Pair _ _ loc).loc = loc
|
||||||
|
(Case _ _ _ _ loc).loc = loc
|
||||||
|
(Enum _ loc).loc = loc
|
||||||
|
(Tag _ loc).loc = loc
|
||||||
|
(Eq _ _ _ loc).loc = loc
|
||||||
|
(DLam _ _ loc).loc = loc
|
||||||
|
(DApp _ _ loc).loc = loc
|
||||||
|
(Nat loc).loc = loc
|
||||||
|
(Zero loc).loc = loc
|
||||||
|
(Succ _ loc).loc = loc
|
||||||
|
(BOX _ _ loc).loc = loc
|
||||||
|
(Box _ loc).loc = loc
|
||||||
|
(V _ loc).loc = loc
|
||||||
|
(Ann _ _ loc).loc = loc
|
||||||
|
(Coe _ _ _ _ loc).loc = loc
|
||||||
|
(Comp _ _ _ _ _ _ _ loc).loc = loc
|
||||||
|
|
||||||
|
export
|
||||||
|
Located PCaseBody where
|
||||||
|
(CasePair _ _ loc).loc = loc
|
||||||
|
(CaseEnum _ loc).loc = loc
|
||||||
|
(CaseNat _ _ loc).loc = loc
|
||||||
|
(CaseBox _ _ loc).loc = loc
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
record PDefinition where
|
record PDefinition where
|
||||||
constructor MkPDef
|
constructor MkPDef
|
||||||
qty : Qty
|
qty : PQty
|
||||||
name : PBaseName
|
name : PBaseName
|
||||||
type : Maybe PTerm
|
type : Maybe PTerm
|
||||||
term : PTerm
|
term : PTerm
|
||||||
|
loc_ : Loc
|
||||||
%name PDefinition def
|
%name PDefinition def
|
||||||
%runElab derive "PDefinition" [Eq, Ord, Show]
|
%runElab derive "PDefinition" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
export Located PDefinition where def.loc = def.loc_
|
||||||
|
|
||||||
mutual
|
mutual
|
||||||
public export
|
public export
|
||||||
record PNamespace where
|
record PNamespace where
|
||||||
constructor MkPNamespace
|
constructor MkPNamespace
|
||||||
name : Mods
|
name : Mods
|
||||||
decls : List PDecl
|
decls : List PDecl
|
||||||
|
loc_ : Loc
|
||||||
%name PNamespace ns
|
%name PNamespace ns
|
||||||
|
|
||||||
public export
|
public export
|
||||||
|
@ -95,125 +163,25 @@ mutual
|
||||||
%name PDecl decl
|
%name PDecl decl
|
||||||
%runElab deriveMutual ["PNamespace", "PDecl"] [Eq, Ord, Show]
|
%runElab deriveMutual ["PNamespace", "PDecl"] [Eq, Ord, Show]
|
||||||
|
|
||||||
|
export Located PNamespace where ns.loc = ns.loc_
|
||||||
|
|
||||||
|
export
|
||||||
|
Located PDecl where
|
||||||
|
(PDef def).loc = def.loc
|
||||||
|
(PNs ns).loc = ns.loc
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PTopLevel = PD PDecl | PLoad String
|
data PTopLevel = PD PDecl | PLoad String Loc
|
||||||
%name PTopLevel t
|
%name PTopLevel t
|
||||||
%runElab derive "PTopLevel" [Eq, Ord, Show]
|
%runElab derive "PTopLevel" [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
toPDimWith : Context' PBaseName d -> Dim d -> PDim
|
Located PTopLevel where
|
||||||
toPDimWith ds (K e) = K e
|
(PD decl).loc = decl.loc
|
||||||
toPDimWith ds (B i) = V $ ds !!! i
|
(PLoad _ loc).loc = loc
|
||||||
|
|
||||||
export
|
|
||||||
toPDim : Dim 0 -> PDim
|
|
||||||
toPDim = toPDimWith [<]
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
fromNat : Nat -> PTerm
|
fromNat : Nat -> Loc -> PTerm
|
||||||
fromNat 0 = Zero
|
fromNat 0 loc = Zero loc
|
||||||
fromNat (S k) = Succ $ fromNat k
|
fromNat (S k) loc = Succ (fromNat k loc) loc
|
||||||
|
|
||||||
|
|
||||||
{-
|
|
||||||
mutual
|
|
||||||
namespace Term
|
|
||||||
export
|
|
||||||
toPTermWith : Context' PBaseName d -> Context' PBaseName n ->
|
|
||||||
Term d n -> PTerm
|
|
||||||
toPTermWith ds ns t =
|
|
||||||
let Element t _ = pushSubsts t in
|
|
||||||
toPTermWith' ds ns t
|
|
||||||
|
|
||||||
private
|
|
||||||
toPTermWith' : Context' PBaseName d -> Context' PBaseName n ->
|
|
||||||
(t : Term d n) -> (0 _ : NotClo t) =>
|
|
||||||
PTerm
|
|
||||||
toPTermWith' ds ns s = case s of
|
|
||||||
TYPE l =>
|
|
||||||
TYPE l
|
|
||||||
Pi qty arg (S [< x] res) =>
|
|
||||||
Pi qty (Just $ show x)
|
|
||||||
(toPTermWith ds ns arg)
|
|
||||||
(toPTermWith ds (ns :< baseStr x) res.term)
|
|
||||||
Lam (S [< x] body) =>
|
|
||||||
Lam (Just $ show x) $
|
|
||||||
toPTermWith ds (ns :< baseStr x) body.term
|
|
||||||
Sig fst (S [< x] snd) =>
|
|
||||||
Sig (Just $ show x)
|
|
||||||
(toPTermWith ds ns fst)
|
|
||||||
(toPTermWith ds (ns :< baseStr x) snd.term)
|
|
||||||
Pair fst snd =>
|
|
||||||
Pair (toPTermWith ds ns fst) (toPTermWith ds ns snd)
|
|
||||||
Enum cases =>
|
|
||||||
Enum $ SortedSet.toList cases
|
|
||||||
Tag tag =>
|
|
||||||
Tag tag
|
|
||||||
Eq (S [< i] ty) l r =>
|
|
||||||
Eq (Just $ show i, toPTermWith (ds :< baseStr i) ns ty.term)
|
|
||||||
(toPTermWith ds ns l) (toPTermWith ds ns r)
|
|
||||||
DLam (S [< i] body) =>
|
|
||||||
DLam (Just $ show i) $ toPTermWith (ds :< baseStr i) ns body.term
|
|
||||||
Nat => Nat
|
|
||||||
Zero => Zero
|
|
||||||
Succ n => Succ $ toPTermWith ds ns n
|
|
||||||
BOX q ty => BOX q $ toPTermWith ds ns ty
|
|
||||||
Box val => Box $ toPTermWith ds ns val
|
|
||||||
E e =>
|
|
||||||
toPTermWith ds ns e
|
|
||||||
|
|
||||||
namespace Elim
|
|
||||||
export
|
|
||||||
toPTermWith : Context' PBaseName d -> Context' PBaseName n ->
|
|
||||||
Elim d n -> PTerm
|
|
||||||
toPTermWith ds ns e =
|
|
||||||
let Element e _ = pushSubsts e in
|
|
||||||
toPTermWith' ds ns e
|
|
||||||
|
|
||||||
private
|
|
||||||
toPTermWith' : Context' PBaseName d -> Context' PBaseName n ->
|
|
||||||
(e : Elim d n) -> (0 _ : NotClo e) =>
|
|
||||||
PTerm
|
|
||||||
toPTermWith' ds ns e = case e of
|
|
||||||
F x =>
|
|
||||||
V $ toPName x
|
|
||||||
B i =>
|
|
||||||
V $ MakePName [<] $ ns !!! i
|
|
||||||
fun :@ arg =>
|
|
||||||
toPTermWith ds ns fun :@ toPTermWith ds ns arg
|
|
||||||
CasePair qty pair (S [< r] ret) (S [< x, y] body) =>
|
|
||||||
Case qty (toPTermWith ds ns pair)
|
|
||||||
(Just $ show r, toPTermWith ds (ns :< baseStr r) ret.term)
|
|
||||||
(CasePair (Just $ show x, Just $ show y) $
|
|
||||||
toPTermWith ds (ns :< baseStr x :< baseStr y) body.term)
|
|
||||||
CaseEnum qty tag (S [< r] ret) arms =>
|
|
||||||
Case qty (toPTermWith ds ns tag)
|
|
||||||
(Just $ baseStr r, toPTermWith ds (ns :< baseStr r) ret.term)
|
|
||||||
(CaseEnum $ mapSnd (toPTermWith ds ns) <$> SortedMap.toList arms)
|
|
||||||
CaseNat qtyNat qtyIH nat (S [< r] ret) zer (S [< p, ih] suc) =>
|
|
||||||
Case qtyNat (toPTermWith ds ns nat)
|
|
||||||
(Just $ baseStr r, toPTermWith ds (ns :< baseStr r) ret.term)
|
|
||||||
(CaseNat (toPTermWith ds ns zer)
|
|
||||||
(Just $ baseStr p, qtyIH, Just $ baseStr ih,
|
|
||||||
toPTermWith ds (ns :< baseStr p :< baseStr ih) suc.term))
|
|
||||||
CaseBox qty box (S [< r] ret) (S [< b] body) =>
|
|
||||||
Case qty (toPTermWith ds ns box)
|
|
||||||
(Just $ baseStr r, toPTermWith ds (ns :< baseStr r) ret.term)
|
|
||||||
(CaseBox (Just $ baseStr b) $
|
|
||||||
toPTermWith ds (ns :< baseStr b) body.term)
|
|
||||||
fun :% arg =>
|
|
||||||
toPTermWith ds ns fun :% toPDimWith ds arg
|
|
||||||
tm :# ty =>
|
|
||||||
toPTermWith ds ns tm :# toPTermWith ds ns ty
|
|
||||||
|
|
||||||
namespace Term
|
|
||||||
export
|
|
||||||
toPTerm : Term 0 0 -> PTerm
|
|
||||||
toPTerm = toPTermWith [<] [<]
|
|
||||||
|
|
||||||
namespace Elim
|
|
||||||
export
|
|
||||||
toPTerm : Elim 0 0 -> PTerm
|
|
||||||
toPTerm = toPTermWith [<] [<]
|
|
||||||
-}
|
|
||||||
|
|
|
@ -14,6 +14,7 @@ modules =
|
||||||
Quox.EffExtra,
|
Quox.EffExtra,
|
||||||
Quox.Decidable,
|
Quox.Decidable,
|
||||||
Quox.No,
|
Quox.No,
|
||||||
|
Quox.Loc,
|
||||||
Quox.OPE,
|
Quox.OPE,
|
||||||
Quox.Pretty,
|
Quox.Pretty,
|
||||||
Quox.Syntax,
|
Quox.Syntax,
|
||||||
|
|
|
@ -37,12 +37,13 @@ ToInfo Failure where
|
||||||
|
|
||||||
|
|
||||||
parameters {c : Bool} {auto _ : Show b}
|
parameters {c : Bool} {auto _ : Show b}
|
||||||
(grm : Grammar c a) (fromP : a -> Either FromParser.Error b)
|
(grm : FileName -> Grammar c a)
|
||||||
|
(fromP : a -> Either FromParser.Error b)
|
||||||
(inp : String)
|
(inp : String)
|
||||||
parameters {default (ltrim inp) label : String}
|
parameters {default (ltrim inp) label : String}
|
||||||
parsesWith : (b -> Bool) -> Test
|
parsesWith : (b -> Bool) -> Test
|
||||||
parsesWith p = test label $ do
|
parsesWith p = test label $ do
|
||||||
pres <- mapFst ParseError $ lexParseWith grm inp
|
pres <- mapFst ParseError $ lexParseWith (grm "‹test›") inp
|
||||||
res <- mapFst FromParser $ fromP pres
|
res <- mapFst FromParser $ fromP pres
|
||||||
unless (p res) $ Left $ WrongResult $ show res
|
unless (p res) $ Left $ WrongResult $ show res
|
||||||
|
|
||||||
|
@ -55,11 +56,11 @@ parameters {c : Bool} {auto _ : Show b}
|
||||||
parameters {default "\{ltrim inp} # fails" label : String}
|
parameters {default "\{ltrim inp} # fails" label : String}
|
||||||
parseFails : Test
|
parseFails : Test
|
||||||
parseFails = test label $ do
|
parseFails = test label $ do
|
||||||
pres <- mapFst ParseError $ lexParseWith grm inp
|
pres <- mapFst ParseError $ lexParseWith (grm "‹test›") inp
|
||||||
either (const $ Right ()) (Left . ExpectedFail . show) $ fromP pres
|
either (const $ Right ()) (Left . ExpectedFail . show) $ fromP pres
|
||||||
|
|
||||||
|
|
||||||
FromString BName where fromString = Just . fromString
|
FromString PatVar where fromString x = PV x Nothing
|
||||||
|
|
||||||
runFromParser : {default empty defs : Definitions} ->
|
runFromParser : {default empty defs : Definitions} ->
|
||||||
Eff FromParserPure a -> Either FromParser.Error a
|
Eff FromParserPure a -> Either FromParser.Error a
|
||||||
|
|
|
@ -4,6 +4,9 @@ import Quox.Parser
|
||||||
import Data.List
|
import Data.List
|
||||||
import Data.String
|
import Data.String
|
||||||
import TAP
|
import TAP
|
||||||
|
import Language.Reflection
|
||||||
|
|
||||||
|
%language ElabReflection
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data Failure =
|
data Failure =
|
||||||
|
@ -29,204 +32,227 @@ ToInfo Failure where
|
||||||
[("type", "ExpectedFail"), ("got", got)]
|
[("type", "ExpectedFail"), ("got", got)]
|
||||||
|
|
||||||
|
|
||||||
parameters {c : Bool} {auto _ : Show a} (grm : Grammar c a)
|
parameters {auto _ : (Show a, Eq a)} {c : Bool} (grm : FileName -> Grammar c a)
|
||||||
(inp : String)
|
parsesWith : String -> (a -> Bool) -> Test
|
||||||
parameters {default (ltrim inp) label : String}
|
parsesWith inp p = test (ltrim inp) $ do
|
||||||
parsesWith : (a -> Bool) -> Test
|
res <- mapFst ParseError $ lexParseWith (grm "‹test›") inp
|
||||||
parsesWith p = test label $ do
|
unless (p res) $ Left $ WrongResult $ show res
|
||||||
res <- mapFst ParseError $ lexParseWith grm inp
|
|
||||||
unless (p res) $ Left $ WrongResult $ show res
|
|
||||||
|
|
||||||
parses : Test
|
parsesAs : String -> a -> Test
|
||||||
parses = parsesWith $ const True
|
parsesAs inp exp = parsesWith inp (== exp)
|
||||||
|
|
||||||
parsesAs : Eq a => a -> Test
|
%macro
|
||||||
parsesAs exp = parsesWith (== exp)
|
parseMatch : String -> TTImp -> Elab Test
|
||||||
|
parseMatch inp pat =
|
||||||
|
parsesWith inp <$> check `(\case ~(pat) => True; _ => False)
|
||||||
|
|
||||||
parameters {default "\{ltrim inp} # fails" label : String}
|
parseFails : String -> Test
|
||||||
parseFails : Test
|
parseFails inp = test "\{ltrim inp} # fails" $ do
|
||||||
parseFails = test label $ do
|
either (const $ Right ()) (Left . ExpectedFail . show) $
|
||||||
either (const $ Right ()) (Left . ExpectedFail . show) $
|
lexParseWith (grm "‹test›") inp
|
||||||
lexParseWith grm inp
|
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
tests : Test
|
tests : Test
|
||||||
tests = "parser" :- [
|
tests = "parser" :- [
|
||||||
"pattern vars" :- [
|
"pattern vars" :- [
|
||||||
parsesAs patVar "_" Nothing,
|
parseMatch patVar "_" `(Unused _),
|
||||||
parsesAs patVar "F" (Just "F"),
|
parseMatch patVar "F" `(PV "F" _),
|
||||||
parseFails patVar "a.b.c"
|
parseFails patVar "a.b.c"
|
||||||
],
|
],
|
||||||
|
|
||||||
"names" :- [
|
"names" :- [
|
||||||
parsesAs qname "x"
|
parsesAs (const qname) "x"
|
||||||
(MakePName [<] "x"),
|
(MakePName [<] "x"),
|
||||||
parsesAs qname "Data.String.length"
|
parsesAs (const qname) "Data.String.length"
|
||||||
(MakePName [< "Data", "String"] "length"),
|
(MakePName [< "Data", "String"] "length"),
|
||||||
parseFails qname "_"
|
parseFails (const qname) "_"
|
||||||
],
|
],
|
||||||
|
|
||||||
"dimensions" :- [
|
"dimensions" :- [
|
||||||
parsesAs dim "0" (K Zero),
|
parseMatch dim "0" `(K Zero _),
|
||||||
parsesAs dim "1" (K One),
|
parseMatch dim "1" `(K One _),
|
||||||
parsesAs dim "𝑖" (V "𝑖"),
|
parseMatch dim "𝑖" `(V "𝑖" _),
|
||||||
parseFails dim "M.x",
|
parseFails dim "M.x",
|
||||||
parseFails dim "_"
|
parseFails dim "_"
|
||||||
],
|
],
|
||||||
|
|
||||||
"quantities" :- [
|
"quantities" :- [
|
||||||
parsesAs qty "0" Zero,
|
parseMatch qty "0" `(PQ Zero _),
|
||||||
parsesAs qty "1" One,
|
parseMatch qty "1" `(PQ One _),
|
||||||
parsesAs qty "ω" Any,
|
parseMatch qty "ω" `(PQ Any _),
|
||||||
parsesAs qty "#" Any,
|
parseMatch qty "#" `(PQ Any _),
|
||||||
parseFails qty "anythingElse",
|
parseFails qty "anythingElse",
|
||||||
parseFails qty "_"
|
parseFails qty "_"
|
||||||
],
|
],
|
||||||
|
|
||||||
"enum types" :- [
|
"enum types" :- [
|
||||||
parsesAs term #"{}"# (Enum []),
|
parseMatch term #"{}"# `(Enum [] _),
|
||||||
parsesAs term #"{a}"# (Enum ["a"]),
|
parseMatch term #"{a}"# `(Enum ["a"] _),
|
||||||
parsesAs term #"{a,}"# (Enum ["a"]),
|
parseMatch term #"{a,}"# `(Enum ["a"] _),
|
||||||
parsesAs term #"{a.b.c.d}"# (Enum ["a.b.c.d"]),
|
parseMatch term #"{a.b.c.d}"# `(Enum ["a.b.c.d"] _),
|
||||||
parsesAs term #"{"hel lo"}"# (Enum ["hel lo"]),
|
parseMatch term #"{"hel lo"}"# `(Enum ["hel lo"] _),
|
||||||
parsesAs term #"{a, b}"# (Enum ["a", "b"]),
|
parseMatch term #"{a, b}"# `(Enum ["a", "b"] _),
|
||||||
parsesAs term #"{a, b,}"# (Enum ["a", "b"]),
|
parseMatch term #"{a, b,}"# `(Enum ["a", "b"] _),
|
||||||
parsesAs term #"{a, b, ","}"# (Enum ["a", "b", ","]),
|
parseMatch term #"{a, b, ","}"# `(Enum ["a", "b", ","] _),
|
||||||
parseFails term #"{,}"#
|
parseFails term #"{,}"#
|
||||||
],
|
],
|
||||||
|
|
||||||
"tags" :- [
|
"tags" :- [
|
||||||
parsesAs term #" 'a "# (Tag "a"),
|
parseMatch term #" 'a "# `(Tag "a" _),
|
||||||
parsesAs term #" 'abc "# (Tag "abc"),
|
parseMatch term #" 'abc "# `(Tag "abc" _),
|
||||||
parsesAs term #" '"abc" "# (Tag "abc"),
|
parseMatch term #" '"abc" "# `(Tag "abc" _),
|
||||||
parsesAs term #" '"a b c" "# (Tag "a b c"),
|
parseMatch term #" '"a b c" "# `(Tag "a b c" _),
|
||||||
parsesAs term #" 'a b c "# (Tag "a" :@ V "b" :@ V "c")
|
note "application to two arguments",
|
||||||
{label = "'a b c # application to two args"}
|
parseMatch term #" 'a b c "#
|
||||||
|
`(App (App (Tag "a" _) (V "b" _) _) (V "c" _) _)
|
||||||
],
|
],
|
||||||
|
|
||||||
"universes" :- [
|
"universes" :- [
|
||||||
parsesAs term "★₀" (TYPE 0),
|
parseMatch term "★₀" `(TYPE 0 _),
|
||||||
parsesAs term "★1" (TYPE 1),
|
parseMatch term "★1" `(TYPE 1 _),
|
||||||
parsesAs term "★ 2" (TYPE 2),
|
parseMatch term "★ 2" `(TYPE 2 _),
|
||||||
parsesAs term "Type₃" (TYPE 3),
|
parseMatch term "Type₃" `(TYPE 3 _),
|
||||||
parsesAs term "Type4" (TYPE 4),
|
parseMatch term "Type4" `(TYPE 4 _),
|
||||||
parsesAs term "Type 100" (TYPE 100),
|
parseMatch term "Type 100" `(TYPE 100 _),
|
||||||
parsesAs term "(Type 1000)" (TYPE 1000),
|
parseMatch term "(Type 1000)" `(TYPE 1000 _),
|
||||||
parseFails term "Type",
|
parseFails term "Type",
|
||||||
parseFails term "★"
|
parseFails term "★"
|
||||||
],
|
],
|
||||||
|
|
||||||
"applications" :- [
|
"applications" :- [
|
||||||
parsesAs term "f" (V "f"),
|
parseMatch term "f"
|
||||||
parsesAs term "f.x.y" (V $ MakePName [< "f", "x"] "y"),
|
`(V "f" _),
|
||||||
parsesAs term "f x" (V "f" :@ V "x"),
|
parseMatch term "f.x.y"
|
||||||
parsesAs term "f x y" (V "f" :@ V "x" :@ V "y"),
|
`(V (MakePName [< "f", "x"] "y") _),
|
||||||
parsesAs term "(f x) y" (V "f" :@ V "x" :@ V "y"),
|
parseMatch term "f x"
|
||||||
parsesAs term "f (g x)" (V "f" :@ (V "g" :@ V "x")),
|
`(App (V "f" _) (V "x" _) _),
|
||||||
parsesAs term "f (g x) y" (V "f" :@ (V "g" :@ V "x") :@ V "y"),
|
parseMatch term "f x y"
|
||||||
parsesAs term "f @p" (V "f" :% V "p"),
|
`(App (App (V "f" _) (V "x" _) _) (V "y" _) _),
|
||||||
parsesAs term "f x @p y" (V "f" :@ V "x" :% V "p" :@ V "y")
|
parseMatch term "(f x) y"
|
||||||
|
`(App (App (V "f" _) (V "x" _) _) (V "y" _) _),
|
||||||
|
parseMatch term "f (g x)"
|
||||||
|
`(App (V "f" _) (App (V "g" _) (V "x" _) _) _),
|
||||||
|
parseMatch term "f (g x) y"
|
||||||
|
`(App (App (V "f" _) (App (V "g" _) (V "x" _) _) _) (V "y" _) _),
|
||||||
|
parseMatch term "f @p"
|
||||||
|
`(DApp (V "f" _) (V "p" _) _),
|
||||||
|
parseMatch term "f x @p y"
|
||||||
|
`(App (DApp (App (V "f" _) (V "x" _) _) (V "p" _) _) (V "y" _) _)
|
||||||
],
|
],
|
||||||
|
|
||||||
"annotations" :- [
|
"annotations" :- [
|
||||||
parsesAs term "f :: A" (V "f" :# V "A"),
|
parseMatch term "f :: A"
|
||||||
parsesAs term "f ∷ A" (V "f" :# V "A"),
|
`(Ann (V "f" _) (V "A" _) _),
|
||||||
parsesAs term "f x y ∷ A B C"
|
parseMatch term "f ∷ A"
|
||||||
((V "f" :@ V "x" :@ V "y") :#
|
`(Ann (V "f" _) (V "A" _) _),
|
||||||
(V "A" :@ V "B" :@ V "C")),
|
parseMatch term "f x y ∷ A B C"
|
||||||
parsesAs term "Type 0 ∷ Type 1 ∷ Type 2"
|
`(Ann (App (App (V "f" _) (V "x" _) _) (V "y" _) _)
|
||||||
(TYPE 0 :# (TYPE 1 :# TYPE 2))
|
(App (App (V "A" _) (V "B" _) _) (V "C" _) _) _),
|
||||||
|
parseMatch term "Type 0 ∷ Type 1 ∷ Type 2"
|
||||||
|
`(Ann (TYPE 0 _) (Ann (TYPE 1 _) (TYPE 2 _) _) _)
|
||||||
],
|
],
|
||||||
|
|
||||||
"binders" :- [
|
"binders" :- [
|
||||||
parsesAs term "1.(x : A) → B x" $
|
parseMatch term "1.(x : A) → B x"
|
||||||
Pi One (Just "x") (V "A") (V "B" :@ V "x"),
|
`(Pi (PQ One _) (PV "x" _) (V "A" _) (App (V "B" _) (V "x" _) _) _),
|
||||||
parsesAs term "1.(x : A) -> B x" $
|
parseMatch term "1.(x : A) -> B x"
|
||||||
Pi One (Just "x") (V "A") (V "B" :@ V "x"),
|
`(Pi (PQ One _) (PV "x" _) (V "A" _) (App (V "B" _) (V "x" _) _) _),
|
||||||
parsesAs term "ω.(x : A) → B x" $
|
parseMatch term "ω.(x : A) → B x"
|
||||||
Pi Any (Just "x") (V "A") (V "B" :@ V "x"),
|
`(Pi (PQ Any _) (PV "x" _) (V "A" _) (App (V "B" _) (V "x" _) _) _),
|
||||||
parsesAs term "#.(x : A) -> B x" $
|
parseMatch term "#.(x : A) -> B x"
|
||||||
Pi Any (Just "x") (V "A") (V "B" :@ V "x"),
|
`(Pi (PQ Any _) (PV "x" _) (V "A" _) (App (V "B" _) (V "x" _) _) _),
|
||||||
parsesAs term "1.(x y : A) -> B x" $
|
parseMatch term "1.(x y : A) -> B x"
|
||||||
Pi One (Just "x") (V "A") $ Pi One (Just "y") (V "A") (V "B" :@ V "x"),
|
`(Pi (PQ One _) (PV "x" _) (V "A" _)
|
||||||
|
(Pi (PQ One _) (PV "y" _) (V "A" _)
|
||||||
|
(App (V "B" _) (V "x" _) _) _) _),
|
||||||
parseFails term "(x : A) → B x",
|
parseFails term "(x : A) → B x",
|
||||||
parsesAs term "1.A → B"
|
parseMatch term "1.A → B"
|
||||||
(Pi One Nothing (V "A") (V "B")),
|
`(Pi (PQ One _) (Unused _) (V "A" _) (V "B" _) _),
|
||||||
parsesAs term "1.(List A) → List B"
|
parseMatch term "1.(List A) → List B"
|
||||||
(Pi One Nothing (V "List" :@ V "A") (V "List" :@ V "B")),
|
`(Pi (PQ One _) (Unused _)
|
||||||
|
(App (V "List" _) (V "A" _) _)
|
||||||
|
(App (V "List" _) (V "B" _) _) _),
|
||||||
parseFails term "1.List A → List B",
|
parseFails term "1.List A → List B",
|
||||||
parsesAs term "(x : A) × B x" $
|
parseMatch term "(x : A) × B x"
|
||||||
Sig (Just "x") (V "A") (V "B" :@ V "x"),
|
`(Sig (PV "x" _) (V "A" _) (App (V "B" _) (V "x" _) _) _),
|
||||||
parsesAs term "(x : A) ** B x" $
|
parseMatch term "(x : A) ** B x"
|
||||||
Sig (Just "x") (V "A") (V "B" :@ V "x"),
|
`(Sig (PV "x" _) (V "A" _) (App (V "B" _) (V "x" _) _) _),
|
||||||
parsesAs term "(x y : A) × B x" $
|
parseMatch term "(x y : A) × B" $
|
||||||
Sig (Just "x") (V "A") $ Sig (Just "y") (V "A") (V "B" :@ V "x"),
|
`(Sig (PV "x" _) (V "A" _) (Sig (PV "y" _) (V "A" _) (V "B" _) _) _),
|
||||||
parseFails term "1.(x : A) × B x",
|
parseFails term "1.(x : A) × B x",
|
||||||
parsesAs term "A × B" $
|
parseMatch term "A × B"
|
||||||
Sig Nothing (V "A") (V "B"),
|
`(Sig (Unused _) (V "A" _) (V "B" _) _),
|
||||||
parsesAs term "A ** B" $
|
parseMatch term "A ** B"
|
||||||
Sig Nothing (V "A") (V "B"),
|
`(Sig (Unused _) (V "A" _) (V "B" _) _),
|
||||||
parsesAs term "A × B × C" $
|
parseMatch term "A × B × C" $
|
||||||
Sig Nothing (V "A") (Sig Nothing (V "B") (V "C")),
|
`(Sig (Unused _) (V "A" _) (Sig (Unused _) (V "B" _) (V "C" _) _) _),
|
||||||
parsesAs term "(A × B) × C" $
|
parseMatch term "(A × B) × C" $
|
||||||
Sig Nothing (Sig Nothing (V "A") (V "B")) (V "C")
|
`(Sig (Unused _) (Sig (Unused _) (V "A" _) (V "B" _) _) (V "C" _) _)
|
||||||
],
|
],
|
||||||
|
|
||||||
"lambdas" :- [
|
"lambdas" :- [
|
||||||
parsesAs term "λ x ⇒ x" $ Lam (Just "x") (V "x"),
|
parseMatch term "λ x ⇒ x"
|
||||||
parsesAs term "λ x ⇒ x" $ Lam (Just "x") (V "x"),
|
`(Lam (PV "x" _) (V "x" _) _),
|
||||||
parsesAs term "fun x => x" $ Lam (Just "x") (V "x"),
|
parseMatch term "fun x => x"
|
||||||
parsesAs term "δ i ⇒ x @i" $ DLam (Just "i") (V "x" :% V "i"),
|
`(Lam (PV "x" _) (V "x" _) _),
|
||||||
parsesAs term "dfun i => x @i" $ DLam (Just "i") (V "x" :% V "i"),
|
parseMatch term "δ i ⇒ x @i"
|
||||||
parsesAs term "λ x y z ⇒ x z y" $
|
`(DLam (PV "i" _) (DApp (V "x" _) (V "i" _) _) _),
|
||||||
Lam (Just "x") $ Lam (Just "y") $ Lam (Just "z") $
|
parseMatch term "dfun i => x @i"
|
||||||
V "x" :@ V "z" :@ V "y"
|
`(DLam (PV "i" _) (DApp (V "x" _) (V "i" _) _) _),
|
||||||
|
parseMatch term "λ x y z ⇒ x z y"
|
||||||
|
`(Lam (PV "x" _)
|
||||||
|
(Lam (PV "y" _)
|
||||||
|
(Lam (PV "z" _)
|
||||||
|
(App (App (V "x" _) (V "z" _) _) (V "y" _) _) _) _) _)
|
||||||
],
|
],
|
||||||
|
|
||||||
"pairs" :- [
|
"pairs" :- [
|
||||||
parsesAs term "(x, y)" $
|
parseMatch term "(x, y)"
|
||||||
Pair (V "x") (V "y"),
|
`(Pair (V "x" _) (V "y" _) _),
|
||||||
parsesAs term "(x, y, z)" $
|
parseMatch term "(x, y, z)"
|
||||||
Pair (V "x") (Pair (V "y") (V "z")),
|
`(Pair (V "x" _) (Pair (V "y" _) (V "z" _) _) _),
|
||||||
parsesAs term "((x, y), z)" $
|
parseMatch term "((x, y), z)"
|
||||||
Pair (Pair (V "x") (V "y")) (V "z"),
|
`(Pair (Pair (V "x" _) (V "y" _) _) (V "z" _) _),
|
||||||
parsesAs term "(f x, g @y)" $
|
parseMatch term "(f x, g @y)"
|
||||||
Pair (V "f" :@ V "x") (V "g" :% V "y"),
|
`(Pair (App (V "f" _) (V "x" _) _) (DApp (V "g" _) (V "y" _) _) _),
|
||||||
parsesAs term "((x : A) × B, 0.(x : C) → D)" $
|
parseMatch term "((x : A) × B, 0.(x : C) → D)"
|
||||||
Pair (Sig (Just "x") (V "A") (V "B"))
|
`(Pair (Sig (PV "x" _) (V "A" _) (V "B" _) _)
|
||||||
(Pi Zero (Just "x") (V "C") (V "D")),
|
(Pi (PQ Zero _) (PV "x" _) (V "C" _) (V "D" _) _) _),
|
||||||
parsesAs term "(λ x ⇒ x, δ i ⇒ e @i)" $
|
parseMatch term "(λ x ⇒ x, δ i ⇒ e @i)"
|
||||||
Pair (Lam (Just "x") (V "x"))
|
`(Pair (Lam (PV "x" _) (V "x" _) _)
|
||||||
(DLam (Just "i") (V "e" :% V "i")),
|
(DLam (PV "i" _) (DApp (V "e" _) (V "i" _) _) _) _),
|
||||||
parsesAs term "(x,)" (V "x"), -- i GUESS
|
parseMatch term "(x,)" `(V "x" _), -- i GUESS
|
||||||
parseFails term "(,y)",
|
parseFails term "(,y)",
|
||||||
parseFails term "(x,,y)"
|
parseFails term "(x,,y)"
|
||||||
],
|
],
|
||||||
|
|
||||||
"equality type" :- [
|
"equality type" :- [
|
||||||
parsesAs term "Eq [i ⇒ A] s t" $
|
parseMatch term "Eq [i ⇒ A] s t"
|
||||||
Eq (Just "i", V "A") (V "s") (V "t"),
|
`(Eq (PV "i" _, V "A" _) (V "s" _) (V "t" _) _),
|
||||||
parsesAs term "Eq [i ⇒ A B (C @i)] (f x y) (g y z)" $
|
parseMatch term "Eq [i ⇒ A (B @i)] (f x) (g y)"
|
||||||
Eq (Just "i", V "A" :@ V "B" :@ (V "C" :% V "i"))
|
`(Eq (PV "i" _, App (V "A" _) (DApp (V "B" _) (V "i" _) _) _)
|
||||||
(V "f" :@ V "x" :@ V "y") (V "g" :@ V "y" :@ V "z"),
|
(App (V "f" _) (V "x" _) _)
|
||||||
parsesAs term "Eq [A] s t" $
|
(App (V "g" _) (V "y" _) _) _),
|
||||||
Eq (Nothing, V "A") (V "s") (V "t"),
|
parseMatch term "Eq [A] s t"
|
||||||
parsesAs term "s ≡ t : A" $
|
`(Eq (Unused _, V "A" _) (V "s" _) (V "t" _) _),
|
||||||
Eq (Nothing, V "A") (V "s") (V "t"),
|
parseMatch term "s ≡ t : A"
|
||||||
parsesAs term "s == t : A" $
|
`(Eq (Unused _, V "A" _) (V "s" _) (V "t" _) _),
|
||||||
Eq (Nothing, V "A") (V "s") (V "t"),
|
parseMatch term "s == t : A"
|
||||||
parsesAs term "f x y ≡ g y z : A B C" $
|
`(Eq (Unused _, V "A" _) (V "s" _) (V "t" _) _),
|
||||||
Eq (Nothing, V "A" :@ V "B" :@ V "C")
|
parseMatch term "f x ≡ g y : A B"
|
||||||
(V "f" :@ V "x" :@ V "y") (V "g" :@ V "y" :@ V "z"),
|
`(Eq (Unused _, App (V "A" _) (V "B" _) _)
|
||||||
parsesAs term "(A × B) ≡ (A' × B') : ★₁" $
|
(App (V "f" _) (V "x" _) _)
|
||||||
Eq (Nothing, TYPE 1)
|
(App (V "g" _) (V "y" _) _) _),
|
||||||
(Sig Nothing (V "A") (V "B")) (Sig Nothing (V "A'") (V "B'")),
|
parseMatch term "(A × B) ≡ (A' × B') : ★₁"
|
||||||
|
`(Eq (Unused _, TYPE 1 _)
|
||||||
|
(Sig (Unused _) (V "A" _) (V "B" _) _)
|
||||||
|
(Sig (Unused _) (V "A'" _) (V "B'" _) _) _),
|
||||||
note "A × (B ≡ A' × B' : ★₁)",
|
note "A × (B ≡ A' × B' : ★₁)",
|
||||||
parsesAs term "A × B ≡ A' × B' : ★₁" $
|
parseMatch term "A × B ≡ A' × B' : ★₁"
|
||||||
Sig Nothing (V "A") $
|
`(Sig (Unused _) (V "A" _)
|
||||||
Eq (Nothing, TYPE 1)
|
(Eq (Unused _, TYPE 1 _)
|
||||||
(V "B") (Sig Nothing (V "A'") (V "B'")),
|
(V "B" _) (Sig (Unused _) (V "A'" _) (V "B'" _) _) _) _),
|
||||||
parseFails term "Eq",
|
parseFails term "Eq",
|
||||||
parseFails term "Eq s t",
|
parseFails term "Eq s t",
|
||||||
parseFails term "s ≡ t",
|
parseFails term "s ≡ t",
|
||||||
|
@ -234,127 +260,144 @@ tests = "parser" :- [
|
||||||
],
|
],
|
||||||
|
|
||||||
"naturals" :- [
|
"naturals" :- [
|
||||||
parsesAs term "ℕ" Nat,
|
parseMatch term "ℕ" `(Nat _),
|
||||||
parsesAs term "Nat" Nat,
|
parseMatch term "Nat" `(Nat _),
|
||||||
parsesAs term "zero" Zero,
|
parseMatch term "zero" `(Zero _),
|
||||||
parsesAs term "succ n" (Succ $ V "n"),
|
parseMatch term "succ n" `(Succ (V "n" _) _),
|
||||||
parsesAs term "3" (fromNat 3),
|
parseMatch term "3"
|
||||||
parsesAs term "succ (succ 5)" (fromNat 7),
|
`(Succ (Succ (Succ (Zero _) _) _) _),
|
||||||
|
parseMatch term "succ (succ 1)"
|
||||||
|
`(Succ (Succ (Succ (Zero _) _) _) _),
|
||||||
parseFails term "succ succ 5",
|
parseFails term "succ succ 5",
|
||||||
parseFails term "succ"
|
parseFails term "succ"
|
||||||
],
|
],
|
||||||
|
|
||||||
"box" :- [
|
"box" :- [
|
||||||
parsesAs term "[1.ℕ]" $ BOX One Nat,
|
parseMatch term "[1.ℕ]"
|
||||||
parsesAs term "[ω. ℕ × ℕ]" $ BOX Any (Sig Nothing Nat Nat),
|
`(BOX (PQ One _) (Nat _) _),
|
||||||
parsesAs term "[a]" $ Box (V "a"),
|
parseMatch term "[ω. ℕ × ℕ]"
|
||||||
parsesAs term "[0]" $ Box Zero,
|
`(BOX (PQ Any _) (Sig (Unused _) (Nat _) (Nat _) _) _),
|
||||||
parsesAs term "[1]" $ Box (Succ Zero)
|
parseMatch term "[a]"
|
||||||
|
`(Box (V "a" _) _),
|
||||||
|
parseMatch term "[0]"
|
||||||
|
`(Box (Zero _) _),
|
||||||
|
parseMatch term "[1]"
|
||||||
|
`(Box (Succ (Zero _) _) _)
|
||||||
],
|
],
|
||||||
|
|
||||||
"coe" :- [
|
"coe" :- [
|
||||||
parsesAs term "coe [A] @p @q x" $
|
parseMatch term "coe [A] @p @q x"
|
||||||
Coe (Nothing, V "A") (V "p") (V "q") (V "x"),
|
`(Coe (Unused _, V "A" _) (V "p" _) (V "q" _) (V "x" _) _),
|
||||||
parsesAs term "coe [i ⇒ A] @p @q x" $
|
parseMatch term "coe [i ⇒ A] @p @q x"
|
||||||
Coe (Just "i", V "A") (V "p") (V "q") (V "x"),
|
`(Coe (PV "i" _, V "A" _) (V "p" _) (V "q" _) (V "x" _) _),
|
||||||
parseFails term "coe [A] @p @q",
|
parseFails term "coe [A] @p @q",
|
||||||
parseFails term "coe A @p @q x",
|
parseFails term "coe A @p @q x",
|
||||||
parseFails term "coe [i ⇒ A] @p q x"
|
parseFails term "coe [i ⇒ A] @p q x"
|
||||||
],
|
],
|
||||||
|
|
||||||
"comp" :- [
|
"comp" :- [
|
||||||
parsesAs term "comp [A] @p @q s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }" $
|
parseMatch term "comp [A] @p @q s @r { 0 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁ }"
|
||||||
Comp (Nothing, V "A") (V "p") (V "q") (V "s") (V "r")
|
`(Comp (Unused _, V "A" _) (V "p" _) (V "q" _) (V "s" _) (V "r" _)
|
||||||
(Just "𝑗", V "s₀") (Just "𝑘", V "s₁"),
|
(PV "𝑗" _, V "s₀" _) (PV "𝑘" _, V "s₁" _) _),
|
||||||
parsesAs term "comp [A] @p @q s @r { 1 𝑗 ⇒ s₀; 0 𝑘 ⇒ s₁; }" $
|
parseMatch term "comp [A] @p @q s @r { 1 𝑗 ⇒ s₀; 0 𝑘 ⇒ s₁; }"
|
||||||
Comp (Nothing, V "A") (V "p") (V "q") (V "s") (V "r")
|
`(Comp (Unused _, V "A" _) (V "p" _) (V "q" _) (V "s" _) (V "r" _)
|
||||||
(Just "𝑘", V "s₁") (Just "𝑗", V "s₀"),
|
(PV "𝑘" _, V "s₁" _) (PV "𝑗" _, V "s₀" _) _),
|
||||||
parseFails term "comp [A] @p @q s @r { 1 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁; }",
|
parseFails term "comp [A] @p @q s @r { 1 𝑗 ⇒ s₀; 1 𝑘 ⇒ s₁; }",
|
||||||
parseFails term "comp [A] @p @q s @r { 0 𝑗 ⇒ s₀ }",
|
parseFails term "comp [A] @p @q s @r { 0 𝑗 ⇒ s₀ }",
|
||||||
parseFails term "comp [A] @p @q s @r { }"
|
parseFails term "comp [A] @p @q s @r { }"
|
||||||
],
|
],
|
||||||
|
|
||||||
"case" :- [
|
"case" :- [
|
||||||
parsesAs term
|
parseMatch term
|
||||||
"case1 f s return x ⇒ A x of { (l, r) ⇒ add l r }" $
|
"case1 f s return x ⇒ A x of { (l, r) ⇒ r l }"
|
||||||
Case One (V "f" :@ V "s")
|
`(Case (PQ One _) (App (V "f" _) (V "s" _) _)
|
||||||
(Just "x", V "A" :@ V "x")
|
(PV "x" _, App (V "A" _) (V "x" _) _)
|
||||||
(CasePair (Just "l", Just "r")
|
(CasePair (PV "l" _, PV "r" _)
|
||||||
(V "add" :@ V "l" :@ V "r")),
|
(App (V "r" _) (V "l" _) _) _) _),
|
||||||
parsesAs term
|
parseMatch term
|
||||||
"case1 f s return x ⇒ A x of { (l, r) ⇒ add l r; }" $
|
"case1 f s return x => A x of { (l, r) ⇒ r l; }"
|
||||||
Case One (V "f" :@ V "s")
|
`(Case (PQ One _) (App (V "f" _) (V "s" _) _)
|
||||||
(Just "x", V "A" :@ V "x")
|
(PV "x" _, App (V "A" _) (V "x" _) _)
|
||||||
(CasePair (Just "l", Just "r")
|
(CasePair (PV "l" _, PV "r" _)
|
||||||
(V "add" :@ V "l" :@ V "r")),
|
(App (V "r" _) (V "l" _) _) _) _),
|
||||||
parsesAs term
|
parseMatch term
|
||||||
"case 1 . f s return x ⇒ A x of { (l, r) ⇒ add l r }" $
|
"case 1 . f s return x ⇒ A x of { (l, r) ⇒ r l }"
|
||||||
Case One (V "f" :@ V "s")
|
`(Case (PQ One _) (App (V "f" _) (V "s" _) _)
|
||||||
(Just "x", V "A" :@ V "x")
|
(PV "x" _, App (V "A" _) (V "x" _) _)
|
||||||
(CasePair (Just "l", Just "r")
|
(CasePair (PV "l" _, PV "r" _)
|
||||||
(V "add" :@ V "l" :@ V "r")),
|
(App (V "r" _) (V "l" _) _) _) _),
|
||||||
parsesAs term
|
parseMatch term
|
||||||
"case1 t return A of { 'x ⇒ p; 'y ⇒ q; 'z ⇒ r }" $
|
"case1 t return A of { 'x ⇒ p; 'y ⇒ q; 'z ⇒ r }"
|
||||||
Case One (V "t")
|
`(Case (PQ One _) (V "t" _)
|
||||||
(Nothing, V "A")
|
(Unused _, V "A" _)
|
||||||
(CaseEnum [("x", V "p"), ("y", V "q"), ("z", V "r")]),
|
(CaseEnum [(PT "x" _, V "p" _),
|
||||||
parsesAs term "caseω t return A of {}" $
|
(PT "y" _, V "q" _),
|
||||||
Case Any (V "t") (Nothing, V "A") (CaseEnum []),
|
(PT "z" _, V "r" _)] _) _),
|
||||||
parsesAs term "case# t return A of {}" $
|
parseMatch term "caseω t return A of {}"
|
||||||
Case Any (V "t") (Nothing, V "A") (CaseEnum []),
|
`(Case (PQ Any _) (V "t" _) (Unused _, V "A" _) (CaseEnum [] _) _),
|
||||||
parsesAs term "caseω n return A of { 0 ⇒ a; succ n' ⇒ b }" $
|
parseMatch term "case# t return A of {}"
|
||||||
Case Any (V "n") (Nothing, V "A") $
|
`(Case (PQ Any _) (V "t" _) (Unused _, V "A" _) (CaseEnum [] _) _),
|
||||||
CaseNat (V "a") (Just "n'", Zero, Nothing, V "b"),
|
parseMatch term "caseω n return A of { 0 ⇒ a; succ n' ⇒ b }"
|
||||||
parsesAs term "caseω n return ℕ of { succ _, 1.ih ⇒ ih; zero ⇒ 0; }" $
|
`(Case (PQ Any _) (V "n" _) (Unused _, V "A" _)
|
||||||
Case Any (V "n") (Nothing, Nat) $
|
(CaseNat (V "a" _) (PV "n'" _, PQ Zero _, Unused _, V "b" _) _) _),
|
||||||
CaseNat Zero (Nothing, One, Just "ih", V "ih"),
|
parseMatch term "caseω n return ℕ of { succ _, 1.ih ⇒ ih; zero ⇒ 0; }"
|
||||||
|
`(Case (PQ Any _) (V "n" _) (Unused _, Nat _)
|
||||||
|
(CaseNat (Zero _) (Unused _, PQ One _, PV "ih" _, V "ih" _) _) _),
|
||||||
parseFails term "caseω n return A of { zero ⇒ a }",
|
parseFails term "caseω n return A of { zero ⇒ a }",
|
||||||
parseFails term "caseω n return ℕ of { succ ⇒ 5 }"
|
parseFails term "caseω n return ℕ of { succ ⇒ 5 }"
|
||||||
],
|
],
|
||||||
|
|
||||||
"definitions" :- [
|
"definitions" :- [
|
||||||
parsesAs definition "defω x : {a} × {b} = ('a, 'b)" $
|
parseMatch definition "defω x : {a} × {b} = ('a, 'b);"
|
||||||
MkPDef Any "x" (Just $ Sig Nothing (Enum ["a"]) (Enum ["b"]))
|
`(MkPDef (PQ Any _) "x"
|
||||||
(Pair (Tag "a") (Tag "b")),
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
parsesAs definition "defω x : {a} × {b} = ('a, 'b)" $
|
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
||||||
MkPDef Any "x" (Just $ Sig Nothing (Enum ["a"]) (Enum ["b"]))
|
parseMatch definition "def# x : {a} ** {b} = ('a, 'b)"
|
||||||
(Pair (Tag "a") (Tag "b")),
|
`(MkPDef (PQ Any _) "x"
|
||||||
parsesAs definition "def# x : {a} ** {b} = ('a, 'b)" $
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
MkPDef Any "x" (Just $ Sig Nothing (Enum ["a"]) (Enum ["b"]))
|
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
||||||
(Pair (Tag "a") (Tag "b")),
|
parseMatch definition "def ω.x : {a} × {b} = ('a, 'b)"
|
||||||
parsesAs definition "def ω.x : {a} × {b} = ('a, 'b)" $
|
`(MkPDef (PQ Any _) "x"
|
||||||
MkPDef Any "x" (Just $ Sig Nothing (Enum ["a"]) (Enum ["b"]))
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
(Pair (Tag "a") (Tag "b")),
|
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
||||||
parsesAs definition "def x : {a} × {b} = ('a, 'b)" $
|
parseMatch definition "def x : {a} × {b} = ('a, 'b)"
|
||||||
MkPDef Any "x" (Just $ Sig Nothing (Enum ["a"]) (Enum ["b"]))
|
`(MkPDef (PQ Any _) "x"
|
||||||
(Pair (Tag "a") (Tag "b")),
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
parsesAs definition "def0 A : ★₀ = {a, b, c}" $
|
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
||||||
MkPDef Zero "A" (Just $ TYPE 0) (Enum ["a", "b", "c"])
|
parseMatch definition "def0 A : ★₀ = {a, b, c}"
|
||||||
|
`(MkPDef (PQ Zero _) "A" (Just $ TYPE 0 _)
|
||||||
|
(Enum ["a", "b", "c"] _) _)
|
||||||
],
|
],
|
||||||
|
|
||||||
"top level" :- [
|
"top level" :- [
|
||||||
parsesAs input "def0 A : ★₀ = {}; def0 B : ★₁ = A;" $
|
parseMatch input "def0 A : ★₀ = {}; def0 B : ★₁ = A;"
|
||||||
[PD $ PDef $ MkPDef Zero "A" (Just $ TYPE 0) (Enum []),
|
`([PD $ PDef $ MkPDef (PQ Zero _) "A" (Just $ TYPE 0 _) (Enum [] _) _,
|
||||||
PD $ PDef $ MkPDef Zero "B" (Just $ TYPE 1) (V "A")],
|
PD $ PDef $ MkPDef (PQ Zero _) "B" (Just $ TYPE 1 _) (V "A" _) _]),
|
||||||
parsesAs input "def0 A : ★₀ = {} def0 B : ★₁ = A" $
|
parseMatch input "def0 A : ★₀ = {} def0 B : ★₁ = A" $
|
||||||
[PD $ PDef $ MkPDef Zero "A" (Just $ TYPE 0) (Enum []),
|
`([PD $ PDef $ MkPDef (PQ Zero _) "A" (Just $ TYPE 0 _) (Enum [] _) _,
|
||||||
PD $ PDef $ MkPDef Zero "B" (Just $ TYPE 1) (V "A")],
|
PD $ PDef $ MkPDef (PQ Zero _) "B" (Just $ TYPE 1 _) (V "A" _) _]),
|
||||||
parsesAs input "" [] {label = "[empty input]"},
|
note "empty input",
|
||||||
|
parsesAs input "" [],
|
||||||
parseFails input ";;;;;;;;;;;;;;;;;;;;;;;;;;",
|
parseFails input ";;;;;;;;;;;;;;;;;;;;;;;;;;",
|
||||||
parsesAs input "namespace a {}" [PD $ PNs $ MkPNamespace [< "a"] []],
|
parseMatch input "namespace a {}"
|
||||||
parsesAs input "namespace a.b.c {}"
|
`([PD $ PNs $ MkPNamespace [< "a"] [] _]),
|
||||||
[PD $ PNs $ MkPNamespace [< "a", "b", "c"] []],
|
parseMatch input "namespace a.b.c {}"
|
||||||
parsesAs input "namespace a {namespace b {}}"
|
`([PD $ PNs $ MkPNamespace [< "a", "b", "c"] [] _]),
|
||||||
[PD $ PNs $ MkPNamespace [< "a"] [PNs $ MkPNamespace [< "b"] []]],
|
parseMatch input "namespace a {namespace b {}}"
|
||||||
parsesAs input "namespace a {def x = 't ∷ {t}}"
|
`([PD $ PNs $ MkPNamespace [< "a"] [PNs $ MkPNamespace [< "b"] [] _] _]),
|
||||||
[PD $ PNs $ MkPNamespace [< "a"]
|
parseMatch input "namespace a {def x = 't ∷ {t}}"
|
||||||
[PDef $ MkPDef Any "x" Nothing (Tag "t" :# Enum ["t"])]],
|
`([PD $ PNs $ MkPNamespace [< "a"]
|
||||||
parsesAs input "namespace a {def x = 't ∷ {t}} def y = a.x"
|
[PDef $ MkPDef (PQ Any _) "x" Nothing
|
||||||
[PD $ PNs $ MkPNamespace [< "a"]
|
(Ann (Tag "t" _) (Enum ["t"] _) _) _] _]),
|
||||||
[PDef $ MkPDef Any "x" Nothing (Tag "t" :# Enum ["t"])],
|
parseMatch input "namespace a {def x = 't ∷ {t}} def y = a.x"
|
||||||
PD $ PDef $ MkPDef Any "y" Nothing (V $ MakePName [< "a"] "x")],
|
`([PD $ PNs $ MkPNamespace [< "a"]
|
||||||
parsesAs input #" load "a.quox"; def b = a.b "#
|
[PDef $ MkPDef (PQ Any _) "x" Nothing
|
||||||
[PLoad "a.quox",
|
(Ann (Tag "t" _) (Enum ["t"] _) _) _] _,
|
||||||
PD $ PDef $ MkPDef Any "b" Nothing (V $ MakePName [< "a"] "b")]
|
PD $ PDef $ MkPDef (PQ Any _) "y" Nothing
|
||||||
|
(V (MakePName [< "a"] "x") _) _]),
|
||||||
|
parseMatch input #" load "a.quox"; def b = a.b "#
|
||||||
|
`([PLoad "a.quox" _,
|
||||||
|
PD $ PDef $ MkPDef (PQ Any _) "b" Nothing
|
||||||
|
(V (MakePName [< "a"] "b") _) _])
|
||||||
]
|
]
|
||||||
]
|
]
|
||||||
|
|
Loading…
Reference in a new issue