add module Parser.Syntax with PTerm and toPTerm
This commit is contained in:
parent
8fc0b414cf
commit
0cae84c75b
4 changed files with 159 additions and 56 deletions
|
@ -1,18 +1,19 @@
|
||||||
module Quox.Parser
|
module Quox.Parser
|
||||||
|
|
||||||
import public Quox.Syntax.Qty.Three
|
|
||||||
import public Quox.Syntax
|
|
||||||
import public Quox.Lexer
|
import public Quox.Lexer
|
||||||
|
import public Quox.Parser.Syntax
|
||||||
|
|
||||||
import Data.Fin
|
import Data.Fin
|
||||||
import Data.Vect
|
import Data.Vect
|
||||||
import public Text.Parser
|
import public Text.Parser
|
||||||
|
|
||||||
|
|
||||||
|
-- [fixme] without this import idris fails to solve the
|
||||||
|
-- IsReserved arguments, which, ??????????????????????
|
||||||
import Derive.Prelude
|
import Derive.Prelude
|
||||||
%hide TT.Name
|
%hide TT.Name
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
%language ElabReflection
|
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
|
@ -20,59 +21,8 @@ public export
|
||||||
Grammar = Core.Grammar () Token
|
Grammar = Core.Grammar () Token
|
||||||
%hide Core.Grammar
|
%hide Core.Grammar
|
||||||
|
|
||||||
|
|
||||||
public export
|
|
||||||
0 BName : Type
|
|
||||||
BName = Maybe BaseName
|
|
||||||
|
|
||||||
public export
|
|
||||||
0 PUniverse : Type
|
|
||||||
PUniverse = Nat
|
|
||||||
|
|
||||||
public export
|
|
||||||
0 PQty : Type
|
|
||||||
PQty = Three
|
|
||||||
|
|
||||||
namespace PDim
|
|
||||||
public export
|
|
||||||
data PDim = K DimConst | V BaseName
|
|
||||||
%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 Name
|
|
||||||
| (:#) PTerm PTerm
|
|
||||||
|
|
||||||
public export
|
|
||||||
data PCaseBody =
|
|
||||||
CasePair (BName, BName) PTerm
|
|
||||||
| CaseEnum (List (TagVal, PTerm))
|
|
||||||
|
|
||||||
%runElab deriveMutual ["PTerm", "PCaseBody"] [Eq, Ord, Show]
|
|
||||||
|
|
||||||
private
|
private
|
||||||
data PArg = T PTerm | D PDim
|
data PArg = T PTerm | D PDim
|
||||||
%runElab derive "PArg" [Eq, Ord, Show]
|
|
||||||
|
|
||||||
|
|
||||||
private
|
private
|
||||||
|
|
152
lib/Quox/Parser/Syntax.idr
Normal file
152
lib/Quox/Parser/Syntax.idr
Normal file
|
@ -0,0 +1,152 @@
|
||||||
|
module Quox.Parser.Syntax
|
||||||
|
|
||||||
|
import public Quox.Syntax
|
||||||
|
import public Quox.Syntax.Qty.Three
|
||||||
|
|
||||||
|
import Derive.Prelude
|
||||||
|
%hide TT.Name
|
||||||
|
|
||||||
|
%default total
|
||||||
|
%language ElabReflection
|
||||||
|
|
||||||
|
|
||||||
|
public export
|
||||||
|
0 BName : Type
|
||||||
|
BName = Maybe BaseName
|
||||||
|
|
||||||
|
public export
|
||||||
|
0 PUniverse : Type
|
||||||
|
PUniverse = Nat
|
||||||
|
|
||||||
|
public export
|
||||||
|
0 PQty : Type
|
||||||
|
PQty = Three
|
||||||
|
|
||||||
|
namespace PDim
|
||||||
|
public export
|
||||||
|
data PDim = K DimConst | V BaseName
|
||||||
|
%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 Name
|
||||||
|
| (:#) PTerm PTerm
|
||||||
|
|
||||||
|
public export
|
||||||
|
data PCaseBody =
|
||||||
|
CasePair (BName, BName) PTerm
|
||||||
|
| CaseEnum (List (TagVal, PTerm))
|
||||||
|
|
||||||
|
%runElab deriveMutual ["PTerm", "PCaseBody"] [Eq, Ord, Show]
|
||||||
|
|
||||||
|
|
||||||
|
export
|
||||||
|
toPDimWith : Context' BaseName 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' BaseName d -> Context' BaseName n ->
|
||||||
|
Term Three d n -> PTerm
|
||||||
|
toPTermWith ds ns t =
|
||||||
|
let Element t _ = pushSubsts t in
|
||||||
|
toPTermWith' ds ns t
|
||||||
|
|
||||||
|
private
|
||||||
|
toPTermWith' : Context' BaseName d -> Context' BaseName n ->
|
||||||
|
(t : Term Three d n) -> (0 _ : NotClo t) =>
|
||||||
|
PTerm
|
||||||
|
toPTermWith' ds ns s = case s of
|
||||||
|
TYPE (U l) =>
|
||||||
|
TYPE l
|
||||||
|
Pi qty arg (S [x] res) =>
|
||||||
|
Pi qty (Just x) (toPTermWith ds ns arg)
|
||||||
|
(toPTermWith ds (ns :< x) res.term)
|
||||||
|
Lam (S [x] body) =>
|
||||||
|
Lam (Just x) $ toPTermWith ds (ns :< x) body.term
|
||||||
|
Sig fst (S [x] snd) =>
|
||||||
|
Sig (Just x) (toPTermWith ds ns fst)
|
||||||
|
(toPTermWith ds (ns :< 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 i, toPTermWith (ds :< i) ns ty.term)
|
||||||
|
(toPTermWith ds ns l) (toPTermWith ds ns r)
|
||||||
|
DLam (S [i] body) =>
|
||||||
|
DLam (Just i) $ toPTermWith (ds :< i) ns body.term
|
||||||
|
E e =>
|
||||||
|
toPTermWith ds ns e
|
||||||
|
|
||||||
|
namespace Elim
|
||||||
|
export
|
||||||
|
toPTermWith : Context' BaseName d -> Context' BaseName n ->
|
||||||
|
Elim Three d n -> PTerm
|
||||||
|
toPTermWith ds ns e =
|
||||||
|
let Element e _ = pushSubsts e in
|
||||||
|
toPTermWith' ds ns e
|
||||||
|
|
||||||
|
private
|
||||||
|
toPTermWith' : Context' BaseName d -> Context' BaseName n ->
|
||||||
|
(e : Elim Three d n) -> (0 _ : NotClo e) =>
|
||||||
|
PTerm
|
||||||
|
toPTermWith' ds ns e = case e of
|
||||||
|
F x =>
|
||||||
|
V x
|
||||||
|
B i =>
|
||||||
|
V $ unq $ 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 r, toPTermWith ds (ns :< r) ret.term)
|
||||||
|
(CasePair (Just x, Just y) $
|
||||||
|
toPTermWith ds (ns :< x :< y) body.term)
|
||||||
|
CaseEnum qty tag (S [r] ret) arms =>
|
||||||
|
Case qty (toPTermWith ds ns tag)
|
||||||
|
(Just r, toPTermWith ds (ns :< 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 [<] [<]
|
|
@ -291,11 +291,11 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
||||||
let bodyctx = extendTyN [< tfst, tsnd.term] ctx
|
let bodyctx = extendTyN [< tfst, tsnd.term] ctx
|
||||||
bodyty = substCasePairRet pairres.type ret
|
bodyty = substCasePairRet pairres.type ret
|
||||||
pisg = pi * sg.fst
|
pisg = pi * sg.fst
|
||||||
bodyout <- popQs [< pisg, pisg] !(checkC bodyctx sg body.term bodyty)
|
bodyout <- checkC bodyctx sg body.term bodyty
|
||||||
-- then Ψ | Γ ⊢ σ · case ⋯ ⇒ ret[pair/p] ⊳ πΣ₁ + Σ₂
|
-- then Ψ | Γ ⊢ σ · case ⋯ ⇒ ret[pair/p] ⊳ πΣ₁ + Σ₂
|
||||||
pure $ InfRes {
|
pure $ InfRes {
|
||||||
type = sub1 ret pair,
|
type = sub1 ret pair,
|
||||||
qout = pi * pairres.qout + bodyout
|
qout = pi * pairres.qout + !(popQs [< pisg, pisg] bodyout)
|
||||||
}
|
}
|
||||||
|
|
||||||
infer' ctx sg (CaseEnum pi t ret arms) {n} = do
|
infer' ctx sg (CaseEnum pi t ret arms) {n} = do
|
||||||
|
|
|
@ -36,4 +36,5 @@ modules =
|
||||||
Quox.Typing,
|
Quox.Typing,
|
||||||
Quox.Typechecker,
|
Quox.Typechecker,
|
||||||
Quox.Lexer,
|
Quox.Lexer,
|
||||||
|
Quox.Parser.Syntax,
|
||||||
Quox.Parser
|
Quox.Parser
|
||||||
|
|
Loading…
Reference in a new issue