Lib/TT.idr is well scoped

This commit is contained in:
2023-07-13 20:06:03 -07:00
parent ed3ee96df9
commit 59f726ab96
6 changed files with 203 additions and 90 deletions

View File

@@ -1,14 +1,23 @@
Parser is in place.
I think we almost have enough of a parser to take another step. Currently trying to make TT well-scoped, considering well-named..
Need to update todos below.
Should I use well-scoped indices or well-scoped names...
Parser:
- [x] import statement - [x] import statement
- [x] def - [x] def
- [x] simple decl - [x] simple decl
- [ ] type definition - [ ] type definition
- [ ] read files - [x] read files
- [ ] write tests
- [ ] type checking / elab
- [ ]
- [ ]
- [ ]
- [ ] symbolic execution - [ ] symbolic execution
- [ ] compilation - [ ] compilation
- [ ] write tests

BIN
papers/prettier.pdf Normal file

Binary file not shown.

View File

@@ -112,7 +112,8 @@ letExpr = do
alts <- startBlock $ someSame $ letAssign alts <- startBlock $ someSame $ letAssign
keyword' "in" keyword' "in"
scope <- term scope <- term
pure $ RLet alts scope
pure $ foldl (\ acc, (n,v) => RLet n RHole v acc) scope alts
where where
letAssign : Parser (Name,Raw) letAssign : Parser (Name,Raw)
letAssign = do letAssign = do
@@ -178,7 +179,7 @@ expBinder = do
sym ")" sym ")"
sym "->" sym "->"
scope <- typeExpr scope <- typeExpr
pure $ RPi name Explicit ty scope pure $ RPi (Just name) Explicit ty scope
impBinder : Parser Raw impBinder : Parser Raw
impBinder = do impBinder = do
@@ -189,7 +190,7 @@ impBinder = do
sym "}" sym "}"
sym "->" sym "->"
scope <- typeExpr scope <- typeExpr
pure $ RPi name Implicit ty scope pure $ RPi (Just name) Implicit ty scope
-- something binder looking -- something binder looking
-- todo sepby space or whatever -- todo sepby space or whatever
@@ -205,7 +206,7 @@ typeExpr = binder
case scope of case scope of
Nothing => pure exp Nothing => pure exp
-- consider Maybe String to represent missing -- consider Maybe String to represent missing
(Just scope) => pure $ RPi "_" Explicit exp scope (Just scope) => pure $ RPi Nothing Explicit exp scope
-- And top level stuff -- And top level stuff

View File

@@ -1,111 +1,150 @@
module Lib.TT module Lib.TT
-- For SourcePos -- For SourcePos
import Lib.Parser.Impl import Lib.Parser.Impl
import Data.Fin
import Data.Vect
public export public export
Name : Type Name : Type
Name = String Name = String
-- Trying to do well-scoped here, so the indices are proven.
export export
data Icit = Implicit | Explicit data Icit = Implicit | Explicit
-- Sorta cribbed from Kovacs %name Icit icit
Ty : Type
-- Idris and Kovacs have Icit at this level.
public export
data Tm
= Local Nat -- IsVar
| Ref String
| Lam Name Icit Tm
| App Tm Tm
| U
| Pi Name Ty Ty
| Let Name Ty Tm Tm
Ty = Tm
public export public export
data Closure : Type data Tm : Nat -> Nat -> Type where
VTy : Type Local : Fin k -> Tm k n
Bnd : Fin n -> Tm k n
Ref : String -> Tm k n
Lam : Name -> Icit -> Tm k (S n) -> Tm k n
App : Tm k n -> Tm k n -> Tm k n
U : Tm k n
Pi : Name -> Icit -> Tm k n -> Tm k (S n) -> Tm k n
Let : Name -> Icit -> Tm k n -> Tm k n -> Tm k (S n) -> Tm k n
-- Closure unpacked in the original %name Tm t, u, v
public export
data Val
= VVar Nat -- level
| VApp Val (Lazy Val)
| VLam Name Icit Closure
| VPi Name (Lazy VTy) Closure
| VU
VTy = Val -- public export
-- data Closure : Nat -> Type
data Val : Nat -> Type
0 Closure : Nat -> Type
-- IS/TypeTheory.idr is calling this a Kripke function space
-- Yaffle, IS/TypeTheory use a function here.
-- Kovacs, Idris use Env and Tm
Closure n = (l : Nat) -> Val (l + n) -> Val (l + n)
public export public export
Env : Type data Val : Nat -> Type where
Env = List Val -- This will be local / flex with spine.
VVar : Fin n -> Val n
VRef : String -> Val n
VApp : Val n -> Lazy (Val n) -> Val n
VLam : Name -> Icit -> Closure n -> Val n
VPi : Name -> Icit -> Lazy (Val n) -> Closure n -> Val n
VU : Val n
-- ||| Env k n holds the evaluation environment.
||| k is the number of levels and n is the size
||| of the environment.
public export
Env : Nat -> Nat -> Type
Env k n = Vect n (Val k)
lvl2Ix : Nat -> Nat -> Nat export
eval : Env k n -> Tm k n -> Val k
data Closure : Type where vapp : Val k -> Val k -> Val k
MkClosure : Env -> Tm -> Closure vapp (VLam _ icit t) u = t 0 u
vapp t u = VApp t u
infixl 8 $$ -- weakenEnv : (l : Nat) -> Env k n -> Env (l + k) n
eval : Env -> Tm -> Val weakenVal : {e : Nat} -> Val k -> Val (e + k)
weakenVal (VVar x) = VVar (shift _ x)
weakenVal (VRef str) = VRef str
weakenVal (VApp x y) = VApp (weakenVal x) (weakenVal y)
weakenVal (VLam str icit f) = VLam str icit
(\g, v => rewrite plusAssociative g e k in f (g + e) (rewrite sym $ plusAssociative g e k in v))
weakenVal (VPi str icit x f) = VPi str icit (weakenVal {e} x)
(\g, v => rewrite plusAssociative g e k in f (g + e) (rewrite sym $ plusAssociative g e k in v))
weakenVal VU = VU
($$) : Closure -> Lazy Val -> Val bind : (e : Nat) -> Val (plus e k) -> Env k n -> Env (e + k) (S n)
(MkClosure env t) $$ u = eval (u :: env) t bind e v env = v :: map weakenVal env
eval env (Local k) = ?hole -- is this weaken or thin?
eval env (Ref x) = ?hole_1 weaken : {e : Nat} -> Tm k (S n) -> Tm (plus e k) (S n)
eval env (Lam x _ t) = ?hole_2 weaken (Local x) = Local (shift _ x)
eval env (App t u) = case (eval env t, eval env u) of weaken (Bnd x) = Bnd x
(VLam _ icit t, u) => t $$ u weaken (Ref str) = Ref str
(t, u) => VApp t u weaken (Lam str x t) = Lam str x (weaken t)
weaken (App t u) = App (weaken t) (weaken u)
weaken U = U
weaken (Pi str x t u) = Pi str x (weaken t) (weaken u)
weaken (Let str x t u v) = Let str x (weaken t) (weaken u) (weaken v)
eval env (Local x) = VVar x -- this is a hole in intrinsic, vfree x in the other
eval env (Ref x) = VRef x
eval env (Bnd n) = index n env
eval env (Lam x icit t) = VLam x icit (\e, u => eval (bind e u env) (weaken t))-- (MkClosure env t)
eval env (App t u) = vapp (eval env t) (eval env u)
eval env U = VU eval env U = VU
eval env (Pi x a b) = VPi x (eval env a) (MkClosure env b) eval env (Pi x icit a b) = VPi x icit (eval env a) (\e, u => eval (bind e u env) (weaken b))
eval env (Let x _ t u) = eval (eval env t :: env) u -- This one we need to make
eval env (Let x icit ty t u) = eval (eval env t :: env) u
quote : Nat -> Val -> Tm vfresh : (l : Nat) -> Val (S l)
quote l (VVar k) = Local (lvl2Ix l k) vfresh l = VVar last
quote : (k : Nat) -> Val k -> Tm 0 k
quote l (VVar k) = Bnd (complement k) -- level to index
quote l (VApp t u) = App (quote l t) (quote l u) quote l (VApp t u) = App (quote l t) (quote l u)
quote l (VLam x icit t) = Lam x icit (quote (l + 1) (t $$ VVar l)) -- so this one is calling the kripke on [x] and a fresh var
quote l (VPi x a b) = Pi x (quote l a) (quote (l+1) (b $$ VVar l)) quote l (VLam x icit t) = Lam x icit (quote (S l) (t 1 (vfresh l)))
quote l VU = ?rhs_4 quote l (VPi x icit a b) = Pi x icit (quote l a) (quote (S l) (b 1 $ vfresh l))
quote l VU = U
quote _ (VRef n) = Ref n
nf : Env -> Tm -> Tm nf : {n : Nat} -> Env 0 n -> Tm 0 n -> Tm 0 0
nf env t = quote (length env) (eval env t) nf env t = quote 0 (eval env t)
---
public export
conv : (lvl : Nat) -> Val -> Val -> Bool
--
public export
Types : Type
Types = List (Name, Lazy VTy)
public export public export
record Ctx where conv : (lvl : Nat) -> Val n -> Val n -> Bool
data BD = Bound | Defined
public export
Types : Nat -> Type
Types n = Vect n (Name, Lazy (Val n))
public export
record Ctx (n : Nat) where
constructor MkCtx constructor MkCtx
env : Env env : Env k n -- for eval
types : Types types : Types n -- name lookup, pp
lvl : Nat bds : Vect n BD -- meta creation
-- For now, we're following Kovacs and using a node for lvl : Nat -- This is n, do we need it?
-- source position. Might switch to FC at some point? -- Kovacs and Weirich use a position node
pos : SourcePos pos : SourcePos
public export %name Ctx ctx
emptyCtx : Ctx
emptyCtx = MkCtx [] [] 0 (0,0)
public export public export
bind : Name -> Lazy VTy -> Ctx -> Ctx emptyCtx : Ctx Z
bind x a (MkCtx env types l pos) = emptyCtx = MkCtx {k=0} [] [] [] 0 (0,0)
MkCtx (VVar l :: env) ((x,a) :: types) (l+1) pos
-- public export
-- bind : Name -> Lazy (Val (k + n)) -> Ctx n -> Ctx (S n)
-- bind x a (MkCtx env types bds l pos) =
-- MkCtx (VVar l :: env) ((x,a) :: types) (Bound :: bds) (l+1) pos
-- public export
-- define : Name -> Val n -> Lazy (Val n) -> Ctx n -> Ctx (S n)
-- define x v ty (MkCtx env types bds l pos) =
-- MkCtx (v :: env) ((x,ty) :: types) (Defined :: bds) (l + 1) pos

View File

@@ -11,6 +11,8 @@ import System
import System.File import System.File
import System.Directory import System.Directory
import Control.App import Control.App
import Syntax
import Lib.Prettier
-- Put stuff in attic -- Put stuff in attic
-- Error printing -- Error printing
@@ -30,8 +32,10 @@ test pa src = do
putStrLn "- Toks" putStrLn "- Toks"
printLn $ toks printLn $ toks
putStrLn "- Parse" putStrLn "- Parse"
let res = parse pa toks let Right res = parse pa toks
printLn res | Left y => putStrLn "Error: \{y}"
printLn $ res
-- let toks2 = layout toks -- let toks2 = layout toks
-- printLn $ map value toks2 -- printLn $ map value toks2
@@ -40,9 +44,13 @@ test pa src = do
testFile : String -> IO () testFile : String -> IO ()
testFile fn = do testFile fn = do
putStrLn ("***" ++ fn) putStrLn ("***" ++ fn)
Right text <- readFile $ "eg/" ++ fn Right src <- readFile $ "eg/" ++ fn
| Left err => printLn err | Left err => printLn err
test parseMod text let toks = tokenise src
let Right res = parse parseMod toks
| Left y => putStrLn "Error: \{y}"
putStrLn $ pretty 80 $ pretty res
olderTests : IO () olderTests : IO ()
olderTests = do olderTests = do

View File

@@ -1,7 +1,9 @@
module Syntax module Syntax
import Data.String import Data.String
import Data.Maybe
import Lib.Parser.Impl import Lib.Parser.Impl
import Lib.Prettier
Name = String Name = String
@@ -14,7 +16,9 @@ public export
data RigCount = Rig0 | RigW data RigCount = Rig0 | RigW
-- I think I got Eq from pi-forall, it uses it for equality args (which are kinda like Prop/Rig0?) -- I think I got Eq from pi-forall, it uses it for equality args (which are kinda like Prop/Rig0?)
public export public export
data Plicity = Implicit | Explicit | Eq data Plicity = Implicit | Explicit -- | Eq
%name Plicity icit
public export public export
data Pattern data Pattern
@@ -29,14 +33,16 @@ data Pattern
public export public export
data CaseAlt = MkAlt Pattern Raw data CaseAlt = MkAlt Pattern Raw
-- TODO redo this with names for documentation
public export public export
data Raw data Raw
= RVar Name = RVar Name
| RLam String Plicity Raw | RLam String Plicity Raw
| RApp Raw Raw Plicity | RApp Raw Raw Plicity
| RU | RU
| RPi Name Plicity Raw Raw | RPi (Maybe Name) Plicity Raw Raw
| RLet (List (Name, Raw)) Raw | RLet Name Raw Raw Raw
| RSrcPos SourcePos Raw | RSrcPos SourcePos Raw
| RAnn Raw Raw | RAnn Raw Raw
@@ -45,6 +51,8 @@ data Raw
| RHole | RHole
| RParseError String | RParseError String
%name Raw tm
-- derive some stuff - I'd like json, eq, show, ... -- derive some stuff - I'd like json, eq, show, ...
data Decl : Type where data Decl : Type where
@@ -118,7 +126,7 @@ Show CaseAlt where
Show Plicity where Show Plicity where
show Implicit = "Implicit" show Implicit = "Implicit"
show Explicit = "Explicit" show Explicit = "Explicit"
show Eq = "Eq" -- show Eq = "Eq"
covering covering
Show Raw where Show Raw where
@@ -126,7 +134,7 @@ Show Raw where
show (RVar name) = foo ["RVar", show name] show (RVar name) = foo ["RVar", show name]
show (RAnn t ty) = foo [ "RAnn", show t, show ty] show (RAnn t ty) = foo [ "RAnn", show t, show ty]
show (RLit x) = foo [ "RLit", show x] show (RLit x) = foo [ "RLit", show x]
show (RLet alts y) = foo [ "Let", show alts, show y] show (RLet x ty v scope) = foo [ "Let", show x, " : ", show ty, " = ", show v, " in ", show scope]
show (RPi str x y z) = foo [ "Pi", show str, show x, show y, show z] show (RPi str x y z) = foo [ "Pi", show str, show x, show y, show z]
show (RApp x y z) = foo [ "App", show x, show y, show z] show (RApp x y z) = foo [ "App", show x, show y, show z]
show (RLam x i y) = foo [ "Lam", show x, show i, show y] show (RLam x i y) = foo [ "Lam", show x, show i, show y]
@@ -135,4 +143,52 @@ Show Raw where
show RU = "U" show RU = "U"
show (RSrcPos pos tm) = show tm show (RSrcPos pos tm) = show tm
export
interface Pretty a where
pretty : a -> Doc
export
Pretty Raw where
pretty = asDoc 0
where
wrap : Plicity -> Doc -> Doc
wrap Implicit x = x
wrap Explicit x = text "{" ++ x ++ text "}"
par : Nat -> Nat -> Doc -> Doc
par p p' d = if p' < p then text "(" ++ d ++ text ")" else d
asDoc : Nat -> Raw -> Doc
asDoc p (RVar str) = text str
asDoc p (RLam str icit x) = par p 0 $ text "\\" ++ wrap icit (text str) <+> text "=>" <+> asDoc 0 x
-- This needs precedence and operators...
asDoc p (RApp x y Explicit) = par p 2 $ asDoc 2 x <+> asDoc 3 y
asDoc p (RApp x y Implicit) = par p 2 $ asDoc 2 x <+> text "{" ++ asDoc 0 y ++ text "}"
asDoc p RU = text "U"
asDoc p (RPi Nothing Implicit ty scope) = par p 1 $ asDoc p ty <+> text "->" <+/> asDoc p scope
asDoc p (RPi (Just x) Implicit ty scope) =
par p 1 $ text "(" <+> text x <+> text ":" <+> asDoc p ty <+> text ")" <+> text "->" <+/> asDoc p scope
asDoc p (RPi nm Explicit ty scope) =
par p 1 $ text "{" <+> text (fromMaybe "_" nm) <+> text ":" <+> asDoc p ty <+> text "}" <+> text "->" <+/> asDoc 1 scope
asDoc p (RLet x v ty scope) =
par p 0 $ text "let" <+> text x <+> text ":" <+> asDoc p ty
<+> text "=" <+> asDoc p v
<+/> text "in" <+> asDoc p scope
asDoc p (RSrcPos x y) = asDoc p y
-- does this exist?
asDoc p (RAnn x y) = text "TODO - RAnn"
asDoc p (RLit x) = text (show x)
asDoc p (RCase x xs) = text "TODO - RCase" --?asDoc p_rhs_9
asDoc p RHole = text "_"
asDoc p (RParseError str) = text "PraseError \{str}"
export
Pretty Module where
pretty (MkModule name imports decls) =
text "module" <+> text name </> stack (map doDecl decls)
where
doDecl : Decl -> Doc
doDecl (TypeSig nm ty) = text nm <+> text ":" <+> nest 2 (pretty ty)
doDecl (Def nm tm) = text nm <+> text "=" <+> nest 2 (pretty tm)
doDecl (DImport nm) = text "import" <+> text nm ++ line
doDecl (Data str x xs) = text "TODO data"