quox/lib/Quox/Parser/Syntax.idr

220 lines
5.8 KiB
Idris
Raw Normal View History

module Quox.Parser.Syntax
import public Quox.Syntax
2023-03-06 05:35:57 -05:00
import public Quox.Definition
import Derive.Prelude
%hide TT.Name
%default total
%language ElabReflection
public export
BName : Type
BName = Maybe String
public export
PUniverse : Type
PUniverse = Nat
namespace PDim
public export
data PDim = K DimConst | V PBaseName
2023-03-05 10:49:50 -05:00
%name PDim p, q
%runElab derive "PDim" [Eq, Ord, Show]
namespace PTerm
mutual
||| terms out of the parser with BVs and bidirectionality still tangled up
public export
data PTerm =
TYPE Nat
2023-04-01 13:16:43 -04:00
| Pi Qty BName PTerm PTerm
| Lam BName PTerm
| (:@) PTerm PTerm
| Sig BName PTerm PTerm
| Pair PTerm PTerm
| Case Qty PTerm (BName, PTerm) PCaseBody
| Enum (List TagVal)
| Tag TagVal
| Eq (BName, PTerm) PTerm PTerm
| DLam BName PTerm
| (:%) PTerm PDim
2023-03-26 08:40:54 -04:00
| Nat
| Zero | Succ PTerm
2023-04-01 13:16:43 -04:00
| BOX Qty PTerm
2023-03-31 13:11:35 -04:00
| Box PTerm
| V PName
| (:#) PTerm PTerm
2023-04-15 09:13:01 -04:00
| Coe (BName, PTerm) PDim PDim PTerm
| Comp (BName, PTerm) PDim PDim PTerm PDim (BName, PTerm) (BName, PTerm)
2023-03-05 10:49:50 -05:00
%name PTerm s, t
public export
data PCaseBody =
CasePair (BName, BName) PTerm
| CaseEnum (List (TagVal, PTerm))
2023-04-01 13:16:43 -04:00
| CaseNat PTerm (BName, Qty, BName, PTerm)
2023-03-31 13:11:35 -04:00
| CaseBox BName PTerm
2023-03-05 10:49:50 -05:00
%name PCaseBody body
%runElab deriveMutual ["PTerm", "PCaseBody"] [Eq, Ord, Show]
2023-03-06 05:35:57 -05:00
public export
record PDefinition where
constructor MkPDef
2023-04-01 13:16:43 -04:00
qty : Qty
name : PBaseName
type : Maybe PTerm
2023-03-06 05:35:57 -05:00
term : PTerm
%name PDefinition def
%runElab derive "PDefinition" [Eq, Ord, Show]
mutual
public export
record PNamespace where
constructor MkPNamespace
name : Mods
decls : List PDecl
%name PNamespace ns
public export
data PDecl =
PDef PDefinition
| PNs PNamespace
%name PDecl decl
%runElab deriveMutual ["PNamespace", "PDecl"] [Eq, Ord, Show]
public export
data PTopLevel = PD PDecl | PLoad String
%name PTopLevel t
%runElab derive "PTopLevel" [Eq, Ord, Show]
2023-03-06 05:35:57 -05:00
export
toPDimWith : Context' PBaseName d -> Dim d -> PDim
toPDimWith ds (K e) = K e
toPDimWith ds (B i) = V $ ds !!! i
export
toPDim : Dim 0 -> PDim
toPDim = toPDimWith [<]
2023-04-03 11:46:23 -04:00
public export
fromNat : Nat -> PTerm
fromNat 0 = Zero
fromNat (S k) = Succ $ fromNat k
2023-04-03 11:46:23 -04:00
{-
mutual
namespace Term
export
toPTermWith : Context' PBaseName d -> Context' PBaseName n ->
2023-04-01 13:16:43 -04:00
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 ->
2023-04-01 13:16:43 -04:00
(t : Term d n) -> (0 _ : NotClo t) =>
PTerm
toPTermWith' ds ns s = case s of
2023-03-05 10:48:29 -05:00
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
2023-03-26 08:40:54 -04:00
Nat => Nat
Zero => Zero
Succ n => Succ $ toPTermWith ds ns n
2023-03-31 13:11:35 -04:00
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 ->
2023-04-01 13:16:43 -04:00
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 ->
2023-04-01 13:16:43 -04:00
(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)
2023-03-26 08:40:54 -04:00
(Just $ baseStr r, toPTermWith ds (ns :< baseStr r) ret.term)
(CaseEnum $ mapSnd (toPTermWith ds ns) <$> SortedMap.toList arms)
2023-03-26 08:40:54 -04:00
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))
2023-03-31 13:11:35 -04:00
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
2023-04-01 13:16:43 -04:00
toPTerm : Term 0 0 -> PTerm
toPTerm = toPTermWith [<] [<]
namespace Elim
export
2023-04-01 13:16:43 -04:00
toPTerm : Elim 0 0 -> PTerm
toPTerm = toPTermWith [<] [<]
2023-04-03 11:46:23 -04:00
-}