module Quox.Parser.Syntax import public Quox.Syntax import public Quox.Syntax.Qty.Three import public Quox.Definition import Derive.Prelude %hide TT.Name %default total %language ElabReflection public export 0 BName : Type BName = Maybe String public export 0 PUniverse : Type PUniverse = Nat public export 0 PQty : Type PQty = Three namespace PDim public export data PDim = K DimConst | V PBaseName %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 | Pi PQty BName PTerm PTerm | Lam BName PTerm | (:@) PTerm PTerm | Sig BName PTerm PTerm | Pair PTerm PTerm | Case PQty PTerm (BName, PTerm) (PCaseBody) | Enum (List TagVal) | Tag TagVal | Eq (BName, PTerm) PTerm PTerm | DLam BName PTerm | (:%) PTerm PDim | V PName | (:#) PTerm PTerm %name PTerm s, t public export data PCaseBody = CasePair (BName, BName) PTerm | CaseEnum (List (TagVal, PTerm)) %name PCaseBody body %runElab deriveMutual ["PTerm", "PCaseBody"] [Eq, Ord, Show] public export record PDefinition where constructor MkPDef qty : PQty name : PName type : Maybe PTerm 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] 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 [<] mutual namespace Term export toPTermWith : Context' PBaseName d -> Context' PBaseName n -> Term Three 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 Three 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 E e => toPTermWith ds ns e namespace Elim export toPTermWith : Context' PBaseName d -> Context' PBaseName n -> Elim Three 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 Three 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 $ show r, toPTermWith ds (ns :< baseStr r) ret.term) (CaseEnum $ mapSnd (toPTermWith ds ns) <$> SortedMap.toList arms) fun :% arg => toPTermWith ds ns fun :% toPDimWith ds arg tm :# ty => toPTermWith ds ns tm :# toPTermWith ds ns ty namespace Term export toPTerm : Term Three 0 0 -> PTerm toPTerm = toPTermWith [<] [<] namespace Elim export toPTerm : Elim Three 0 0 -> PTerm toPTerm = toPTermWith [<] [<]