crude but effective stratification
This commit is contained in:
parent
e4a20cc632
commit
42aa07c9c8
31 changed files with 817 additions and 582 deletions
|
@ -121,15 +121,29 @@ namespace Char
|
|||
isOther = isOther . genCat
|
||||
|
||||
|
||||
export
|
||||
isSupDigit : Char -> Bool
|
||||
isSupDigit ch = ch `elem` unpack "⁰¹²³⁴⁵⁶⁷⁸⁹"
|
||||
|
||||
export
|
||||
isSubDigit : Char -> Bool
|
||||
isSubDigit ch = ch `elem` unpack "₀₁₂₃₄₅₆₇₈₉"
|
||||
|
||||
export
|
||||
isAsciiDigit : Char -> Bool
|
||||
isAsciiDigit ch = '0' <= ch && ch <= '9'
|
||||
|
||||
export
|
||||
isIdStart : Char -> Bool
|
||||
isIdStart ch =
|
||||
ch == '_' || isLetter ch || isNumber ch && not ('0' <= ch && ch <= '9')
|
||||
(ch == '_' || isLetter ch || isNumber ch) &&
|
||||
not (isSupDigit ch || isAsciiDigit ch)
|
||||
|
||||
export
|
||||
isIdCont : Char -> Bool
|
||||
isIdCont ch =
|
||||
isIdStart ch || ch == '\'' || ch == '-' || isMark ch || isNumber ch
|
||||
(isIdStart ch || ch == '\'' || ch == '-' || isMark ch || isNumber ch) &&
|
||||
not (isSupDigit ch)
|
||||
|
||||
export
|
||||
isIdConnector : Char -> Bool
|
||||
|
|
85
lib/Quox/Displace.idr
Normal file
85
lib/Quox/Displace.idr
Normal file
|
@ -0,0 +1,85 @@
|
|||
module Quox.Displace
|
||||
|
||||
import Quox.Syntax
|
||||
|
||||
|
||||
parameters (k : Universe)
|
||||
namespace Term
|
||||
export doDisplace : Term d n -> Term d n
|
||||
export doDisplaceS : ScopeTermN s d n -> ScopeTermN s d n
|
||||
export doDisplaceDS : DScopeTermN s d n -> DScopeTermN s d n
|
||||
|
||||
namespace Elim
|
||||
export doDisplace : Elim d n -> Elim d n
|
||||
|
||||
namespace Term
|
||||
doDisplace (TYPE l loc) = TYPE (k + l) loc
|
||||
doDisplace (Pi qty arg res loc) =
|
||||
Pi qty (doDisplace arg) (doDisplaceS res) loc
|
||||
doDisplace (Lam body loc) = Lam (doDisplaceS body) loc
|
||||
doDisplace (Sig fst snd loc) = Sig (doDisplace fst) (doDisplaceS snd) loc
|
||||
doDisplace (Pair fst snd loc) = Pair (doDisplace fst) (doDisplace snd) loc
|
||||
doDisplace (Enum cases loc) = Enum cases loc
|
||||
doDisplace (Tag tag loc) = Tag tag loc
|
||||
doDisplace (Eq ty l r loc) =
|
||||
Eq (doDisplaceDS ty) (doDisplace l) (doDisplace r) loc
|
||||
doDisplace (DLam body loc) = DLam (doDisplaceDS body) loc
|
||||
doDisplace (Nat loc) = Nat loc
|
||||
doDisplace (Zero loc) = Zero loc
|
||||
doDisplace (Succ p loc) = Succ (doDisplace p) loc
|
||||
doDisplace (BOX qty ty loc) = BOX qty (doDisplace ty) loc
|
||||
doDisplace (Box val loc) = Box (doDisplace val) loc
|
||||
doDisplace (E e) = E (doDisplace e)
|
||||
doDisplace (CloT (Sub t th)) =
|
||||
CloT (Sub (doDisplace t) (map doDisplace th))
|
||||
doDisplace (DCloT (Sub t th)) =
|
||||
DCloT (Sub (doDisplace t) th)
|
||||
|
||||
doDisplaceS (S names (Y body)) = S names $ Y $ doDisplace body
|
||||
doDisplaceS (S names (N body)) = S names $ N $ doDisplace body
|
||||
|
||||
doDisplaceDS (S names (Y body)) = S names $ Y $ doDisplace body
|
||||
doDisplaceDS (S names (N body)) = S names $ N $ doDisplace body
|
||||
|
||||
namespace Elim
|
||||
doDisplace (F x u loc) = F x (k + u) loc
|
||||
doDisplace (B i loc) = B i loc
|
||||
doDisplace (App fun arg loc) = App (doDisplace fun) (doDisplace arg) loc
|
||||
doDisplace (CasePair qty pair ret body loc) =
|
||||
CasePair qty (doDisplace pair) (doDisplaceS ret) (doDisplaceS body) loc
|
||||
doDisplace (CaseEnum qty tag ret arms loc) =
|
||||
CaseEnum qty (doDisplace tag) (doDisplaceS ret) (map doDisplace arms) loc
|
||||
doDisplace (CaseNat qty qtyIH nat ret zero succ loc) =
|
||||
CaseNat qty qtyIH (doDisplace nat) (doDisplaceS ret)
|
||||
(doDisplace zero) (doDisplaceS succ) loc
|
||||
doDisplace (CaseBox qty box ret body loc) =
|
||||
CaseBox qty (doDisplace box) (doDisplaceS ret) (doDisplaceS body) loc
|
||||
doDisplace (DApp fun arg loc) =
|
||||
DApp (doDisplace fun) arg loc
|
||||
doDisplace (Ann tm ty loc) =
|
||||
Ann (doDisplace tm) (doDisplace ty) loc
|
||||
doDisplace (Coe ty p q val loc) =
|
||||
Coe (doDisplaceDS ty) p q (doDisplace val) loc
|
||||
doDisplace (Comp ty p q val r zero one loc) =
|
||||
Comp (doDisplace ty) p q (doDisplace val) r
|
||||
(doDisplaceDS zero) (doDisplaceDS one) loc
|
||||
doDisplace (TypeCase ty ret arms def loc) =
|
||||
TypeCase (doDisplace ty) (doDisplace ret)
|
||||
(map doDisplaceS arms) (doDisplace def) loc
|
||||
doDisplace (CloE (Sub e th)) =
|
||||
CloE (Sub (doDisplace e) (map doDisplace th))
|
||||
doDisplace (DCloE (Sub e th)) =
|
||||
DCloE (Sub (doDisplace e) th)
|
||||
|
||||
|
||||
namespace Term
|
||||
export
|
||||
displace : Universe -> Term d n -> Term d n
|
||||
displace 0 t = t
|
||||
displace u t = doDisplace u t
|
||||
|
||||
namespace Elim
|
||||
export
|
||||
displace : Universe -> Elim d n -> Elim d n
|
||||
displace 0 t = t
|
||||
displace u t = doDisplace u t
|
|
@ -398,11 +398,11 @@ parameters (defs : Definitions)
|
|||
(0 ne : NotRedex defs e) -> (0 nf : NotRedex defs f) ->
|
||||
Equal_ ()
|
||||
|
||||
compare0' ctx e@(F x {}) f@(F y {}) _ _ =
|
||||
unless (x == y) $ clashE e.loc ctx e f
|
||||
compare0' ctx e@(F x u _) f@(F y v _) _ _ =
|
||||
unless (x == y && u == v) $ clashE e.loc ctx e f
|
||||
compare0' ctx e@(F {}) f _ _ = clashE e.loc ctx e f
|
||||
|
||||
compare0' ctx e@(B i {}) f@(B j {}) _ _ =
|
||||
compare0' ctx e@(B i _) f@(B j _) _ _ =
|
||||
unless (i == j) $ clashE e.loc ctx e f
|
||||
compare0' ctx e@(B {}) f _ _ = clashE e.loc ctx e f
|
||||
|
||||
|
@ -538,6 +538,7 @@ parameters (defs : Definitions)
|
|||
compare0' _ (Comp {r = B i _, _}) _ _ _ = absurd i
|
||||
compare0' _ _ (Comp {r = K _ _, _}) _ nf = void $ absurd $ noOr2 nf
|
||||
|
||||
-- (type case equality purely structural)
|
||||
compare0' ctx (TypeCase ty1 ret1 arms1 def1 eloc)
|
||||
(TypeCase ty2 ret2 arms2 def2 floc) ne _ =
|
||||
local_ Equal $ do
|
||||
|
@ -551,6 +552,11 @@ parameters (defs : Definitions)
|
|||
(lookupPrecise k arms1) (lookupPrecise k arms2) def1
|
||||
compare0' ctx e@(TypeCase {}) f _ _ = clashE e.loc ctx e f
|
||||
|
||||
-- Ψ | Γ ⊢ s <: f ⇐ A
|
||||
-- --------------------------
|
||||
-- Ψ | Γ ⊢ (s ∷ A) <: f ⇒ A
|
||||
--
|
||||
-- and vice versa
|
||||
compare0' ctx (Ann s a _) f _ _ = Term.compare0 ctx a s (E f)
|
||||
compare0' ctx e (Ann t b _) _ _ = Term.compare0 ctx b (E e) t
|
||||
compare0' ctx e@(Ann {}) f _ _ = clashE e.loc ctx e f
|
||||
|
|
|
@ -124,7 +124,7 @@ fromList = fromPName . fromListP
|
|||
|
||||
export
|
||||
syntaxChars : List Char
|
||||
syntaxChars = ['(', ')', '[', ']', '{', '}', '"', '\'', ',', '.', ';']
|
||||
syntaxChars = ['(', ')', '[', ']', '{', '}', '"', '\'', ',', '.', ';', '^']
|
||||
|
||||
export
|
||||
isSymStart, isSymCont : Char -> Bool
|
||||
|
|
|
@ -83,15 +83,16 @@ avoidDim ds loc x =
|
|||
fromName (const $ throw $ DimNameInTerm loc x.base) (pure . fromPName) ds x
|
||||
|
||||
private
|
||||
resolveName : Mods -> Loc -> Name -> Eff FromParserPure (Term d n)
|
||||
resolveName ns loc x =
|
||||
resolveName : Mods -> Loc -> Name -> Maybe Universe ->
|
||||
Eff FromParserPure (Term d n)
|
||||
resolveName ns loc x u =
|
||||
let here = addMods ns x in
|
||||
if isJust $ lookup here !(getAt DEFS) then
|
||||
pure $ FT here loc
|
||||
pure $ FT here (fromMaybe 0 u) loc
|
||||
else do
|
||||
let ns :< _ = ns
|
||||
| _ => throw $ TermNotInScope loc x
|
||||
resolveName ns loc x
|
||||
resolveName ns loc x u
|
||||
|
||||
export
|
||||
fromPatVar : PatVar -> BindName
|
||||
|
@ -107,6 +108,17 @@ fromPTagVal : PTagVal -> TagVal
|
|||
fromPTagVal (PT t _) = t
|
||||
|
||||
|
||||
private
|
||||
fromV : Context' PatVar d -> Context' PatVar n ->
|
||||
PName -> Maybe Universe -> Loc -> Eff FromParserPure (Term d n)
|
||||
fromV ds ns x u loc = fromName bound free ns x where
|
||||
bound : Var n -> Eff FromParserPure (Term d n)
|
||||
bound i = do whenJust u $ \u => throw $ DisplacedBoundVar loc x
|
||||
pure $ E $ B i loc
|
||||
free : PName -> Eff FromParserPure (Term d n)
|
||||
free x = do x <- avoidDim ds loc x
|
||||
resolveName !(getAt NS) loc x u
|
||||
|
||||
mutual
|
||||
export
|
||||
fromPTermWith : Context' PatVar d -> Context' PatVar n ->
|
||||
|
@ -199,9 +211,7 @@ mutual
|
|||
<*> fromPTermTScope ds ns [< b] body
|
||||
<*> pure loc
|
||||
|
||||
V x loc =>
|
||||
fromName (\i => pure $ E $ B i loc)
|
||||
(resolveName !(getAt NS) loc <=< avoidDim ds loc) ns x
|
||||
V x u loc => fromV ds ns x u loc
|
||||
|
||||
Ann s a loc =>
|
||||
map E $ Ann
|
||||
|
|
|
@ -24,6 +24,7 @@ data Error =
|
|||
| DimNotInScope Loc PBaseName
|
||||
| QtyNotGlobal Loc Qty
|
||||
| DimNameInTerm Loc PBaseName
|
||||
| DisplacedBoundVar Loc PName
|
||||
| WrapTypeError TypeError
|
||||
| LoadError Loc String FileError
|
||||
| WrapParseError String ParseError
|
||||
|
@ -89,6 +90,11 @@ parameters (showContext : Bool)
|
|||
(sep ["dimension" <++> !(hl DVar $ text i),
|
||||
"used in a term context"])
|
||||
|
||||
prettyError (DisplacedBoundVar loc x) = pure $
|
||||
vappend !(prettyLoc loc)
|
||||
(sep ["local variable" <++> !(hl TVar $ text $ toDotsP x),
|
||||
"cannot be displaced"])
|
||||
|
||||
prettyError (WrapTypeError err) =
|
||||
Typing.prettyError showContext $ trimContext 2 err
|
||||
|
||||
|
|
|
@ -21,7 +21,8 @@ import Derive.Prelude
|
|||
||| @ Nat nat literal
|
||||
||| @ String string literal
|
||||
||| @ Tag tag literal
|
||||
||| @ TYPE "Type" or "★" with subscript
|
||||
||| @ TYPE "Type" or "★" with ascii nat directly after
|
||||
||| @ Sup superscript or ^ number (displacement, or universe for ★)
|
||||
public export
|
||||
data Token =
|
||||
Reserved String
|
||||
|
@ -30,6 +31,7 @@ data Token =
|
|||
| Str String
|
||||
| Tag String
|
||||
| TYPE Nat
|
||||
| Sup Nat
|
||||
%runElab derive "Token" [Eq, Ord, Show]
|
||||
|
||||
-- token or whitespace
|
||||
|
@ -94,21 +96,33 @@ fromSub c = case c of
|
|||
'₀' => '0'; '₁' => '1'; '₂' => '2'; '₃' => '3'; '₄' => '4'
|
||||
'₅' => '5'; '₆' => '6'; '₇' => '7'; '₈' => '8'; '₉' => '9'; _ => c
|
||||
|
||||
private %inline
|
||||
fromSup : Char -> Char
|
||||
fromSup c = case c of
|
||||
'⁰' => '0'; '¹' => '1'; '²' => '2'; '³' => '3'; '⁴' => '4'
|
||||
'⁵' => '5'; '⁶' => '6'; '⁷' => '7'; '⁸' => '8'; '⁹' => '9'; _ => c
|
||||
|
||||
private %inline
|
||||
subToNat : String -> Nat
|
||||
subToNat = cast . pack . map fromSub . unpack
|
||||
|
||||
private %inline
|
||||
supToNat : String -> Nat
|
||||
supToNat = cast . pack . map fromSup . unpack
|
||||
|
||||
-- ★0, Type0. base ★/Type is a Reserved
|
||||
private
|
||||
universe : Tokenizer TokenW
|
||||
universe = universeWith "★" <|> universeWith "Type" where
|
||||
universeWith : String -> Tokenizer TokenW
|
||||
universeWith pfx =
|
||||
let len = length pfx in
|
||||
match (exact pfx <+> some (range '0' '9'))
|
||||
(TYPE . cast . drop len) <|>
|
||||
match (exact pfx <+> some (range '₀' '₉'))
|
||||
(TYPE . subToNat . drop len)
|
||||
match (exact pfx <+> digits) (TYPE . cast . drop len)
|
||||
|
||||
private
|
||||
sup : Tokenizer TokenW
|
||||
sup = match (some $ pred isSupDigit) (Sup . supToNat)
|
||||
<|> match (is '^' <+> digits) (Sup . cast . drop 1)
|
||||
|
||||
|
||||
private %inline
|
||||
|
@ -219,7 +233,7 @@ tokens = choice $
|
|||
blockComment (exact "{-") (exact "-}")] <+>
|
||||
[universe] <+> -- ★ᵢ takes precedence over bare ★
|
||||
map resTokenizer reserved <+>
|
||||
[nat, string, tag, name]
|
||||
[sup, nat, string, tag, name]
|
||||
|
||||
export
|
||||
lex : String -> Either Error (List (WithBounds Token))
|
||||
|
|
|
@ -106,10 +106,14 @@ export
|
|||
strLit : Grammar True String
|
||||
strLit = terminalMatch "string literal" `(Str s) `(s)
|
||||
|
||||
||| single-token universe, like ★₀ or Type1
|
||||
||| single-token universe, like ★0 or Type1
|
||||
export
|
||||
universe1 : Grammar True Universe
|
||||
universe1 = terminalMatch "universe" `(TYPE u) `(u)
|
||||
universeTok : Grammar True Universe
|
||||
universeTok = terminalMatch "universe" `(TYPE u) `(u)
|
||||
|
||||
export
|
||||
super : Grammar True Nat
|
||||
super = terminalMatch "superscript number or '^'" `(Sup n) `(n)
|
||||
|
||||
||| possibly-qualified name
|
||||
export
|
||||
|
@ -134,6 +138,11 @@ qtyVal = terminalMatchN "quantity"
|
|||
[(`(Nat 0), `(Zero)), (`(Nat 1), `(One)), (`(Reserved "ω"), `(Any))]
|
||||
|
||||
|
||||
||| optional superscript number
|
||||
export
|
||||
displacement : Grammar False (Maybe Universe)
|
||||
displacement = optional super
|
||||
|
||||
||| quantity (0, 1, or ω)
|
||||
export
|
||||
qty : FileName -> Grammar True PQty
|
||||
|
@ -263,6 +272,10 @@ tupleTerm fname = withLoc fname $ do
|
|||
terms <- delimSep1 "(" ")" "," $ assert_total term fname
|
||||
pure $ \loc => foldr1 (\s, t => Pair s t loc) terms
|
||||
|
||||
export
|
||||
universe1 : Grammar True Universe
|
||||
universe1 = universeTok <|> res "★" *> super
|
||||
|
||||
||| argument/atomic term: single-token terms, or those with delimiters e.g.
|
||||
||| `[t]`
|
||||
export
|
||||
|
@ -275,7 +288,7 @@ termArg fname = withLoc fname $
|
|||
<|> Nat <$ res "ℕ"
|
||||
<|> Zero <$ res "zero"
|
||||
<|> [|fromNat nat|]
|
||||
<|> [|V qname|]
|
||||
<|> [|V qname displacement|]
|
||||
<|> const <$> tupleTerm fname
|
||||
|
||||
export
|
||||
|
@ -345,7 +358,10 @@ compTerm fname = withLoc fname $ do
|
|||
|
||||
export
|
||||
splitUniverseTerm : FileName -> Grammar True PTerm
|
||||
splitUniverseTerm fname = withLoc fname $ resC "★" *> mustWork [|TYPE nat|]
|
||||
splitUniverseTerm fname =
|
||||
withLoc fname $ resC "★" *> mustWork [|TYPE $ nat <|> super|]
|
||||
-- having super here looks redundant, but when parsing a non-atomic term
|
||||
-- this branch will be taken first
|
||||
|
||||
export
|
||||
eqTerm : FileName -> Grammar True PTerm
|
||||
|
|
|
@ -87,7 +87,7 @@ namespace PTerm
|
|||
| BOX PQty PTerm Loc
|
||||
| Box PTerm Loc
|
||||
|
||||
| V PName Loc
|
||||
| V PName (Maybe Universe) Loc
|
||||
| Ann PTerm PTerm Loc
|
||||
|
||||
| Coe (PatVar, PTerm) PDim PDim PTerm Loc
|
||||
|
@ -124,7 +124,7 @@ Located PTerm where
|
|||
(Succ _ loc).loc = loc
|
||||
(BOX _ _ loc).loc = loc
|
||||
(Box _ loc).loc = loc
|
||||
(V _ loc).loc = loc
|
||||
(V _ _ loc).loc = loc
|
||||
(Ann _ _ loc).loc = loc
|
||||
(Coe _ _ _ _ loc).loc = loc
|
||||
(Comp _ _ _ _ _ _ _ loc).loc = loc
|
||||
|
|
|
@ -38,7 +38,7 @@ data HL
|
|||
= Delim
|
||||
| Free | TVar | TVarErr
|
||||
| Dim | DVar | DVarErr
|
||||
| Qty
|
||||
| Qty | Universe
|
||||
| Syntax
|
||||
| Tag
|
||||
%runElab derive "HL" [Eq, Ord, Show]
|
||||
|
@ -74,16 +74,17 @@ runPrettyWith prec flavor highlight indent act =
|
|||
|
||||
export %inline
|
||||
toSGR : HL -> List SGR
|
||||
toSGR Delim = []
|
||||
toSGR TVar = [SetForeground BrightYellow]
|
||||
toSGR TVarErr = [SetForeground BrightYellow, SetStyle SingleUnderline]
|
||||
toSGR Dim = [SetForeground BrightGreen]
|
||||
toSGR DVar = [SetForeground BrightGreen]
|
||||
toSGR DVarErr = [SetForeground BrightGreen, SetStyle SingleUnderline]
|
||||
toSGR Qty = [SetForeground BrightMagenta]
|
||||
toSGR Free = [SetForeground BrightBlue]
|
||||
toSGR Syntax = [SetForeground BrightCyan]
|
||||
toSGR Tag = [SetForeground BrightRed]
|
||||
toSGR Delim = []
|
||||
toSGR Free = [SetForeground BrightBlue]
|
||||
toSGR TVar = [SetForeground BrightYellow]
|
||||
toSGR TVarErr = [SetForeground BrightYellow, SetStyle SingleUnderline]
|
||||
toSGR Dim = [SetForeground BrightGreen]
|
||||
toSGR DVar = [SetForeground BrightGreen]
|
||||
toSGR DVarErr = [SetForeground BrightGreen, SetStyle SingleUnderline]
|
||||
toSGR Qty = [SetForeground BrightMagenta]
|
||||
toSGR Universe = [SetForeground BrightRed]
|
||||
toSGR Syntax = [SetForeground BrightCyan]
|
||||
toSGR Tag = [SetForeground BrightRed]
|
||||
|
||||
export %inline
|
||||
highlightSGR : HL -> Highlight
|
||||
|
@ -205,9 +206,13 @@ withPrec : PPrec -> Eff Pretty a -> Eff Pretty a
|
|||
withPrec = localAt_ PREC
|
||||
|
||||
|
||||
export
|
||||
prettyName : Name -> Doc opts
|
||||
prettyName = text . toDots
|
||||
|
||||
export
|
||||
prettyFree : {opts : _} -> Name -> Eff Pretty (Doc opts)
|
||||
prettyFree = hl Free . text . toDots
|
||||
prettyFree = hl Free . prettyName
|
||||
|
||||
export
|
||||
prettyBind' : BindName -> Doc opts
|
||||
|
|
|
@ -3,6 +3,7 @@ module Quox.Reduce
|
|||
import Quox.No
|
||||
import Quox.Syntax
|
||||
import Quox.Definition
|
||||
import Quox.Displace
|
||||
import Quox.Typing.Context
|
||||
import Quox.Typing.Error
|
||||
import Data.SnocVect
|
||||
|
@ -237,9 +238,9 @@ parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
|
|||
export covering
|
||||
computeElimType : (e : Elim d n) -> (0 ne : No (isRedexE defs e)) =>
|
||||
WhnfM (Term d n)
|
||||
computeElimType (F {x, loc}) = do
|
||||
computeElimType (F {x, u, loc}) = do
|
||||
let Just def = lookup x defs | Nothing => throw $ NotInScope loc x
|
||||
pure $ def.type
|
||||
pure $ displace u def.type
|
||||
computeElimType (B {i, _}) = pure $ ctx.tctx !! i
|
||||
computeElimType (App {fun = f, arg = s, loc}) {ne} = do
|
||||
Pi {arg, res, _} <- whnf0 defs ctx =<< computeElimType f {ne = noOr1 ne}
|
||||
|
@ -253,10 +254,10 @@ parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
|
|||
Eq {ty, _} <- whnf0 defs ctx =<< computeElimType f {ne = noOr1 ne}
|
||||
| t => throw $ ExpectedEq loc ctx.names t
|
||||
pure $ dsub1 ty p
|
||||
computeElimType (Ann {ty, _}) = pure ty
|
||||
computeElimType (Coe {ty, q, _}) = pure $ dsub1 ty q
|
||||
computeElimType (Comp {ty, _}) = pure ty
|
||||
computeElimType (TypeCase {ret, _}) = pure ret
|
||||
computeElimType (Ann {ty, _}) = pure ty
|
||||
|
||||
parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext (S d) n)
|
||||
||| for π.(x : A) → B, returns (A, B);
|
||||
|
@ -477,7 +478,7 @@ reduceTypeCase defs ctx ty u ret arms def loc = case ty of
|
|||
|
||||
||| pushes a coercion inside a whnf-ed term
|
||||
private covering
|
||||
pushCoe : {n, d : Nat} -> (defs : Definitions) -> WhnfContext d n ->
|
||||
pushCoe : {d, n : Nat} -> (defs : Definitions) -> WhnfContext d n ->
|
||||
BindName ->
|
||||
(ty : Term (S d) n) -> (0 tynf : No (isRedexT defs ty)) =>
|
||||
Dim d -> Dim d ->
|
||||
|
@ -524,8 +525,7 @@ pushCoe defs ctx i ty p q s loc =
|
|||
snd' = E $ CoeT i tsnd' p q snd snd.loc
|
||||
pure $
|
||||
Element (Ann (Pair fst' snd' pairLoc)
|
||||
(Sig (tfst // one q) (tsnd // one q) sigLoc) loc)
|
||||
Ah
|
||||
(Sig (tfst // one q) (tsnd // one q) sigLoc) loc) Ah
|
||||
|
||||
-- η expand, like for Lam
|
||||
--
|
||||
|
@ -569,9 +569,9 @@ where
|
|||
|
||||
export covering
|
||||
Whnf Elim Reduce.isRedexE where
|
||||
whnf defs ctx (F x loc) with (lookupElim x defs) proof eq
|
||||
_ | Just y = whnf defs ctx $ setLoc loc y
|
||||
_ | Nothing = pure $ Element (F x loc) $ rewrite eq in Ah
|
||||
whnf defs ctx (F x u loc) with (lookupElim x defs) proof eq
|
||||
_ | Just y = whnf defs ctx $ setLoc loc $ displace u y
|
||||
_ | Nothing = pure $ Element (F x u loc) $ rewrite eq in Ah
|
||||
|
||||
whnf _ _ (B i loc) = pure $ nred $ B i loc
|
||||
|
||||
|
@ -659,10 +659,10 @@ Whnf Elim Reduce.isRedexE where
|
|||
Right nb =>
|
||||
pure $ Element (CaseBox pi box ret body caseLoc) $ boxnf `orNo` nb
|
||||
|
||||
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @0 ⇝ t ∷ A‹0/𝑗›
|
||||
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @1 ⇝ u ∷ A‹1/𝑗›
|
||||
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @𝑘 ⇝ s‹𝑘/𝑖› ∷ A‹𝑘/𝑗›
|
||||
-- (if 𝑘 is a variable)
|
||||
-- e : Eq (𝑗 ⇒ A) t u ⊢ e @0 ⇝ t ∷ A‹0/𝑗›
|
||||
-- e : Eq (𝑗 ⇒ A) t u ⊢ e @1 ⇝ u ∷ A‹1/𝑗›
|
||||
--
|
||||
-- ((δ 𝑖 ⇒ s) ∷ Eq (𝑗 ⇒ A) t u) @𝑘 ⇝ s‹𝑘/𝑖› ∷ A‹𝑘/𝑗›
|
||||
whnf defs ctx (DApp f p appLoc) = do
|
||||
Element f fnf <- whnf defs ctx f
|
||||
case nchoose $ isDLamHead f of
|
||||
|
|
|
@ -134,8 +134,11 @@ mutual
|
|||
||| first argument `d` is dimension scope size, second `n` is term scope size
|
||||
public export
|
||||
data Elim : (d, n : Nat) -> Type where
|
||||
||| free variable
|
||||
F : (x : Name) -> (loc : Loc) -> Elim d n
|
||||
||| free variable, possibly with a displacement (see @crude, or @mugen for a
|
||||
||| more abstract and formalised take)
|
||||
|||
|
||||
||| e.g. if f : ★₀ → ★₁, then f¹ : ★₁ → ★₂
|
||||
F : (x : Name) -> (u : Universe) -> (loc : Loc) -> Elim d n
|
||||
||| bound variable
|
||||
B : (i : Var n) -> (loc : Loc) -> Elim d n
|
||||
|
||||
|
@ -318,8 +321,8 @@ Eq0 {ty, l, r, loc} = Eq {ty = SN ty, l, r, loc}
|
|||
|
||||
||| same as `F` but as a term
|
||||
public export %inline
|
||||
FT : Name -> (loc : Loc) -> Term d n
|
||||
FT x loc = E $ F x loc
|
||||
FT : Name -> Universe -> Loc -> Term d n
|
||||
FT x u loc = E $ F x u loc
|
||||
|
||||
||| abbreviation for a bound variable like `BV 4` instead of
|
||||
||| `B (VS (VS (VS (VS VZ))))`
|
||||
|
@ -357,7 +360,7 @@ typeCase1Y ty ret k ns body loc = typeCase ty ret [(k ** SY ns body)] def loc
|
|||
|
||||
export
|
||||
Located (Elim d n) where
|
||||
(F _ loc).loc = loc
|
||||
(F _ _ loc).loc = loc
|
||||
(B _ loc).loc = loc
|
||||
(App _ _ loc).loc = loc
|
||||
(CasePair _ _ _ _ loc).loc = loc
|
||||
|
@ -404,7 +407,7 @@ Located1 f => Located (Scoped s f n) where
|
|||
|
||||
export
|
||||
Relocatable (Elim d n) where
|
||||
setLoc loc (F x _) = F x loc
|
||||
setLoc loc (F x u _) = F x u loc
|
||||
setLoc loc (B i _) = B i loc
|
||||
setLoc loc (App fun arg _) = App fun arg loc
|
||||
setLoc loc (CasePair qty pair ret body _) =
|
||||
|
|
|
@ -14,7 +14,7 @@ import Derive.Prelude
|
|||
|
||||
export
|
||||
prettyUniverse : {opts : _} -> Universe -> Eff Pretty (Doc opts)
|
||||
prettyUniverse = hl Syntax . text . show
|
||||
prettyUniverse = hl Universe . text . show
|
||||
|
||||
|
||||
export
|
||||
|
@ -38,6 +38,14 @@ subscript = pack . map sub . unpack where
|
|||
'0' => '₀'; '1' => '₁'; '2' => '₂'; '3' => '₃'; '4' => '₄'
|
||||
'5' => '₅'; '6' => '₆'; '7' => '₇'; '8' => '₈'; '9' => '₉'; _ => c
|
||||
|
||||
private
|
||||
superscript : String -> String
|
||||
superscript = pack . map sup . unpack where
|
||||
sup : Char -> Char
|
||||
sup c = case c of
|
||||
'0' => '⁰'; '1' => '¹'; '2' => '²'; '3' => '³'; '4' => '⁴'
|
||||
'5' => '⁵'; '6' => '⁶'; '7' => '⁷'; '8' => '⁸'; '9' => '⁹'; _ => c
|
||||
|
||||
|
||||
private
|
||||
PiBind : Nat -> Nat -> Type
|
||||
|
@ -368,11 +376,19 @@ where
|
|||
header cs = sep <$> traverse (\(s, xs) => header1 s (toList xs)) (toList cs)
|
||||
|
||||
|
||||
prettyDisp : {opts : _} -> Universe -> Eff Pretty (Maybe (Doc opts))
|
||||
prettyDisp 0 = pure Nothing
|
||||
prettyDisp u = map Just $ hl Universe =<<
|
||||
ifUnicode (text $ superscript $ show u) (text $ "^" ++ show u)
|
||||
|
||||
|
||||
prettyTerm dnames tnames (TYPE l _) =
|
||||
hl Syntax =<<
|
||||
case !(askAt FLAVOR) of
|
||||
Unicode => pure $ text $ "★" ++ subscript (show l)
|
||||
Ascii => prettyAppD (text "Type") [text $ show l]
|
||||
case !(askAt FLAVOR) of
|
||||
Unicode => do
|
||||
star <- hl Syntax "★"
|
||||
level <- hl Universe $ text $ superscript $ show l
|
||||
pure $ hcat [star, level]
|
||||
Ascii => [|hl Syntax "Type" <++> hl Universe (text $ show l)|]
|
||||
|
||||
prettyTerm dnames tnames (Pi qty arg res _) =
|
||||
parensIfM Outer =<< do
|
||||
|
@ -455,8 +471,10 @@ prettyTerm dnames tnames t0@(CloT (Sub t ph)) =
|
|||
prettyTerm dnames tnames t0@(DCloT (Sub t ph)) =
|
||||
prettyTerm dnames tnames $ assert_smaller t0 $ pushSubstsWith' ph id t
|
||||
|
||||
prettyElim dnames tnames (F x _) =
|
||||
prettyFree x
|
||||
prettyElim dnames tnames (F x u _) = do
|
||||
x <- prettyFree x
|
||||
u <- prettyDisp u
|
||||
pure $ maybe x (x <+>) u
|
||||
|
||||
prettyElim dnames tnames (B i _) =
|
||||
prettyTBind $ tnames !!! i
|
||||
|
|
|
@ -39,7 +39,7 @@ subDArgs e th = DCloE $ Sub e th
|
|||
export
|
||||
CanDSubst Elim where
|
||||
e // Shift SZ = e
|
||||
F x loc // _ = F x loc
|
||||
F x u loc // _ = F x u loc
|
||||
B i loc // _ = B i loc
|
||||
e@(DApp {}) // th = subDArgs e th
|
||||
DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th
|
||||
|
@ -73,7 +73,7 @@ export %inline FromVar (Term d) where fromVarLoc = E .: fromVar
|
|||
||| - otherwise, wraps in a new closure
|
||||
export
|
||||
CanSubstSelf (Elim d) where
|
||||
F x loc // _ = F x loc
|
||||
F x u loc // _ = F x u loc
|
||||
B i loc // th = getLoc th i loc
|
||||
CloE (Sub e ph) // th = assert_total CloE $ Sub e $ ph . th
|
||||
e // th = case force th of
|
||||
|
@ -292,8 +292,8 @@ mutual
|
|||
|
||||
export
|
||||
PushSubsts Elim Subst.isCloE where
|
||||
pushSubstsWith th ph (F x loc) =
|
||||
nclo $ F x loc
|
||||
pushSubstsWith th ph (F x u loc) =
|
||||
nclo $ F x u loc
|
||||
pushSubstsWith th ph (B i loc) =
|
||||
let res = getLoc ph i loc in
|
||||
case nchoose $ isCloE res of
|
||||
|
|
|
@ -90,8 +90,8 @@ mutual
|
|||
|
||||
private
|
||||
tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1)
|
||||
tightenE p (F x loc) =
|
||||
pure $ F x loc
|
||||
tightenE p (F x u loc) =
|
||||
pure $ F x u loc
|
||||
tightenE p (B i loc) =
|
||||
B <$> tighten p i <*> pure loc
|
||||
tightenE p (App fun arg loc) =
|
||||
|
@ -204,8 +204,8 @@ mutual
|
|||
|
||||
export
|
||||
dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n)
|
||||
dtightenE p (F x loc) =
|
||||
pure $ F x loc
|
||||
dtightenE p (F x u loc) =
|
||||
pure $ F x u loc
|
||||
dtightenE p (B i loc) =
|
||||
pure $ B i loc
|
||||
dtightenE p (App fun arg loc) =
|
||||
|
|
|
@ -2,6 +2,7 @@ module Quox.Typechecker
|
|||
|
||||
import public Quox.Typing
|
||||
import public Quox.Equal
|
||||
import Quox.Displace
|
||||
|
||||
import Data.List
|
||||
import Data.SnocVect
|
||||
|
@ -107,8 +108,14 @@ mutual
|
|||
TC (CheckResult' n)
|
||||
checkC ctx sg subj ty =
|
||||
wrapErr (WhileChecking ctx sg.fst subj ty) $
|
||||
let Element subj nc = pushSubsts subj in
|
||||
check' ctx sg subj ty
|
||||
checkCNoWrap ctx sg subj ty
|
||||
|
||||
export covering %inline
|
||||
checkCNoWrap : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n ->
|
||||
TC (CheckResult' n)
|
||||
checkCNoWrap ctx sg subj ty =
|
||||
let Element subj nc = pushSubsts subj in
|
||||
check' ctx sg subj ty
|
||||
|
||||
||| "Ψ | Γ ⊢₀ s ⇐ ★ᵢ"
|
||||
|||
|
||||
|
@ -324,14 +331,14 @@ mutual
|
|||
(subj : Elim d n) -> (0 nc : NotClo subj) =>
|
||||
TC (InferResult' d n)
|
||||
|
||||
infer' ctx sg (F x loc) = do
|
||||
infer' ctx sg (F x u loc) = do
|
||||
-- if π·x : A {≔ s} in global context
|
||||
g <- lookupFree x loc !defs
|
||||
-- if σ ≤ π
|
||||
expectCompatQ loc sg.fst g.qty.fst
|
||||
-- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎
|
||||
let Val d = ctx.dimLen; Val n = ctx.termLen
|
||||
pure $ InfRes {type = g.type, qout = zeroFor ctx}
|
||||
pure $ InfRes {type = displace u g.type, qout = zeroFor ctx}
|
||||
|
||||
infer' ctx sg (B i _) =
|
||||
-- if x : A ∈ Γ
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue