add source locations to inner syntax

This commit is contained in:
rhiannon morris 2023-05-02 03:06:25 +02:00
parent 30fa93ab4e
commit d5f4a012c5
35 changed files with 3210 additions and 2482 deletions

View File

@ -33,4 +33,4 @@ def sym : 0.(A : ★₀) → 0.(x y : A) → 1.(x ≡ y : A) → y ≡ x : A =
def trans : 0.(A : ★₀) → 0.(x y z : A) → def trans : 0.(A : ★₀) → 0.(x y z : A) →
ω.(x ≡ y : A) → ω.(y ≡ z : A) → x ≡ z : A = ω.(x ≡ y : A) → ω.(y ≡ z : A) → x ≡ z : A =
λ A x y z eq1 eq2 ⇒ δ 𝑖 λ A x y z eq1 eq2 ⇒ δ 𝑖
comp [A] @0 @1 (eq1 @𝑖) @𝑖 { 0 _ ⇒ eq1 @0; 1 𝑗 ⇒ eq2 @𝑗 }; comp [A] @0 @1 (eq1 @𝑖) @𝑖 { 0 _ ⇒ eq1 @0; 1 𝑗 ⇒ eq2 @0 };

View File

@ -20,9 +20,10 @@ main : IO ()
main = do main = do
seen <- newIORef SortedSet.empty seen <- newIORef SortedSet.empty
defs <- newIORef SortedMap.empty defs <- newIORef SortedMap.empty
suf <- newIORef $ the Nat 0
for_ (drop 1 !getArgs) $ \file => do for_ (drop 1 !getArgs) $ \file => do
putStrLn "checking \{file}" putStrLn "checking \{file}"
Right res <- fromParserIO ["."] seen defs $ loadProcessFile file Right res <- fromParserIO ["."] seen suf defs $ loadProcessFile noLoc file
| Left err => die $ prettyError True True err | Left err => die $ prettyError True True err
for_ res $ \(name, def) => do for_ res $ \(name, def) => do
putDoc $ map termHL $ nest 2 $ putDoc $ map termHL $ nest 2 $

View File

@ -40,6 +40,10 @@ public export
NContext : Nat -> Type NContext : Nat -> Type
NContext = Context' BaseName NContext = Context' BaseName
public export
BContext : Nat -> Type
BContext = Context' BindName
public export public export
unsnoc : Context tm (S n) -> (Context tm n, tm n) unsnoc : Context tm (S n) -> (Context tm n, tm n)
@ -183,6 +187,10 @@ parameters {auto _ : Applicative f}
traverse f [<] = pure [<] traverse f [<] = pure [<]
traverse f (tel :< x) = [|traverse f tel :< f x|] traverse f (tel :< x) = [|traverse f tel :< f x|]
export %inline
traverse' : (a -> f b) -> Telescope' a from to -> f (Telescope' b from to)
traverse' f = traverse f
infixl 3 `app` infixl 3 `app`
||| like `(<*>)` but with effects ||| like `(<*>)` but with effects
export export
@ -197,6 +205,7 @@ parameters {auto _ : Applicative f}
sequence : Telescope (f . tm) from to -> f (Telescope tm from to) sequence : Telescope (f . tm) from to -> f (Telescope tm from to)
sequence = traverse id sequence = traverse id
parameters {0 tm1, tm2 : Nat -> Type} parameters {0 tm1, tm2 : Nat -> Type}
(f : forall n. tm1 n -> tm2 n) (f : forall n. tm1 n -> tm2 n)
export %inline export %inline
@ -207,6 +216,8 @@ parameters {0 tm1, tm2 : Nat -> Type}
(<$>) : Telescope tm1 from to -> Telescope tm2 from to (<$>) : Telescope tm1 from to -> Telescope tm2 from to
(<$>) = map (<$>) = map
export %inline export %inline
(<*>) : Telescope (\n => tm1 n -> tm2 n) from to -> (<*>) : Telescope (\n => tm1 n -> tm2 n) from to ->
Telescope tm1 from to -> Telescope tm2 from to Telescope tm1 from to -> Telescope tm2 from to
@ -311,3 +322,9 @@ export %inline
export %inline export %inline
(forall n. PrettyHL (tm n)) => PrettyHL (Telescope tm from to) where (forall n. PrettyHL (tm n)) => PrettyHL (Telescope tm from to) where
prettyM tel = separate (hl Delim ";") <$> traverse prettyM (toList tel) prettyM tel = separate (hl Delim ";") <$> traverse prettyM (toList tel)
namespace BContext
export
toNames : BContext n -> SnocList BaseName
toNames = foldl (\xs, x => xs :< x.name) [<]

View File

@ -3,6 +3,7 @@ module Quox.Definition
import public Quox.No import public Quox.No
import public Quox.Syntax import public Quox.Syntax
import public Data.SortedMap import public Data.SortedMap
import public Quox.Loc
import Control.Eff import Control.Eff
import Decidable.Decidable import Decidable.Decidable
@ -25,14 +26,18 @@ record Definition where
qty : GQty qty : GQty
type0 : Term 0 0 type0 : Term 0 0
body0 : DefBody body0 : DefBody
loc_ : Loc
public export %inline public export %inline
mkPostulate : GQty -> (type0 : Term 0 0) -> Definition mkPostulate : GQty -> (type0 : Term 0 0) -> Loc -> Definition
mkPostulate qty type0 = MkDef {qty, type0, body0 = Postulate} mkPostulate qty type0 loc_ = MkDef {qty, type0, body0 = Postulate, loc_}
public export %inline public export %inline
mkDef : GQty -> (type0, term0 : Term 0 0) -> Definition mkDef : GQty -> (type0, term0 : Term 0 0) -> Loc -> Definition
mkDef qty type0 term0 = MkDef {qty, type0, body0 = Concrete term0} mkDef qty type0 term0 loc_ = MkDef {qty, type0, body0 = Concrete term0, loc_}
export Located Definition where def.loc = def.loc_
export Relocatable Definition where setLoc loc = {loc_ := loc}
parameters {d, n : Nat} parameters {d, n : Nat}
@ -46,7 +51,7 @@ parameters {d, n : Nat}
public export %inline public export %inline
toElim : Definition -> Maybe $ Elim d n toElim : Definition -> Maybe $ Elim d n
toElim def = pure $ !def.term :# def.type toElim def = pure $ Ann !def.term def.type def.loc
public export %inline public export %inline
@ -62,9 +67,13 @@ Definitions : Type
Definitions = SortedMap Name Definition Definitions = SortedMap Name Definition
public export public export
0 DefsReader : Type -> Type DefsReader : Type -> Type
DefsReader = ReaderL DEFS Definitions DefsReader = ReaderL DEFS Definitions
public export
DefsState : Type -> Type
DefsState = StateL DEFS Definitions
export export
defs : Has DefsReader fs => Eff fs Definitions defs : Has DefsReader fs => Eff fs Definitions
defs = askAt DEFS defs = askAt DEFS

View File

@ -9,20 +9,41 @@ import Quox.EffExtra
public export public export
0 EqModeState : Type -> Type EqModeState : Type -> Type
EqModeState = State EqMode EqModeState = State EqMode
public export public export
0 EqualEff : List (Type -> Type) Equal : Type -> Type
EqualEff = [ErrorEff, EqModeState] Equal = Eff [ErrorEff, DefsReader, NameGen]
public export public export
0 Equal : Type -> Type Equal_ : Type -> Type
Equal = Eff $ EqualEff Equal_ = Eff [ErrorEff, NameGen, EqModeState]
export export
runEqual : EqMode -> Equal a -> Either Error a runEqualWith_ : EqMode -> NameSuf -> Equal_ a -> (Either Error a, NameSuf)
runEqual mode = extract . runExcept . evalState mode runEqualWith_ mode suf act =
extract $
runNameGenWith suf $
runExcept $
evalState mode act
export
runEqual_ : EqMode -> Equal_ a -> Either Error a
runEqual_ mode act = fst $ runEqualWith_ mode 0 act
export
runEqualWith : NameSuf -> Definitions -> Equal a -> (Either Error a, NameSuf)
runEqualWith suf defs act =
extract $
runStateAt GEN suf $
runReaderAt DEFS defs $
runExcept act
export
runEqual : Definitions -> Equal a -> Either Error a
runEqual defs act = fst $ runEqualWith 0 defs act
export %inline export %inline
@ -30,22 +51,22 @@ mode : Has EqModeState fs => Eff fs EqMode
mode = get mode = get
parameters (ctx : EqContext n) parameters (loc : Loc) (ctx : EqContext n)
private %inline private %inline
clashT : Term 0 n -> Term 0 n -> Term 0 n -> Equal a clashT : Term 0 n -> Term 0 n -> Term 0 n -> Equal_ a
clashT ty s t = throw $ ClashT ctx !mode ty s t clashT ty s t = throw $ ClashT loc ctx !mode ty s t
private %inline private %inline
clashTy : Term 0 n -> Term 0 n -> Equal a clashTy : Term 0 n -> Term 0 n -> Equal_ a
clashTy s t = throw $ ClashTy ctx !mode s t clashTy s t = throw $ ClashTy loc ctx !mode s t
private %inline private %inline
clashE : Elim 0 n -> Elim 0 n -> Equal a clashE : Elim 0 n -> Elim 0 n -> Equal_ a
clashE e f = throw $ ClashE ctx !mode e f clashE e f = throw $ ClashE loc ctx !mode e f
private %inline private %inline
wrongType : Term 0 n -> Term 0 n -> Equal a wrongType : Term 0 n -> Term 0 n -> Equal_ a
wrongType ty s = throw $ WrongType ctx ty s wrongType ty s = throw $ WrongType loc ctx ty s
public export %inline public export %inline
@ -62,8 +83,8 @@ sameTyCon (Enum {}) (Enum {}) = True
sameTyCon (Enum {}) _ = False sameTyCon (Enum {}) _ = False
sameTyCon (Eq {}) (Eq {}) = True sameTyCon (Eq {}) (Eq {}) = True
sameTyCon (Eq {}) _ = False sameTyCon (Eq {}) _ = False
sameTyCon Nat Nat = True sameTyCon (Nat {}) (Nat {}) = True
sameTyCon Nat _ = False sameTyCon (Nat {}) _ = False
sameTyCon (BOX {}) (BOX {}) = True sameTyCon (BOX {}) (BOX {}) = True
sameTyCon (BOX {}) _ = False sameTyCon (BOX {}) _ = False
sameTyCon (E {}) (E {}) = True sameTyCon (E {}) (E {}) = True
@ -80,37 +101,39 @@ sameTyCon (E {}) _ = False
||| * an enum type is a subsingleton if it has zero or one tags. ||| * an enum type is a subsingleton if it has zero or one tags.
||| * a box type is a subsingleton if its content is ||| * a box type is a subsingleton if its content is
public export covering public export covering
isSubSing : Has ErrorEff fs => {n : Nat} -> isSubSing : {n : Nat} -> Definitions -> EqContext n -> Term 0 n -> Equal_ Bool
Definitions -> EqContext n -> Term 0 n -> Eff fs Bool
isSubSing defs ctx ty0 = do isSubSing defs ctx ty0 = do
Element ty0 nc <- whnf defs ctx ty0 Element ty0 nc <- whnf defs ctx ty0.loc ty0
case ty0 of case ty0 of
TYPE _ => pure False TYPE {} => pure False
Pi _ arg res => isSubSing defs (extendTy Zero res.name arg ctx) res.term Pi {arg, res, _} =>
Sig fst snd => isSubSing defs ctx fst `andM` isSubSing defs (extendTy Zero res.name arg ctx) res.term
isSubSing defs (extendTy Zero snd.name fst ctx) snd.term Sig {fst, snd, _} =>
Enum tags => pure $ length (SortedSet.toList tags) <= 1 isSubSing defs ctx fst `andM`
Eq {} => pure True isSubSing defs (extendTy Zero snd.name fst ctx) snd.term
Nat => pure False Enum {cases, _} =>
BOX _ ty => isSubSing defs ctx ty pure $ length (SortedSet.toList cases) <= 1
E (s :# _) => isSubSing defs ctx s Eq {} => pure True
E _ => pure False Nat {} => pure False
Lam _ => pure False BOX {ty, _} => isSubSing defs ctx ty
Pair _ _ => pure False E (Ann {tm, _}) => isSubSing defs ctx tm
Tag _ => pure False E _ => pure False
DLam _ => pure False Lam {} => pure False
Zero => pure False Pair {} => pure False
Succ _ => pure False Tag {} => pure False
Box _ => pure False DLam {} => pure False
Zero {} => pure False
Succ {} => pure False
Box {} => pure False
export export
ensureTyCon : Has ErrorEff fs => ensureTyCon : Has ErrorEff fs =>
(ctx : EqContext n) -> (t : Term 0 n) -> (loc : Loc) -> (ctx : EqContext n) -> (t : Term 0 n) ->
Eff fs (So (isTyConE t)) Eff fs (So (isTyConE t))
ensureTyCon ctx t = case nchoose $ isTyConE t of ensureTyCon loc ctx t = case nchoose $ isTyConE t of
Left y => pure y Left y => pure y
Right n => throw $ NotType (toTyContext ctx) (t // shift0 ctx.dimLen) Right n => throw $ NotType loc (toTyContext ctx) (t // shift0 ctx.dimLen)
||| performs the minimum work required to recompute the type of an elim. ||| performs the minimum work required to recompute the type of an elim.
||| |||
@ -118,10 +141,10 @@ ensureTyCon ctx t = case nchoose $ isTyConE t of
private covering private covering
computeElimTypeE : (defs : Definitions) -> EqContext n -> computeElimTypeE : (defs : Definitions) -> EqContext n ->
(e : Elim 0 n) -> (0 ne : NotRedex defs e) => (e : Elim 0 n) -> (0 ne : NotRedex defs e) =>
Equal (Term 0 n) Equal_ (Term 0 n)
computeElimTypeE defs ectx e = computeElimTypeE defs ectx e =
let Val n = ectx.termLen in let Val n = ectx.termLen in
rethrow $ computeElimType defs (toWhnfContext ectx) e lift $ computeElimType defs (toWhnfContext ectx) e
parameters (defs : Definitions) parameters (defs : Definitions)
mutual mutual
@ -131,55 +154,56 @@ parameters (defs : Definitions)
||| |||
||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠ ||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠
export covering %inline export covering %inline
compare0 : EqContext n -> (ty, s, t : Term 0 n) -> Equal () compare0 : EqContext n -> (ty, s, t : Term 0 n) -> Equal_ ()
compare0 ctx ty s t = compare0 ctx ty s t =
wrapErr (WhileComparingT ctx !mode ty s t) $ do wrapErr (WhileComparingT ctx !mode ty s t) $ do
let Val n = ctx.termLen let Val n = ctx.termLen
Element ty _ <- whnf defs ctx ty Element ty' _ <- whnf defs ctx ty.loc ty
Element s _ <- whnf defs ctx s Element s' _ <- whnf defs ctx s.loc s
Element t _ <- whnf defs ctx t Element t' _ <- whnf defs ctx t.loc t
tty <- ensureTyCon ctx ty tty <- ensureTyCon ty.loc ctx ty'
compare0' ctx ty s t compare0' ctx ty' s' t'
||| converts an elim "Γ ⊢ e" to "Γ, x ⊢ e x", for comparing with ||| converts an elim "Γ ⊢ e" to "Γ, x ⊢ e x", for comparing with
||| a lambda "Γ ⊢ λx ⇒ t" that has been converted to "Γ, x ⊢ t". ||| a lambda "Γ ⊢ λx ⇒ t" that has been converted to "Γ, x ⊢ t".
private %inline private %inline
toLamBody : Elim d n -> Term d (S n) toLamBody : Elim d n -> Term d (S n)
toLamBody e = E $ weakE 1 e :@ BVT 0 toLamBody e = E $ App (weakE 1 e) (BVT 0 e.loc) e.loc
private covering private covering
compare0' : EqContext n -> compare0' : EqContext n ->
(ty, s, t : Term 0 n) -> (ty, s, t : Term 0 n) ->
(0 _ : NotRedex defs ty) => (0 _ : So (isTyConE ty)) => (0 _ : NotRedex defs ty) => (0 _ : So (isTyConE ty)) =>
(0 _ : NotRedex defs s) => (0 _ : NotRedex defs t) => (0 _ : NotRedex defs s) => (0 _ : NotRedex defs t) =>
Equal () Equal_ ()
compare0' ctx (TYPE _) s t = compareType ctx s t compare0' ctx (TYPE {}) s t = compareType ctx s t
compare0' ctx ty@(Pi {qty, arg, res}) s t {n} = local_ Equal $ compare0' ctx ty@(Pi {qty, arg, res, _}) s t {n} = local_ Equal $
case (s, t) of case (s, t) of
-- Γ, x : A ⊢ s = t : B -- Γ, x : A ⊢ s = t : B
-- ------------------------------------------- -- -------------------------------------------
-- Γ ⊢ (λ x ⇒ s) = (λ x ⇒ t) : (π·x : A) → B -- Γ ⊢ (λ x ⇒ s) = (λ x ⇒ t) : (π·x : A) → B
(Lam b1, Lam b2) => compare0 ctx' res.term b1.term b2.term (Lam b1 {}, Lam b2 {}) =>
compare0 ctx' res.term b1.term b2.term
-- Γ, x : A ⊢ s = e x : B -- Γ, x : A ⊢ s = e x : B
-- ----------------------------------- -- -----------------------------------
-- Γ ⊢ (λ x ⇒ s) = e : (π·x : A) → B -- Γ ⊢ (λ x ⇒ s) = e : (π·x : A) → B
(E e, Lam b) => eta e b (E e, Lam b {}) => eta s.loc e b
(Lam b, E e) => eta e b (Lam b {}, E e) => eta s.loc e b
(E e, E f) => Elim.compare0 ctx e f (E e, E f) => Elim.compare0 ctx e f
(Lam _, t) => wrongType ctx ty t (Lam {}, t) => wrongType t.loc ctx ty t
(E _, t) => wrongType ctx ty t (E _, t) => wrongType t.loc ctx ty t
(s, _) => wrongType ctx ty s (s, _) => wrongType s.loc ctx ty s
where where
ctx' : EqContext (S n) ctx' : EqContext (S n)
ctx' = extendTy qty res.name arg ctx ctx' = extendTy qty res.name arg ctx
eta : Elim 0 n -> ScopeTerm 0 n -> Equal () eta : Loc -> Elim 0 n -> ScopeTerm 0 n -> Equal_ ()
eta e (S _ (Y b)) = compare0 ctx' res.term (toLamBody e) b eta _ e (S _ (Y b)) = compare0 ctx' res.term (toLamBody e) b
eta e (S _ (N _)) = clashT ctx ty s t eta loc e (S _ (N _)) = clashT loc ctx ty s t
compare0' ctx ty@(Sig {fst, snd, _}) s t = local_ Equal $ compare0' ctx ty@(Sig {fst, snd, _}) s t = local_ Equal $
case (s, t) of case (s, t) of
@ -188,34 +212,35 @@ parameters (defs : Definitions)
-- Γ ⊢ (s₁, t₁) = (s₂,t₂) : (x : A) × B -- Γ ⊢ (s₁, t₁) = (s₂,t₂) : (x : A) × B
-- --
-- [todo] η for π ≥ 0 maybe -- [todo] η for π ≥ 0 maybe
(Pair sFst sSnd, Pair tFst tSnd) => do (Pair sFst sSnd {}, Pair tFst tSnd {}) => do
compare0 ctx fst sFst tFst compare0 ctx fst sFst tFst
compare0 ctx (sub1 snd (sFst :# fst)) sSnd tSnd compare0 ctx (sub1 snd (Ann sFst fst fst.loc)) sSnd tSnd
(E e, E f) => Elim.compare0 ctx e f (E e, E f) => Elim.compare0 ctx e f
(Pair {}, E _) => clashT ctx ty s t (Pair {}, E _) => clashT s.loc ctx ty s t
(E _, Pair {}) => clashT ctx ty s t (E _, Pair {}) => clashT s.loc ctx ty s t
(Pair {}, t) => wrongType ctx ty t (Pair {}, t) => wrongType t.loc ctx ty t
(E _, t) => wrongType ctx ty t (E _, t) => wrongType t.loc ctx ty t
(s, _) => wrongType ctx ty s (s, _) => wrongType s.loc ctx ty s
compare0' ctx ty@(Enum tags) s t = local_ Equal $ compare0' ctx ty@(Enum {}) s t = local_ Equal $
case (s, t) of case (s, t) of
-- -------------------- -- --------------------
-- Γ ⊢ `t = `t : {ts} -- Γ ⊢ `t = `t : {ts}
-- --
-- t ∈ ts is in the typechecker, not here, ofc -- t ∈ ts is in the typechecker, not here, ofc
(Tag t1, Tag t2) => unless (t1 == t2) $ clashT ctx ty s t (Tag t1 {}, Tag t2 {}) =>
(E e, E f) => Elim.compare0 ctx e f unless (t1 == t2) $ clashT s.loc ctx ty s t
(E e, E f) => Elim.compare0 ctx e f
(Tag _, E _) => clashT ctx ty s t (Tag {}, E _) => clashT s.loc ctx ty s t
(E _, Tag _) => clashT ctx ty s t (E _, Tag {}) => clashT s.loc ctx ty s t
(Tag _, t) => wrongType ctx ty t (Tag {}, t) => wrongType t.loc ctx ty t
(E _, t) => wrongType ctx ty t (E _, t) => wrongType t.loc ctx ty t
(s, _) => wrongType ctx ty s (s, _) => wrongType s.loc ctx ty s
compare0' _ (Eq {}) _ _ = compare0' _ (Eq {}) _ _ =
-- ✨ uip ✨ -- ✨ uip ✨
@ -224,84 +249,85 @@ parameters (defs : Definitions)
-- Γ ⊢ e = f : Eq [i ⇒ A] s t -- Γ ⊢ e = f : Eq [i ⇒ A] s t
pure () pure ()
compare0' ctx Nat s t = local_ Equal $ compare0' ctx nat@(Nat {}) s t = local_ Equal $
case (s, t) of case (s, t) of
-- --------------- -- ---------------
-- Γ ⊢ 0 = 0 : -- Γ ⊢ 0 = 0 :
(Zero, Zero) => pure () (Zero {}, Zero {}) => pure ()
-- Γ ⊢ m = n : -- Γ ⊢ s = t :
-- ------------------------- -- -------------------------
-- Γ ⊢ succ m = succ n : -- Γ ⊢ succ s = succ t :
(Succ m, Succ n) => compare0 ctx Nat m n (Succ s' {}, Succ t' {}) => compare0 ctx nat s' t'
(E e, E f) => Elim.compare0 ctx e f (E e, E f) => Elim.compare0 ctx e f
(Zero, Succ _) => clashT ctx Nat s t (Zero {}, Succ {}) => clashT s.loc ctx nat s t
(Zero, E _) => clashT ctx Nat s t (Zero {}, E _) => clashT s.loc ctx nat s t
(Succ _, Zero) => clashT ctx Nat s t (Succ {}, Zero {}) => clashT s.loc ctx nat s t
(Succ _, E _) => clashT ctx Nat s t (Succ {}, E _) => clashT s.loc ctx nat s t
(E _, Zero) => clashT ctx Nat s t (E _, Zero {}) => clashT s.loc ctx nat s t
(E _, Succ _) => clashT ctx Nat s t (E _, Succ {}) => clashT s.loc ctx nat s t
(Zero, t) => wrongType ctx Nat t (Zero {}, t) => wrongType t.loc ctx nat t
(Succ _, t) => wrongType ctx Nat t (Succ {}, t) => wrongType t.loc ctx nat t
(E _, t) => wrongType ctx Nat t (E _, t) => wrongType t.loc ctx nat t
(s, _) => wrongType ctx Nat s (s, _) => wrongType s.loc ctx nat s
compare0' ctx ty@(BOX q ty') s t = local_ Equal $ compare0' ctx ty@(BOX q ty' {}) s t = local_ Equal $
case (s, t) of case (s, t) of
-- Γ ⊢ s = t : A -- Γ ⊢ s = t : A
-- ----------------------- -- -----------------------
-- Γ ⊢ [s] = [t] : [π.A] -- Γ ⊢ [s] = [t] : [π.A]
(Box s, Box t) => compare0 ctx ty' s t (Box s' {}, Box t' {}) => compare0 ctx ty' s' t'
(E e, E f) => Elim.compare0 ctx e f (E e, E f) => Elim.compare0 ctx e f
(Box _, t) => wrongType ctx ty t (Box {}, t) => wrongType t.loc ctx ty t
(E _, t) => wrongType ctx ty t (E _, t) => wrongType t.loc ctx ty t
(s, _) => wrongType ctx ty s (s, _) => wrongType s.loc ctx ty s
compare0' ctx ty@(E _) s t = do compare0' ctx ty@(E _) s t = do
-- a neutral type can only be inhabited by neutral values -- a neutral type can only be inhabited by neutral values
-- e.g. an abstract value in an abstract type, bound variables, … -- e.g. an abstract value in an abstract type, bound variables, …
E e <- pure s | _ => wrongType ctx ty s let E e = s | _ => wrongType s.loc ctx ty s
E f <- pure t | _ => wrongType ctx ty t E f = t | _ => wrongType t.loc ctx ty t
Elim.compare0 ctx e f Elim.compare0 ctx e f
||| compares two types, using the current variance `mode` for universes. ||| compares two types, using the current variance `mode` for universes.
||| fails if they are not types, even if they would happen to be equal. ||| fails if they are not types, even if they would happen to be equal.
export covering %inline export covering %inline
compareType : EqContext n -> (s, t : Term 0 n) -> Equal () compareType : EqContext n -> (s, t : Term 0 n) -> Equal_ ()
compareType ctx s t = do compareType ctx s t = do
let Val n = ctx.termLen let Val n = ctx.termLen
Element s _ <- whnf defs ctx s Element s' _ <- whnf defs ctx s.loc s
Element t _ <- whnf defs ctx t Element t' _ <- whnf defs ctx t.loc t
ts <- ensureTyCon ctx s ts <- ensureTyCon s.loc ctx s'
tt <- ensureTyCon ctx t tt <- ensureTyCon t.loc ctx t'
st <- either pure (const $ clashTy ctx s t) $ nchoose $ sameTyCon s t st <- either pure (const $ clashTy s.loc ctx s' t') $
compareType' ctx s t nchoose $ sameTyCon s' t'
compareType' ctx s' t'
private covering private covering
compareType' : EqContext n -> (s, t : Term 0 n) -> compareType' : EqContext n -> (s, t : Term 0 n) ->
(0 _ : NotRedex defs s) => (0 _ : So (isTyConE s)) => (0 _ : NotRedex defs s) => (0 _ : So (isTyConE s)) =>
(0 _ : NotRedex defs t) => (0 _ : So (isTyConE t)) => (0 _ : NotRedex defs t) => (0 _ : So (isTyConE t)) =>
(0 _ : So (sameTyCon s t)) => (0 _ : So (sameTyCon s t)) =>
Equal () Equal_ ()
-- equality is the same as subtyping, except with the -- equality is the same as subtyping, except with the
-- "≤" in the TYPE rule being replaced with "=" -- "≤" in the TYPE rule being replaced with "="
compareType' ctx (TYPE k) (TYPE l) = compareType' ctx a@(TYPE k {}) (TYPE l {}) =
-- 𝓀 -- 𝓀
-- ---------------------- -- ----------------------
-- Γ ⊢ Type 𝓀 <: Type -- Γ ⊢ Type 𝓀 <: Type
expectModeU !mode k l expectModeU a.loc !mode k l
compareType' ctx (Pi {qty = sQty, arg = sArg, res = sRes, _}) compareType' ctx a@(Pi {qty = sQty, arg = sArg, res = sRes, _})
(Pi {qty = tQty, arg = tArg, res = tRes, _}) = do (Pi {qty = tQty, arg = tArg, res = tRes, _}) = do
-- Γ ⊢ A₁ :> A₂ Γ, x : A₁ ⊢ B₁ <: B₂ -- Γ ⊢ A₁ :> A₂ Γ, x : A₁ ⊢ B₁ <: B₂
-- ---------------------------------------- -- ----------------------------------------
-- Γ ⊢ (π·x : A₁) → B₁ <: (π·x : A₂) → B₂ -- Γ ⊢ (π·x : A₁) → B₁ <: (π·x : A₂) → B₂
expectEqualQ sQty tQty expectEqualQ a.loc sQty tQty
local flip $ compareType ctx sArg tArg -- contra local flip $ compareType ctx sArg tArg -- contra
compareType (extendTy Zero sRes.name sArg ctx) sRes.term tRes.term compareType (extendTy Zero sRes.name sArg ctx) sRes.term tRes.term
@ -325,21 +351,21 @@ parameters (defs : Definitions)
Term.compare0 ctx sTy.zero sl tl Term.compare0 ctx sTy.zero sl tl
Term.compare0 ctx sTy.one sr tr Term.compare0 ctx sTy.one sr tr
compareType' ctx s@(Enum tags1) t@(Enum tags2) = do compareType' ctx s@(Enum tags1 {}) t@(Enum tags2 {}) = do
-- ------------------ -- ------------------
-- Γ ⊢ {ts} <: {ts} -- Γ ⊢ {ts} <: {ts}
-- --
-- no subtyping based on tag subsets, since that would need -- no subtyping based on tag subsets, since that would need
-- a runtime coercion -- a runtime coercion
unless (tags1 == tags2) $ clashTy ctx s t unless (tags1 == tags2) $ clashTy s.loc ctx s t
compareType' ctx Nat Nat = compareType' ctx (Nat {}) (Nat {}) =
-- ------------ -- ------------
-- Γ ⊢ <: -- Γ ⊢ <:
pure () pure ()
compareType' ctx (BOX pi a) (BOX rh b) = do compareType' ctx (BOX pi a loc) (BOX rh b {}) = do
expectEqualQ pi rh expectEqualQ loc pi rh
compareType ctx a b compareType ctx a b
compareType' ctx (E e) (E f) = do compareType' ctx (E e) (E f) = do
@ -347,13 +373,17 @@ parameters (defs : Definitions)
-- has been inlined by whnf -- has been inlined by whnf
Elim.compare0 ctx e f Elim.compare0 ctx e f
-- Ψ | Γ ⊢₀ e ⇒ Eq [𝑖 ⇒ A] s t
-- -----------------------------
-- Ψ | Γ ⊢ e @0 = s ⇒ A[0/𝑖]
-- Ψ | Γ ⊢ e @1 = s ⇒ A[1/𝑖]
private covering private covering
replaceEnd : EqContext n -> replaceEnd : EqContext n ->
(e : Elim 0 n) -> DimConst -> (0 ne : NotRedex defs e) -> (e : Elim 0 n) -> Loc -> DimConst -> Loc ->
Equal (Elim 0 n) (0 ne : NotRedex defs e) -> Equal_ (Elim 0 n)
replaceEnd ctx e p ne = do replaceEnd ctx e eloc p ploc ne = do
(ty, l, r) <- expectEq defs ctx !(computeElimTypeE defs ctx e) (ty, l, r) <- expectEq defs ctx eloc !(computeElimTypeE defs ctx e)
pure $ ends l r p :# dsub1 ty (K p) pure $ Ann (ends l r p) (dsub1 ty (K p ploc)) eloc
namespace Elim namespace Elim
-- [fixme] the following code ends up repeating a lot of work in the -- [fixme] the following code ends up repeating a lot of work in the
@ -364,133 +394,179 @@ parameters (defs : Definitions)
||| ⚠ **assumes that they have both been typechecked, and have ||| ⚠ **assumes that they have both been typechecked, and have
||| equal types.** ⚠ ||| equal types.** ⚠
export covering %inline export covering %inline
compare0 : EqContext n -> (e, f : Elim 0 n) -> Equal () compare0 : EqContext n -> (e, f : Elim 0 n) -> Equal_ ()
compare0 ctx e f = compare0 ctx e f =
wrapErr (WhileComparingE ctx !mode e f) $ do wrapErr (WhileComparingE ctx !mode e f) $ do
let Val n = ctx.termLen let Val n = ctx.termLen
Element e ne <- whnf defs ctx e Element e' ne <- whnf defs ctx e.loc e
Element f nf <- whnf defs ctx f Element f' nf <- whnf defs ctx f.loc f
-- [fixme] there is a better way to do this "isSubSing" stuff for sure unless !(isSubSing defs ctx =<< computeElimTypeE defs ctx e') $
unless !(isSubSing defs ctx !(computeElimTypeE defs ctx e)) $ compare0' ctx e' f' ne nf
compare0' ctx e f ne nf
private covering private covering
compare0' : EqContext n -> compare0' : EqContext n ->
(e, f : Elim 0 n) -> (e, f : Elim 0 n) ->
(0 ne : NotRedex defs e) -> (0 nf : NotRedex defs f) -> (0 ne : NotRedex defs e) -> (0 nf : NotRedex defs f) ->
Equal () Equal_ ()
-- replace applied equalities with the appropriate end first -- replace applied equalities with the appropriate end first
-- e.g. e : Eq [i ⇒ A] s t ⊢ e 𝟎 = s : A𝟎/i -- (see `replaceEnd`)
-- compare0' ctx (DApp e (K p ploc) loc) f ne nf =
-- [todo] maybe have typed whnf and do this (and η???) there instead compare0 ctx !(replaceEnd ctx e loc p ploc $ noOr1 ne) f
compare0' ctx (e :% K p) f ne nf = compare0' ctx e (DApp f (K q qloc) loc) ne nf =
compare0 ctx !(replaceEnd ctx e p $ noOr1 ne) f compare0 ctx e !(replaceEnd ctx f loc q qloc $ noOr1 nf)
compare0' ctx e (f :% K q) ne nf =
compare0 ctx e !(replaceEnd ctx f q $ noOr1 nf)
compare0' ctx e@(F x) f@(F y) _ _ = unless (x == y) $ clashE ctx e f compare0' ctx e@(F x {}) f@(F y {}) _ _ =
compare0' ctx e@(F _) f _ _ = clashE ctx e f unless (x == y) $ 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) _ _ = unless (i == j) $ clashE ctx e f compare0' ctx e@(B i {}) f@(B j {}) _ _ =
compare0' ctx e@(B _) f _ _ = clashE ctx e f unless (i == j) $ clashE e.loc ctx e f
compare0' ctx e@(B {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (e :@ s) (f :@ t) ne nf = -- Ψ | Γ ⊢ e = f ⇒ π.(x : A) → B
-- Ψ | Γ ⊢ s = t ⇐ A
-- -------------------------------
-- Ψ | Γ ⊢ e s = f t ⇒ B[s∷A/x]
compare0' ctx (App e s eloc) (App f t floc) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
(_, arg, _) <- expectPi defs ctx =<< (_, arg, _) <- expectPi defs ctx eloc =<<
computeElimTypeE defs ctx e @{noOr1 ne} computeElimTypeE defs ctx e @{noOr1 ne}
Term.compare0 ctx arg s t Term.compare0 ctx arg s t
compare0' ctx e@(_ :@ _) f _ _ = clashE ctx e f compare0' ctx e@(App {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (CasePair epi e eret ebody) -- Ψ | Γ ⊢ e = f ⇒ (x : A) × B
(CasePair fpi f fret fbody) ne nf = -- Ψ | Γ, 0.p : (x : A) × B ⊢ Q = R
-- Ψ | Γ, x : A, y : B ⊢ s = t ⇐ Q[((x, y) ∷ (x : A) × B)/p]
-- -----------------------------------------------------------
-- Ψ | Γ ⊢ caseπ e return Q of { (x, y) ⇒ s }
-- = caseπ f return R of { (x, y) ⇒ t } ⇒ Q[e/p]
compare0' ctx (CasePair epi e eret ebody eloc)
(CasePair fpi f fret fbody {}) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
(fst, snd) <- expectSig defs ctx ety (fst, snd) <- expectSig defs ctx eloc ety
let [< x, y] = ebody.names let [< x, y] = ebody.names
Term.compare0 (extendTyN [< (epi, x, fst), (epi, y, snd.term)] ctx) Term.compare0 (extendTyN [< (epi, x, fst), (epi, y, snd.term)] ctx)
(substCasePairRet ety eret) (substCasePairRet ebody.names ety eret)
ebody.term fbody.term ebody.term fbody.term
expectEqualQ epi fpi expectEqualQ e.loc epi fpi
compare0' ctx e@(CasePair {}) f _ _ = clashE ctx e f compare0' ctx e@(CasePair {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (CaseEnum epi e eret earms) -- Ψ | Γ ⊢ e = f ⇒ {𝐚s}
(CaseEnum fpi f fret farms) ne nf = -- Ψ | Γ, x : {𝐚s} ⊢ Q = R
-- Ψ | Γ ⊢ sᵢ = tᵢ ⇐ Q[𝐚ᵢ∷{𝐚s}]
-- --------------------------------------------------
-- Ψ | Γ ⊢ caseπ e return Q of { '𝐚ᵢ ⇒ sᵢ }
-- = caseπ f return R of { '𝐚ᵢ ⇒ tᵢ } ⇒ Q[e/x]
compare0' ctx (CaseEnum epi e eret earms eloc)
(CaseEnum fpi f fret farms floc) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
for_ !(expectEnum defs ctx ety) $ \t => for_ !(expectEnum defs ctx eloc ety) $ \t => do
compare0 ctx (sub1 eret $ Tag t :# ety) l <- lookupArm eloc t earms
!(lookupArm t earms) !(lookupArm t farms) r <- lookupArm floc t farms
expectEqualQ epi fpi compare0 ctx (sub1 eret $ Ann (Tag t l.loc) ety l.loc) l r
expectEqualQ eloc epi fpi
where where
lookupArm : TagVal -> CaseEnumArms d n -> Equal (Term d n) lookupArm : Loc -> TagVal -> CaseEnumArms d n -> Equal_ (Term d n)
lookupArm t arms = case lookup t arms of lookupArm loc t arms = case lookup t arms of
Just arm => pure arm Just arm => pure arm
Nothing => throw $ TagNotIn t (fromList $ keys arms) Nothing => throw $ TagNotIn loc t (fromList $ keys arms)
compare0' ctx e@(CaseEnum {}) f _ _ = clashE ctx e f compare0' ctx e@(CaseEnum {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (CaseNat epi epi' e eret ezer esuc) -- Ψ | Γ ⊢ e = f ⇒
(CaseNat fpi fpi' f fret fzer fsuc) ne nf = -- Ψ | Γ, x : ⊢ Q = R
-- Ψ | Γ ⊢ s₀ = t₀ ⇐ Q[(0 ∷ )/x]
-- Ψ | Γ, x : , y : Q ⊢ s₁ = t₁ ⇐ Q[(succ x ∷ )/x]
-- -----------------------------------------------------
-- Ψ | Γ ⊢ caseπ e return Q of { 0 ⇒ s₀; x, π.y ⇒ s₁ }
-- = caseπ f return R of { 0 ⇒ t₀; x, π.y ⇒ t₁ }
-- ⇒ Q[e/x]
compare0' ctx (CaseNat epi epi' e eret ezer esuc eloc)
(CaseNat fpi fpi' f fret fzer fsuc floc) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
compare0 ctx (sub1 eret (Zero :# Nat)) ezer fzer compare0 ctx
(sub1 eret (Ann (Zero ezer.loc) (Nat ezer.loc) ezer.loc))
ezer fzer
let [< p, ih] = esuc.names let [< p, ih] = esuc.names
compare0 (extendTyN [< (epi, p, Nat), (epi', ih, eret.term)] ctx) compare0
(substCaseSuccRet eret) (extendTyN [< (epi, p, Nat p.loc), (epi', ih, eret.term)] ctx)
esuc.term fsuc.term (substCaseSuccRet esuc.names eret) esuc.term fsuc.term
expectEqualQ epi fpi expectEqualQ e.loc epi fpi
expectEqualQ epi' fpi' expectEqualQ e.loc epi' fpi'
compare0' ctx e@(CaseNat {}) f _ _ = clashE ctx e f compare0' ctx e@(CaseNat {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (CaseBox epi e eret ebody) -- Ψ | Γ ⊢ e = f ⇒ [ρ. A]
(CaseBox fpi f fret fbody) ne nf = -- Ψ | Γ, x : [ρ. A] ⊢ Q = R
-- Ψ | Γ, x : A ⊢ s = t ⇐ Q[([x] ∷ [ρ. A])/x]
-- --------------------------------------------------
-- Ψ | Γ ⊢ caseπ e return Q of { [x] ⇒ s }
-- = caseπ f return R of { [x] ⇒ t } ⇒ Q[e/x]
compare0' ctx (CaseBox epi e eret ebody eloc)
(CaseBox fpi f fret fbody floc) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
(q, ty) <- expectBOX defs ctx ety (q, ty) <- expectBOX defs ctx eloc ety
compare0 (extendTy (epi * q) ebody.name ty ctx) compare0 (extendTy (epi * q) ebody.name ty ctx)
(substCaseBoxRet ety eret) (substCaseBoxRet ebody.name ety eret)
ebody.term fbody.term ebody.term fbody.term
expectEqualQ epi fpi expectEqualQ eloc epi fpi
compare0' ctx e@(CaseBox {}) f _ _ = clashE ctx e f compare0' ctx e@(CaseBox {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (s :# a) (t :# b) _ _ = -- Ψ | Γ ⊢ s <: t : B
-- --------------------------------
-- Ψ | Γ ⊢ (s ∷ A) <: (t ∷ B) ⇒ B
--
-- and similar for :> and A
compare0' ctx (Ann s a _) (Ann t b _) _ _ =
let ty = case !mode of Super => a; _ => b in let ty = case !mode of Super => a; _ => b in
Term.compare0 ctx ty s t Term.compare0 ctx ty s t
compare0' ctx (Coe ty1 p1 q1 (E val1)) (Coe ty2 p2 q2 (E val2)) ne nf = do -- Ψ | Γ ⊢ Ap₁/𝑖 <: Bp₂/𝑖
-- Ψ | Γ ⊢ Aq₁/𝑖 <: Bq₂/𝑖
-- Ψ | Γ ⊢ e <: f ⇒ _
-- (non-neutral forms have the coercion already pushed in)
-- -----------------------------------------------------------
-- Ψ | Γ ⊢ coe [𝑖 ⇒ A] @p₁ @q₁ e
-- <: coe [𝑖 ⇒ B] @p₂ @q₂ f ⇒ Bq₂/𝑖
compare0' ctx (Coe ty1 p1 q1 (E val1) _)
(Coe ty2 p2 q2 (E val2) _) ne nf = do
compareType ctx (dsub1 ty1 p1) (dsub1 ty2 p2) compareType ctx (dsub1 ty1 p1) (dsub1 ty2 p2)
compareType ctx (dsub1 ty1 q1) (dsub1 ty2 q2) compareType ctx (dsub1 ty1 q1) (dsub1 ty2 q2)
compare0 ctx val1 val2 compare0 ctx val1 val2
compare0' ctx e@(Coe {}) f _ _ = clashE ctx e f compare0' ctx e@(Coe {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (Comp _ _ _ _ (K _) _ _) _ ne _ = void $ absurd $ noOr2 ne -- (no neutral compositions in a closed dctx)
compare0' ctx (Comp _ _ _ _ (B i) _ _) _ _ _ = absurd i compare0' _ (Comp {r = K e _, _}) _ ne _ = void $ absurd $ noOr2 ne
compare0' ctx _ (Comp _ _ _ _ (K _) _ _) _ nf = void $ absurd $ noOr2 nf compare0' _ (Comp {r = B i _, _}) _ _ _ = absurd i
compare0' _ _ (Comp {r = K _ _, _}) _ nf = void $ absurd $ noOr2 nf
compare0' ctx (TypeCase ty1 ret1 arms1 def1) compare0' ctx (TypeCase ty1 ret1 arms1 def1 eloc)
(TypeCase ty2 ret2 arms2 def2) (TypeCase ty2 ret2 arms2 def2 floc) ne _ =
ne _ =
local_ Equal $ do local_ Equal $ do
compare0 ctx ty1 ty2 compare0 ctx ty1 ty2
u <- expectTYPE defs ctx =<< computeElimTypeE defs ctx ty1 @{noOr1 ne} u <- expectTYPE defs ctx eloc =<<
computeElimTypeE defs ctx ty1 @{noOr1 ne}
compareType ctx ret1 ret2 compareType ctx ret1 ret2
compareType ctx def1 def2 compareType ctx def1 def2
for_ allKinds $ \k => for_ allKinds $ \k =>
compareArm ctx k ret1 u compareArm ctx k ret1 u
(lookupPrecise k arms1) (lookupPrecise k arms2) def1 (lookupPrecise k arms1) (lookupPrecise k arms2) def1
compare0' ctx e@(TypeCase {}) f _ _ = clashE ctx e f compare0' ctx e@(TypeCase {}) f _ _ = clashE e.loc ctx e f
compare0' ctx (s :# a) f _ _ = Term.compare0 ctx a s (E f) compare0' ctx (Ann s a _) f _ _ = Term.compare0 ctx a s (E f)
compare0' ctx e (t :# b) _ _ = Term.compare0 ctx b (E e) t compare0' ctx e (Ann t b _) _ _ = Term.compare0 ctx b (E e) t
compare0' ctx e@(_ :# _) f _ _ = clashE ctx e f compare0' ctx e@(Ann {}) f _ _ = clashE e.loc ctx e f
||| compare two type-case branches, which came from the arms of the given ||| compare two type-case branches, which came from the arms of the given
||| kind. `ret` is the return type of the case expression, and `u` is the ||| kind. `ret` is the return type of the case expression, and `u` is the
@ -500,7 +576,7 @@ parameters (defs : Definitions)
(ret : Term 0 n) -> (u : Universe) -> (ret : Term 0 n) -> (u : Universe) ->
(b1, b2 : Maybe (TypeCaseArmBody k 0 n)) -> (b1, b2 : Maybe (TypeCaseArmBody k 0 n)) ->
(def : Term 0 n) -> (def : Term 0 n) ->
Equal () Equal_ ()
compareArm {b1 = Nothing, b2 = Nothing, _} = pure () compareArm {b1 = Nothing, b2 = Nothing, _} = pure ()
compareArm ctx k ret u b1 b2 def = compareArm ctx k ret u b1 b2 def =
let def = SN def in let def = SN def in
@ -510,22 +586,22 @@ parameters (defs : Definitions)
compareArm_ : EqContext n -> (k : TyConKind) -> compareArm_ : EqContext n -> (k : TyConKind) ->
(ret : Term 0 n) -> (u : Universe) -> (ret : Term 0 n) -> (u : Universe) ->
(b1, b2 : TypeCaseArmBody k 0 n) -> (b1, b2 : TypeCaseArmBody k 0 n) ->
Equal () Equal_ ()
compareArm_ ctx KTYPE ret u b1 b2 = compareArm_ ctx KTYPE ret u b1 b2 =
compare0 ctx ret b1.term b2.term compare0 ctx ret b1.term b2.term
compareArm_ ctx KPi ret u b1 b2 = do compareArm_ ctx KPi ret u b1 b2 = do
let [< a, b] = b1.names let [< a, b] = b1.names
ctx = extendTyN ctx = extendTyN
[< (Zero, a, TYPE u), [< (Zero, a, TYPE u a.loc),
(Zero, b, Arr Zero (BVT 0) (TYPE u))] ctx (Zero, b, Arr Zero (BVT 0 b.loc) (TYPE u b.loc) b.loc)] ctx
compare0 ctx (weakT 2 ret) b1.term b2.term compare0 ctx (weakT 2 ret) b1.term b2.term
compareArm_ ctx KSig ret u b1 b2 = do compareArm_ ctx KSig ret u b1 b2 = do
let [< a, b] = b1.names let [< a, b] = b1.names
ctx = extendTyN ctx = extendTyN
[< (Zero, a, TYPE u), [< (Zero, a, TYPE u a.loc),
(Zero, b, Arr Zero (BVT 0) (TYPE u))] ctx (Zero, b, Arr Zero (BVT 0 b.loc) (TYPE u b.loc) b.loc)] ctx
compare0 ctx (weakT 2 ret) b1.term b2.term compare0 ctx (weakT 2 ret) b1.term b2.term
compareArm_ ctx KEnum ret u b1 b2 = compareArm_ ctx KEnum ret u b1 b2 =
@ -534,70 +610,76 @@ parameters (defs : Definitions)
compareArm_ ctx KEq ret u b1 b2 = do compareArm_ ctx KEq ret u b1 b2 = do
let [< a0, a1, a, l, r] = b1.names let [< a0, a1, a, l, r] = b1.names
ctx = extendTyN ctx = extendTyN
[< (Zero, a0, TYPE u), [< (Zero, a0, TYPE u a0.loc),
(Zero, a1, TYPE u), (Zero, a1, TYPE u a1.loc),
(Zero, a, Eq0 (TYPE u) (BVT 1) (BVT 0)), (Zero, a, Eq0 (TYPE u a.loc)
(Zero, l, BVT 2), (BVT 1 a0.loc) (BVT 0 a1.loc) a.loc),
(Zero, r, BVT 2)] ctx (Zero, l, BVT 2 a0.loc),
(Zero, r, BVT 2 a1.loc)] ctx
compare0 ctx (weakT 5 ret) b1.term b2.term compare0 ctx (weakT 5 ret) b1.term b2.term
compareArm_ ctx KNat ret u b1 b2 = compareArm_ ctx KNat ret u b1 b2 =
compare0 ctx ret b1.term b2.term compare0 ctx ret b1.term b2.term
compareArm_ ctx KBOX ret u b1 b2 = do compareArm_ ctx KBOX ret u b1 b2 = do
let ctx = extendTy Zero b1.name (TYPE u) ctx let ctx = extendTy Zero b1.name (TYPE u b1.name.loc) ctx
compare0 ctx (weakT 1 ret) b1.term b1.term compare0 ctx (weakT 1 ret) b1.term b1.term
parameters {auto _ : (Has DefsReader fs, Has ErrorEff fs)} (ctx : TyContext d n) parameters (loc : Loc) (ctx : TyContext d n)
-- [todo] only split on the dvars that are actually used anywhere in -- [todo] only split on the dvars that are actually used anywhere in
-- the calls to `splits` -- the calls to `splits`
parameters (mode : EqMode) parameters (mode : EqMode)
private
fromEqual_ : Equal_ a -> Equal a
fromEqual_ act = lift $ evalState mode act
private
eachFace : Applicative f => (EqContext n -> DSubst d 0 -> f ()) -> f ()
eachFace act =
for_ (splits loc ctx.dctx) $ \th => act (makeEqContext ctx th) th
private
runCompare : (Definitions -> EqContext n -> DSubst d 0 -> Equal_ ()) ->
Equal ()
runCompare act = fromEqual_ $ eachFace $ act !defs
namespace Term namespace Term
export covering export covering
compare : (ty, s, t : Term d n) -> Eff fs () compare : (ty, s, t : Term d n) -> Equal ()
compare ty s t = compare ty s t = runCompare $ \defs, ectx, th =>
map fst $ runState @{Z} mode $ compare0 defs ectx (ty // th) (s // th) (t // th)
for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in
lift $ compare0 !defs ectx (ty // th) (s // th) (t // th)
export covering export covering
compareType : (s, t : Term d n) -> Eff fs () compareType : (s, t : Term d n) -> Equal ()
compareType s t = compareType s t = runCompare $ \defs, ectx, th =>
map fst $ runState @{Z} mode $ compareType defs ectx (s // th) (t // th)
for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in
lift $ compareType !defs ectx (s // th) (t // th)
namespace Elim namespace Elim
||| you don't have to pass the type in but the arguments must still be ||| you don't have to pass the type in but the arguments must still be
||| of the same type!! ||| of the same type!!
export covering %inline export covering
compare : (e, f : Elim d n) -> Eff fs () compare : (e, f : Elim d n) -> Equal ()
compare e f = compare e f = runCompare $ \defs, ectx, th =>
map fst $ runState @{Z} mode $ compare0 defs ectx (e // th) (f // th)
for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in
lift $ compare0 !defs ectx (e // th) (f // th)
namespace Term namespace Term
export covering %inline export covering %inline
equal, sub, super : (ty, s, t : Term d n) -> Eff fs () equal, sub, super : (ty, s, t : Term d n) -> Equal ()
equal = compare Equal equal = compare Equal
sub = compare Sub sub = compare Sub
super = compare Super super = compare Super
export covering %inline export covering %inline
equalType, subtype, supertype : (s, t : Term d n) -> Eff fs () equalType, subtype, supertype : (s, t : Term d n) -> Equal ()
equalType = compareType Equal equalType = compareType Equal
subtype = compareType Sub subtype = compareType Sub
supertype = compareType Super supertype = compareType Super
namespace Elim namespace Elim
export covering %inline export covering %inline
equal, sub, super : (e, f : Elim d n) -> Eff fs () equal, sub, super : (e, f : Elim d n) -> Equal ()
equal = compare Equal equal = compare Equal
sub = compare Sub sub = compare Sub
super = compare Super super = compare Super

View File

@ -1,7 +1,8 @@
||| file locations ||| file locations
module Quox.Loc module Quox.Loc
import Text.Bounded import public Text.Bounded
import Data.SortedMap
import Derive.Prelude import Derive.Prelude
%default total %default total
@ -11,53 +12,110 @@ public export
FileName : Type FileName : Type
FileName = String FileName = String
%runElab derive "Bounds" [Ord]
public export public export
record Loc' where data Loc_ = NoLoc | YesLoc FileName Bounds
%name Loc_ loc
%runElab derive "Loc_" [Eq, Ord, Show]
||| a wrapper for locations which are always considered equal
public export
record Loc where
constructor L constructor L
fname : FileName val : Loc_
startLine, startCol, endLine, endCol : Int %name Loc loc
%name Loc' loc %runElab derive "Loc" [Show]
%runElab derive "Loc'" [Eq, Ord, Show]
public export export %inline Eq Loc where _ == _ = True
Loc : Type export %inline Ord Loc where compare _ _ = EQ
Loc = Maybe Loc'
export public export %inline
noLoc : Loc
noLoc = L NoLoc
public export %inline
makeLoc : FileName -> Bounds -> Loc makeLoc : FileName -> Bounds -> Loc
makeLoc fname (MkBounds {startLine, startCol, endLine, endCol}) = makeLoc = L .: YesLoc
Just $ L {fname, startLine, startCol, endLine, endCol}
export export
onlyStart_ : Loc_ -> Loc_
onlyStart_ NoLoc = NoLoc
onlyStart_ (YesLoc fname (MkBounds sl sc _ _)) =
YesLoc fname $ MkBounds sl sc sl sc
export %inline
onlyStart : Loc -> Loc onlyStart : Loc -> Loc
onlyStart Nothing = Nothing onlyStart = {val $= onlyStart_}
onlyStart (Just (L fname sl sc _ _)) = Just $ L fname sl sc sl sc
export export
onlyEnd_ : Loc_ -> Loc_
onlyEnd_ NoLoc = NoLoc
onlyEnd_ (YesLoc fname (MkBounds _ _ el ec)) =
YesLoc fname $ MkBounds el ec el ec
export %inline
onlyEnd : Loc -> Loc onlyEnd : Loc -> Loc
onlyEnd Nothing = Nothing onlyEnd = {val $= onlyEnd_}
onlyEnd (Just (L fname _ _ el ec)) = Just $ L fname el ec el ec
export export
extend : Loc -> Bounds -> Loc extend_ : Loc_ -> Bounds -> Loc_
extend Nothing _ = Nothing extend_ NoLoc _ = NoLoc
extend (Just (L fname sl1 sc1 el1 ec1)) (MkBounds sl2 sc2 el2 ec2) = extend_ (YesLoc fname (MkBounds sl1 sc1 el1 ec1)) (MkBounds sl2 sc2 el2 ec2) =
let (sl, sc) = (sl1, sc1) `min` (sl2, sc2) let (sl, sc) = (sl1, sc1) `min` (sl2, sc2)
(el, ec) = (el1, ec1) `max` (el2, ec2) (el, ec) = (el1, ec1) `max` (el2, ec2)
in in
Just $ L fname sl sc el ec YesLoc fname $ MkBounds sl sc el ec
export
extend : Loc -> Bounds -> Loc
extend l b = L $ extend_ l.val b
export export
extend' : Loc -> Maybe Bounds -> Loc extend' : Loc -> Maybe Bounds -> Loc
extend' l b = maybe l (extend l) b extend' l b = maybe l (extend l) b
namespace Loc_
export
(.bounds) : Loc_ -> Maybe Bounds
(YesLoc _ b).bounds = Just b
(NoLoc).bounds = Nothing
namespace Loc
export
(.bounds) : Loc -> Maybe Bounds
l.bounds = l.val.bounds
export %inline
extendL : Loc -> Loc -> Loc
extendL l1 l2 = l1 `extend'` l2.bounds
infixr 1 `or_`, `or`
export %inline
or_ : Loc_ -> Loc_ -> Loc_
or_ l1@(YesLoc {}) _ = l1
or_ NoLoc l2 = l2
export %inline
or : Loc -> Loc -> Loc
or (L l1) (L l2) = L $ l1 `or_` l2
public export public export
interface Located a where (.loc) : a -> Loc interface Located a where (.loc) : a -> Loc
export public export
(.bounds) : Loc -> Maybe Bounds 0 Located1 : (a -> Type) -> Type
(Just (L {fname, startLine, startCol, endLine, endCol})).bounds = Located1 f = forall x. Located (f x)
Just $ MkBounds {startLine, startCol, endLine, endCol}
(Nothing).bounds = Nothing public export
interface Located a => Relocatable a where setLoc : Loc -> a -> a
public export
0 Relocatable1 : (a -> Type) -> Type
Relocatable1 f = forall x. Relocatable (f x)

View File

@ -1,8 +1,10 @@
module Quox.Name module Quox.Name
import Quox.Loc
import Quox.CharExtra import Quox.CharExtra
import public Data.SnocList import public Data.SnocList
import Data.List import Data.List
import Control.Eff
import Text.Lexer import Text.Lexer
import Derive.Prelude import Derive.Prelude
@ -12,16 +14,22 @@ import Derive.Prelude
%language ElabReflection %language ElabReflection
public export
NameSuf : Type
NameSuf = Nat
public export public export
data BaseName data BaseName
= UN String -- user-given name = UN String -- user-given name
| Unused -- "_" | MN String NameSuf -- machine-generated name
| Unused -- "_"
%runElab derive "BaseName" [Eq, Ord] %runElab derive "BaseName" [Eq, Ord]
export export
baseStr : BaseName -> String baseStr : BaseName -> String
baseStr (UN x) = x baseStr (UN x) = x
baseStr Unused = "_" baseStr (MN x i) = "\{x}#\{show i}"
baseStr Unused = "_"
export Show BaseName where show = baseStr export Show BaseName where show = baseStr
export FromString BaseName where fromString = UN export FromString BaseName where fromString = UN
@ -83,6 +91,17 @@ export FromString PName where fromString = MakePName [<]
export FromString Name where fromString = fromPBaseName export FromString Name where fromString = fromPBaseName
public export
record BindName where
constructor BN
name : BaseName
loc_ : Loc
%runElab derive "BindName" [Eq, Ord, Show]
export Located BindName where n.loc = n.loc_
export Relocatable BindName where setLoc loc (BN x _) = BN x loc
export export
toDotsP : PName -> String toDotsP : PName -> String
toDotsP x = fastConcat $ cast $ map (<+> ".") x.mods :< x.base toDotsP x = fastConcat $ cast $ map (<+> ".") x.mods :< x.base
@ -140,3 +159,41 @@ isName str =
case scan name [] (unpack str) of case scan name [] (unpack str) of
Just (_, []) => True Just (_, []) => True
_ => False _ => False
public export
data GenTag = GEN
public export
NameGen : Type -> Type
NameGen = StateL GEN NameSuf
export
runNameGenWith : Has NameGen fs =>
NameSuf -> Eff fs a -> Eff (fs - NameGen) (a, NameSuf)
runNameGenWith = runStateAt GEN
export
runNameGen : Has NameGen fs => Eff fs a -> Eff (fs - NameGen) a
runNameGen = map fst . runNameGenWith 0
||| generate a fresh name with the given base
export
mn : Has NameGen fs => PBaseName -> Eff fs BaseName
mn base = do
i <- getAt GEN
modifyAt GEN S
pure $ MN base i
||| generate a fresh binding name with the given base and
||| (optionally) location `loc`
export
mnb : Has NameGen fs =>
PBaseName -> {default noLoc loc : Loc} -> Eff fs BindName
mnb base = pure $ BN !(mn base) loc
export
fresh : Has NameGen fs => BindName -> Eff fs BindName
fresh (BN (UN str) loc) = mnb str {loc}
fresh (BN (MN str k) loc) = mnb str {loc}
fresh (BN Unused loc) = mnb "x" {loc}

View File

@ -1,4 +1,4 @@
||| take freshly-parsed input, translate it to core, and check it ||| take freshly-parsed input, scope check, type check, add to env
module Quox.Parser.FromParser module Quox.Parser.FromParser
import Quox.Parser.Syntax import Quox.Parser.Syntax
@ -41,19 +41,19 @@ data StateTag = NS | SEEN
public export public export
FromParserPure : List (Type -> Type) FromParserPure : List (Type -> Type)
FromParserPure = FromParserPure =
[Except Error, StateL DEFS Definitions, StateL NS Mods] [Except Error, DefsState, StateL NS Mods, NameGen]
public export public export
FromParserEff : List (Type -> Type) LoadFile' : List (Type -> Type)
FromParserEff = LoadFile' = [IO, StateL SEEN SeenFiles, Reader IncludePath]
[Except Error, StateL DEFS Definitions, StateL NS Mods,
Reader IncludePath, StateL SEEN SeenFiles, IO]
public export public export
FromParser : Type -> Type LoadFile : List (Type -> Type)
FromParser = Eff FromParserEff LoadFile = LoadFile' ++ [Except Error]
-- [todo] put the locs in the core ast, obv public export
FromParserIO : List (Type -> Type)
FromParserIO = FromParserPure ++ LoadFile'
parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a) parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
@ -70,31 +70,32 @@ parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
export export
fromPDimWith : Has (Except Error) fs => fromPDimWith : Has (Except Error) fs =>
Context' PatVar d -> PDim -> Eff fs (Dim d) Context' PatVar d -> PDim -> Eff fs (Dim d)
fromPDimWith ds (K e _) = pure $ K e fromPDimWith ds (K e loc) = pure $ K e loc
fromPDimWith ds (V i _) = fromPDimWith ds (V i loc) =
fromBaseName (pure . B) (const $ throw $ DimNotInScope i) ds i fromBaseName (\i => pure $ B i loc)
(const $ throw $ DimNotInScope loc i) ds i
private private
avoidDim : Has (Except Error) fs => avoidDim : Has (Except Error) fs =>
Context' PatVar d -> PName -> Eff fs Name Context' PatVar d -> Loc -> PName -> Eff fs Name
avoidDim ds x = avoidDim ds loc x =
fromName (const $ throw $ DimNameInTerm x.base) (pure . fromPName) ds x fromName (const $ throw $ DimNameInTerm loc x.base) (pure . fromPName) ds x
private private
resolveName : Mods -> Name -> Eff FromParserPure (Term d n) resolveName : Mods -> Loc -> Name -> Eff FromParserPure (Term d n)
resolveName ns x = resolveName ns loc x =
let here = addMods ns x in let here = addMods ns x in
if isJust $ lookup here !(getAt DEFS) then if isJust $ lookup here !(getAt DEFS) then
pure $ FT here pure $ FT here loc
else do else do
let ns :< _ = ns let ns :< _ = ns
| _ => throw $ TermNotInScope x | _ => throw $ TermNotInScope loc x
resolveName ns x resolveName ns loc x
export export
fromPatVar : PatVar -> BaseName fromPatVar : PatVar -> BindName
fromPatVar (Unused _) = Unused fromPatVar (Unused loc) = BN Unused loc
fromPatVar (PV x _) = UN x fromPatVar (PV x loc) = BN (UN x) loc
export export
fromPQty : PQty -> Qty fromPQty : PQty -> Qty
@ -110,93 +111,112 @@ mutual
fromPTermWith : Context' PatVar d -> Context' PatVar n -> fromPTermWith : Context' PatVar d -> Context' PatVar n ->
PTerm -> Eff FromParserPure (Term d n) PTerm -> Eff FromParserPure (Term d n)
fromPTermWith ds ns t0 = case t0 of fromPTermWith ds ns t0 = case t0 of
TYPE k _ => TYPE k loc =>
pure $ TYPE k pure $ TYPE k loc
Pi pi x s t _ => Pi pi x s t loc =>
Pi (fromPQty pi) Pi (fromPQty pi)
<$> fromPTermWith ds ns s <$> fromPTermWith ds ns s
<*> fromPTermTScope ds ns [< x] t <*> fromPTermTScope ds ns [< x] t
<*> pure loc
Lam x s _ => Lam x s loc =>
Lam <$> fromPTermTScope ds ns [< x] s Lam <$> fromPTermTScope ds ns [< x] s <*> pure loc
App s t _ => App s t loc =>
map E $ (:@) <$> fromPTermElim ds ns s <*> fromPTermWith ds ns t map E $ App
<$> fromPTermElim ds ns s
<*> fromPTermWith ds ns t
<*> pure loc
Sig x s t _ => Sig x s t loc =>
Sig <$> fromPTermWith ds ns s <*> fromPTermTScope ds ns [< x] t Sig <$> fromPTermWith ds ns s
<*> fromPTermTScope ds ns [< x] t
<*> pure loc
Pair s t _ => Pair s t loc =>
Pair <$> fromPTermWith ds ns s <*> fromPTermWith ds ns t Pair <$> fromPTermWith ds ns s <*> fromPTermWith ds ns t <*> pure loc
Case pi pair (r, ret) (CasePair (x, y) body _) _ => Case pi pair (r, ret) (CasePair (x, y) body _) loc =>
map E $ CasePair (fromPQty pi) map E $ CasePair (fromPQty pi)
<$> fromPTermElim ds ns pair <$> fromPTermElim ds ns pair
<*> fromPTermTScope ds ns [< r] ret <*> fromPTermTScope ds ns [< r] ret
<*> fromPTermTScope ds ns [< x, y] body <*> fromPTermTScope ds ns [< x, y] body
<*> pure loc
Case pi tag (r, ret) (CaseEnum arms _) _ => Case pi tag (r, ret) (CaseEnum arms _) loc =>
map E $ CaseEnum (fromPQty pi) map E $ CaseEnum (fromPQty pi)
<$> fromPTermElim ds ns tag <$> fromPTermElim ds ns tag
<*> fromPTermTScope ds ns [< r] ret <*> fromPTermTScope ds ns [< r] ret
<*> assert_total fromPTermEnumArms ds ns arms <*> assert_total fromPTermEnumArms ds ns arms
<*> pure loc
Nat _ => pure Nat Nat loc => pure $ Nat loc
Zero _ => pure Zero Zero loc => pure $ Zero loc
Succ n _ => [|Succ $ fromPTermWith ds ns n|] Succ n loc => [|Succ (fromPTermWith ds ns n) (pure loc)|]
Case pi nat (r, ret) (CaseNat zer (s, pi', ih, suc) _) _ => Case pi nat (r, ret) (CaseNat zer (s, pi', ih, suc) _) loc =>
map E $ CaseNat (fromPQty pi) (fromPQty pi') map E $ CaseNat (fromPQty pi) (fromPQty pi')
<$> fromPTermElim ds ns nat <$> fromPTermElim ds ns nat
<*> fromPTermTScope ds ns [< r] ret <*> fromPTermTScope ds ns [< r] ret
<*> fromPTermWith ds ns zer <*> fromPTermWith ds ns zer
<*> fromPTermTScope ds ns [< s, ih] suc <*> fromPTermTScope ds ns [< s, ih] suc
<*> pure loc
Enum strs _ => Enum strs loc =>
let set = SortedSet.fromList strs in let set = SortedSet.fromList strs in
if length strs == length (SortedSet.toList set) then if length strs == length (SortedSet.toList set) then
pure $ Enum set pure $ Enum set loc
else else
throw $ DuplicatesInEnum strs throw $ DuplicatesInEnum loc strs
Tag str _ => pure $ Tag str Tag str loc => pure $ Tag str loc
Eq (i, ty) s t _ => Eq (i, ty) s t loc =>
Eq <$> fromPTermDScope ds ns [< i] ty Eq <$> fromPTermDScope ds ns [< i] ty
<*> fromPTermWith ds ns s <*> fromPTermWith ds ns s
<*> fromPTermWith ds ns t <*> fromPTermWith ds ns t
<*> pure loc
DLam i s _ => DLam i s loc =>
DLam <$> fromPTermDScope ds ns [< i] s DLam <$> fromPTermDScope ds ns [< i] s <*> pure loc
DApp s p _ => DApp s p loc =>
map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p map E $ DApp
<$> fromPTermElim ds ns s
<*> fromPDimWith ds p
<*> pure loc
BOX q ty _ => BOX (fromPQty q) <$> fromPTermWith ds ns ty BOX q ty loc => BOX (fromPQty q) <$> fromPTermWith ds ns ty <*> pure loc
Box val _ => Box <$> fromPTermWith ds ns val Box val loc => Box <$> fromPTermWith ds ns val <*> pure loc
Case pi box (r, ret) (CaseBox b body _) _ => Case pi box (r, ret) (CaseBox b body _) loc =>
map E $ CaseBox (fromPQty pi) map E $ CaseBox (fromPQty pi)
<$> fromPTermElim ds ns box <$> fromPTermElim ds ns box
<*> fromPTermTScope ds ns [< r] ret <*> fromPTermTScope ds ns [< r] ret
<*> fromPTermTScope ds ns [< b] body <*> fromPTermTScope ds ns [< b] body
<*> pure loc
V x _ => V x loc =>
fromName (pure . E . B) (resolveName !(getAt NS) <=< avoidDim ds) ns x fromName (\i => pure $ E $ B i loc)
(resolveName !(getAt NS) loc <=< avoidDim ds loc) ns x
Ann s a _ => Ann s a loc =>
map E $ (:#) <$> fromPTermWith ds ns s <*> fromPTermWith ds ns a map E $ Ann
<$> fromPTermWith ds ns s
<*> fromPTermWith ds ns a
<*> pure loc
Coe (i, ty) p q val _ => Coe (i, ty) p q val loc =>
map E $ Coe map E $ Coe
<$> fromPTermDScope ds ns [< i] ty <$> fromPTermDScope ds ns [< i] ty
<*> fromPDimWith ds p <*> fromPDimWith ds p
<*> fromPDimWith ds q <*> fromPDimWith ds q
<*> fromPTermWith ds ns val <*> fromPTermWith ds ns val
<*> pure loc
Comp (i, ty) p q val r (j0, val0) (j1, val1) _ => Comp (i, ty) p q val r (j0, val0) (j1, val1) loc =>
map E $ CompH' map E $ CompH'
<$> fromPTermDScope ds ns [< i] ty <$> fromPTermDScope ds ns [< i] ty
<*> fromPDimWith ds p <*> fromPDimWith ds p
@ -205,6 +225,7 @@ mutual
<*> fromPDimWith ds r <*> fromPDimWith ds r
<*> fromPTermDScope ds ns [< j0] val0 <*> fromPTermDScope ds ns [< j0] val0
<*> fromPTermDScope ds ns [< j1] val1 <*> fromPTermDScope ds ns [< j1] val1
<*> pure loc
private private
fromPTermEnumArms : Context' PatVar d -> Context' PatVar n -> fromPTermEnumArms : Context' PatVar d -> Context' PatVar n ->
@ -221,7 +242,7 @@ mutual
case !(fromPTermWith ds ns e) of case !(fromPTermWith ds ns e) of
E e => pure e E e => pure e
t => let ctx = MkNameContexts (map fromPatVar ds) (map fromPatVar ns) in t => let ctx = MkNameContexts (map fromPatVar ds) (map fromPatVar ns) in
throw $ AnnotationNeeded ctx t throw $ AnnotationNeeded t.loc ctx t
-- [todo] use SN if the var is named but still unused -- [todo] use SN if the var is named but still unused
private private
@ -251,10 +272,10 @@ fromPTerm = fromPTermWith [<] [<]
export export
globalPQty : (q : Qty) -> Eff [Except Error] (So $ isGlobal q) globalPQty : Loc -> (q : Qty) -> Eff [Except Error] (So $ isGlobal q)
globalPQty pi = case choose $ isGlobal pi of globalPQty loc pi = case choose $ isGlobal pi of
Left y => pure y Left y => pure y
Right _ => throw $ QtyNotGlobal pi Right _ => throw $ QtyNotGlobal loc pi
export export
@ -262,30 +283,31 @@ fromPBaseNameNS : PBaseName -> Eff [StateL NS Mods] Name
fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name
private private
injTC : TC a -> Eff FromParserPure a liftTC : TC a -> Eff FromParserPure a
injTC act = rethrow $ mapFst WrapTypeError $ runTC !(getAt DEFS) act liftTC act = do
res <- lift $ runExcept $ runReaderAt DEFS !(getAt DEFS) act
rethrow $ mapFst WrapTypeError res
export covering export covering
fromPDef : PDefinition -> Eff FromParserPure NDefinition fromPDef : PDefinition -> Eff FromParserPure NDefinition
fromPDef (MkPDef qty pname ptype pterm _) = do fromPDef (MkPDef qty pname ptype pterm defLoc) = do
name <- lift $ fromPBaseNameNS pname name <- lift $ fromPBaseNameNS pname
let qty = fromPQty qty qtyGlobal <- lift $ globalPQty qty.loc qty.val
qtyGlobal <- lift $ globalPQty qty let gqty = Element qty.val qtyGlobal
let gqty = Element qty qtyGlobal sqty = globalToSubj gqty
let sqty = globalToSubj gqty
type <- lift $ traverse fromPTerm ptype type <- lift $ traverse fromPTerm ptype
term <- lift $ fromPTerm pterm term <- lift $ fromPTerm pterm
case type of case type of
Just type => do Just type => do
injTC $ checkTypeC empty type Nothing liftTC $ checkTypeC empty type Nothing
injTC $ ignore $ checkC empty sqty term type liftTC $ ignore $ checkC empty sqty term type
let def = mkDef gqty type term let def = mkDef gqty type term defLoc
modifyAt DEFS $ insert name def modifyAt DEFS $ insert name def
pure (name, def) pure (name, def)
Nothing => do Nothing => do
let E elim = term | t => throw $ AnnotationNeeded empty t let E elim = term | _ => throw $ AnnotationNeeded term.loc empty term
res <- injTC $ inferC empty sqty elim res <- liftTC $ inferC empty sqty elim
let def = mkDef gqty res.type term let def = mkDef gqty res.type term defLoc
modifyAt DEFS $ insert name def modifyAt DEFS $ insert name def
pure (name, def) pure (name, def)
@ -296,27 +318,23 @@ fromPDecl (PNs ns) =
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
public export
LoadFile : List (Type -> Type)
LoadFile = [IO, StateL SEEN SeenFiles, Reader IncludePath, Except Error]
export covering export covering
loadFile : String -> Eff LoadFile (Maybe String) loadFile : Loc -> String -> Eff LoadFile (Maybe String)
loadFile file = loadFile loc file =
if contains file !(getAt SEEN) then if contains file !(getAt SEEN) then
pure Nothing pure Nothing
else do else do
Just ifile <- firstExists (map (</> file) !ask) Just ifile <- firstExists (map (</> file) !ask)
| Nothing => throw $ LoadError file FileNotFound | Nothing => throw $ LoadError loc file FileNotFound
case !(readFile ifile) of case !(readFile ifile) of
Right res => modifyAt SEEN (insert file) $> Just res Right res => modifyAt SEEN (insert file) $> Just res
Left err => throw $ LoadError ifile err Left err => throw $ LoadError loc ifile err
mutual mutual
export covering export covering
loadProcessFile : String -> FromParser (List NDefinition) loadProcessFile : Loc -> String -> Eff FromParserIO (List NDefinition)
loadProcessFile file = loadProcessFile loc file =
case !(lift $ loadFile file) of case !(lift $ loadFile loc file) of
Just inp => do Just inp => do
tl <- either (throw . WrapParseError file) pure $ lexParseInput file inp tl <- either (throw . WrapParseError file) pure $ lexParseInput file inp
concat <$> traverse fromPTopLevel tl concat <$> traverse fromPTopLevel tl
@ -324,26 +342,29 @@ mutual
||| populates the `defs` field of the state ||| populates the `defs` field of the state
export covering export covering
fromPTopLevel : PTopLevel -> FromParser (List NDefinition) fromPTopLevel : PTopLevel -> Eff FromParserIO (List NDefinition)
fromPTopLevel (PD decl) = lift $ fromPDecl decl fromPTopLevel (PD decl) = lift $ fromPDecl decl
fromPTopLevel (PLoad file _) = loadProcessFile file fromPTopLevel (PLoad file loc) = loadProcessFile loc file
export export
fromParserPure : Definitions -> fromParserPure : NameSuf -> Definitions ->
Eff FromParserPure a -> Eff FromParserPure a ->
Either Error (a, Definitions) (Either Error (a, Definitions), NameSuf)
fromParserPure defs act = fromParserPure suf defs act =
extract $ extract $
runStateAt GEN suf $
runExcept $ runExcept $
evalStateAt NS [<] $ evalStateAt NS [<] $
runStateAt DEFS defs act runStateAt DEFS defs act
export export
fromParserIO : (MonadRec io, HasIO io) => fromParserIO : (MonadRec io, HasIO io) =>
IncludePath -> IORef SeenFiles -> IORef Definitions -> IncludePath ->
FromParser a -> io (Either Error a) IORef SeenFiles -> IORef NameSuf -> IORef Definitions ->
fromParserIO inc seen defs act = Eff FromParserIO a -> io (Either Error a)
fromParserIO inc seen suf defs act =
runIO $ runIO $
runStateIORefAt GEN suf $
runExcept $ runExcept $
evalStateAt NS [<] $ evalStateAt NS [<] $
runStateIORefAt SEEN seen $ runStateIORefAt SEEN seen $

View File

@ -18,33 +18,24 @@ ParseError = Parser.Error
public export public export
data Error = data Error =
AnnotationNeeded (NameContexts d n) (Term d n) AnnotationNeeded Loc (NameContexts d n) (Term d n)
| DuplicatesInEnum (List TagVal) | DuplicatesInEnum Loc (List TagVal)
| TermNotInScope Name | TermNotInScope Loc Name
| DimNotInScope PBaseName | DimNotInScope Loc PBaseName
| QtyNotGlobal Qty | QtyNotGlobal Loc Qty
| DimNameInTerm PBaseName | DimNameInTerm Loc PBaseName
| WrapTypeError TypeError | WrapTypeError TypeError
| LoadError String FileError | LoadError Loc String FileError
| WrapParseError String ParseError | WrapParseError String ParseError
parameters (unicode, showContext : Bool) parameters (unicode, showContext : Bool)
export
prettyBounds : String -> Bounds -> Doc HL
prettyBounds file (MkBounds l1 c1 l2 c2) =
hcat [hl Free $ pretty file, hl Delim ":",
hl TVar $ pretty l1, hl Delim ":",
hl DVar $ pretty c1, hl Delim "-",
hl TVar $ pretty l2, hl Delim ":",
hl DVar $ pretty c2]
export export
prettyParseError1 : String -> ParsingError _ -> Doc HL prettyParseError1 : String -> ParsingError _ -> Doc HL
prettyParseError1 file (Error msg Nothing) = prettyParseError1 file (Error msg Nothing) =
pretty msg pretty msg
prettyParseError1 file (Error msg (Just bounds)) = prettyParseError1 file (Error msg (Just bounds)) =
asep [prettyBounds file bounds <+> hl Delim ":", pretty msg] hsep [prettyLoc $ makeLoc file bounds, pretty msg]
export export
prettyParseError : String -> ParseError -> Doc HL prettyParseError : String -> ParseError -> Doc HL
@ -56,33 +47,38 @@ parameters (unicode, showContext : Bool)
export export
prettyError : Error -> Doc HL prettyError : Error -> Doc HL
prettyError (AnnotationNeeded ctx tm) = prettyError (AnnotationNeeded loc ctx tm) =
sep ["the term", prettyTerm unicode ctx.dnames ctx.tnames tm, sep [prettyLoc loc <++> "the term",
prettyTerm unicode ctx.dnames ctx.tnames tm,
"needs a type annotation"] "needs a type annotation"]
-- [todo] print the original PTerm instead -- [todo] print the original PTerm instead
prettyError (DuplicatesInEnum tags) = prettyError (DuplicatesInEnum loc tags) =
sep ["duplicate tags in enum type", braces $ fillSep $ map pretty tags] sep [prettyLoc loc <++> "duplicate tags in enum type",
braces $ fillSep $ map pretty tags]
prettyError (DimNotInScope i) = prettyError (DimNotInScope loc i) =
sep ["dimension", pretty0 unicode $ DV $ fromString i, "not in scope"] sep [prettyLoc loc <++> "dimension",
pretty0 unicode $ DV $ fromString i, "not in scope"]
prettyError (TermNotInScope x) = prettyError (TermNotInScope loc x) =
sep ["term variable", pretty0 unicode $ F x {d = 0, n = 0}, "not in scope"] sep [prettyLoc loc <++> "term variable",
hl Free $ pretty0 unicode x, "not in scope"]
prettyError (QtyNotGlobal pi) = prettyError (QtyNotGlobal loc pi) =
sep ["quantity", pretty0 unicode pi, sep [prettyLoc loc <++> "quantity", pretty0 unicode pi,
"can't be used on a top level declaration"] "can't be used on a top level declaration"]
prettyError (DimNameInTerm i) = prettyError (DimNameInTerm loc i) =
sep ["dimension variable", pretty0 unicode $ DV $ fromString i, sep [prettyLoc loc <++> "dimension variable",
"used in a term context"] pretty0 unicode $ DV $ fromString i, "used in a term context"]
prettyError (WrapTypeError err) = prettyError (WrapTypeError err) =
Typing.prettyError unicode showContext $ trimContext 2 err Typing.prettyError unicode showContext $ trimContext 2 err
prettyError (LoadError str err) = prettyError (LoadError loc str err) =
vsep [hsep ["couldn't load file", pretty str], fromString $ show err] vsep [hsep [prettyLoc loc, "couldn't load file", pretty str],
fromString $ show err]
prettyError (WrapParseError file err) = prettyError (WrapParseError file err) =
prettyParseError file err prettyParseError file err

View File

@ -36,7 +36,7 @@ lexParseWith grm input = do
export export
withLoc : {c : Bool} -> FileName -> (Grammar c (Loc -> a)) -> Grammar c a withLoc : {c : Bool} -> FileName -> (Grammar c (Loc -> a)) -> Grammar c a
withLoc fname act = bounds act <&> \res => withLoc fname act = bounds act <&> \res =>
if res.isIrrelevant then res.val Nothing if res.isIrrelevant then res.val noLoc
else res.val $ makeLoc fname res.bounds else res.val $ makeLoc fname res.bounds
export export
@ -241,40 +241,40 @@ casePat fname = withLoc fname $
<|> delim "[" "]" [|PBox (patVar fname)|] <|> delim "[" "]" [|PBox (patVar fname)|]
<|> fatalError "invalid pattern" <|> fatalError "invalid pattern"
export covering export
term : FileName -> Grammar True PTerm term : FileName -> Grammar True PTerm
-- defined after all the subterm parsers -- defined after all the subterm parsers
export covering export
typeLine : FileName -> Grammar True (PatVar, PTerm) typeLine : FileName -> Grammar True (PatVar, PTerm)
typeLine fname = do typeLine fname = do
resC "[" resC "["
mustWork $ do mustWork $ do
i <- patVar fname <* resC "" <|> unused fname i <- patVar fname <* resC "" <|> unused fname
t <- term fname <* needRes "]" t <- assert_total term fname <* needRes "]"
pure (i, t) pure (i, t)
||| box term `[t]` or type `[π.A]` ||| box term `[t]` or type `[π.A]`
export covering export
boxTerm : FileName -> Grammar True PTerm boxTerm : FileName -> Grammar True PTerm
boxTerm fname = withLoc fname $ do boxTerm fname = withLoc fname $ do
res "["; commit res "["; commit
q <- optional $ qty fname <* res "." q <- optional $ qty fname <* res "."
t <- mustWork $ term fname <* needRes "]" t <- mustWork $ assert_total term fname <* needRes "]"
pure $ maybe (Box t) (\q => BOX q t) q pure $ maybe (Box t) (\q => BOX q t) q
||| tuple term like `(a, b)`, or parenthesised single term. ||| tuple term like `(a, b)`, or parenthesised single term.
||| allows terminating comma. more than two elements are nested on the right: ||| allows terminating comma. more than two elements are nested on the right:
||| `(a, b, c, d) = (a, (b, (c, d)))`. ||| `(a, b, c, d) = (a, (b, (c, d)))`.
export covering export
tupleTerm : FileName -> Grammar True PTerm tupleTerm : FileName -> Grammar True PTerm
tupleTerm fname = withLoc fname $ do tupleTerm fname = withLoc fname $ do
terms <- delimSep1 "(" ")" "," $ term fname terms <- delimSep1 "(" ")" "," $ assert_total term fname
pure $ \loc => foldr1 (\s, t => Pair s t loc) terms pure $ \loc => foldr1 (\s, t => Pair s t loc) terms
||| argument/atomic term: single-token terms, or those with delimiters e.g. ||| argument/atomic term: single-token terms, or those with delimiters e.g.
||| `[t]` ||| `[t]`
export covering export
termArg : FileName -> Grammar True PTerm termArg : FileName -> Grammar True PTerm
termArg fname = withLoc fname $ termArg fname = withLoc fname $
[|TYPE universe1|] [|TYPE universe1|]
@ -287,7 +287,7 @@ termArg fname = withLoc fname $
<|> [|V qname|] <|> [|V qname|]
<|> const <$> tupleTerm fname <|> const <$> tupleTerm fname
export covering export
coeTerm : FileName -> Grammar True PTerm coeTerm : FileName -> Grammar True PTerm
coeTerm fname = withLoc fname $ do coeTerm fname = withLoc fname $ do
resC "coe" resC "coe"
@ -298,9 +298,11 @@ public export
CompBranch : Type CompBranch : Type
CompBranch = (DimConst, PatVar, PTerm) CompBranch = (DimConst, PatVar, PTerm)
export covering export
compBranch : FileName -> Grammar True CompBranch compBranch : FileName -> Grammar True CompBranch
compBranch fname = [|(,,) dimConst (patVar fname) (needRes "" *> term fname)|] compBranch fname =
[|(,,) dimConst (patVar fname)
(needRes "" *> assert_total term fname)|]
private private
checkCompTermBody : (PatVar, PTerm) -> PDim -> PDim -> PTerm -> PDim -> checkCompTermBody : (PatVar, PTerm) -> PDim -> PDim -> PTerm -> PDim ->
@ -313,7 +315,7 @@ checkCompTermBody a p q s r (e0, s0) (e1, s1) bounds =
(_, _) => (_, _) =>
fatalLoc bounds "body of 'comp' needs one 0 case and one 1 case" fatalLoc bounds "body of 'comp' needs one 0 case and one 1 case"
export covering export
compTerm : FileName -> Grammar True PTerm compTerm : FileName -> Grammar True PTerm
compTerm fname = withLoc fname $ do compTerm fname = withLoc fname $ do
resC "comp" resC "comp"
@ -328,27 +330,27 @@ compTerm fname = withLoc fname $ do
let body = bounds $ mergeBounds bodyStart bodyEnd let body = bounds $ mergeBounds bodyStart bodyEnd
checkCompTermBody a p q s r s0 s1 body checkCompTermBody a p q s r s0 s1 body
export covering export
splitUniverseTerm : FileName -> Grammar True PTerm splitUniverseTerm : FileName -> Grammar True PTerm
splitUniverseTerm fname = withLoc fname $ resC "" *> mustWork [|TYPE nat|] splitUniverseTerm fname = withLoc fname $ resC "" *> mustWork [|TYPE nat|]
export covering export
eqTerm : FileName -> Grammar True PTerm eqTerm : FileName -> Grammar True PTerm
eqTerm fname = withLoc fname $ eqTerm fname = withLoc fname $
resC "Eq" *> mustWork [|Eq (typeLine fname) (termArg fname) (termArg fname)|] resC "Eq" *> mustWork [|Eq (typeLine fname) (termArg fname) (termArg fname)|]
export covering export
succTerm : FileName -> Grammar True PTerm succTerm : FileName -> Grammar True PTerm
succTerm fname = withLoc fname $ succTerm fname = withLoc fname $
resC "succ" *> mustWork [|Succ (termArg fname)|] resC "succ" *> mustWork [|Succ (termArg fname)|]
||| a dimension argument with an `@` prefix, or ||| a dimension argument with an `@` prefix, or
||| a term argument with no prefix ||| a term argument with no prefix
export covering export
anyArg : FileName -> Grammar True (Either PDim PTerm) anyArg : FileName -> Grammar True (Either PDim PTerm)
anyArg fname = dimArg fname <||> termArg fname anyArg fname = dimArg fname <||> termArg fname
export covering export
normalAppTerm : FileName -> Grammar True PTerm normalAppTerm : FileName -> Grammar True PTerm
normalAppTerm fname = withLoc fname $ do normalAppTerm fname = withLoc fname $ do
head <- termArg fname head <- termArg fname
@ -360,7 +362,7 @@ where ap : Loc -> PTerm -> Either PDim PTerm -> PTerm
||| application term `f x @y z`, or other terms that look like application ||| application term `f x @y z`, or other terms that look like application
||| like `succ` or `coe`. ||| like `succ` or `coe`.
export covering export
appTerm : FileName -> Grammar True PTerm appTerm : FileName -> Grammar True PTerm
appTerm fname = appTerm fname =
coeTerm fname coeTerm fname
@ -370,53 +372,55 @@ appTerm fname =
<|> succTerm fname <|> succTerm fname
<|> normalAppTerm fname <|> normalAppTerm fname
export covering export
infixEqTerm : FileName -> Grammar True PTerm infixEqTerm : FileName -> Grammar True PTerm
infixEqTerm fname = withLoc fname $ do infixEqTerm fname = withLoc fname $ do
l <- appTerm fname; commit l <- appTerm fname; commit
rest <- optional $ rest <- optional $ res "" *>
res "" *> [|(,) (term fname) (needRes ":" *> appTerm fname)|] [|(,) (assert_total term fname) (needRes ":" *> appTerm fname)|]
let u = Unused $ onlyStart l.loc let u = Unused $ onlyStart l.loc
pure $ \loc => maybe l (\rest => Eq (u, snd rest) l (fst rest) loc) rest pure $ \loc => maybe l (\rest => Eq (u, snd rest) l (fst rest) loc) rest
export covering export
annTerm : FileName -> Grammar True PTerm annTerm : FileName -> Grammar True PTerm
annTerm fname = withLoc fname $ do annTerm fname = withLoc fname $ do
tm <- infixEqTerm fname; commit tm <- infixEqTerm fname; commit
ty <- optional $ res "" *> term fname ty <- optional $ res "" *> assert_total term fname
pure $ \loc => maybe tm (\ty => Ann tm ty loc) ty pure $ \loc => maybe tm (\ty => Ann tm ty loc) ty
export covering export
lamTerm : FileName -> Grammar True PTerm lamTerm : FileName -> Grammar True PTerm
lamTerm fname = withLoc fname $ do lamTerm fname = withLoc fname $ do
k <- DLam <$ res "δ" <|> Lam <$ res "λ" k <- DLam <$ res "δ" <|> Lam <$ res "λ"
mustWork $ do mustWork $ do
xs <- some $ patVar fname; needRes "" xs <- some $ patVar fname; needRes ""
body <- term fname; commit body <- assert_total term fname; commit
pure $ \loc => foldr (\x, s => k x s loc) body xs pure $ \loc => foldr (\x, s => k x s loc) body xs
-- [todo] fix the backtracking in e.g. (F x y z × B) -- [todo] fix the backtracking in e.g. (F x y z × B)
export covering export
properBinders : FileName -> Grammar True (List1 PatVar, PTerm) properBinders : FileName -> Grammar True (List1 PatVar, PTerm)
properBinders fname = do properBinders fname = assert_total $ do
-- putting assert_total directly on `term`, in this one function,
-- doesn't work. i cannot tell why
res "(" res "("
xs <- some $ patVar fname; resC ":" xs <- some $ patVar fname; resC ":"
t <- term fname; needRes ")" t <- term fname; needRes ")"
pure (xs, t) pure (xs, t)
export covering export
piTerm : FileName -> Grammar True PTerm piTerm : FileName -> Grammar True PTerm
piTerm fname = withLoc fname $ do piTerm fname = withLoc fname $ do
q <- qty fname; resC "." q <- qty fname; resC "."
dom <- piBinder; needRes "" dom <- piBinder; needRes ""
cod <- term fname; commit cod <- assert_total term fname; commit
pure $ \loc => foldr (\x, t => Pi q x (snd dom) t loc) cod (fst dom) pure $ \loc => foldr (\x, t => Pi q x (snd dom) t loc) cod (fst dom)
where where
piBinder : Grammar True (List1 PatVar, PTerm) piBinder : Grammar True (List1 PatVar, PTerm)
piBinder = properBinders fname piBinder = properBinders fname
<|> [|(,) [|singleton $ unused fname|] (termArg fname)|] <|> [|(,) [|singleton $ unused fname|] (termArg fname)|]
export covering export
sigmaTerm : FileName -> Grammar True PTerm sigmaTerm : FileName -> Grammar True PTerm
sigmaTerm fname = sigmaTerm fname =
(properBinders fname >>= continueDep) (properBinders fname >>= continueDep)
@ -440,9 +444,10 @@ public export
PCaseArm : Type PCaseArm : Type
PCaseArm = (PCasePat, PTerm) PCaseArm = (PCasePat, PTerm)
export covering export
caseArm : FileName -> Grammar True PCaseArm caseArm : FileName -> Grammar True PCaseArm
caseArm fname = [|(,) (casePat fname) (needRes "" *> term fname)|] caseArm fname =
[|(,) (casePat fname) (needRes "" *> assert_total term fname)|]
export export
checkCaseArms : Loc -> List PCaseArm -> Grammar False PCaseBody checkCaseArms : Loc -> List PCaseArm -> Grammar False PCaseBody
@ -468,30 +473,30 @@ checkCaseArms loc ((PBox x _, rhs) :: rest) =
if null rest then pure $ CaseBox x rhs loc if null rest then pure $ CaseBox x rhs loc
else fatalError "unexpected pattern after box" else fatalError "unexpected pattern after box"
export covering export
caseBody : FileName -> Grammar True PCaseBody caseBody : FileName -> Grammar True PCaseBody
caseBody fname = do caseBody fname = do
body <- bounds $ delimSep "{" "}" ";" $ caseArm fname body <- bounds $ delimSep "{" "}" ";" $ caseArm fname
let loc = makeLoc fname body.bounds let loc = makeLoc fname body.bounds
checkCaseArms loc body.val checkCaseArms loc body.val
export covering export
caseReturn : FileName -> Grammar True (PatVar, PTerm) caseReturn : FileName -> Grammar True (PatVar, PTerm)
caseReturn fname = do caseReturn fname = do
x <- patVar fname <* resC "" <|> unused fname x <- patVar fname <* resC "" <|> unused fname
ret <- term fname ret <- assert_total term fname
pure (x, ret) pure (x, ret)
export covering export
caseTerm : FileName -> Grammar True PTerm caseTerm : FileName -> Grammar True PTerm
caseTerm fname = withLoc fname $ do caseTerm fname = withLoc fname $ do
qty <- caseIntro fname; commit qty <- caseIntro fname; commit
head <- mustWork $ term fname; needRes "return" head <- mustWork $ assert_total term fname; needRes "return"
ret <- mustWork $ caseReturn fname; needRes "of" ret <- mustWork $ caseReturn fname; needRes "of"
body <- mustWork $ caseBody fname body <- mustWork $ caseBody fname
pure $ Case qty head ret body pure $ Case qty head ret body
-- export covering -- export
-- term : FileName -> Grammar True PTerm -- term : FileName -> Grammar True PTerm
term fname = lamTerm fname term fname = lamTerm fname
<|> caseTerm fname <|> caseTerm fname
@ -499,7 +504,7 @@ term fname = lamTerm fname
<|> sigmaTerm fname <|> sigmaTerm fname
export covering export
decl : FileName -> Grammar True PDecl decl : FileName -> Grammar True PDecl
||| `def` alone means `defω` ||| `def` alone means `defω`
@ -512,7 +517,7 @@ defIntro fname =
let any = PQ Any $ makeLoc fname pos.bounds let any = PQ Any $ makeLoc fname pos.bounds
option any $ qty fname <* needRes "." option any $ qty fname <* needRes "."
export covering export
definition : FileName -> Grammar True PDefinition definition : FileName -> Grammar True PDefinition
definition fname = withLoc fname $ do definition fname = withLoc fname $ do
qty <- defIntro fname qty <- defIntro fname
@ -522,7 +527,7 @@ definition fname = withLoc fname $ do
optRes ";" optRes ";"
pure $ MkPDef qty name type term pure $ MkPDef qty name type term
export covering export
namespace_ : FileName -> Grammar True PNamespace namespace_ : FileName -> Grammar True PNamespace
namespace_ fname = withLoc fname $ do namespace_ fname = withLoc fname $ do
ns <- resC "namespace" *> qname; needRes "{" ns <- resC "namespace" *> qname; needRes "{"
@ -531,28 +536,28 @@ namespace_ fname = withLoc fname $ do
where where
nsInner : Grammar True (List PDecl) nsInner : Grammar True (List PDecl)
nsInner = [] <$ resC "}" nsInner = [] <$ resC "}"
<|> [|(decl fname <* commit) :: nsInner|] <|> [|(assert_total decl fname <* commit) :: assert_total nsInner|]
decl fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|] decl fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|]
export covering export
load : FileName -> Grammar True PTopLevel load : FileName -> Grammar True PTopLevel
load fname = withLoc fname $ load fname = withLoc fname $
resC "load" *> mustWork [|PLoad strLit|] <* optRes ";" resC "load" *> mustWork [|PLoad strLit|] <* optRes ";"
export covering export
topLevel : FileName -> Grammar True PTopLevel topLevel : FileName -> Grammar True PTopLevel
topLevel fname = load fname <|> [|PD $ decl fname|] topLevel fname = load fname <|> [|PD $ decl fname|]
export covering export
input : FileName -> Grammar False (List PTopLevel) input : FileName -> Grammar False (List PTopLevel)
input fname = [] <$ eof input fname = [] <$ eof
<|> [|(topLevel fname <* commit) :: input fname|] <|> [|(topLevel fname <* commit) :: assert_total input fname|]
export covering export
lexParseTerm : FileName -> String -> Either Error PTerm lexParseTerm : FileName -> String -> Either Error PTerm
lexParseTerm = lexParseWith . term lexParseTerm = lexParseWith . term
export covering export
lexParseInput : FileName -> String -> Either Error (List PTopLevel) lexParseInput : FileName -> String -> Either Error (List PTopLevel)
lexParseInput = lexParseWith . input lexParseInput = lexParseWith . input

View File

@ -33,11 +33,14 @@ isUnused _ = False
public export public export
data PQty = PQ Qty Loc record PQty where
constructor PQ
val : Qty
loc_ : Loc
%name PQty qty %name PQty qty
%runElab derive "PQty" [Eq, Ord, Show] %runElab derive "PQty" [Eq, Ord, Show]
export Located PQty where (PQ _ loc).loc = loc export Located PQty where q.loc = q.loc_
namespace PDim namespace PDim
public export public export

View File

@ -1,5 +1,6 @@
module Quox.Pretty module Quox.Pretty
import Quox.Loc
import Quox.Name import Quox.Name
import public Text.PrettyPrint.Prettyprinter.Doc import public Text.PrettyPrint.Prettyprinter.Doc
@ -200,6 +201,7 @@ pretty0 unicode = pretty0With unicode [<] [<]
export PrettyHL BaseName where prettyM = pure . pretty . baseStr export PrettyHL BaseName where prettyM = pure . pretty . baseStr
export PrettyHL Name where prettyM = pure . pretty . toDots export PrettyHL Name where prettyM = pure . pretty . toDots
export PrettyHL BindName where prettyM = prettyM . name
export export
@ -278,3 +280,14 @@ epretty @{p} x = Evidence a (p, x)
public export data Lit = L (Doc HL) public export data Lit = L (Doc HL)
export PrettyHL Lit where prettyM (L doc) = pure doc export PrettyHL Lit where prettyM (L doc) = pure doc
export
prettyLoc : Loc -> Doc HL
prettyLoc (L NoLoc) = hl TVarErr "no location" <+> hl Delim ":"
prettyLoc (L (YesLoc file (MkBounds l1 c1 l2 c2))) =
hcat [hl Free $ pretty file, hl Delim ":",
hl TVar $ pretty l1, hl Delim ":",
hl DVar $ pretty c1, hl Delim "-",
hl TVar $ pretty l2, hl Delim ":",
hl DVar $ pretty c2, hl Delim ":"]

File diff suppressed because it is too large Load Diff

View File

@ -1,5 +1,6 @@
module Quox.Syntax.Dim module Quox.Syntax.Dim
import Quox.Loc
import Quox.Name import Quox.Name
import Quox.Syntax.Var import Quox.Syntax.Var
import Quox.Syntax.Subst import Quox.Syntax.Subst
@ -17,7 +18,6 @@ import Derive.Prelude
public export public export
data DimConst = Zero | One data DimConst = Zero | One
%name DimConst e %name DimConst e
%runElab derive "DimConst" [Eq, Ord, Show] %runElab derive "DimConst" [Eq, Ord, Show]
||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`. ||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`.
@ -26,23 +26,42 @@ ends : Lazy a -> Lazy a -> DimConst -> a
ends l r Zero = l ends l r Zero = l
ends l r One = r ends l r One = r
export Uninhabited (Zero = One) where uninhabited _ impossible
export Uninhabited (One = Zero) where uninhabited _ impossible
public export
DecEq DimConst where
decEq Zero Zero = Yes Refl
decEq Zero One = No absurd
decEq One Zero = No absurd
decEq One One = Yes Refl
public export public export
data Dim : Nat -> Type where data Dim : Nat -> Type where
K : DimConst -> Dim d K : DimConst -> Loc -> Dim d
B : Var d -> Dim d B : Var d -> Loc -> Dim d
%name Dim.Dim p, q %name Dim.Dim p, q
%runElab deriveIndexed "Dim" [Eq, Ord, Show] %runElab deriveIndexed "Dim" [Eq, Ord, Show]
export
Located (Dim d) where
(K _ loc).loc = loc
(B _ loc).loc = loc
export
Relocatable (Dim d) where
setLoc loc (K e _) = K e loc
setLoc loc (B i _) = B i loc
export export
PrettyHL DimConst where PrettyHL DimConst where
prettyM = pure . hl Dim . ends "0" "1" prettyM = pure . hl Dim . ends "0" "1"
export export
PrettyHL (Dim n) where PrettyHL (Dim n) where
prettyM (K e) = prettyM e prettyM (K e _) = prettyM e
prettyM (B i) = prettyVar DVar DVarErr (!ask).dnames i prettyM (B i _) = prettyVar DVar DVarErr (!ask).dnames i
export export
prettyDim : (dnames : NContext d) -> Dim d -> Doc HL prettyDim : (dnames : NContext d) -> Dim d -> Doc HL
@ -60,13 +79,13 @@ prettyDim dnames p =
||| - `x` otherwise. ||| - `x` otherwise.
public export public export
endsOr : Lazy a -> Lazy a -> Lazy a -> Dim n -> a endsOr : Lazy a -> Lazy a -> Lazy a -> Dim n -> a
endsOr l r x (K e) = ends l r e endsOr l r x (K e _) = ends l r e
endsOr l r x (B _) = x endsOr l r x (B _ _) = x
public export %inline public export %inline
toConst : Dim 0 -> DimConst toConst : Dim 0 -> DimConst
toConst (K e) = e toConst (K e _) = e
public export public export
@ -81,52 +100,55 @@ prettyDSubst th =
!(ifUnicode "" "<") !(ifUnicode "" ">") th !(ifUnicode "" "<") !(ifUnicode "" ">") th
public export FromVar Dim where fromVar = B public export FromVar Dim where fromVarLoc = B
export export
CanShift Dim where CanShift Dim where
K e // _ = K e K e loc // _ = K e loc
B i // by = B (i // by) B i loc // by = B (i // by) loc
export export
CanSubstSelf Dim where CanSubstSelf Dim where
K e // _ = K e K e loc // _ = K e loc
B i // th = th !! i B i loc // th = getLoc th i loc
export Uninhabited (Zero = One) where uninhabited _ impossible export Uninhabited (B i loc1 = K e loc2) where uninhabited _ impossible
export Uninhabited (One = Zero) where uninhabited _ impossible export Uninhabited (K e loc1 = B i loc2) where uninhabited _ impossible
export Uninhabited (B i = K e) where uninhabited _ impossible
export Uninhabited (K e = B i) where uninhabited _ impossible
public export %inline Injective Dim.B where injective Refl = Refl
public export %inline Injective Dim.K where injective Refl = Refl
public export public export
DecEq DimConst where data Eqv : Dim d1 -> Dim d2 -> Type where
decEq Zero Zero = Yes Refl EK : K e _ `Eqv` K e _
decEq Zero One = No absurd EB : i `Eqv` j -> B i _ `Eqv` B j _
decEq One Zero = No absurd
decEq One One = Yes Refl export Uninhabited (K e l1 `Eqv` B i l2) where uninhabited _ impossible
export Uninhabited (B i l1 `Eqv` K e l2) where uninhabited _ impossible
export
injectiveB : B i loc1 `Eqv` B j loc2 -> i `Eqv` j
injectiveB (EB e) = e
export
injectiveK : K e loc1 `Eqv` K f loc2 -> e = f
injectiveK EK = Refl
public export public export
DecEq (Dim d) where decEqv : Dec2 Dim.Eqv
decEq (K e) (K f) with (decEq e f) decEqv (K e _) (K f _) = case decEq e f of
_ | Yes prf = Yes $ cong K prf Yes Refl => Yes EK
_ | No contra = No $ contra . injective No n => No $ n . injectiveK
decEq (K e) (B j) = No absurd decEqv (B i _) (B j _) = case decEqv i j of
decEq (B i) (K f) = No absurd Yes y => Yes $ EB y
decEq (B i) (B j) with (decEq i j) No n => No $ \(EB y) => n y
_ | Yes prf = Yes $ cong B prf decEqv (B _ _) (K _ _) = No absurd
_ | No contra = No $ contra . injective decEqv (K _ _) (B _ _) = No absurd
||| abbreviation for a bound variable like `BV 4` instead of ||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))` ||| `B (VS (VS (VS (VS VZ))))`
public export %inline public export %inline
BV : (i : Nat) -> (0 _ : LT i d) => Dim d BV : (i : Nat) -> (0 _ : LT i d) => (loc : Loc) -> Dim d
BV i = B $ V i BV i loc = B (V i) loc
export export

View File

@ -72,7 +72,7 @@ toMaybe (Just x) = Just x
export export
fromGround' : Context' DimConst d -> DimEq' d fromGround' : Context' DimConst d -> DimEq' d
fromGround' [<] = [<] fromGround' [<] = [<]
fromGround' (ctx :< e) = fromGround' ctx :< Just (K e) fromGround' (ctx :< e) = fromGround' ctx :< Just (K e noLoc)
export export
fromGround : Context' DimConst d -> DimEq d fromGround : Context' DimConst d -> DimEq d
@ -98,8 +98,8 @@ get' : DimEq' d -> Var d -> Maybe (Dim d)
get' = getWith $ \p, by => map (// by) p get' = getWith $ \p, by => map (// by) p
public export %inline public export %inline
getVar : DimEq' d -> Var d -> Dim d getVar : DimEq' d -> Var d -> Loc -> Dim d
getVar eqs i = fromMaybe (B i) $ get' eqs i getVar eqs i loc = fromMaybe (B i loc) $ get' eqs i
public export %inline public export %inline
getShift' : Shift len out -> DimEq' len -> Var len -> Maybe (Dim out) getShift' : Shift len out -> DimEq' len -> Var len -> Maybe (Dim out)
@ -107,8 +107,8 @@ getShift' = getShiftWith $ \p, by => map (// by) p
public export %inline public export %inline
get : DimEq' d -> Dim d -> Dim d get : DimEq' d -> Dim d -> Dim d
get _ (K e) = K e get _ (K e loc) = K e loc
get eqs (B i) = getVar eqs i get eqs (B i loc) = getVar eqs i loc
public export %inline public export %inline
@ -126,7 +126,7 @@ C eqs :<? d = C $ eqs :< map (get eqs) d
private %inline private %inline
ifVar : Var d -> Dim d -> Maybe (Dim d) -> Maybe (Dim d) ifVar : Var d -> Dim d -> Maybe (Dim d) -> Maybe (Dim d)
ifVar i p = map $ \q => if q == B i then p else q ifVar i p = map $ \q => if q == B i noLoc then p else q
-- (using decEq instead of (==) because of the proofs below) -- (using decEq instead of (==) because of the proofs below)
private %inline private %inline
@ -135,39 +135,43 @@ checkConst e f eqs = if isYes $ e `decEq` f then C eqs else ZeroIsOne
export export
setConst : Var d -> DimConst -> DimEq' d -> DimEq d setConst : Var d -> DimConst -> Loc -> DimEq' d -> DimEq d
setConst VZ e (eqs :< Nothing) = C $ eqs :< Just (K e) setConst VZ e loc (eqs :< Nothing) =
setConst VZ e (eqs :< Just (K f)) = checkConst e f $ eqs :< Just (K f) C $ eqs :< Just (K e loc)
setConst VZ e (eqs :< Just (B i)) = setConst i e eqs :<? Just (K e) setConst VZ e _ (eqs :< Just (K f loc)) =
setConst (VS i) e (eqs :< p) = setConst i e eqs :<? ifVar i (K e) p checkConst e f $ eqs :< Just (K f loc)
setConst VZ e loc (eqs :< Just (B i _)) =
setConst i e loc eqs :<? Just (K e loc)
setConst (VS i) e loc (eqs :< p) =
setConst i e loc eqs :<? ifVar i (K e loc) p
mutual mutual
private private
setVar' : (i, j : Var d) -> (0 _ : i `LT` j) -> DimEq' d -> DimEq d setVar' : (i, j : Var d) -> (0 _ : i `LT` j) -> Loc -> DimEq' d -> DimEq d
setVar' VZ (VS i) LTZ (eqs :< Nothing) = setVar' VZ (VS i) LTZ loc (eqs :< Nothing) =
C eqs :<? Just (B i) C eqs :<? Just (B i loc)
setVar' VZ (VS i) LTZ (eqs :< Just (K e)) = setVar' VZ (VS i) LTZ loc (eqs :< Just (K e eloc)) =
setConst i e eqs :<? Just (K e) setConst i e loc eqs :<? Just (K e eloc)
setVar' VZ (VS i) LTZ (eqs :< Just (B j)) = setVar' VZ (VS i) LTZ loc (eqs :< Just (B j jloc)) =
setVar i j eqs :<? Just (B (max i j)) setVar i j loc jloc eqs :<? Just (if j > i then B j jloc else B i loc)
setVar' (VS i) (VS j) (LTS lt) (eqs :< p) = setVar' (VS i) (VS j) (LTS lt) loc (eqs :< p) =
setVar' i j lt eqs :<? ifVar i (B j) p setVar' i j lt loc eqs :<? ifVar i (B j loc) p
export %inline export %inline
setVar : (i, j : Var d) -> DimEq' d -> DimEq d setVar : (i, j : Var d) -> Loc -> Loc -> DimEq' d -> DimEq d
setVar i j eqs with (compareP i j) | (compare i.nat j.nat) setVar i j li lj eqs with (compareP i j) | (compare i.nat j.nat)
setVar i j eqs | IsLT lt | LT = setVar' i j lt eqs setVar i j li lj eqs | IsLT lt | LT = setVar' i j lt lj eqs
setVar i i eqs | IsEQ | EQ = C eqs setVar i i li lj eqs | IsEQ | EQ = C eqs
setVar i j eqs | IsGT gt | GT = setVar' j i gt eqs setVar i j li lj eqs | IsGT gt | GT = setVar' j i gt li eqs
export %inline export %inline
set : (p, q : Dim d) -> DimEq d -> DimEq d set : (p, q : Dim d) -> DimEq d -> DimEq d
set _ _ ZeroIsOne = ZeroIsOne set _ _ ZeroIsOne = ZeroIsOne
set (K e) (K f) (C eqs) = checkConst e f eqs set (K e eloc) (K f floc) (C eqs) = checkConst e f eqs
set (K e) (B i) (C eqs) = setConst i e eqs set (K e eloc) (B i iloc) (C eqs) = setConst i e eloc eqs
set (B i) (K e) (C eqs) = setConst i e eqs set (B i iloc) (K e eloc) (C eqs) = setConst i e eloc eqs
set (B i) (B j) (C eqs) = setVar i j eqs set (B i iloc) (B j jloc) (C eqs) = setVar i j iloc jloc eqs
public export %inline public export %inline
@ -175,25 +179,26 @@ Split : Nat -> Type
Split d = (DimEq' d, DSubst (S d) d) Split d = (DimEq' d, DSubst (S d) d)
export %inline export %inline
split1 : DimConst -> DimEq' (S d) -> Maybe (Split d) split1 : DimConst -> Loc -> DimEq' (S d) -> Maybe (Split d)
split1 e eqs = case setConst VZ e eqs of split1 e loc eqs = case setConst VZ e loc eqs of
ZeroIsOne => Nothing ZeroIsOne => Nothing
C (eqs :< _) => Just (eqs, K e ::: id) C (eqs :< _) => Just (eqs, K e loc ::: id)
export %inline export %inline
split : DimEq' (S d) -> List (Split d) split : Loc -> DimEq' (S d) -> List (Split d)
split eqs = toList (split1 Zero eqs) <+> toList (split1 One eqs) split loc eqs = toList (split1 Zero loc eqs) <+> toList (split1 One loc eqs)
export export
splits' : DimEq' d -> List (DSubst d 0) splits' : Loc -> DimEq' d -> List (DSubst d 0)
splits' [<] = [id] splits' _ [<] = [id]
splits' eqs@(_ :< _) = [th . ph | (eqs', th) <- split eqs, ph <- splits' eqs'] splits' loc eqs@(_ :< _) =
[th . ph | (eqs', th) <- split loc eqs, ph <- splits' loc eqs']
||| the Loc is put into each of the DimConsts
export %inline export %inline
splits : DimEq d -> List (DSubst d 0) splits : Loc -> DimEq d -> List (DSubst d 0)
splits ZeroIsOne = [] splits _ ZeroIsOne = []
splits (C eqs) = splits' eqs splits loc (C eqs) = splits' loc eqs
private private
@ -208,16 +213,16 @@ newGet' d i = newGetShift d i SZ
export export
0 newGet : (d : Nat) -> (p : Dim d) -> get (new' {d}) p = p 0 newGet : (d : Nat) -> (p : Dim d) -> get (new' {d}) p = p
newGet d (K e) = Refl newGet d (K e _) = Refl
newGet d (B i) = rewrite newGet' d i in Refl newGet d (B i _) = rewrite newGet' d i in Refl
export export
0 setSelf : (p : Dim d) -> (eqs : DimEq d) -> set p p eqs = eqs 0 setSelf : (p : Dim d) -> (eqs : DimEq d) -> set p p eqs = eqs
setSelf p ZeroIsOne = Refl setSelf p ZeroIsOne = Refl
setSelf (K Zero) (C eqs) = Refl setSelf (K Zero _) (C eqs) = Refl
setSelf (K One) (C eqs) = Refl setSelf (K One _) (C eqs) = Refl
setSelf (B i) (C eqs) with (compareP i i) | (compare i.nat i.nat) setSelf (B i _) (C eqs) with (compareP i i) | (compare i.nat i.nat)
_ | IsLT lt | LT = absurd lt _ | IsLT lt | LT = absurd lt
_ | IsEQ | EQ = Refl _ | IsEQ | EQ = Refl
_ | IsGT gt | GT = absurd gt _ | IsGT gt | GT = absurd gt
@ -250,7 +255,7 @@ PrettyHL (DimEq' d) where
go [<] = pure [<] go [<] = pure [<]
go (eqs :< Nothing) = local {dnames $= tail} $ go eqs go (eqs :< Nothing) = local {dnames $= tail} $ go eqs
go (eqs :< Just p) = do go (eqs :< Just p) = do
eq <- prettyCst (BV {d = 1} 0) (weakD 1 p) eq <- prettyCst (BV {d = 1} 0 noLoc) (weakD 1 p)
eqs <- local {dnames $= tail} $ go eqs eqs <- local {dnames $= tail} $ go eqs
pure $ eqs :< eq pure $ eqs :< eq
@ -262,16 +267,16 @@ PrettyHL (DimEq d) where
prettyM (C eqs) = prettyM eqs prettyM (C eqs) = prettyM eqs
export export
prettyDimEq : NContext d -> DimEq d -> Doc HL prettyDimEq : BContext d -> DimEq d -> Doc HL
prettyDimEq ds = pretty0With False (toSnocList' ds) [<] prettyDimEq ds = pretty0With False (toNames ds) [<]
public export public export
wf' : DimEq' d -> Bool wf' : DimEq' d -> Bool
wf' [<] = True wf' [<] = True
wf' (eqs :< Nothing) = wf' eqs wf' (eqs :< Nothing) = wf' eqs
wf' (eqs :< Just (K e)) = wf' eqs wf' (eqs :< Just (K e _)) = wf' eqs
wf' (eqs :< Just (B i)) = isNothing (get' eqs i) && wf' eqs wf' (eqs :< Just (B i _)) = isNothing (get' eqs i) && wf' eqs
public export public export
wf : DimEq d -> Bool wf : DimEq d -> Bool

View File

@ -35,27 +35,28 @@ public export
data Eqv : Shift from1 to1 -> Shift from2 to2 -> Type where data Eqv : Shift from1 to1 -> Shift from2 to2 -> Type where
EqSZ : SZ `Eqv` SZ EqSZ : SZ `Eqv` SZ
EqSS : by `Eqv` bz -> SS by `Eqv` SS bz EqSS : by `Eqv` bz -> SS by `Eqv` SS bz
%name Eqv e %name Shift.Eqv e
||| two equivalent shifts are equal if they have the same indices. using (by : Shift from to, bz : Shift from to)
export ||| two equivalent shifts are equal if they have the same indices.
0 fromEqv : by `Eqv` bz -> by = bz export
fromEqv EqSZ = Refl 0 fromEqv : by `Eqv` bz -> by = bz
fromEqv (EqSS e) = cong SS $ fromEqv e fromEqv EqSZ = Refl
fromEqv (EqSS e) = cong SS $ fromEqv e
||| two equal shifts are equivalent. ||| two equal shifts are equivalent.
export export
0 toEqv : by = bz -> by `Eqv` bz 0 toEqv : by = bz -> by `Eqv` bz
toEqv Refl {by = SZ} = EqSZ toEqv Refl {by = SZ} = EqSZ
toEqv Refl {by = (SS by)} = EqSS $ toEqv Refl toEqv Refl {by = (SS by)} = EqSS $ toEqv Refl
export export
eqLen : Shift from1 to -> Shift from2 to -> Maybe (from1 = from2) cmpLen : Shift from1 to -> Shift from2 to -> Either Ordering (from1 = from2)
eqLen SZ SZ = Just Refl cmpLen SZ SZ = Right Refl
eqLen SZ (SS by) = Nothing cmpLen SZ (SS by) = Left LT
eqLen (SS by) SZ = Nothing cmpLen (SS by) SZ = Left GT
eqLen (SS by) (SS bz) = eqLen by bz cmpLen (SS by) (SS bz) = cmpLen by bz
export export
0 shiftDiff : (by : Shift from to) -> to = by.nat + from 0 shiftDiff : (by : Shift from to) -> to = by.nat + from

View File

@ -48,12 +48,16 @@ interface FromVar term => CanSubstSelf term where
(//) : term from -> Lazy (Subst term from to) -> term to (//) : term from -> Lazy (Subst term from to) -> term to
infixl 8 !!
public export public export
(!!) : FromVar term => Subst term from to -> Var from -> term to getLoc : FromVar term => Subst term from to -> Var from -> Loc -> term to
(Shift by) !! i = fromVar $ shift by i getLoc (Shift by) i loc = fromVarLoc (shift by i) loc
(t ::: th) !! VZ = t getLoc (t ::: th) VZ _ = t
(t ::: th) !! (VS i) = th !! i getLoc (t ::: th) (VS i) loc = getLoc th i loc
-- infixl 8 !!
-- public export
-- (!!) : FromVar term => Subst term from to -> Var from -> term to
-- th !! i = getLoc th i noLoc
public export public export
@ -160,12 +164,16 @@ PrettyHL (f to) => PrettyHL (Subst f from to) where
prettyM th = prettySubstM prettyM (!ask).tnames TVar "[" "]" th prettyM th = prettySubstM prettyM (!ask).tnames TVar "[" "]" th
||| whether two substitutions with the same codomain have the same shape
||| (the same number of terms and the same shift at the end). if so, they
||| also have the same domain
export export
eqShape : Subst env from1 to -> Subst env from2 to -> Maybe (from1 = from2) cmpShape : Subst env from1 to -> Subst env from2 to ->
eqShape (Shift by) (Shift bz) = eqLen by bz Either Ordering (from1 = from2)
eqShape (Shift by) (t ::: th) = Nothing cmpShape (Shift by) (Shift bz) = cmpLen by bz
eqShape (t ::: th) (Shift by) = Nothing cmpShape (Shift _) (_ ::: _) = Left LT
eqShape (t ::: th) (x ::: ph) = cong S <$> eqShape th ph cmpShape (_ ::: _) (Shift _) = Left GT
cmpShape (_ ::: th) (_ ::: ph) = cong S <$> cmpShape th ph
public export public export
@ -175,13 +183,20 @@ record WithSubst tm env n where
subst : Lazy (Subst env from n) subst : Lazy (Subst env from n)
export export
(forall n. Eq (tm n), Eq (env n)) => Eq (WithSubst tm env n) where (Eq (env n), forall n. Eq (tm n)) => Eq (WithSubst tm env n) where
Sub t1 s1 == Sub t2 s2 = Sub t1 s1 == Sub t2 s2 =
case eqShape s1 s2 of case cmpShape s1 s2 of
Just Refl => t1 == t2 && s1 == s2 Left _ => False
Nothing => False Right Refl => t1 == t2 && s1 == s2
export
(Ord (env n), forall n. Ord (tm n)) => Ord (WithSubst tm env n) where
Sub t1 s1 `compare` Sub t2 s2 =
case cmpShape s1 s2 of
Left o => o
Right Refl => compare (t1, s1) (t2, s2)
export %hint export %hint
ShowWithSubst : (forall n. Show (tm n), Show (env n)) => ShowWithSubst : (Show (env n), forall n. Show (tm n)) =>
Show (WithSubst tm env n) Show (WithSubst tm env n)
ShowWithSubst = deriveShow ShowWithSubst = deriveShow

View File

@ -7,6 +7,7 @@ import public Quox.Syntax.Qty
import public Quox.Syntax.Dim import public Quox.Syntax.Dim
import public Quox.Syntax.Term.TyConKind import public Quox.Syntax.Term.TyConKind
import public Quox.Name import public Quox.Name
import public Quox.Loc
import public Quox.Context import public Quox.Context
import Quox.Pretty import Quox.Pretty
@ -63,7 +64,7 @@ ShowScopedBody = deriveShow
public export public export
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
constructor S constructor S
names : NContext s names : BContext s
body : ScopedBody s f n body : ScopedBody s f n
%name Scoped body %name Scoped body
@ -88,38 +89,38 @@ mutual
public export public export
data Term : (d, n : Nat) -> Type where data Term : (d, n : Nat) -> Type where
||| type of types ||| type of types
TYPE : (l : Universe) -> Term d n TYPE : (l : Universe) -> (loc : Loc) -> Term d n
||| function type ||| function type
Pi : (qty : Qty) -> (arg : Term d n) -> Pi : (qty : Qty) -> (arg : Term d n) ->
(res : ScopeTerm d n) -> Term d n (res : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| function term ||| function term
Lam : (body : ScopeTerm d n) -> Term d n Lam : (body : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| pair type ||| pair type
Sig : (fst : Term d n) -> (snd : ScopeTerm d n) -> Term d n Sig : (fst : Term d n) -> (snd : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| pair value ||| pair value
Pair : (fst, snd : Term d n) -> Term d n Pair : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
||| enumeration type ||| enumeration type
Enum : (cases : SortedSet TagVal) -> Term d n Enum : (cases : SortedSet TagVal) -> (loc : Loc) -> Term d n
||| enumeration value ||| enumeration value
Tag : (tag : TagVal) -> Term d n Tag : (tag : TagVal) -> (loc : Loc) -> Term d n
||| equality type ||| equality type
Eq : (ty : DScopeTerm d n) -> (l, r : Term d n) -> Term d n Eq : (ty : DScopeTerm d n) -> (l, r : Term d n) -> (loc : Loc) -> Term d n
||| equality term ||| equality term
DLam : (body : DScopeTerm d n) -> Term d n DLam : (body : DScopeTerm d n) -> (loc : Loc) -> Term d n
||| natural numbers (temporary until 𝐖 gets added) ||| natural numbers (temporary until 𝐖 gets added)
Nat : Term d n Nat : (loc : Loc) -> Term d n
-- [todo] can these be elims? -- [todo] can these be elims?
Zero : Term d n Zero : (loc : Loc) -> Term d n
Succ : (p : Term d n) -> Term d n Succ : (p : Term d n) -> (loc : Loc) -> Term d n
||| "box" (package a value up with a certain quantity) ||| "box" (package a value up with a certain quantity)
BOX : (qty : Qty) -> (ty : Term d n) -> Term d n BOX : (qty : Qty) -> (ty : Term d n) -> (loc : Loc) -> Term d n
Box : (val : Term d n) -> Term d n Box : (val : Term d n) -> (loc : Loc) -> Term d n
||| elimination ||| elimination
E : (e : Elim d n) -> Term d n E : (e : Elim d n) -> Term d n
@ -134,12 +135,12 @@ mutual
public export public export
data Elim : (d, n : Nat) -> Type where data Elim : (d, n : Nat) -> Type where
||| free variable ||| free variable
F : (x : Name) -> Elim d n F : (x : Name) -> (loc : Loc) -> Elim d n
||| bound variable ||| bound variable
B : (i : Var n) -> Elim d n B : (i : Var n) -> (loc : Loc) -> Elim d n
||| term application ||| term application
(:@) : (fun : Elim d n) -> (arg : Term d n) -> Elim d n App : (fun : Elim d n) -> (arg : Term d n) -> (loc : Loc) -> Elim d n
||| pair destruction ||| pair destruction
||| |||
@ -148,12 +149,14 @@ mutual
CasePair : (qty : Qty) -> (pair : Elim d n) -> CasePair : (qty : Qty) -> (pair : Elim d n) ->
(ret : ScopeTerm d n) -> (ret : ScopeTerm d n) ->
(body : ScopeTermN 2 d n) -> (body : ScopeTermN 2 d n) ->
(loc : Loc) ->
Elim d n Elim d n
||| enum matching ||| enum matching
CaseEnum : (qty : Qty) -> (tag : Elim d n) -> CaseEnum : (qty : Qty) -> (tag : Elim d n) ->
(ret : ScopeTerm d n) -> (ret : ScopeTerm d n) ->
(arms : CaseEnumArms d n) -> (arms : CaseEnumArms d n) ->
(loc : Loc) ->
Elim d n Elim d n
||| nat matching ||| nat matching
@ -161,33 +164,36 @@ mutual
(ret : ScopeTerm d n) -> (ret : ScopeTerm d n) ->
(zero : Term d n) -> (zero : Term d n) ->
(succ : ScopeTermN 2 d n) -> (succ : ScopeTermN 2 d n) ->
(loc : Loc) ->
Elim d n Elim d n
||| unboxing ||| unboxing
CaseBox : (qty : Qty) -> (box : Elim d n) -> CaseBox : (qty : Qty) -> (box : Elim d n) ->
(ret : ScopeTerm d n) -> (ret : ScopeTerm d n) ->
(body : ScopeTerm d n) -> (body : ScopeTerm d n) ->
(loc : Loc) ->
Elim d n Elim d n
||| dim application ||| dim application
(:%) : (fun : Elim d n) -> (arg : Dim d) -> Elim d n DApp : (fun : Elim d n) -> (arg : Dim d) -> (loc : Loc) -> Elim d n
||| type-annotated term ||| type-annotated term
(:#) : (tm, ty : Term d n) -> Elim d n Ann : (tm, ty : Term d n) -> (loc : Loc) -> Elim d n
||| coerce a value along a type equality, or show its coherence ||| coerce a value along a type equality, or show its coherence
||| [@xtt; §2.1.1] ||| [@xtt; §2.1.1]
Coe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> Coe : (ty : DScopeTerm d n) -> (p, q : Dim d) ->
(val : Term d n) -> Elim d n (val : Term d n) -> (loc : Loc) -> Elim d n
||| "generalised composition" [@xtt; §2.1.2] ||| "generalised composition" [@xtt; §2.1.2]
Comp : (ty : Term d n) -> (p, q : Dim d) -> Comp : (ty : Term d n) -> (p, q : Dim d) ->
(val : Term d n) -> (r : Dim d) -> (val : Term d n) -> (r : Dim d) ->
(zero, one : DScopeTerm d n) -> Elim d n (zero, one : DScopeTerm d n) -> (loc : Loc) -> Elim d n
||| match on types. needed for b.s. of coercions [@xtt; §2.2] ||| match on types. needed for b.s. of coercions [@xtt; §2.2]
TypeCase : (ty : Elim d n) -> (ret : Term d n) -> TypeCase : (ty : Elim d n) -> (ret : Term d n) ->
(arms : TypeCaseArms d n) -> (def : Term d n) -> (arms : TypeCaseArms d n) -> (def : Term d n) ->
(loc : Loc) ->
Elim d n Elim d n
||| term closure/suspended substitution ||| term closure/suspended substitution
@ -244,88 +250,211 @@ mutual
||| scope which ignores all its binders ||| scope which ignores all its binders
public export %inline public export %inline
SN : {s : Nat} -> f n -> Scoped s f n SN : {s : Nat} -> f n -> Scoped s f n
SN = S (replicate s Unused) . N SN = S (replicate s $ BN Unused noLoc) . N
||| scope which uses its binders ||| scope which uses its binders
public export %inline public export %inline
SY : NContext s -> f (s + n) -> Scoped s f n SY : BContext s -> f (s + n) -> Scoped s f n
SY ns = S ns . Y SY ns = S ns . Y
public export %inline public export %inline
name : Scoped 1 f n -> BaseName name : Scoped 1 f n -> BindName
name (S [< x] _) = x name (S [< x] _) = x
public export %inline public export %inline
(.name) : Scoped 1 f n -> BaseName (.name) : Scoped 1 f n -> BindName
s.name = name s s.name = name s
||| more convenient Pi ||| more convenient Pi
public export %inline public export %inline
PiY : (qty : Qty) -> (x : BaseName) -> PiY : (qty : Qty) -> (x : BindName) ->
(arg : Term d n) -> (res : Term d (S n)) -> Term d n (arg : Term d n) -> (res : Term d (S n)) -> (loc : Loc) -> Term d n
PiY {qty, x, arg, res} = Pi {qty, arg, res = SY [< x] res} PiY {qty, x, arg, res, loc} = Pi {qty, arg, res = SY [< x] res, loc}
||| more convenient Lam
public export %inline
LamY : (x : BindName) -> (body : Term d (S n)) -> (loc : Loc) -> Term d n
LamY {x, body, loc} = Lam {body = SY [< x] body, loc}
public export %inline
LamN : (body : Term d n) -> (loc : Loc) -> Term d n
LamN {body, loc} = Lam {body = SN body, loc}
||| non dependent function type ||| non dependent function type
public export %inline public export %inline
Arr : (qty : Qty) -> (arg, res : Term d n) -> Term d n Arr : (qty : Qty) -> (arg, res : Term d n) -> (loc : Loc) -> Term d n
Arr {qty, arg, res} = Pi {qty, arg, res = SN res} Arr {qty, arg, res, loc} = Pi {qty, arg, res = SN res, loc}
||| more convenient Sig ||| more convenient Sig
public export %inline public export %inline
SigY : (x : BaseName) -> (fst : Term d n) -> SigY : (x : BindName) -> (fst : Term d n) ->
(snd : Term d (S n)) -> Term d n (snd : Term d (S n)) -> (loc : Loc) -> Term d n
SigY {x, fst, snd} = Sig {fst, snd = SY [< x] snd} SigY {x, fst, snd, loc} = Sig {fst, snd = SY [< x] snd, loc}
||| non dependent pair type ||| non dependent pair type
public export %inline public export %inline
And : (fst, snd : Term d n) -> Term d n And : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
And {fst, snd} = Sig {fst, snd = SN snd} And {fst, snd, loc} = Sig {fst, snd = SN snd, loc}
||| more convenient Eq ||| more convenient Eq
public export %inline public export %inline
EqY : (i : BaseName) -> (ty : Term (S d) n) -> EqY : (i : BindName) -> (ty : Term (S d) n) ->
(l, r : Term d n) -> Term d n (l, r : Term d n) -> (loc : Loc) -> Term d n
EqY {i, ty, l, r} = Eq {ty = SY [< i] ty, l, r} EqY {i, ty, l, r, loc} = Eq {ty = SY [< i] ty, l, r, loc}
||| more convenient DLam
public export %inline
DLamY : (i : BindName) -> (body : Term (S d) n) -> (loc : Loc) -> Term d n
DLamY {i, body, loc} = DLam {body = SY [< i] body, loc}
public export %inline
DLamN : (body : Term d n) -> (loc : Loc) -> Term d n
DLamN {body, loc} = DLam {body = SN body, loc}
||| non dependent equality type ||| non dependent equality type
public export %inline public export %inline
Eq0 : (ty, l, r : Term d n) -> Term d n Eq0 : (ty, l, r : Term d n) -> (loc : Loc) -> Term d n
Eq0 {ty, l, r} = Eq {ty = SN ty, l, r} Eq0 {ty, l, r, loc} = Eq {ty = SN ty, l, r, loc}
||| same as `F` but as a term ||| same as `F` but as a term
public export %inline public export %inline
FT : Name -> Term d n FT : Name -> (loc : Loc) -> Term d n
FT = E . F FT x loc = E $ F x loc
||| abbreviation for a bound variable like `BV 4` instead of ||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))` ||| `B (VS (VS (VS (VS VZ))))`
public export %inline public export %inline
BV : (i : Nat) -> (0 _ : LT i n) => Elim d n BV : (i : Nat) -> (0 _ : LT i n) => (loc : Loc) -> Elim d n
BV i = B $ V i BV i loc = B (V i) loc
||| same as `BV` but as a term ||| same as `BV` but as a term
public export %inline public export %inline
BVT : (i : Nat) -> (0 _ : LT i n) => Term d n BVT : (i : Nat) -> (0 _ : LT i n) => (loc : Loc) -> Term d n
BVT i = E $ BV i BVT i loc = E $ BV i loc
public export public export
makeNat : Nat -> Term d n makeNat : Nat -> Loc -> Term d n
makeNat 0 = Zero makeNat 0 loc = Zero loc
makeNat (S k) = Succ $ makeNat k makeNat (S k) loc = Succ (makeNat k loc) loc
public export %inline public export %inline
enum : List TagVal -> Term d n enum : List TagVal -> Loc -> Term d n
enum = Enum . SortedSet.fromList enum ts loc = Enum (SortedSet.fromList ts) loc
public export %inline public export %inline
typeCase : Elim d n -> Term d n -> typeCase : Elim d n -> Term d n ->
List (TypeCaseArm d n) -> Term d n -> Elim d n List (TypeCaseArm d n) -> Term d n -> Loc -> Elim d n
typeCase ty ret arms def = TypeCase ty ret (fromList arms) def typeCase ty ret arms def loc = TypeCase ty ret (fromList arms) def loc
public export %inline public export %inline
typeCase1Y : Elim d n -> Term d n -> typeCase1Y : Elim d n -> Term d n ->
(k : TyConKind) -> NContext (arity k) -> Term d (arity k + n) -> (k : TyConKind) -> BContext (arity k) -> Term d (arity k + n) ->
{default Nat def : Term d n} -> (loc : Loc) ->
{default (Nat loc) def : Term d n} ->
Elim d n Elim d n
typeCase1Y ty ret k ns body {def} = typeCase1Y ty ret k ns body loc = typeCase ty ret [(k ** SY ns body)] def loc
typeCase ty ret [(k ** SY ns body)] def
export
Located (Elim d n) where
(F _ loc).loc = loc
(B _ loc).loc = loc
(App _ _ loc).loc = loc
(CasePair _ _ _ _ loc).loc = loc
(CaseEnum _ _ _ _ loc).loc = loc
(CaseNat _ _ _ _ _ _ loc).loc = loc
(CaseBox _ _ _ _ loc).loc = loc
(DApp _ _ loc).loc = loc
(Ann _ _ loc).loc = loc
(Coe _ _ _ _ loc).loc = loc
(Comp _ _ _ _ _ _ _ loc).loc = loc
(TypeCase _ _ _ _ loc).loc = loc
(CloE (Sub e _)).loc = e.loc
(DCloE (Sub e _)).loc = e.loc
export
Located (Term d n) where
(TYPE _ loc).loc = loc
(Pi _ _ _ loc).loc = loc
(Lam _ loc).loc = loc
(Sig _ _ loc).loc = loc
(Pair _ _ loc).loc = loc
(Enum _ loc).loc = loc
(Tag _ loc).loc = loc
(Eq _ _ _ loc).loc = loc
(DLam _ loc).loc = loc
(Nat loc).loc = loc
(Zero loc).loc = loc
(Succ _ loc).loc = loc
(BOX _ _ loc).loc = loc
(Box _ loc).loc = loc
(E e).loc = e.loc
(CloT (Sub t _)).loc = t.loc
(DCloT (Sub t _)).loc = t.loc
export
Located1 f => Located (ScopedBody s f n) where
(Y t).loc = t.loc
(N t).loc = t.loc
export
Located1 f => Located (Scoped s f n) where
t.loc = t.body.loc
export
Relocatable (Elim d n) where
setLoc loc (F x _) = F x loc
setLoc loc (B i _) = B i loc
setLoc loc (App fun arg _) = App fun arg loc
setLoc loc (CasePair qty pair ret body _) =
CasePair qty pair ret body loc
setLoc loc (CaseEnum qty tag ret arms _) =
CaseEnum qty tag ret arms loc
setLoc loc (CaseNat qty qtyIH nat ret zero succ _) =
CaseNat qty qtyIH nat ret zero succ loc
setLoc loc (CaseBox qty box ret body _) =
CaseBox qty box ret body loc
setLoc loc (DApp fun arg _) =
DApp fun arg loc
setLoc loc (Ann tm ty _) =
Ann tm ty loc
setLoc loc (Coe ty p q val _) =
Coe ty p q val loc
setLoc loc (Comp ty p q val r zero one _) =
Comp ty p q val r zero one loc
setLoc loc (TypeCase ty ret arms def _) =
TypeCase ty ret arms def loc
setLoc loc (CloE (Sub term subst)) =
CloE $ Sub (setLoc loc term) subst
setLoc loc (DCloE (Sub term subst)) =
DCloE $ Sub (setLoc loc term) subst
export
Relocatable (Term d n) where
setLoc loc (TYPE l _) = TYPE l loc
setLoc loc (Pi qty arg res _) = Pi qty arg res loc
setLoc loc (Lam body _) = Lam body loc
setLoc loc (Sig fst snd _) = Sig fst snd loc
setLoc loc (Pair fst snd _) = Pair fst snd loc
setLoc loc (Enum cases _) = Enum cases loc
setLoc loc (Tag tag _) = Tag tag loc
setLoc loc (Eq ty l r _) = Eq ty l r loc
setLoc loc (DLam body _) = DLam body loc
setLoc loc (Nat _) = Nat loc
setLoc loc (Zero _) = Zero loc
setLoc loc (Succ p _) = Succ p loc
setLoc loc (BOX qty ty _) = BOX qty ty loc
setLoc loc (Box val _) = Box val loc
setLoc loc (E e) = E $ setLoc loc e
setLoc loc (CloT (Sub term subst)) = CloT $ Sub (setLoc loc term) subst
setLoc loc (DCloT (Sub term subst)) = DCloT $ Sub (setLoc loc term) subst
export
Relocatable1 f => Relocatable (ScopedBody s f n) where
setLoc loc (Y body) = Y $ setLoc loc body
setLoc loc (N body) = N $ setLoc loc body
export
Relocatable1 f => Relocatable (Scoped s f n) where
setLoc loc (S names body) = S (setLoc loc <$> names) (setLoc loc body)

View File

@ -82,8 +82,8 @@ PrettyHL a => PrettyHL (Binder a) where
export export
prettyBindType : PrettyHL a => PrettyHL b => prettyBindType : PrettyHL a => PrettyHL b =>
Pretty.HasEnv m => Pretty.HasEnv m =>
Maybe Qty -> BaseName -> a -> Doc HL -> b -> m (Doc HL) Maybe Qty -> BindName -> a -> Doc HL -> b -> m (Doc HL)
prettyBindType q x s arr t = do prettyBindType q (BN x _) s arr t = do
bind <- case q of bind <- case q of
Nothing => pretty0M $ MkBinder x s Nothing => pretty0M $ MkBinder x s
Just q => pretty0M $ MkWithQty q $ MkBinder x s Just q => pretty0M $ MkWithQty q $ MkBinder x s
@ -92,14 +92,15 @@ prettyBindType q x s arr t = do
export export
prettyArm : PrettyHL a => Pretty.HasEnv m => prettyArm : PrettyHL a => Pretty.HasEnv m =>
BinderSort -> SnocList BaseName -> Doc HL -> a -> m (Doc HL) BinderSort -> SnocList BindName -> Doc HL -> a -> m (Doc HL)
prettyArm sort xs pat body = do prettyArm sort xs pat body = do
let xs = map name xs
body <- withPrec Outer $ unders sort xs $ prettyM body body <- withPrec Outer $ unders sort xs $ prettyM body
pure $ hang 2 $ sep [pat <++> !darrowD, body] pure $ hang 2 $ sep [pat <++> !darrowD, body]
export export
prettyLams : PrettyHL a => Pretty.HasEnv m => prettyLams : PrettyHL a => Pretty.HasEnv m =>
Maybe (Doc HL) -> BinderSort -> SnocList BaseName -> a -> Maybe (Doc HL) -> BinderSort -> SnocList BindName -> a ->
m (Doc HL) m (Doc HL)
prettyLams lam sort names body = do prettyLams lam sort names body = do
let var = case sort of T => TVar; D => DVar let var = case sort of T => TVar; D => DVar
@ -109,14 +110,15 @@ prettyLams lam sort names body = do
public export public export
data TypeLine a = MkTypeLine BaseName a data TypeLine a = MkTypeLine BindName a
export export
PrettyHL a => PrettyHL (TypeLine a) where PrettyHL a => PrettyHL (TypeLine a) where
prettyM (MkTypeLine Unused ty) =
bracks <$> pretty0M ty
prettyM (MkTypeLine i ty) = prettyM (MkTypeLine i ty) =
map bracks $ withPrec Outer $ prettyLams Nothing D [< i] ty if i.name == Unused then
bracks <$> pretty0M ty
else
map bracks $ withPrec Outer $ prettyLams Nothing D [< i] ty
export export
@ -142,28 +144,28 @@ prettyTuple = map (parens . align . separate commaD) . traverse prettyM
export export
prettyArms : PrettyHL a => Pretty.HasEnv m => prettyArms : PrettyHL a => Pretty.HasEnv m =>
BinderSort -> List (SnocList BaseName, Doc HL, a) -> m (Doc HL) BinderSort -> List (SnocList BindName, Doc HL, a) -> m (Doc HL)
prettyArms s = prettyArms s =
map (braces . aseparate semiD) . map (braces . aseparate semiD) .
traverse (\(xs, l, r) => prettyArm s xs l r) traverse (\(xs, l, r) => prettyArm s xs l r)
export export
prettyCase' : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) => prettyCase' : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) =>
Doc HL -> a -> BaseName -> b -> Doc HL -> a -> BindName -> b ->
List (SnocList BaseName, Doc HL, c) -> List (SnocList BindName, Doc HL, c) ->
m (Doc HL) m (Doc HL)
prettyCase' intro elim r ret arms = do prettyCase' intro elim r ret arms = do
elim <- pretty0M elim elim <- pretty0M elim
ret <- case r of ret <- case r.name of
Unused => under T r $ pretty0M ret Unused => under T r.name $ pretty0M ret
_ => prettyLams Nothing T [< r] ret _ => prettyLams Nothing T [< r] ret
arms <- prettyArms T arms arms <- prettyArms T arms
pure $ asep [intro <++> elim, returnD <++> ret, ofD <++> arms] pure $ asep [intro <++> elim, returnD <++> ret, ofD <++> arms]
export export
prettyCase : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) => prettyCase : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) =>
Qty -> a -> BaseName -> b -> Qty -> a -> BindName -> b ->
List (SnocList BaseName, Doc HL, c) -> List (SnocList BindName, Doc HL, c) ->
m (Doc HL) m (Doc HL)
prettyCase pi elim r ret arms = do prettyCase pi elim r ret arms = do
caseq <- (caseD <+>) <$> prettySuffix pi caseq <- (caseD <+>) <$> prettySuffix pi
@ -197,14 +199,14 @@ prettyBoxVal : PrettyHL a => Pretty.HasEnv m => a -> m (Doc HL)
prettyBoxVal val = bracks <$> pretty0M val prettyBoxVal val = bracks <$> pretty0M val
export export
prettyCompPat : Pretty.HasEnv m => DimConst -> BaseName -> m (Doc HL) prettyCompPat : Pretty.HasEnv m => DimConst -> BindName -> m (Doc HL)
prettyCompPat e j = hsep <$> sequence [pretty0M e, pretty0M $ DV j] prettyCompPat e j = hsep <$> sequence [pretty0M e, pretty0M $ DV j.name]
export export
toNatLit : Term d n -> Maybe Nat toNatLit : Term d n -> Maybe Nat
toNatLit Zero = Just 0 toNatLit (Zero _) = Just 0
toNatLit (Succ n) = [|S $ toNatLit n|] toNatLit (Succ n _) = [|S $ toNatLit n|]
toNatLit _ = Nothing toNatLit _ = Nothing
private private
eterm : Term d n -> Exists (Term d) eterm : Term d n -> Exists (Term d)
@ -216,69 +218,69 @@ parameters (showSubsts : Bool)
export covering export covering
[TermSubst] PrettyHL (Term d n) using ElimSubst [TermSubst] PrettyHL (Term d n) using ElimSubst
where where
prettyM (TYPE l) = prettyM (TYPE l _) =
pure $ !typeD <+> hl Syntax !(prettyUnivSuffix l) pure $ !typeD <+> hl Syntax !(prettyUnivSuffix l)
prettyM (Pi qty s (S _ (N t))) = do prettyM (Pi qty s (S _ (N t)) _) = do
dom <- pretty0M $ MkWithQty qty s dom <- pretty0M $ MkWithQty qty s
cod <- withPrec AnnR $ prettyM t cod <- withPrec AnnR $ prettyM t
parensIfM AnnR $ asep [dom <++> !arrowD, cod] parensIfM AnnR $ asep [dom <++> !arrowD, cod]
prettyM (Pi qty s (S [< x] (Y t))) = prettyM (Pi qty s (S [< x] (Y t)) _) =
prettyBindType (Just qty) x s !arrowD t prettyBindType (Just qty) x s !arrowD t
prettyM (Lam (S x t)) = prettyM (Lam (S x t) _) =
let GotLams {names, body, _} = getLams' x t.term Refl in let GotLams {names, body, _} = getLams' x t.term Refl in
prettyLams (Just !lamD) T (toSnocList' names) body prettyLams (Just !lamD) T (toSnocList' names) body
prettyM (Sig s (S _ (N t))) = do prettyM (Sig s (S _ (N t)) _) = do
s <- withPrec InTimes $ prettyM s s <- withPrec InTimes $ prettyM s
t <- withPrec Times $ prettyM t t <- withPrec Times $ prettyM t
parensIfM Times $ asep [s <++> !timesD, t] parensIfM Times $ asep [s <++> !timesD, t]
prettyM (Sig s (S [< x] (Y t))) = prettyM (Sig s (S [< x] (Y t)) _) =
prettyBindType Nothing x s !timesD t prettyBindType Nothing x s !timesD t
prettyM (Pair s t) = prettyM (Pair s t _) =
let GotPairs {init, last, _} = getPairs' [< s] t in let GotPairs {init, last, _} = getPairs' [< s] t in
prettyTuple $ toList $ init :< last prettyTuple $ toList $ init :< last
prettyM (Enum tags) = prettyM (Enum tags _) =
pure $ delims "{" "}" . aseparate comma $ map prettyTagBare $ pure $ delims "{" "}" . aseparate comma $ map prettyTagBare $
Prelude.toList tags Prelude.toList tags
prettyM (Tag t) = prettyM (Tag t _) =
pure $ prettyTag t pure $ prettyTag t
prettyM (Eq (S _ (N ty)) l r) = do prettyM (Eq (S _ (N ty)) l r _) = do
l <- withPrec InEq $ prettyM l l <- withPrec InEq $ prettyM l
r <- withPrec InEq $ prettyM r r <- withPrec InEq $ prettyM r
ty <- withPrec InEq $ prettyM ty ty <- withPrec InEq $ prettyM ty
parensIfM Eq $ asep [l <++> !eqndD, r <++> colonD, ty] parensIfM Eq $ asep [l <++> !eqndD, r <++> colonD, ty]
prettyM (Eq (S [< i] (Y ty)) l r) = do prettyM (Eq (S [< i] (Y ty)) l r _) = do
prettyApps Nothing (L eqD) prettyApps Nothing (L eqD)
[epretty $ MkTypeLine i ty, epretty l, epretty r] [epretty $ MkTypeLine i ty, epretty l, epretty r]
prettyM (DLam (S i t)) = prettyM (DLam (S i t) _) =
let GotDLams {names, body, _} = getDLams' i t.term Refl in let GotDLams {names, body, _} = getDLams' i t.term Refl in
prettyLams (Just !dlamD) D (toSnocList' names) body prettyLams (Just !dlamD) D (toSnocList' names) body
prettyM Nat = natD prettyM (Nat _) = natD
prettyM Zero = pure $ hl Syntax "0" prettyM (Zero _) = pure $ hl Syntax "0"
prettyM (Succ n) = prettyM (Succ n _) =
case toNatLit n of case toNatLit n of
Just n => pure $ hl Syntax $ pretty $ S n Just n => pure $ hl Syntax $ pretty $ S n
Nothing => prettyApps Nothing (L succD) [n] Nothing => prettyApps Nothing (L succD) [n]
prettyM (BOX pi ty) = do prettyM (BOX pi ty _) = do
pi <- pretty0M pi pi <- pretty0M pi
ty <- pretty0M ty ty <- pretty0M ty
pure $ bracks $ hcat [pi, dotD, align ty] pure $ bracks $ hcat [pi, dotD, align ty]
prettyM (Box val) = prettyBoxVal val prettyM (Box val _) = prettyBoxVal val
prettyM (E e) = prettyM e prettyM (E e) = prettyM e
@ -299,49 +301,49 @@ parameters (showSubsts : Bool)
export covering export covering
[ElimSubst] PrettyHL (Elim d n) using TermSubst [ElimSubst] PrettyHL (Elim d n) using TermSubst
where where
prettyM (F x) = prettyM (F x _) =
hl' Free <$> prettyM x hl' Free <$> prettyM x
prettyM (B i) = prettyM (B i _) =
prettyVar TVar TVarErr (!ask).tnames i prettyVar TVar TVarErr (!ask).tnames i
prettyM (e :@ s) = prettyM (App e s _) =
let GotArgs {fun, args, _} = getArgs' e [s] in let GotArgs {fun, args, _} = getArgs' e [s] in
prettyApps Nothing fun args prettyApps Nothing fun args
prettyM (CasePair pi p (S [< r] ret) (S [< x, y] body)) = do prettyM (CasePair pi p (S [< r] ret) (S [< x, y] body) _) = do
pat <- parens . separate commaD <$> traverse (hlF TVar . prettyM) [x, y] pat <- parens . separate commaD <$> traverse (hlF TVar . prettyM) [x, y]
prettyCase pi p r ret.term [([< x, y], pat, body.term)] prettyCase pi p r ret.term [([< x, y], pat, body.term)]
prettyM (CaseEnum pi t (S [< r] ret) arms) = prettyM (CaseEnum pi t (S [< r] ret) arms _) =
prettyCase pi t r ret.term prettyCase pi t r ret.term
[([<], prettyTag t, b) | (t, b) <- SortedMap.toList arms] [([<], prettyTag t, b) | (t, b) <- SortedMap.toList arms]
prettyM (CaseNat pi pi' nat (S [< r] ret) zer (S [< s, ih] suc)) = prettyM (CaseNat pi pi' nat (S [< r] ret) zer (S [< s, ih] suc) _) =
prettyCase pi nat r ret.term prettyCase pi nat r ret.term
[([<], zeroD, eterm zer), [([<], zeroD, eterm zer),
([< s, ih], !succPat, eterm suc.term)] ([< s, ih], !succPat, eterm suc.term)]
where where
succPat : m (Doc HL) succPat : m (Doc HL)
succPat = case (ih, pi') of succPat = case (ih, pi') of
(Unused, Zero) => pure $ succD <++> !(pretty0M s) (BN Unused _, Zero) => pure $ succD <++> !(pretty0M s)
_ => pure $ asep [succD <++> !(pretty0M s) <+> comma, _ => pure $ asep [succD <++> !(pretty0M s) <+> comma,
!(pretty0M $ MkWithQty pi' ih)] !(pretty0M $ MkWithQty pi' ih)]
prettyM (CaseBox pi box (S [< r] ret) (S [< u] body)) = prettyM (CaseBox pi box (S [< r] ret) (S [< u] body) _) =
prettyCase pi box r ret.term prettyCase pi box r ret.term
[([< u], !(prettyBoxVal $ TV u), body.term)] [([< u], !(prettyBoxVal $ TV u.name), body.term)]
prettyM (e :% d) = prettyM (DApp e d _) =
let GotDArgs {fun, args, _} = getDArgs' e [d] in let GotDArgs {fun, args, _} = getDArgs' e [d] in
prettyApps (Just "@") fun args prettyApps (Just "@") fun args
prettyM (s :# a) = do prettyM (Ann s a _) = do
s <- withPrec AnnL $ prettyM s s <- withPrec AnnL $ prettyM s
a <- withPrec AnnR $ prettyM a a <- withPrec AnnR $ prettyM a
parensIfM AnnR $ hang 2 $ s <++> !annD <%%> a parensIfM AnnR $ hang 2 $ s <++> !annD <%%> a
prettyM (Coe (S [< i] ty) p q val) = prettyM (Coe (S [< i] ty) p q val _) =
let ty = case ty of let ty = case ty of
Y ty => epretty $ MkTypeLine i ty Y ty => epretty $ MkTypeLine i ty
N ty => epretty ty N ty => epretty ty
@ -352,9 +354,9 @@ parameters (showSubsts : Bool)
(Just "@", epretty q), (Just "@", epretty q),
(Nothing, epretty val)] (Nothing, epretty val)]
prettyM (Comp ty p q val r (S [< z] zero) (S [< o] one)) = do prettyM (Comp ty p q val r (S [< z] zero) (S [< o] one) _) = do
apps <- prettyApps' (L compD) apps <- prettyApps' (L compD)
[(Nothing, epretty $ MkTypeLine Unused ty), [(Nothing, epretty $ MkTypeLine (BN Unused noLoc) ty),
(Just "@", epretty p), (Just "@", epretty p),
(Just "@", epretty q), (Just "@", epretty q),
(Nothing, epretty val), (Nothing, epretty val),
@ -364,29 +366,30 @@ parameters (showSubsts : Bool)
([< o], !(prettyCompPat One o), one.term)] ([< o], !(prettyCompPat One o), one.term)]
pure $ apps <++> arms pure $ apps <++> arms
prettyM (TypeCase ty ret arms def) = do prettyM (TypeCase ty ret arms def _) = do
arms <- traverse fromArm (toList arms) arms <- traverse fromArm (toList arms)
prettyCase' typecaseD ty Unused ret $ prettyCase' typecaseD ty (BN Unused noLoc) ret $
arms ++ [([<], hl Syntax "_", eterm def)] arms ++ [([<], hl Syntax "_", eterm def)]
where where
v : BaseName -> Doc HL v : BindName -> Doc HL
v = pretty0 True . TV v = pretty0 True . TV . name
tyCasePat : (k : TyConKind) -> NContext (arity k) -> m (Doc HL) tyCasePat : (k : TyConKind) -> BContext (arity k) -> m (Doc HL)
tyCasePat KTYPE [<] = typeD tyCasePat KTYPE [<] = typeD
tyCasePat KPi [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b] tyCasePat KPi [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b]
tyCasePat KSig [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b] tyCasePat KSig [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b]
tyCasePat KEnum [<] = pure $ hl Syntax "{}" tyCasePat KEnum [<] = pure $ hl Syntax "{}"
tyCasePat KEq vars = prettyApps Nothing (L eqD) $ map TV $ toList' vars
tyCasePat KNat [<] = natD tyCasePat KNat [<] = natD
tyCasePat KBOX [< a] = pure $ bracks $ v a tyCasePat KBOX [< a] = pure $ bracks $ v a
tyCasePat KEq vars =
prettyApps Nothing (L eqD) $ map (TV . name) $ toList' vars
fromArm : TypeCaseArm d n -> fromArm : TypeCaseArm d n ->
m (SnocList BaseName, Doc HL, Exists (Term d)) m (SnocList BindName, Doc HL, Exists (Term d))
fromArm (k ** S ns t) = fromArm (k ** S ns t) =
pure (toSnocList' ns, !(tyCasePat k ns), eterm t.term) pure (toSnocList' ns, !(tyCasePat k ns), eterm t.term)
prettyM (CloE (Sub e th)) = prettyM (CloE (Sub e th)) =
if showSubsts then if showSubsts then
parensIfM SApp . hang 2 =<< parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM e) <%> prettyTSubst th|] [|withPrec SApp (prettyM e) <%> prettyTSubst th|]
@ -414,7 +417,7 @@ PrettyHL (Elim d n) where prettyM = prettyM @{ElimSubst False}
export covering export covering
prettyTerm : (unicode : Bool) -> prettyTerm : (unicode : Bool) ->
(dnames : NContext d) -> (tnames : NContext n) -> (dnames : BContext d) -> (tnames : BContext n) ->
Term d n -> Doc HL Term d n -> Doc HL
prettyTerm unicode dnames tnames term = prettyTerm unicode dnames tnames term =
pretty0With unicode (toSnocList' dnames) (toSnocList' tnames) term pretty0With unicode (toNames dnames) (toNames tnames) term

View File

@ -13,8 +13,8 @@ import public Data.Vect
public export %inline public export %inline
isLam : Term {} -> Bool isLam : Term {} -> Bool
isLam (Lam _) = True isLam (Lam {}) = True
isLam _ = False isLam _ = False
public export public export
0 NotLam : Pred $ Term {} 0 NotLam : Pred $ Term {}
@ -23,8 +23,8 @@ NotLam = No . isLam
public export %inline public export %inline
isDLam : Term {} -> Bool isDLam : Term {} -> Bool
isDLam (DLam _) = True isDLam (DLam {}) = True
isDLam _ = False isDLam _ = False
public export public export
0 NotDLam : Pred $ Term {} 0 NotDLam : Pred $ Term {}
@ -43,7 +43,7 @@ NotPair = No . isPair
public export %inline public export %inline
isApp : Elim {} -> Bool isApp : Elim {} -> Bool
isApp (_ :@ _) = True isApp (App {}) = True
isApp _ = False isApp _ = False
public export public export
@ -53,19 +53,21 @@ NotApp = No . isApp
public export %inline public export %inline
isDApp : Elim {} -> Bool isDApp : Elim {} -> Bool
isDApp (_ :% _) = True isDApp (DApp {}) = True
isDApp _ = False isDApp _ = False
public export public export
0 NotDApp : Pred $ Elim {} 0 NotDApp : Pred $ Elim {}
NotDApp = No . isDApp NotDApp = No . isDApp
infixl 9 :@@ -- infixl 9 :@@
||| apply multiple arguments at once -- ||| apply multiple arguments at once
public export %inline -- public export %inline
(:@@) : Elim d n -> List (Term d n) -> Elim d n -- (:@@) : Elim d n -> List (Term d n) -> Elim d n
f :@@ ss = foldl (:@) f ss -- f :@@ ss = foldl app f ss where
-- app : Elim d n -> Term d n -> Elim d n
-- app f s = App f s (f.loc `extend'` s.loc.bounds)
public export public export
record GetArgs d n where record GetArgs d n where
@ -85,7 +87,7 @@ mutual
getArgsNc : (e : Elim d n) -> (0 nc : NotClo e) => getArgsNc : (e : Elim d n) -> (0 nc : NotClo e) =>
List (Term d n) -> GetArgs d n List (Term d n) -> GetArgs d n
getArgsNc fun args = case nchoose $ isApp fun of getArgsNc fun args = case nchoose $ isApp fun of
Left y => let f :@ a = fun in getArgs' f (a :: args) Left y => let App f a _ = fun in getArgs' f (a :: args)
Right n => GotArgs {fun, args, notApp = n} Right n => GotArgs {fun, args, notApp = n}
||| splits an application into its head and arguments. if it's not an ||| splits an application into its head and arguments. if it's not an
@ -96,11 +98,13 @@ getArgs : Elim d n -> GetArgs d n
getArgs e = getArgs' e [] getArgs e = getArgs' e []
infixl 9 :%% -- infixl 9 :%%
||| apply multiple dimension arguments at once -- ||| apply multiple dimension arguments at once
public export %inline -- public export %inline
(:%%) : Elim d n -> List (Dim d) -> Elim d n -- (:%%) : Elim d n -> List (Dim d) -> Elim d n
f :%% ss = foldl (:%) f ss -- f :%% ss = foldl dapp f ss where
-- dapp : Elim d n -> Dim d -> Elim d n
-- dapp f p = DApp f p (f.loc `extend'` p.loc.bounds)
public export public export
record GetDArgs d n where record GetDArgs d n where
@ -120,7 +124,7 @@ mutual
getDArgsNc : (e : Elim d n) -> (0 nc : NotClo e) => getDArgsNc : (e : Elim d n) -> (0 nc : NotClo e) =>
List (Dim d) -> GetDArgs d n List (Dim d) -> GetDArgs d n
getDArgsNc fun args = case nchoose $ isDApp fun of getDArgsNc fun args = case nchoose $ isDApp fun of
Left y => let f :% d = fun in getDArgs' f (d :: args) Left y => let DApp f d _ = fun in getDArgs' f (d :: args)
Right n => GotDArgs {fun, args, notDApp = n} Right n => GotDArgs {fun, args, notDApp = n}
||| splits a dimension application into its head and arguments. if it's not an ||| splits a dimension application into its head and arguments. if it's not an
@ -130,33 +134,33 @@ getDArgs : Elim d n -> GetDArgs d n
getDArgs e = getDArgs' e [] getDArgs e = getDArgs' e []
infixr 1 :\\ -- infixr 1 :\\
public export -- public export
absN : NContext m -> Term d (m + n) -> Term d n -- absN : BContext m -> Term d (m + n) -> Term d n
absN [<] s = s -- absN [<] s = s
absN (xs :< x) s = absN xs $ Lam $ ST [< x] s -- absN (xs :< x) s = absN xs $ Lam (ST [< x] s) ?absloc
public export %inline -- public export %inline
(:\\) : NContext m -> Term d (m + n) -> Term d n -- (:\\) : BContext m -> Term d (m + n) -> Term d n
(:\\) = absN -- (:\\) = absN
infixr 1 :\\% -- infixr 1 :\\%
public export -- public export
dabsN : NContext m -> Term (m + d) n -> Term d n -- dabsN : BContext m -> Term (m + d) n -> Term d n
dabsN [<] s = s -- dabsN [<] s = s
dabsN (xs :< x) s = dabsN xs $ DLam $ DST [< x] s -- dabsN (xs :< x) s = dabsN xs $ DLam (DST [< x] s) ?dabsLoc
public export %inline -- public export %inline
(:\\%) : NContext m -> Term (m + d) n -> Term d n -- (:\\%) : BContext m -> Term (m + d) n -> Term d n
(:\\%) = dabsN -- (:\\%) = dabsN
public export public export
record GetLams d n where record GetLams d n where
constructor GotLams constructor GotLams
{0 lams, rest : Nat} {0 lams, rest : Nat}
names : NContext lams names : BContext lams
body : Term d rest body : Term d rest
0 eq : lams + n = rest 0 eq : lams + n = rest
0 notLam : NotLam body 0 notLam : NotLam body
@ -164,7 +168,7 @@ record GetLams d n where
mutual mutual
export %inline export %inline
getLams' : forall lams, rest. getLams' : forall lams, rest.
NContext lams -> Term d rest -> (0 eq : lams + n = rest) -> BContext lams -> Term d rest -> (0 eq : lams + n = rest) ->
GetLams d n GetLams d n
getLams' xs s0 eq = getLams' xs s0 eq =
let Element s nc = pushSubsts s0 in let Element s nc = pushSubsts s0 in
@ -172,12 +176,12 @@ mutual
private private
getLamsNc : forall lams, rest. getLamsNc : forall lams, rest.
NContext lams -> BContext lams ->
(t : Term d rest) -> (0 nc : NotClo t) => (t : Term d rest) -> (0 nc : NotClo t) =>
(0 eq : lams + n = rest) -> (0 eq : lams + n = rest) ->
GetLams d n GetLams d n
getLamsNc xs s Refl = case nchoose $ isLam s of getLamsNc xs s Refl = case nchoose $ isLam s of
Left y => let Lam (S [< x] body) = s in Left y => let Lam (S [< x] body) _ = s in
getLams' (xs :< x) (assert_smaller s body.term) Refl getLams' (xs :< x) (assert_smaller s body.term) Refl
Right n => GotLams xs s Refl n Right n => GotLams xs s Refl n
@ -190,7 +194,7 @@ public export
record GetDLams d n where record GetDLams d n where
constructor GotDLams constructor GotDLams
{0 lams, rest : Nat} {0 lams, rest : Nat}
names : NContext lams names : BContext lams
body : Term rest n body : Term rest n
0 eq : lams + d = rest 0 eq : lams + d = rest
0 notDLam : NotDLam body 0 notDLam : NotDLam body
@ -198,7 +202,7 @@ record GetDLams d n where
mutual mutual
export %inline export %inline
getDLams' : forall lams, rest. getDLams' : forall lams, rest.
NContext lams -> Term rest n -> (0 eq : lams + d = rest) -> BContext lams -> Term rest n -> (0 eq : lams + d = rest) ->
GetDLams d n GetDLams d n
getDLams' xs s0 eq = getDLams' xs s0 eq =
let Element s nc = pushSubsts s0 in let Element s nc = pushSubsts s0 in
@ -206,12 +210,12 @@ mutual
private private
getDLamsNc : forall lams, rest. getDLamsNc : forall lams, rest.
NContext lams -> BContext lams ->
(t : Term rest n) -> (0 nc : NotClo t) => (t : Term rest n) -> (0 nc : NotClo t) =>
(0 eq : lams + d = rest) -> (0 eq : lams + d = rest) ->
GetDLams d n GetDLams d n
getDLamsNc is s Refl = case nchoose $ isDLam s of getDLamsNc is s Refl = case nchoose $ isDLam s of
Left y => let DLam (S [< i] body) = s in Left y => let DLam (S [< i] body) _ = s in
getDLams' (is :< i) (assert_smaller s body.term) Refl getDLams' (is :< i) (assert_smaller s body.term) Refl
Right n => GotDLams is s Refl n Right n => GotDLams is s Refl n
@ -238,7 +242,7 @@ mutual
(t : Term d n) -> (0 nc : NotClo t) => (t : Term d n) -> (0 nc : NotClo t) =>
GetPairs d n GetPairs d n
getPairsNc ss t = case nchoose $ isPair t of getPairsNc ss t = case nchoose $ isPair t of
Left y => let Pair s t = t in Left y => let Pair s t _ = t in
getPairs' (ss :< s) t getPairs' (ss :< s) t
Right n => GotPairs ss t n Right n => GotPairs ss t n

View File

@ -20,14 +20,14 @@ namespace CanDSubst
export export
CanDSubst Term where CanDSubst Term where
s // Shift SZ = s s // Shift SZ = s
TYPE l // _ = TYPE l TYPE l loc // _ = TYPE l loc
DCloT (Sub s ph) // th = DCloT $ Sub s $ ph . th DCloT (Sub s ph) // th = DCloT $ Sub s $ ph . th
s // th = DCloT $ Sub s th s // th = DCloT $ Sub s th
private private
subDArgs : Elim dfrom n -> DSubst dfrom dto -> Elim dto n subDArgs : Elim dfrom n -> DSubst dfrom dto -> Elim dto n
subDArgs (f :% d) th = subDArgs f th :% (d // th) subDArgs (DApp f d loc) th = DApp (subDArgs f th) (d // th) loc
subDArgs e th = DCloE $ Sub e th subDArgs e th = DCloE $ Sub e th
||| does the minimal reasonable work: ||| does the minimal reasonable work:
||| - deletes the closure around a term variable ||| - deletes the closure around a term variable
@ -39,9 +39,9 @@ subDArgs e th = DCloE $ Sub e th
export export
CanDSubst Elim where CanDSubst Elim where
e // Shift SZ = e e // Shift SZ = e
F x // _ = F x F x loc // _ = F x loc
B i // _ = B i B i loc // _ = B i loc
f :% d // th = subDArgs (f :% d) th e@(DApp {}) // th = subDArgs e th
DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th
e // th = DCloE $ Sub e th e // th = DCloE $ Sub e th
@ -61,8 +61,8 @@ namespace DSubst.DScopeTermN
S ns (N body) // th = S ns $ N $ body // th S ns (N body) // th = S ns $ N $ body // th
export %inline FromVar (Elim d) where fromVar = B export %inline FromVar (Elim d) where fromVarLoc = B
export %inline FromVar (Term d) where fromVar = E . fromVar export %inline FromVar (Term d) where fromVarLoc = E .: fromVar
||| does the minimal reasonable work: ||| does the minimal reasonable work:
@ -73,8 +73,8 @@ export %inline FromVar (Term d) where fromVar = E . fromVar
||| - otherwise, wraps in a new closure ||| - otherwise, wraps in a new closure
export export
CanSubstSelf (Elim d) where CanSubstSelf (Elim d) where
F x // _ = F x F x loc // _ = F x loc
B i // th = th !! i B i loc // th = getLoc th i loc
CloE (Sub e ph) // th = assert_total CloE $ Sub e $ ph . th CloE (Sub e ph) // th = assert_total CloE $ Sub e $ ph . th
e // th = case force th of e // th = case force th of
Shift SZ => e Shift SZ => e
@ -93,7 +93,7 @@ namespace CanTSubst
||| - otherwise, wraps in a new closure ||| - otherwise, wraps in a new closure
export export
CanTSubst Term where CanTSubst Term where
TYPE l // _ = TYPE l TYPE l loc // _ = TYPE l loc
E e // th = E $ e // th E e // th = E $ e // th
CloT (Sub s ph) // th = CloT $ Sub s $ ph . th CloT (Sub s ph) // th = CloT $ Sub s $ ph . th
s // th = case force th of s // th = case force th of
@ -192,12 +192,12 @@ dsub1 t p = dsubN t [< p]
public export %inline public export %inline
(.zero) : DScopeTerm d n -> Term d n (.zero) : DScopeTerm d n -> {default noLoc loc : Loc} -> Term d n
body.zero = dsub1 body $ K Zero body.zero = dsub1 body $ K Zero loc
public export %inline public export %inline
(.one) : DScopeTerm d n -> Term d n (.one) : DScopeTerm d n -> {default noLoc loc : Loc} -> Term d n
body.one = dsub1 body $ K One body.one = dsub1 body $ K One loc
public export public export
@ -251,29 +251,34 @@ mutual
mutual mutual
export export
PushSubsts Term Subst.isCloT where PushSubsts Term Subst.isCloT where
pushSubstsWith th ph (TYPE l) = pushSubstsWith th ph (TYPE l loc) =
nclo $ TYPE l nclo $ TYPE l loc
pushSubstsWith th ph (Pi qty a body) = pushSubstsWith th ph (Pi qty a body loc) =
nclo $ Pi qty (a // th // ph) (body // th // ph) nclo $ Pi qty (a // th // ph) (body // th // ph) loc
pushSubstsWith th ph (Lam body) = pushSubstsWith th ph (Lam body loc) =
nclo $ Lam (body // th // ph) nclo $ Lam (body // th // ph) loc
pushSubstsWith th ph (Sig a b) = pushSubstsWith th ph (Sig a b loc) =
nclo $ Sig (a // th // ph) (b // th // ph) nclo $ Sig (a // th // ph) (b // th // ph) loc
pushSubstsWith th ph (Pair s t) = pushSubstsWith th ph (Pair s t loc) =
nclo $ Pair (s // th // ph) (t // th // ph) nclo $ Pair (s // th // ph) (t // th // ph) loc
pushSubstsWith th ph (Enum tags) = pushSubstsWith th ph (Enum tags loc) =
nclo $ Enum tags nclo $ Enum tags loc
pushSubstsWith th ph (Tag tag) = pushSubstsWith th ph (Tag tag loc) =
nclo $ Tag tag nclo $ Tag tag loc
pushSubstsWith th ph (Eq ty l r) = pushSubstsWith th ph (Eq ty l r loc) =
nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph) nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph) loc
pushSubstsWith th ph (DLam body) = pushSubstsWith th ph (DLam body loc) =
nclo $ DLam (body // th // ph) nclo $ DLam (body // th // ph) loc
pushSubstsWith _ _ Nat = nclo Nat pushSubstsWith _ _ (Nat loc) =
pushSubstsWith _ _ Zero = nclo Zero nclo $ Nat loc
pushSubstsWith th ph (Succ n) = nclo $ Succ $ n // th // ph pushSubstsWith _ _ (Zero loc) =
pushSubstsWith th ph (BOX pi ty) = nclo $ BOX pi $ ty // th // ph nclo $ Zero loc
pushSubstsWith th ph (Box val) = nclo $ Box $ val // th // ph pushSubstsWith th ph (Succ n loc) =
nclo $ Succ (n // th // ph) loc
pushSubstsWith th ph (BOX pi ty loc) =
nclo $ BOX pi (ty // th // ph) loc
pushSubstsWith th ph (Box val loc) =
nclo $ Box (val // th // ph) loc
pushSubstsWith th ph (E e) = pushSubstsWith th ph (E e) =
let Element e nc = pushSubstsWith th ph e in nclo $ E e let Element e nc = pushSubstsWith th ph e in nclo $ E e
pushSubstsWith th ph (CloT (Sub s ps)) = pushSubstsWith th ph (CloT (Sub s ps)) =
@ -283,38 +288,38 @@ mutual
export export
PushSubsts Elim Subst.isCloE where PushSubsts Elim Subst.isCloE where
pushSubstsWith th ph (F x) = pushSubstsWith th ph (F x loc) =
nclo $ F x nclo $ F x loc
pushSubstsWith th ph (B i) = pushSubstsWith th ph (B i loc) =
let res = ph !! i in let res = getLoc ph i loc in
case nchoose $ isCloE res of case nchoose $ isCloE res of
Left yes => assert_total pushSubsts res Left yes => assert_total pushSubsts res
Right no => Element res no Right no => Element res no
pushSubstsWith th ph (f :@ s) = pushSubstsWith th ph (App f s loc) =
nclo $ (f // th // ph) :@ (s // th // ph) nclo $ App (f // th // ph) (s // th // ph) loc
pushSubstsWith th ph (CasePair pi p r b) = pushSubstsWith th ph (CasePair pi p r b loc) =
nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph) nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (CaseEnum pi t r arms) = pushSubstsWith th ph (CaseEnum pi t r arms loc) =
nclo $ CaseEnum pi (t // th // ph) (r // th // ph) nclo $ CaseEnum pi (t // th // ph) (r // th // ph)
(map (\b => b // th // ph) arms) (map (\b => b // th // ph) arms) loc
pushSubstsWith th ph (CaseNat pi pi' n r z s) = pushSubstsWith th ph (CaseNat pi pi' n r z s loc) =
nclo $ CaseNat pi pi' (n // th // ph) (r // th // ph) nclo $ CaseNat pi pi' (n // th // ph) (r // th // ph)
(z // th // ph) (s // th // ph) (z // th // ph) (s // th // ph) loc
pushSubstsWith th ph (CaseBox pi x r b) = pushSubstsWith th ph (CaseBox pi x r b loc) =
nclo $ CaseBox pi (x // th // ph) (r // th // ph) (b // th // ph) nclo $ CaseBox pi (x // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (f :% d) = pushSubstsWith th ph (DApp f d loc) =
nclo $ (f // th // ph) :% (d // th) nclo $ DApp (f // th // ph) (d // th) loc
pushSubstsWith th ph (s :# a) = pushSubstsWith th ph (Ann s a loc) =
nclo $ (s // th // ph) :# (a // th // ph) nclo $ Ann (s // th // ph) (a // th // ph) loc
pushSubstsWith th ph (Coe ty p q val) = pushSubstsWith th ph (Coe ty p q val loc) =
nclo $ Coe (ty // th // ph) (p // th) (q // th) (val // th // ph) nclo $ Coe (ty // th // ph) (p // th) (q // th) (val // th // ph) loc
pushSubstsWith th ph (Comp ty p q val r zero one) = pushSubstsWith th ph (Comp ty p q val r zero one loc) =
nclo $ Comp (ty // th // ph) (p // th) (q // th) nclo $ Comp (ty // th // ph) (p // th) (q // th)
(val // th // ph) (r // th) (val // th // ph) (r // th)
(zero // th // ph) (one // th // ph) (zero // th // ph) (one // th // ph) loc
pushSubstsWith th ph (TypeCase ty ret arms def) = pushSubstsWith th ph (TypeCase ty ret arms def loc) =
nclo $ TypeCase (ty // th // ph) (ret // th // ph) nclo $ TypeCase (ty // th // ph) (ret // th // ph)
(map (\t => t // th // ph) arms) (def // th // ph) (map (\t => t // th // ph) arms) (def // th // ph) loc
pushSubstsWith th ph (CloE (Sub e ps)) = pushSubstsWith th ph (CloE (Sub e ps)) =
pushSubstsWith th (comp th ps ph) e pushSubstsWith th (comp th ps ph) e
pushSubstsWith th ph (DCloE (Sub e ps)) = pushSubstsWith th ph (DCloE (Sub e ps)) =
@ -323,14 +328,19 @@ mutual
private %inline private %inline
CompHY : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> CompHY : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(r : Dim d) -> (zero, one : DScopeTerm d n) -> Elim d n (r : Dim d) -> (zero, one : DScopeTerm d n) -> (loc : Loc) -> Elim d n
CompHY {ty, p, q, val, r, zero, one} = CompHY {ty, p, q, val, r, zero, one, loc} =
let ty' = SY ty.names $ ty.term // (B VZ ::: shift 2) in -- [fixme] maintain location of existing B VZ
let ty' = SY ty.names $ ty.term // (B VZ noLoc ::: shift 2) in
Comp { Comp {
ty = dsub1 ty q, p, q, ty = dsub1 ty q, p, q,
val = E $ Coe ty p q val, r, val = E $ Coe ty p q val val.loc, r,
zero = SY zero.names $ E $ Coe ty' (B VZ) (weakD 1 q) zero.term, -- [fixme] better locations for these vars?
one = SY one.names $ E $ Coe ty' (B VZ) (weakD 1 q) one.term zero = SY zero.names $ E $
Coe ty' (B VZ zero.loc) (weakD 1 q) zero.term zero.loc,
one = SY one.names $ E $
Coe ty' (B VZ one.loc) (weakD 1 q) one.term one.loc,
loc
} }
public export %inline public export %inline
@ -338,26 +348,28 @@ CompH' : (ty : DScopeTerm d n) ->
(p, q : Dim d) -> (val : Term d n) -> (r : Dim d) -> (p, q : Dim d) -> (val : Term d n) -> (r : Dim d) ->
(zero : DScopeTerm d n) -> (zero : DScopeTerm d n) ->
(one : DScopeTerm d n) -> (one : DScopeTerm d n) ->
(loc : Loc) ->
Elim d n Elim d n
CompH' {ty, p, q, val, r, zero, one} = CompH' {ty, p, q, val, r, zero, one, loc} =
case dsqueeze ty of case dsqueeze ty of
S _ (N ty) => Comp {ty, p, q, val, r, zero, one} S _ (N ty) => Comp {ty, p, q, val, r, zero, one, loc}
S _ (Y _) => CompHY {ty, p, q, val, r, zero, one} S _ (Y _) => CompHY {ty, p, q, val, r, zero, one, loc}
||| heterogeneous composition, using Comp and Coe (and subst) ||| heterogeneous composition, using Comp and Coe (and subst)
||| |||
||| comp [i ⇒ A] @p @q s { (r=0) j ⇒ t₀; (r=1) j ⇒ t₁ } ||| comp [i ⇒ A] @p @q s @r { 0 j ⇒ t₀; 1 j ⇒ t₁ }
||| ≔ ||| ≔
||| comp [Aq/i] @p @q (coe [i ⇒ A] @p @q s) { ||| comp [Aq/i] @p @q (coe [i ⇒ A] @p @q s) @r {
||| (r=0) j ⇒ coe [i ⇒ A] @j @q t₀; ||| 0 j ⇒ coe [i ⇒ A] @j @q t₀;
||| (r=1) j ⇒ coe [i ⇒ A] @j @q t₁ ||| 1 j ⇒ coe [i ⇒ A] @j @q t₁
||| } ||| }
public export %inline public export %inline
CompH : (i : BaseName) -> (ty : Term (S d) n) -> CompH : (i : BindName) -> (ty : Term (S d) n) ->
(p, q : Dim d) -> (val : Term d n) -> (r : Dim d) -> (p, q : Dim d) -> (val : Term d n) -> (r : Dim d) ->
(j0 : BaseName) -> (zero : Term (S d) n) -> (j0 : BindName) -> (zero : Term (S d) n) ->
(j1 : BaseName) -> (one : Term (S d) n) -> (j1 : BindName) -> (one : Term (S d) n) ->
(loc : Loc) ->
Elim d n Elim d n
CompH {i, ty, p, q, val, r, j0, zero, j1, one} = CompH {i, ty, p, q, val, r, j0, zero, j1, one, loc} =
CompH' {ty = SY [< i] ty, p, q, val, r, CompH' {ty = SY [< i] ty, p, q, val, r,
zero = SY [< j0] zero, one = SY [< j0] one} zero = SY [< j0] zero, one = SY [< j0] one, loc}

View File

@ -18,6 +18,12 @@ Tighten (Shift from) where
tighten (Keep p) (SS by) = [|SS $ tighten p by|] tighten (Keep p) (SS by) = [|SS $ tighten p by|]
export
Tighten Dim where
tighten p (K e loc) = pure $ K e loc
tighten p (B i loc) = B <$> tighten p i <*> pure loc
export export
tightenSub : (forall m, n. OPE m n -> env n -> Maybe (env m)) -> tightenSub : (forall m, n. OPE m n -> env n -> Maybe (env m)) ->
OPE to1 to2 -> Subst env from to2 -> Maybe (Subst env from to1) OPE to1 to2 -> Subst env from to2 -> Maybe (Subst env from to1)
@ -46,23 +52,35 @@ tightenDScope f p (S names (N body)) = S names . N <$> f p body
mutual mutual
private private
tightenT : OPE n1 n2 -> Term d n2 -> Maybe (Term d n1) tightenT : OPE n1 n2 -> Term d n2 -> Maybe (Term d n1)
tightenT p (TYPE l) = pure $ TYPE l tightenT p (TYPE l loc) = pure $ TYPE l loc
tightenT p (Pi qty arg res) = tightenT p (Pi qty arg res loc) =
Pi qty <$> tightenT p arg <*> tightenS p res Pi qty <$> tightenT p arg <*> tightenS p res <*> pure loc
tightenT p (Lam body) = Lam <$> tightenS p body tightenT p (Lam body loc) =
tightenT p (Sig fst snd) = Sig <$> tightenT p fst <*> tightenS p snd Lam <$> tightenS p body <*> pure loc
tightenT p (Pair fst snd) = Pair <$> tightenT p fst <*> tightenT p snd tightenT p (Sig fst snd loc) =
tightenT p (Enum cases) = pure $ Enum cases Sig <$> tightenT p fst <*> tightenS p snd <*> pure loc
tightenT p (Tag tag) = pure $ Tag tag tightenT p (Pair fst snd loc) =
tightenT p (Eq ty l r) = Pair <$> tightenT p fst <*> tightenT p snd <*> pure loc
Eq <$> tightenDS p ty <*> tightenT p l <*> tightenT p r tightenT p (Enum cases loc) =
tightenT p (DLam body) = DLam <$> tightenDS p body pure $ Enum cases loc
tightenT p Nat = pure Nat tightenT p (Tag tag loc) =
tightenT p Zero = pure Zero pure $ Tag tag loc
tightenT p (Succ s) = Succ <$> tightenT p s tightenT p (Eq ty l r loc) =
tightenT p (BOX qty ty) = BOX qty <$> tightenT p ty Eq <$> tightenDS p ty <*> tightenT p l <*> tightenT p r <*> pure loc
tightenT p (Box val) = Box <$> tightenT p val tightenT p (DLam body loc) =
tightenT p (E e) = assert_total $ E <$> tightenE p e DLam <$> tightenDS p body <*> pure loc
tightenT p (Nat loc) =
pure $ Nat loc
tightenT p (Zero loc) =
pure $ Zero loc
tightenT p (Succ s loc) =
Succ <$> tightenT p s <*> pure loc
tightenT p (BOX qty ty loc) =
BOX qty <$> tightenT p ty <*> pure loc
tightenT p (Box val loc) =
Box <$> tightenT p val <*> pure loc
tightenT p (E e) =
assert_total $ E <$> tightenE p e
tightenT p (CloT (Sub tm th)) = do tightenT p (CloT (Sub tm th)) = do
th <- assert_total $ tightenSub tightenE p th th <- assert_total $ tightenSub tightenE p th
pure $ CloT $ Sub tm th pure $ CloT $ Sub tm th
@ -72,45 +90,57 @@ mutual
private private
tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1) tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1)
tightenE p (F x) = pure $ F x tightenE p (F x loc) =
tightenE p (B i) = [|B $ tighten p i|] pure $ F x loc
tightenE p (fun :@ arg) = [|tightenE p fun :@ tightenT p arg|] tightenE p (B i loc) =
tightenE p (CasePair qty pair ret body) = B <$> tighten p i <*> pure loc
tightenE p (App fun arg loc) =
App <$> tightenE p fun <*> tightenT p arg <*> pure loc
tightenE p (CasePair qty pair ret body loc) =
CasePair qty <$> tightenE p pair CasePair qty <$> tightenE p pair
<*> tightenS p ret <*> tightenS p ret
<*> tightenS p body <*> tightenS p body
tightenE p (CaseEnum qty tag ret arms) = <*> pure loc
tightenE p (CaseEnum qty tag ret arms loc) =
CaseEnum qty <$> tightenE p tag CaseEnum qty <$> tightenE p tag
<*> tightenS p ret <*> tightenS p ret
<*> traverse (tightenT p) arms <*> traverse (tightenT p) arms
tightenE p (CaseNat qty qtyIH nat ret zero succ) = <*> pure loc
tightenE p (CaseNat qty qtyIH nat ret zero succ loc) =
CaseNat qty qtyIH CaseNat qty qtyIH
<$> tightenE p nat <$> tightenE p nat
<*> tightenS p ret <*> tightenS p ret
<*> tightenT p zero <*> tightenT p zero
<*> tightenS p succ <*> tightenS p succ
tightenE p (CaseBox qty box ret body) = <*> pure loc
tightenE p (CaseBox qty box ret body loc) =
CaseBox qty <$> tightenE p box CaseBox qty <$> tightenE p box
<*> tightenS p ret <*> tightenS p ret
<*> tightenS p body <*> tightenS p body
tightenE p (fun :% arg) = (:% arg) <$> tightenE p fun <*> pure loc
tightenE p (tm :# ty) = [|tightenT p tm :# tightenT p ty|] tightenE p (DApp fun arg loc) =
tightenE p (Coe ty q0 q1 val) = DApp <$> tightenE p fun <*> pure arg <*> pure loc
tightenE p (Ann tm ty loc) =
Ann <$> tightenT p tm <*> tightenT p ty <*> pure loc
tightenE p (Coe ty q0 q1 val loc) =
Coe <$> tightenDS p ty Coe <$> tightenDS p ty
<*> pure q0 <*> pure q1 <*> pure q0 <*> pure q1
<*> tightenT p val <*> tightenT p val
tightenE p (Comp ty q0 q1 val r zero one) = <*> pure loc
tightenE p (Comp ty q0 q1 val r zero one loc) =
Comp <$> tightenT p ty Comp <$> tightenT p ty
<*> pure q0 <*> pure q1 <*> pure q0 <*> pure q1
<*> tightenT p val <*> tightenT p val
<*> pure r <*> pure r
<*> tightenDS p zero <*> tightenDS p zero
<*> tightenDS p one <*> tightenDS p one
tightenE p (TypeCase ty ret arms def) = <*> pure loc
tightenE p (TypeCase ty ret arms def loc) =
TypeCase <$> tightenE p ty TypeCase <$> tightenE p ty
<*> tightenT p ret <*> tightenT p ret
<*> traverse (tightenS p) arms <*> traverse (tightenS p) arms
<*> tightenT p def <*> tightenT p def
<*> pure loc
tightenE p (CloE (Sub el th)) = do tightenE p (CloE (Sub el th)) = do
th <- assert_total $ tightenSub tightenE p th th <- assert_total $ tightenSub tightenE p th
pure $ CloE $ Sub el th pure $ CloE $ Sub el th
@ -130,35 +160,40 @@ mutual
export Tighten (Elim d) where tighten p e = tightenE p e export Tighten (Elim d) where tighten p e = tightenE p e
export Tighten (Term d) where tighten p t = tightenT p t export Tighten (Term d) where tighten p t = tightenT p t
export
Tighten Dim where
tighten p (K e) = pure $ K e
tighten p (B i) = B <$> tighten p i
mutual mutual
export export
dtightenT : OPE d1 d2 -> Term d2 n -> Maybe (Term d1 n) dtightenT : OPE d1 d2 -> Term d2 n -> Maybe (Term d1 n)
dtightenT p (TYPE l) = pure $ TYPE l dtightenT p (TYPE l loc) =
dtightenT p (Pi qty arg res) = pure $ TYPE l loc
Pi qty <$> dtightenT p arg <*> dtightenS p res dtightenT p (Pi qty arg res loc) =
dtightenT p (Lam body) = Pi qty <$> dtightenT p arg <*> dtightenS p res <*> pure loc
Lam <$> dtightenS p body dtightenT p (Lam body loc) =
dtightenT p (Sig fst snd) = Lam <$> dtightenS p body <*> pure loc
Sig <$> dtightenT p fst <*> dtightenS p snd dtightenT p (Sig fst snd loc) =
dtightenT p (Pair fst snd) = Sig <$> dtightenT p fst <*> dtightenS p snd <*> pure loc
Pair <$> dtightenT p fst <*> dtightenT p snd dtightenT p (Pair fst snd loc) =
dtightenT p (Enum cases) = pure $ Enum cases Pair <$> dtightenT p fst <*> dtightenT p snd <*> pure loc
dtightenT p (Tag tag) = pure $ Tag tag dtightenT p (Enum cases loc) =
dtightenT p (Eq ty l r) = pure $ Enum cases loc
Eq <$> dtightenDS p ty <*> dtightenT p l <*> dtightenT p r dtightenT p (Tag tag loc) =
dtightenT p (DLam body) = DLam <$> dtightenDS p body pure $ Tag tag loc
dtightenT p Nat = pure Nat dtightenT p (Eq ty l r loc) =
dtightenT p Zero = pure Zero Eq <$> dtightenDS p ty <*> dtightenT p l <*> dtightenT p r <*> pure loc
dtightenT p (Succ s) = Succ <$> dtightenT p s dtightenT p (DLam body loc) =
dtightenT p (BOX qty ty) = BOX qty <$> dtightenT p ty DLam <$> dtightenDS p body <*> pure loc
dtightenT p (Box val) = Box <$> dtightenT p val dtightenT p (Nat loc) =
dtightenT p (E e) = assert_total $ E <$> dtightenE p e pure $ Nat loc
dtightenT p (Zero loc) =
pure $ Zero loc
dtightenT p (Succ s loc) =
Succ <$> dtightenT p s <*> pure loc
dtightenT p (BOX qty ty loc) =
BOX qty <$> dtightenT p ty <*> pure loc
dtightenT p (Box val loc) =
Box <$> dtightenT p val <*> pure loc
dtightenT p (E e) =
assert_total $ E <$> dtightenE p e
dtightenT p (CloT (Sub tm th)) = do dtightenT p (CloT (Sub tm th)) = do
tm <- dtightenT p tm tm <- dtightenT p tm
th <- assert_total $ traverse (dtightenE p) th th <- assert_total $ traverse (dtightenE p) th
@ -169,38 +204,48 @@ mutual
export export
dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n) dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n)
dtightenE p (F x) = pure $ F x dtightenE p (F x loc) =
dtightenE p (B i) = pure $ B i pure $ F x loc
dtightenE p (fun :@ arg) = [|dtightenE p fun :@ dtightenT p arg|] dtightenE p (B i loc) =
dtightenE p (CasePair qty pair ret body) = pure $ B i loc
dtightenE p (App fun arg loc) =
App <$> dtightenE p fun <*> dtightenT p arg <*> pure loc
dtightenE p (CasePair qty pair ret body loc) =
CasePair qty <$> dtightenE p pair CasePair qty <$> dtightenE p pair
<*> dtightenS p ret <*> dtightenS p ret
<*> dtightenS p body <*> dtightenS p body
dtightenE p (CaseEnum qty tag ret arms) = <*> pure loc
dtightenE p (CaseEnum qty tag ret arms loc) =
CaseEnum qty <$> dtightenE p tag CaseEnum qty <$> dtightenE p tag
<*> dtightenS p ret <*> dtightenS p ret
<*> traverse (dtightenT p) arms <*> traverse (dtightenT p) arms
dtightenE p (CaseNat qty qtyIH nat ret zero succ) = <*> pure loc
dtightenE p (CaseNat qty qtyIH nat ret zero succ loc) =
CaseNat qty qtyIH CaseNat qty qtyIH
<$> dtightenE p nat <$> dtightenE p nat
<*> dtightenS p ret <*> dtightenS p ret
<*> dtightenT p zero <*> dtightenT p zero
<*> dtightenS p succ <*> dtightenS p succ
dtightenE p (CaseBox qty box ret body) = <*> pure loc
dtightenE p (CaseBox qty box ret body loc) =
CaseBox qty <$> dtightenE p box CaseBox qty <$> dtightenE p box
<*> dtightenS p ret <*> dtightenS p ret
<*> dtightenS p body <*> dtightenS p body
dtightenE p (fun :% arg) = [|dtightenE p fun :% tighten p arg|] <*> pure loc
dtightenE p (tm :# ty) = [|dtightenT p tm :# dtightenT p ty|] dtightenE p (DApp fun arg loc) =
dtightenE p (Coe ty q0 q1 val) = DApp <$> dtightenE p fun <*> tighten p arg <*> pure loc
[|Coe (dtightenDS p ty) (tighten p q0) (tighten p q1) (dtightenT p val)|] dtightenE p (Ann tm ty loc) =
dtightenE p (Comp ty q0 q1 val r zero one) = Ann <$> dtightenT p tm <*> dtightenT p ty <*> pure loc
dtightenE p (Coe ty q0 q1 val loc) =
[|Coe (dtightenDS p ty) (tighten p q0) (tighten p q1) (dtightenT p val)
(pure loc)|]
dtightenE p (Comp ty q0 q1 val r zero one loc) =
[|Comp (dtightenT p ty) (tighten p q0) (tighten p q1) [|Comp (dtightenT p ty) (tighten p q0) (tighten p q1)
(dtightenT p val) (tighten p r) (dtightenT p val) (tighten p r)
(dtightenDS p zero) (dtightenDS p one)|] (dtightenDS p zero) (dtightenDS p one) (pure loc)|]
dtightenE p (TypeCase ty ret arms def) = dtightenE p (TypeCase ty ret arms def loc) =
[|TypeCase (dtightenE p ty) (dtightenT p ret) [|TypeCase (dtightenE p ty) (dtightenT p ret)
(traverse (dtightenS p) arms) (dtightenT p def)|] (traverse (dtightenS p) arms) (dtightenT p def) (pure loc)|]
dtightenE p (CloE (Sub el th)) = do dtightenE p (CloE (Sub el th)) = do
el <- dtightenE p el el <- dtightenE p el
th <- assert_total $ traverse (dtightenE p) th th <- assert_total $ traverse (dtightenE p) th
@ -226,46 +271,55 @@ export [ElimD] Tighten (\d => Elim d n) where tighten p e = dtightenE p e
-- versions of SY, etc, that try to tighten and use SN automatically -- versions of SY, etc, that try to tighten and use SN automatically
public export public export
ST : Tighten f => {s : Nat} -> NContext s -> f (s + n) -> Scoped s f n ST : Tighten f => {s : Nat} -> BContext s -> f (s + n) -> Scoped s f n
ST names body = ST names body =
case tightenN s body of case tightenN s body of
Just body => S names $ N body Just body => S names $ N body
Nothing => S names $ Y body Nothing => S names $ Y body
public export public export
DST : {s : Nat} -> NContext s -> Term (s + d) n -> DScopeTermN s d n DST : {s : Nat} -> BContext s -> Term (s + d) n -> DScopeTermN s d n
DST names body = DST names body =
case tightenN @{TermD} s body of case tightenN @{TermD} s body of
Just body => S names $ N body Just body => S names $ N body
Nothing => S names $ Y body Nothing => S names $ Y body
public export %inline public export %inline
PiT : (qty : Qty) -> (x : BaseName) -> PiT : (qty : Qty) -> (x : BindName) ->
(arg : Term d n) -> (res : Term d (S n)) -> Term d n (arg : Term d n) -> (res : Term d (S n)) -> (loc : Loc) -> Term d n
PiT {qty, x, arg, res} = Pi {qty, arg, res = ST [< x] res} PiT {qty, x, arg, res, loc} = Pi {qty, arg, res = ST [< x] res, loc}
public export %inline public export %inline
SigT : (x : BaseName) -> (fst : Term d n) -> LamT : (x : BindName) -> (body : Term d (S n)) -> (loc : Loc) -> Term d n
(snd : Term d (S n)) -> Term d n LamT {x, body, loc} = Lam {body = ST [< x] body, loc}
SigT {x, fst, snd} = Sig {fst, snd = ST [< x] snd}
public export %inline public export %inline
EqT : (i : BaseName) -> (ty : Term (S d) n) -> SigT : (x : BindName) -> (fst : Term d n) ->
(l, r : Term d n) -> Term d n (snd : Term d (S n)) -> (loc : Loc) -> Term d n
EqT {i, ty, l, r} = Eq {ty = DST [< i] ty, l, r} SigT {x, fst, snd, loc} = Sig {fst, snd = ST [< x] snd, loc}
public export %inline public export %inline
CoeT : (i : BaseName) -> (ty : Term (S d) n) -> EqT : (i : BindName) -> (ty : Term (S d) n) ->
(p, q : Dim d) -> (val : Term d n) -> Elim d n (l, r : Term d n) -> (loc : Loc) -> Term d n
CoeT {i, ty, p, q, val} = Coe {ty = DST [< i] ty, p, q, val} EqT {i, ty, l, r, loc} = Eq {ty = DST [< i] ty, l, r, loc}
public export %inline
DLamT : (i : BindName) -> (body : Term (S d) n) -> (loc : Loc) -> Term d n
DLamT {i, body, loc} = DLam {body = DST [< i] body, loc}
public export %inline
CoeT : (i : BindName) -> (ty : Term (S d) n) ->
(p, q : Dim d) -> (val : Term d n) -> (loc : Loc) -> Elim d n
CoeT {i, ty, p, q, val, loc} = Coe {ty = DST [< i] ty, p, q, val, loc}
public export %inline public export %inline
typeCase1T : Elim d n -> Term d n -> typeCase1T : Elim d n -> Term d n ->
(k : TyConKind) -> NContext (arity k) -> Term d (arity k + n) -> (k : TyConKind) -> BContext (arity k) -> Term d (arity k + n) ->
{default Nat def : Term d n} -> (loc : Loc) ->
{default (Nat loc) def : Term d n} ->
Elim d n Elim d n
typeCase1T ty ret k ns body {def} = typeCase1T ty ret k ns body loc {def} =
typeCase ty ret [(k ** ST ns body)] def typeCase ty ret [(k ** ST ns body)] def loc
export export

View File

@ -1,6 +1,7 @@
module Quox.Syntax.Var module Quox.Syntax.Var
import Quox.Name import public Quox.Loc
import public Quox.Name
import Quox.Pretty import Quox.Pretty
import Quox.OPE import Quox.OPE
@ -42,6 +43,23 @@ export Uninhabited (VZ = VS i) where uninhabited _ impossible
export Uninhabited (VS i = VZ) where uninhabited _ impossible export Uninhabited (VS i = VZ) where uninhabited _ impossible
public export
data Eqv : Var m -> Var n -> Type where
EZ : VZ `Eqv` VZ
ES : i `Eqv` j -> VS i `Eqv` VS j
%name Var.Eqv e
export
decEqv : Dec2 Eqv
decEqv VZ VZ = Yes EZ
decEqv VZ (VS i) = No $ \case _ impossible
decEqv (VS i) VZ = No $ \case _ impossible
decEqv (VS i) (VS j) =
case decEqv i j of
Yes y => Yes $ ES y
No n => No $ \(ES y) => n y
private private
lookupS : Nat -> SnocList a -> Maybe a lookupS : Nat -> SnocList a -> Maybe a
lookupS _ [<] = Nothing lookupS _ [<] = Nothing
@ -148,9 +166,13 @@ weakIsSpec p i = toNatInj $ trans (weakCorrect p i) (sym $ weakSpecCorrect p i)
public export public export
interface FromVar f where %inline fromVar : Var n -> f n interface FromVar f where %inline fromVarLoc : Var n -> Loc -> f n
public export FromVar Var where fromVar = id public export %inline
fromVar : FromVar f => Var n -> {default noLoc loc : Loc} -> f n
fromVar x = fromVarLoc x loc
public export FromVar Var where fromVarLoc x _ = x
export export
tabulateV : {0 tm : Nat -> Type} -> (forall n. Var n -> tm n) -> tabulateV : {0 tm : Nat -> Type} -> (forall n. Var n -> tm n) ->

View File

@ -13,26 +13,30 @@ import Quox.EffExtra
public export public export
0 TCEff : List (Type -> Type) 0 TCEff : List (Type -> Type)
TCEff = [ErrorEff, DefsReader] TCEff = [ErrorEff, DefsReader, NameGen]
public export public export
0 TC : Type -> Type 0 TC : Type -> Type
TC = Eff TCEff TC = Eff TCEff
export export
runTC : Definitions -> TC a -> Either Error a runTCWith : NameSuf -> Definitions -> TC a -> (Either Error a, NameSuf)
runTC defs = runTCWith = runEqualWith
extract . runExcept . runReaderAt DEFS defs
export export
popQs : Has ErrorEff fs => QContext s -> QOutput (s + n) -> Eff fs (QOutput n) runTC : Definitions -> TC a -> Either Error a
popQs [<] qout = pure qout runTC = runEqual
popQs (pis :< pi) (qout :< rh) = do expectCompatQ rh pi; popQs pis qout
export %inline
popQ : Has ErrorEff fs => Qty -> QOutput (S n) -> Eff fs (QOutput n) parameters (loc : Loc)
popQ pi = popQs [< pi] export
popQs : Has ErrorEff fs => QContext s -> QOutput (s + n) -> Eff fs (QOutput n)
popQs [<] qout = pure qout
popQs (pis :< pi) (qout :< rh) = do expectCompatQ loc rh pi; popQs pis qout
export %inline
popQ : Has ErrorEff fs => Qty -> QOutput (S n) -> Eff fs (QOutput n)
popQ pi = popQs [< pi]
export export
lubs1 : List1 (QOutput n) -> Maybe (QOutput n) lubs1 : List1 (QOutput n) -> Maybe (QOutput n)
@ -47,29 +51,32 @@ lubs ctx [] = Just $ zeroFor ctx
lubs ctx (x :: xs) = lubs1 $ x ::: xs lubs ctx (x :: xs) = lubs1 $ x ::: xs
||| context extension with no names or quantities
private
CtxExtension0' : Nat -> Nat -> Nat -> Type
CtxExtension0' s d n = Context (Term d . (+ n)) s
private
addNames0 : Context (Term d . (+ n)) s -> NContext s -> CtxExtension d n (s + n)
addNames0 [<] [<] = [<]
addNames0 (ts :< t) (xs :< x) = addNames0 ts xs :< (Zero, x, t)
export export
typecaseTel : (k : TyConKind) -> Universe -> CtxExtension0' (arity k) d n typecaseTel : (k : TyConKind) -> BContext (arity k) -> Universe ->
typecaseTel k u = case k of CtxExtension d n (arity k + n)
typecaseTel k xs u = case k of
KTYPE => [<] KTYPE => [<]
-- A : ★ᵤ, B : 0.A → ★ᵤ -- A : ★ᵤ, B : 0.A → ★ᵤ
KPi => [< TYPE u, Arr Zero (BVT 0) (TYPE u)] KPi =>
KSig => [< TYPE u, Arr Zero (BVT 0) (TYPE u)] let [< a, b] = xs in
[< (Zero, a, TYPE u a.loc),
(Zero, b, Arr Zero (BVT 0 b.loc) (TYPE u b.loc) b.loc)]
KSig =>
let [< a, b] = xs in
[< (Zero, a, TYPE u a.loc),
(Zero, b, Arr Zero (BVT 0 b.loc) (TYPE u b.loc) b.loc)]
KEnum => [<] KEnum => [<]
-- A₀ : ★ᵤ, A₁ : ★ᵤ, A : (A₀ ≡ A₁ : ★ᵤ), L : A₀, R : A₀ -- A₀ : ★ᵤ, A₁ : ★ᵤ, A : (A₀ ≡ A₁ : ★ᵤ), L : A₀, R : A₀
KEq => [< TYPE u, TYPE u, Eq0 (TYPE u) (BVT 1) (BVT 0), BVT 2, BVT 2] KEq =>
let [< a0, a1, a, l, r] = xs in
[< (Zero, a0, TYPE u a0.loc),
(Zero, a1, TYPE u a1.loc),
(Zero, a, Eq0 (TYPE u a.loc) (BVT 1 a.loc) (BVT 0 a.loc) a.loc),
(Zero, l, BVT 2 l.loc),
(Zero, r, BVT 2 r.loc)]
KNat => [<] KNat => [<]
-- A : ★ᵤ -- A : ★ᵤ
KBOX => [< TYPE u] KBOX => let [< a] = xs in [< (Zero, a, TYPE u a.loc)]
mutual mutual
@ -149,8 +156,8 @@ mutual
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n -> (subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
TC (CheckResult' n) TC (CheckResult' n)
toCheckType ctx sg t ty = do toCheckType ctx sg t ty = do
u <- expectTYPE !defs ctx ty u <- expectTYPE !defs ctx ty.loc ty
expectEqualQ Zero sg.fst expectEqualQ t.loc Zero sg.fst
checkTypeNoWrap ctx t (Just u) checkTypeNoWrap ctx t (Just u)
pure $ zeroFor ctx pure $ zeroFor ctx
@ -159,70 +166,70 @@ mutual
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n -> (subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
TC (CheckResult' n) TC (CheckResult' n)
check' ctx sg t@(TYPE _) ty = toCheckType ctx sg t ty check' ctx sg t@(TYPE {}) ty = toCheckType ctx sg t ty
check' ctx sg t@(Pi {}) ty = toCheckType ctx sg t ty check' ctx sg t@(Pi {}) ty = toCheckType ctx sg t ty
check' ctx sg (Lam body) ty = do check' ctx sg (Lam body loc) ty = do
(qty, arg, res) <- expectPi !defs ctx ty (qty, arg, res) <- expectPi !defs ctx ty.loc ty
-- if Ψ | Γ, x : A ⊢ σ · t ⇐ B ⊳ Σ, ρ·x -- if Ψ | Γ, x : A ⊢ σ · t ⇐ B ⊳ Σ, ρ·x
-- with ρ ≤ σπ -- with ρ ≤ σπ
let qty' = sg.fst * qty let qty' = sg.fst * qty
qout <- checkC (extendTy qty' body.name arg ctx) sg body.term res.term qout <- checkC (extendTy qty' body.name arg ctx) sg body.term res.term
-- then Ψ | Γ ⊢ σ · (λx ⇒ t) ⇐ (π·x : A) → B ⊳ Σ -- then Ψ | Γ ⊢ σ · (λx ⇒ t) ⇐ (π·x : A) → B ⊳ Σ
popQ qty' qout popQ loc qty' qout
check' ctx sg t@(Sig {}) ty = toCheckType ctx sg t ty check' ctx sg t@(Sig {}) ty = toCheckType ctx sg t ty
check' ctx sg (Pair fst snd) ty = do check' ctx sg (Pair fst snd loc) ty = do
(tfst, tsnd) <- expectSig !defs ctx ty (tfst, tsnd) <- expectSig !defs ctx ty.loc ty
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁
qfst <- checkC ctx sg fst tfst qfst <- checkC ctx sg fst tfst
let tsnd = sub1 tsnd (fst :# tfst) let tsnd = sub1 tsnd (Ann fst tfst fst.loc)
-- if Ψ | Γ ⊢ σ · t ⇐ B[s] ⊳ Σ₂ -- if Ψ | Γ ⊢ σ · t ⇐ B[s] ⊳ Σ₂
qsnd <- checkC ctx sg snd tsnd qsnd <- checkC ctx sg snd tsnd
-- then Ψ | Γ ⊢ σ · (s, t) ⇐ (x : A) × B ⊳ Σ₁ + Σ₂ -- then Ψ | Γ ⊢ σ · (s, t) ⇐ (x : A) × B ⊳ Σ₁ + Σ₂
pure $ qfst + qsnd pure $ qfst + qsnd
check' ctx sg t@(Enum _) ty = toCheckType ctx sg t ty check' ctx sg t@(Enum {}) ty = toCheckType ctx sg t ty
check' ctx sg (Tag t) ty = do check' ctx sg (Tag t loc) ty = do
tags <- expectEnum !defs ctx ty tags <- expectEnum !defs ctx ty.loc ty
-- if t ∈ ts -- if t ∈ ts
unless (t `elem` tags) $ throw $ TagNotIn t tags unless (t `elem` tags) $ throw $ TagNotIn loc t tags
-- then Ψ | Γ ⊢ σ · t ⇐ {ts} ⊳ 𝟎 -- then Ψ | Γ ⊢ σ · t ⇐ {ts} ⊳ 𝟎
pure $ zeroFor ctx pure $ zeroFor ctx
check' ctx sg t@(Eq {}) ty = toCheckType ctx sg t ty check' ctx sg t@(Eq {}) ty = toCheckType ctx sg t ty
check' ctx sg (DLam body) ty = do check' ctx sg (DLam body loc) ty = do
(ty, l, r) <- expectEq !defs ctx ty (ty, l, r) <- expectEq !defs ctx ty.loc ty
let ctx' = extendDim body.name ctx let ctx' = extendDim body.name ctx
ty = ty.term ty = ty.term
body = body.term body = body.term
-- if Ψ, i | Γ ⊢ σ · t ⇐ A ⊳ Σ -- if Ψ, i | Γ ⊢ σ · t ⇐ A ⊳ Σ
qout <- checkC ctx' sg body ty qout <- checkC ctx' sg body ty
-- if Ψ, i, i = 0 | Γ ⊢ t = l : A -- if Ψ, i, i = 0 | Γ ⊢ t = l : A
equal (eqDim (BV 0) (K Zero) ctx') ty body (dweakT 1 l) lift $ equal loc (eqDim (B VZ loc) (K Zero loc) ctx') ty body (dweakT 1 l)
-- if Ψ, i, i = 1 | Γ ⊢ t = r : A -- if Ψ, i, i = 1 | Γ ⊢ t = r : A
equal (eqDim (BV 0) (K One) ctx') ty body (dweakT 1 r) lift $ equal loc (eqDim (B VZ loc) (K One loc) ctx') ty body (dweakT 1 r)
-- then Ψ | Γ ⊢ σ · (δ i ⇒ t) ⇐ Eq [i ⇒ A] l r ⊳ Σ -- then Ψ | Γ ⊢ σ · (δ i ⇒ t) ⇐ Eq [i ⇒ A] l r ⊳ Σ
pure qout pure qout
check' ctx sg Nat ty = toCheckType ctx sg Nat ty check' ctx sg t@(Nat {}) ty = toCheckType ctx sg t ty
check' ctx sg Zero ty = do check' ctx sg (Zero {}) ty = do
expectNat !defs ctx ty expectNat !defs ctx ty.loc ty
pure $ zeroFor ctx pure $ zeroFor ctx
check' ctx sg (Succ n) ty = do check' ctx sg (Succ n {}) ty = do
expectNat !defs ctx ty expectNat !defs ctx ty.loc ty
checkC ctx sg n Nat checkC ctx sg n ty
check' ctx sg t@(BOX {}) ty = toCheckType ctx sg t ty check' ctx sg t@(BOX {}) ty = toCheckType ctx sg t ty
check' ctx sg (Box val) ty = do check' ctx sg (Box val loc) ty = do
(q, ty) <- expectBOX !defs ctx ty (q, ty) <- expectBOX !defs ctx ty.loc ty
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ -- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ
valout <- checkC ctx sg val ty valout <- checkC ctx sg val ty
-- then Ψ | Γ ⊢ σ · [s] ⇐ [π.A] ⊳ πΣ -- then Ψ | Γ ⊢ σ · [s] ⇐ [π.A] ⊳ πΣ
@ -232,7 +239,7 @@ mutual
-- if Ψ | Γ ⊢ σ · e ⇒ A' ⊳ Σ -- if Ψ | Γ ⊢ σ · e ⇒ A' ⊳ Σ
infres <- inferC ctx sg e infres <- inferC ctx sg e
-- if Ψ | Γ ⊢ A' <: A -- if Ψ | Γ ⊢ A' <: A
subtype ctx infres.type ty lift $ subtype e.loc ctx infres.type ty
-- then Ψ | Γ ⊢ σ · e ⇐ A ⊳ Σ -- then Ψ | Γ ⊢ σ · e ⇐ A ⊳ Σ
pure infres.qout pure infres.qout
@ -241,13 +248,13 @@ mutual
(subj : Term d n) -> (0 nc : NotClo subj) => (subj : Term d n) -> (0 nc : NotClo subj) =>
Maybe Universe -> TC () Maybe Universe -> TC ()
checkType' ctx (TYPE k) u = do checkType' ctx (TYPE k loc) u = do
-- if 𝓀 < then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type -- if 𝓀 < then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type
case u of case u of
Just l => unless (k < l) $ throw $ BadUniverse k l Just l => unless (k < l) $ throw $ BadUniverse loc k l
Nothing => pure () Nothing => pure ()
checkType' ctx (Pi qty arg res) u = do checkType' ctx (Pi qty arg res _) u = do
-- if Ψ | Γ ⊢₀ A ⇐ Type -- if Ψ | Γ ⊢₀ A ⇐ Type
checkTypeC ctx arg u checkTypeC ctx arg u
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type -- if Ψ | Γ, x : A ⊢₀ B ⇐ Type
@ -255,9 +262,9 @@ mutual
-- then Ψ | Γ ⊢₀ (π·x : A) → B ⇐ Type -- then Ψ | Γ ⊢₀ (π·x : A) → B ⇐ Type
checkType' ctx t@(Lam {}) u = checkType' ctx t@(Lam {}) u =
throw $ NotType ctx t throw $ NotType t.loc ctx t
checkType' ctx (Sig fst snd) u = do checkType' ctx (Sig fst snd _) u = do
-- if Ψ | Γ ⊢₀ A ⇐ Type -- if Ψ | Γ ⊢₀ A ⇐ Type
checkTypeC ctx fst u checkTypeC ctx fst u
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type -- if Ψ | Γ, x : A ⊢₀ B ⇐ Type
@ -265,15 +272,15 @@ mutual
-- then Ψ | Γ ⊢₀ (x : A) × B ⇐ Type -- then Ψ | Γ ⊢₀ (x : A) × B ⇐ Type
checkType' ctx t@(Pair {}) u = checkType' ctx t@(Pair {}) u =
throw $ NotType ctx t throw $ NotType t.loc ctx t
checkType' ctx (Enum _) u = pure () checkType' ctx (Enum {}) u = pure ()
-- Ψ | Γ ⊢₀ {ts} ⇐ Type -- Ψ | Γ ⊢₀ {ts} ⇐ Type
checkType' ctx t@(Tag {}) u = checkType' ctx t@(Tag {}) u =
throw $ NotType ctx t throw $ NotType t.loc ctx t
checkType' ctx (Eq t l r) u = do checkType' ctx (Eq t l r _) u = do
-- if Ψ, i | Γ ⊢₀ A ⇐ Type -- if Ψ, i | Γ ⊢₀ A ⇐ Type
case t.body of case t.body of
Y t' => checkTypeC (extendDim t.name ctx) t' u Y t' => checkTypeC (extendDim t.name ctx) t' u
@ -285,50 +292,31 @@ mutual
-- then Ψ | Γ ⊢₀ Eq [i ⇒ A] l r ⇐ Type -- then Ψ | Γ ⊢₀ Eq [i ⇒ A] l r ⇐ Type
checkType' ctx t@(DLam {}) u = checkType' ctx t@(DLam {}) u =
throw $ NotType ctx t throw $ NotType t.loc ctx t
checkType' ctx Nat u = pure () checkType' ctx (Nat {}) u = pure ()
checkType' ctx Zero u = throw $ NotType ctx Zero checkType' ctx t@(Zero {}) u = throw $ NotType t.loc ctx t
checkType' ctx t@(Succ _) u = throw $ NotType ctx t checkType' ctx t@(Succ {}) u = throw $ NotType t.loc ctx t
checkType' ctx (BOX q ty) u = checkType ctx ty u checkType' ctx (BOX q ty _) u = checkType ctx ty u
checkType' ctx t@(Box _) u = throw $ NotType ctx t checkType' ctx t@(Box {}) u = throw $ NotType t.loc ctx t
checkType' ctx (E e) u = do checkType' ctx (E e) u = do
-- if Ψ | Γ ⊢₀ E ⇒ Type -- if Ψ | Γ ⊢₀ E ⇒ Type
infres <- inferC ctx szero e infres <- inferC ctx szero e
-- if Ψ | Γ ⊢ Type <: Type 𝓀 -- if Ψ | Γ ⊢ Type <: Type 𝓀
case u of case u of
Just u => subtype ctx infres.type (TYPE u) Just u => lift $ subtype e.loc ctx infres.type (TYPE u noLoc)
Nothing => ignore $ Nothing => ignore $ expectTYPE !defs ctx e.loc infres.type
expectTYPE !defs ctx infres.type
-- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀 -- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀
private covering
check0ScopeN : {s : Nat} ->
TyContext d n -> CtxExtension0' s d n ->
ScopeTermN s d n -> Term d n -> TC ()
check0ScopeN ctx ext (S _ (N body)) ty = check0 ctx body ty
check0ScopeN ctx ext (S names (Y body)) ty =
check0 (extendTyN (addNames0 ext names) ctx) body (weakT s ty)
private covering
check0Scope : TyContext d n -> Term d n ->
ScopeTerm d n -> Term d n -> TC ()
check0Scope ctx t = check0ScopeN ctx [< t]
private covering
checkTypeScopeN : TyContext d n -> CtxExtension0' s d n ->
ScopeTermN s d n -> Maybe Universe -> TC ()
checkTypeScopeN ctx ext (S _ (N body)) u = checkType ctx body u
checkTypeScopeN ctx ext (S names (Y body)) u =
checkType (extendTyN (addNames0 ext names) ctx) body u
private covering private covering
checkTypeScope : TyContext d n -> Term d n -> checkTypeScope : TyContext d n -> Term d n ->
ScopeTerm d n -> Maybe Universe -> TC () ScopeTerm d n -> Maybe Universe -> TC ()
checkTypeScope ctx s = checkTypeScopeN ctx [< s] checkTypeScope ctx s (S _ (N body)) u = checkType ctx body u
checkTypeScope ctx s (S [< x] (Y body)) u =
checkType (extendTy Zero x s ctx) body u
private covering private covering
@ -336,20 +324,20 @@ mutual
(subj : Elim d n) -> (0 nc : NotClo subj) => (subj : Elim d n) -> (0 nc : NotClo subj) =>
TC (InferResult' d n) TC (InferResult' d n)
infer' ctx sg (F x) = do infer' ctx sg (F x loc) = do
-- if π·x : A {≔ s} in global context -- if π·x : A {≔ s} in global context
g <- lookupFree x !defs g <- lookupFree x loc !defs
-- if σ ≤ π -- if σ ≤ π
expectCompatQ sg.fst g.qty.fst expectCompatQ loc sg.fst g.qty.fst
-- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎 -- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎
let Val d = ctx.dimLen; Val n = ctx.termLen let Val d = ctx.dimLen; Val n = ctx.termLen
pure $ InfRes {type = g.type, qout = zeroFor ctx} pure $ InfRes {type = g.type, qout = zeroFor ctx}
infer' ctx sg (B i) = infer' ctx sg (B i _) =
-- if x : A ∈ Γ -- if x : A ∈ Γ
-- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ (𝟎, σ·x, 𝟎) -- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ (𝟎, σ·x, 𝟎)
pure $ lookupBound sg.fst i ctx.tctx pure $ lookupBound sg.fst i ctx.tctx
where where
lookupBound : forall n. Qty -> Var n -> TContext d n -> InferResult' d n lookupBound : forall n. Qty -> Var n -> TContext d n -> InferResult' d n
lookupBound pi VZ (ctx :< type) = lookupBound pi VZ (ctx :< type) =
InfRes {type = weakT 1 type, qout = zeroFor ctx :< pi} InfRes {type = weakT 1 type, qout = zeroFor ctx :< pi}
@ -357,19 +345,19 @@ mutual
let InfRes {type, qout} = lookupBound pi i ctx in let InfRes {type, qout} = lookupBound pi i ctx in
InfRes {type = weakT 1 type, qout = qout :< Zero} InfRes {type = weakT 1 type, qout = qout :< Zero}
infer' ctx sg (fun :@ arg) = do infer' ctx sg (App fun arg loc) = do
-- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁
funres <- inferC ctx sg fun funres <- inferC ctx sg fun
(qty, argty, res) <- expectPi !defs ctx funres.type (qty, argty, res) <- expectPi !defs ctx fun.loc funres.type
-- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂ -- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂
argout <- checkC ctx (subjMult sg qty) arg argty argout <- checkC ctx (subjMult sg qty) arg argty
-- then Ψ | Γ ⊢ σ · f s ⇒ B[s] ⊳ Σ₁ + πΣ₂ -- then Ψ | Γ ⊢ σ · f s ⇒ B[s] ⊳ Σ₁ + πΣ₂
pure $ InfRes { pure $ InfRes {
type = sub1 res $ arg :# argty, type = sub1 res $ Ann arg argty arg.loc,
qout = funres.qout + qty * argout qout = funres.qout + qty * argout
} }
infer' ctx sg (CasePair pi pair ret body) = do infer' ctx sg (CasePair pi pair ret body loc) = do
-- no check for 1 ≤ π, since pairs have a single constructor. -- no check for 1 ≤ π, since pairs have a single constructor.
-- e.g. at 0 the components are also 0 in the body -- e.g. at 0 the components are also 0 in the body
-- --
@ -377,27 +365,28 @@ mutual
pairres <- inferC ctx sg pair pairres <- inferC ctx sg pair
-- if Ψ | Γ, p : (x : A) × B ⊢₀ ret ⇐ Type -- if Ψ | Γ, p : (x : A) × B ⊢₀ ret ⇐ Type
checkTypeC (extendTy Zero ret.name pairres.type ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name pairres.type ctx) ret.term Nothing
(tfst, tsnd) <- expectSig !defs ctx pairres.type (tfst, tsnd) <- expectSig !defs ctx pair.loc pairres.type
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐ -- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐
-- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y -- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y
-- with ρ₁, ρ₂ ≤ πσ -- with ρ₁, ρ₂ ≤ πσ
let [< x, y] = body.names let [< x, y] = body.names
pisg = pi * sg.fst pisg = pi * sg.fst
bodyctx = extendTyN [< (pisg, x, tfst), (pisg, y, tsnd.term)] ctx bodyctx = extendTyN [< (pisg, x, tfst), (pisg, y, tsnd.term)] ctx
bodyty = substCasePairRet pairres.type ret bodyty = substCasePairRet body.names pairres.type ret
bodyout <- checkC bodyctx sg body.term bodyty >>= popQs [< pisg, pisg] bodyout <- checkC bodyctx sg body.term bodyty >>=
popQs loc [< pisg, pisg]
-- 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 + bodyout
} }
infer' ctx sg (CaseEnum pi t ret arms) {d, n} = do infer' ctx sg (CaseEnum pi t ret arms loc) {d, n} = do
-- if Ψ | Γ ⊢ σ · t ⇒ {ts} ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · t ⇒ {ts} ⊳ Σ₁
tres <- inferC ctx sg t tres <- inferC ctx sg t
ttags <- expectEnum !defs ctx tres.type ttags <- expectEnum !defs ctx t.loc tres.type
-- if 1 ≤ π, OR there is only zero or one option -- if 1 ≤ π, OR there is only zero or one option
unless (length (SortedSet.toList ttags) <= 1) $ expectCompatQ One pi unless (length (SortedSet.toList ttags) <= 1) $ expectCompatQ loc One pi
-- if Ψ | Γ, x : {ts} ⊢₀ A ⇐ Type -- if Ψ | Γ, x : {ts} ⊢₀ A ⇐ Type
checkTypeC (extendTy Zero ret.name tres.type ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name tres.type ctx) ret.term Nothing
-- if for each "a ⇒ s" in arms, -- if for each "a ⇒ s" in arms,
@ -405,109 +394,109 @@ mutual
-- with Σ₂ = lubs Σᵢ -- with Σ₂ = lubs Σᵢ
let arms = SortedMap.toList arms let arms = SortedMap.toList arms
let armTags = SortedSet.fromList $ map fst arms let armTags = SortedSet.fromList $ map fst arms
unless (ttags == armTags) $ throw $ BadCaseEnum ttags armTags unless (ttags == armTags) $ throw $ BadCaseEnum loc ttags armTags
armres <- for arms $ \(a, s) => armres <- for arms $ \(a, s) =>
checkC ctx sg s (sub1 ret (Tag a :# tres.type)) checkC ctx sg s $ sub1 ret $ Ann (Tag a s.loc) tres.type s.loc
let Just armout = lubs ctx armres let Just armout = lubs ctx armres
| _ => throw $ BadQtys "case arms" ctx $ | _ => throw $ BadQtys loc "case arms" ctx $
zipWith (\qs, (t, rhs) => (qs, Tag t)) armres arms zipWith (\qs, (t, rhs) => (qs, Tag t noLoc)) armres arms
pure $ InfRes { pure $ InfRes {
type = sub1 ret t, type = sub1 ret t,
qout = pi * tres.qout + armout qout = pi * tres.qout + armout
} }
infer' ctx sg (CaseNat pi pi' n ret zer suc) = do infer' ctx sg (CaseNat pi pi' n ret zer suc loc) = do
-- if 1 ≤ π -- if 1 ≤ π
expectCompatQ One pi expectCompatQ loc One pi
-- if Ψ | Γ ⊢ σ · n ⇒ ⊳ Σn -- if Ψ | Γ ⊢ σ · n ⇒ ⊳ Σn
nres <- inferC ctx sg n nres <- inferC ctx sg n
expectNat !defs ctx nres.type let nat = nres.type
expectNat !defs ctx n.loc nat
-- if Ψ | Γ, n : ⊢₀ A ⇐ Type -- if Ψ | Γ, n : ⊢₀ A ⇐ Type
checkTypeC (extendTy Zero ret.name Nat ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name nat ctx) ret.term Nothing
-- if Ψ | Γ ⊢ σ · zer ⇐ A[0 ∷ /n] ⊳ Σz -- if Ψ | Γ ⊢ σ · zer ⇐ A[0 ∷ /n] ⊳ Σz
zerout <- checkC ctx sg zer (sub1 ret (Zero :# Nat)) zerout <- checkC ctx sg zer $ sub1 ret $ Ann (Zero zer.loc) nat zer.loc
-- if Ψ | Γ, n : , ih : A ⊢ σ · suc ⇐ A[succ p ∷ /n] ⊳ Σs, ρ₁.p, ρ₂.ih -- if Ψ | Γ, n : , ih : A ⊢ σ · suc ⇐ A[succ p ∷ /n] ⊳ Σs, ρ₁.p, ρ₂.ih
-- with ρ₂ ≤ π'σ, (ρ₁ + ρ₂) ≤ πσ -- with ρ₂ ≤ π'σ, (ρ₁ + ρ₂) ≤ πσ
let [< p, ih] = suc.names let [< p, ih] = suc.names
pisg = pi * sg.fst pisg = pi * sg.fst
sucCtx = extendTyN [< (pisg, p, Nat), (pi', ih, ret.term)] ctx sucCtx = extendTyN [< (pisg, p, Nat p.loc), (pi', ih, ret.term)] ctx
sucType = substCaseSuccRet ret sucType = substCaseSuccRet suc.names ret
sucout :< qp :< qih <- checkC sucCtx sg suc.term sucType sucout :< qp :< qih <- checkC sucCtx sg suc.term sucType
expectCompatQ qih (pi' * sg.fst) expectCompatQ loc qih (pi' * sg.fst)
-- [fixme] better error here -- [fixme] better error here
expectCompatQ (qp + qih) pisg expectCompatQ loc (qp + qih) pisg
-- then Ψ | Γ ⊢ caseπ ⋯ ⇒ A[n] ⊳ πΣn + Σz + ωΣs -- then Ψ | Γ ⊢ caseπ ⋯ ⇒ A[n] ⊳ πΣn + Σz + ωΣs
pure $ InfRes { pure $ InfRes {
type = sub1 ret n, type = sub1 ret n,
qout = pi * nres.qout + zerout + Any * sucout qout = pi * nres.qout + zerout + Any * sucout
} }
infer' ctx sg (CaseBox pi box ret body) = do infer' ctx sg (CaseBox pi box ret body loc) = do
-- if Ψ | Γ ⊢ σ · b ⇒ [ρ.A] ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · b ⇒ [ρ.A] ⊳ Σ₁
boxres <- inferC ctx sg box boxres <- inferC ctx sg box
(q, ty) <- expectBOX !defs ctx boxres.type (q, ty) <- expectBOX !defs ctx box.loc boxres.type
-- if Ψ | Γ, x : [ρ.A] ⊢₀ R ⇐ Type -- if Ψ | Γ, x : [ρ.A] ⊢₀ R ⇐ Type
checkTypeC (extendTy Zero ret.name boxres.type ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name boxres.type ctx) ret.term Nothing
-- if Ψ | Γ, x : A ⊢ t ⇐ R[[x] ∷ [ρ.A/x]] ⊳ Σ₂, ς·x -- if Ψ | Γ, x : A ⊢ t ⇐ R[[x] ∷ [ρ.A/x]] ⊳ Σ₂, ς·x
-- with ς ≤ ρπσ -- with ς ≤ ρπσ
let qpisg = q * pi * sg.fst let qpisg = q * pi * sg.fst
bodyCtx = extendTy qpisg body.name ty ctx bodyCtx = extendTy qpisg body.name ty ctx
bodyType = substCaseBoxRet ty ret bodyType = substCaseBoxRet body.name ty ret
bodyout <- checkC bodyCtx sg body.term bodyType >>= popQ qpisg bodyout <- checkC bodyCtx sg body.term bodyType >>= popQ loc qpisg
-- then Ψ | Γ ⊢ caseπ ⋯ ⇒ R[b/x] ⊳ Σ₁ + Σ₂ -- then Ψ | Γ ⊢ caseπ ⋯ ⇒ R[b/x] ⊳ Σ₁ + Σ₂
pure $ InfRes { pure $ InfRes {
type = sub1 ret box, type = sub1 ret box,
qout = boxres.qout + bodyout qout = boxres.qout + bodyout
} }
infer' ctx sg (fun :% dim) = do infer' ctx sg (DApp fun dim loc) = do
-- if Ψ | Γ ⊢ σ · f ⇒ Eq [𝑖 ⇒ A] l r ⊳ Σ -- if Ψ | Γ ⊢ σ · f ⇒ Eq [𝑖 ⇒ A] l r ⊳ Σ
InfRes {type, qout} <- inferC ctx sg fun InfRes {type, qout} <- inferC ctx sg fun
ty <- fst <$> expectEq !defs ctx type ty <- fst <$> expectEq !defs ctx fun.loc type
-- then Ψ | Γ ⊢ σ · f p ⇒ Ap/𝑖 ⊳ Σ -- then Ψ | Γ ⊢ σ · f p ⇒ Ap/𝑖 ⊳ Σ
pure $ InfRes {type = dsub1 ty dim, qout} pure $ InfRes {type = dsub1 ty dim, qout}
infer' ctx sg (Coe (S [< i] ty) p q val) = do infer' ctx sg (Coe ty p q val loc) = do
let ty = ty.term checkType (extendDim ty.name ctx) ty.term Nothing
checkType (extendDim i ctx) ty Nothing qout <- checkC ctx sg val $ dsub1 ty p
qout <- checkC ctx sg val (ty // one p) pure $ InfRes {type = dsub1 ty q, qout}
pure $ InfRes {type = ty // one q, qout}
infer' ctx sg (Comp ty p q val r (S [< j0] val0) (S [< j1] val1)) = do infer' ctx sg (Comp ty p q val r (S [< j0] val0) (S [< j1] val1) loc) = do
checkType ctx ty Nothing checkType ctx ty Nothing
qout <- checkC ctx sg val ty qout <- checkC ctx sg val ty
let ty' = dweakT 1 ty; val' = dweakT 1 val; p' = weakD 1 p let ty' = dweakT 1 ty; val' = dweakT 1 val; p' = weakD 1 p
ctx0 = extendDim j0 $ eqDim r (K Zero) ctx ctx0 = extendDim j0 $ eqDim r (K Zero j0.loc) ctx
val0 = val0.term val0 = val0.term
qout0 <- check ctx0 sg val0 ty' qout0 <- check ctx0 sg val0 ty'
equal (eqDim (BV 0) p' ctx0) ty' val0 val' lift $ equal loc (eqDim (B VZ p.loc) p' ctx0) ty' val0 val'
let ctx1 = extendDim j0 $ eqDim r (K One) ctx let ctx1 = extendDim j1 $ eqDim r (K One j1.loc) ctx
val1 = val1.term val1 = val1.term
qout1 <- check ctx1 sg val1 ty' qout1 <- check ctx1 sg val1 ty'
equal (eqDim (BV 0) p' ctx1) ty' val1 val' lift $ equal loc (eqDim (B VZ p.loc) p' ctx1) ty' val1 val'
let qout0' = toMaybe $ map (, val0 // one p) qout0 let qout0' = toMaybe $ map (, val0 // one p) qout0
qout1' = toMaybe $ map (, val1 // one p) qout1 qout1' = toMaybe $ map (, val1 // one p) qout1
qouts = (qout, val) :: catMaybes [qout0', qout1'] qouts = (qout, val) :: catMaybes [qout0', qout1']
let Just qout = lubs ctx $ map fst qouts let Just qout = lubs ctx $ map fst qouts
| Nothing => throw $ BadQtys "composition" ctx qouts | Nothing => throw $ BadQtys loc "composition" ctx qouts
pure $ InfRes {type = ty, qout} pure $ InfRes {type = ty, qout}
infer' ctx sg (TypeCase ty ret arms def) = do infer' ctx sg (TypeCase ty ret arms def loc) = do
-- if σ = 0 -- if σ = 0
expectEqualQ Zero sg.fst expectEqualQ loc Zero sg.fst
-- if Ψ, Γ ⊢₀ e ⇒ Type u -- if Ψ, Γ ⊢₀ e ⇒ Type u
u <- expectTYPE !defs ctx . type =<< inferC ctx szero ty u <- expectTYPE !defs ctx ty.loc . type =<< inferC ctx szero ty
-- if Ψ, Γ ⊢₀ C ⇐ Type (non-dependent return type) -- if Ψ, Γ ⊢₀ C ⇐ Type (non-dependent return type)
checkTypeC ctx ret Nothing checkTypeC ctx ret Nothing
-- if Ψ, Γ' ⊢₀ A ⇐ C for each rhs A -- if Ψ, Γ' ⊢₀ A ⇐ C for each rhs A
for_ allKinds $ \k => for_ allKinds $ \k =>
for_ (lookupPrecise k arms) $ \(S names t) => for_ (lookupPrecise k arms) $ \(S names t) =>
check0 (extendTyN (addNames0 (typecaseTel k u) names) ctx) check0 (extendTyN (typecaseTel k names u) ctx)
t.term (weakT (arity k) ret) t.term (weakT (arity k) ret)
-- then Ψ, Γ ⊢₀ type-case ⋯ ⇒ C -- then Ψ, Γ ⊢₀ type-case ⋯ ⇒ C
pure $ InfRes {type = ret, qout = zeroFor ctx} pure $ InfRes {type = ret, qout = zeroFor ctx}
infer' ctx sg (term :# type) = do infer' ctx sg (Ann term type loc) = do
-- if Ψ | Γ ⊢₀ A ⇐ Type -- if Ψ | Γ ⊢₀ A ⇐ Type
checkTypeC ctx type Nothing checkTypeC ctx type Nothing
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ -- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ

View File

@ -39,111 +39,122 @@ InferResult eqs d n = IfConsistent eqs $ InferResult' d n
export export
lookupFree : Has ErrorEff fs => Name -> Definitions -> Eff fs Definition lookupFree : Has ErrorEff fs => Name -> Loc -> Definitions -> Eff fs Definition
lookupFree x defs = maybe (throw $ NotInScope x) pure $ lookup x defs lookupFree x loc defs = maybe (throw $ NotInScope loc x) pure $ lookup x defs
public export public export
substCasePairRet : Term d n -> ScopeTerm d n -> Term d (2 + n) substCasePairRet : BContext 2 -> Term d n -> ScopeTerm d n -> Term d (2 + n)
substCasePairRet dty retty = substCasePairRet [< x, y] dty retty =
let arg = Pair (BVT 1) (BVT 0) :# (dty // fromNat 2) in let tm = Pair (BVT 1 x.loc) (BVT 0 y.loc) $ x.loc `extendL` y.loc
arg = Ann tm (dty // fromNat 2) tm.loc
in
retty.term // (arg ::: shift 2) retty.term // (arg ::: shift 2)
public export public export
substCaseSuccRet : ScopeTerm d n -> Term d (2 + n) substCaseSuccRet : BContext 2 -> ScopeTerm d n -> Term d (2 + n)
substCaseSuccRet retty = retty.term // (Succ (BVT 1) :# Nat ::: shift 2) substCaseSuccRet [< p, ih] retty =
let arg = Ann (Succ (BVT 1 p.loc) p.loc) (Nat noLoc) $ p.loc `extendL` ih.loc
in
retty.term // (arg ::: shift 2)
public export public export
substCaseBoxRet : Term d n -> ScopeTerm d n -> Term d (S n) substCaseBoxRet : BindName -> Term d n -> ScopeTerm d n -> Term d (S n)
substCaseBoxRet dty retty = substCaseBoxRet x dty retty =
retty.term // (Box (BVT 0) :# weakT 1 dty ::: shift 1) let arg = Ann (Box (BVT 0 x.loc) x.loc) (weakT 1 dty) x.loc in
retty.term // (arg ::: shift 1)
parameters (defs : Definitions) {auto _ : Has ErrorEff fs} parameters (defs : Definitions) {auto _ : (Has ErrorEff fs, Has NameGen fs)}
namespace TyContext namespace TyContext
parameters (ctx : TyContext d n) parameters (ctx : TyContext d n) (loc : Loc)
export covering export covering
whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex => whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
tm d n -> Eff fs (NonRedex tm d n defs) tm d n -> Eff fs (NonRedex tm d n defs)
whnf = let Val n = ctx.termLen; Val d = ctx.dimLen in whnf tm = do
rethrow . whnf defs (toWhnfContext ctx) let Val n = ctx.termLen; Val d = ctx.dimLen
res <- lift $ runExcept $ whnf defs (toWhnfContext ctx) tm
rethrow res
private covering %macro private covering %macro
expect : (forall d, n. NameContexts d n -> Term d n -> Error) -> expect : (forall d, n. Loc -> NameContexts d n -> Term d n -> Error) ->
TTImp -> TTImp -> Elab (Term d n -> Eff fs a) TTImp -> TTImp -> Elab (Term d n -> Eff fs a)
expect k l r = do expect k l r = do
f <- check `(\case ~(l) => Just ~(r); _ => Nothing) f <- check `(\case ~(l) => Just ~(r); _ => Nothing)
pure $ \t => maybe (throw $ k ctx.names t) pure . f . fst =<< whnf t pure $ \t => maybe (throw $ k loc ctx.names t) pure . f . fst =<< whnf t
export covering %inline export covering %inline
expectTYPE : Term d n -> Eff fs Universe expectTYPE : Term d n -> Eff fs Universe
expectTYPE = expect ExpectedTYPE `(TYPE l) `(l) expectTYPE = expect ExpectedTYPE `(TYPE {l, _}) `(l)
export covering %inline export covering %inline
expectPi : Term d n -> Eff fs (Qty, Term d n, ScopeTerm d n) expectPi : Term d n -> Eff fs (Qty, Term d n, ScopeTerm d n)
expectPi = expect ExpectedPi `(Pi {qty, arg, res}) `((qty, arg, res)) expectPi = expect ExpectedPi `(Pi {qty, arg, res, _}) `((qty, arg, res))
export covering %inline export covering %inline
expectSig : Term d n -> Eff fs (Term d n, ScopeTerm d n) expectSig : Term d n -> Eff fs (Term d n, ScopeTerm d n)
expectSig = expect ExpectedSig `(Sig {fst, snd}) `((fst, snd)) expectSig = expect ExpectedSig `(Sig {fst, snd, _}) `((fst, snd))
export covering %inline export covering %inline
expectEnum : Term d n -> Eff fs (SortedSet TagVal) expectEnum : Term d n -> Eff fs (SortedSet TagVal)
expectEnum = expect ExpectedEnum `(Enum ts) `(ts) expectEnum = expect ExpectedEnum `(Enum {cases, _}) `(cases)
export covering %inline export covering %inline
expectEq : Term d n -> Eff fs (DScopeTerm d n, Term d n, Term d n) expectEq : Term d n -> Eff fs (DScopeTerm d n, Term d n, Term d n)
expectEq = expect ExpectedEq `(Eq {ty, l, r}) `((ty, l, r)) expectEq = expect ExpectedEq `(Eq {ty, l, r, _}) `((ty, l, r))
export covering %inline export covering %inline
expectNat : Term d n -> Eff fs () expectNat : Term d n -> Eff fs ()
expectNat = expect ExpectedNat `(Nat) `(()) expectNat = expect ExpectedNat `(Nat {}) `(())
export covering %inline export covering %inline
expectBOX : Term d n -> Eff fs (Qty, Term d n) expectBOX : Term d n -> Eff fs (Qty, Term d n)
expectBOX = expect ExpectedBOX `(BOX {qty, ty}) `((qty, ty)) expectBOX = expect ExpectedBOX `(BOX {qty, ty, _}) `((qty, ty))
namespace EqContext namespace EqContext
parameters (ctx : EqContext n) parameters (ctx : EqContext n) (loc : Loc)
export covering export covering
whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex => whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
tm 0 n -> Eff fs (NonRedex tm 0 n defs) tm 0 n -> Eff fs (NonRedex tm 0 n defs)
whnf = let Val n = ctx.termLen in rethrow . whnf defs (toWhnfContext ctx) whnf tm = do
let Val n = ctx.termLen
res <- lift $ runExcept $ whnf defs (toWhnfContext ctx) tm
rethrow res
private covering %macro private covering %macro
expect : (forall d, n. NameContexts d n -> Term d n -> Error) -> expect : (forall d, n. Loc -> NameContexts d n -> Term d n -> Error) ->
TTImp -> TTImp -> Elab (Term 0 n -> Eff fs a) TTImp -> TTImp -> Elab (Term 0 n -> Eff fs a)
expect k l r = do expect k l r = do
f <- check `(\case ~(l) => Just ~(r); _ => Nothing) f <- check `(\case ~(l) => Just ~(r); _ => Nothing)
pure $ \t => pure $ \t =>
let err = throw $ k ctx.names (t // shift0 ctx.dimLen) in let err = throw $ k loc ctx.names (t // shift0 ctx.dimLen) in
maybe err pure . f . fst =<< whnf t maybe err pure . f . fst =<< whnf t
export covering %inline export covering %inline
expectTYPE : Term 0 n -> Eff fs Universe expectTYPE : Term 0 n -> Eff fs Universe
expectTYPE = expect ExpectedTYPE `(TYPE l) `(l) expectTYPE = expect ExpectedTYPE `(TYPE {l, _}) `(l)
export covering %inline export covering %inline
expectPi : Term 0 n -> Eff fs (Qty, Term 0 n, ScopeTerm 0 n) expectPi : Term 0 n -> Eff fs (Qty, Term 0 n, ScopeTerm 0 n)
expectPi = expect ExpectedPi `(Pi {qty, arg, res}) `((qty, arg, res)) expectPi = expect ExpectedPi `(Pi {qty, arg, res, _}) `((qty, arg, res))
export covering %inline export covering %inline
expectSig : Term 0 n -> Eff fs (Term 0 n, ScopeTerm 0 n) expectSig : Term 0 n -> Eff fs (Term 0 n, ScopeTerm 0 n)
expectSig = expect ExpectedSig `(Sig {fst, snd}) `((fst, snd)) expectSig = expect ExpectedSig `(Sig {fst, snd, _}) `((fst, snd))
export covering %inline export covering %inline
expectEnum : Term 0 n -> Eff fs (SortedSet TagVal) expectEnum : Term 0 n -> Eff fs (SortedSet TagVal)
expectEnum = expect ExpectedEnum `(Enum ts) `(ts) expectEnum = expect ExpectedEnum `(Enum {cases, _}) `(cases)
export covering %inline export covering %inline
expectEq : Term 0 n -> Eff fs (DScopeTerm 0 n, Term 0 n, Term 0 n) expectEq : Term 0 n -> Eff fs (DScopeTerm 0 n, Term 0 n, Term 0 n)
expectEq = expect ExpectedEq `(Eq {ty, l, r}) `((ty, l, r)) expectEq = expect ExpectedEq `(Eq {ty, l, r, _}) `((ty, l, r))
export covering %inline export covering %inline
expectNat : Term 0 n -> Eff fs () expectNat : Term 0 n -> Eff fs ()
expectNat = expect ExpectedNat `(Nat) `(()) expectNat = expect ExpectedNat `(Nat {}) `(())
export covering %inline export covering %inline
expectBOX : Term 0 n -> Eff fs (Qty, Term 0 n) expectBOX : Term 0 n -> Eff fs (Qty, Term 0 n)
expectBOX = expect ExpectedBOX `(BOX {qty, ty}) `((qty, ty)) expectBOX = expect ExpectedBOX `(BOX {qty, ty, _}) `((qty, ty))

View File

@ -31,9 +31,9 @@ record TyContext d n where
{auto dimLen : Singleton d} {auto dimLen : Singleton d}
{auto termLen : Singleton n} {auto termLen : Singleton n}
dctx : DimEq d dctx : DimEq d
dnames : NContext d dnames : BContext d
tctx : TContext d n tctx : TContext d n
tnames : NContext n tnames : BContext n
qtys : QContext n -- only used for printing qtys : QContext n -- only used for printing
%name TyContext ctx %name TyContext ctx
@ -44,9 +44,9 @@ record EqContext n where
{dimLen : Nat} {dimLen : Nat}
{auto termLen : Singleton n} {auto termLen : Singleton n}
dassign : DimAssign dimLen -- only used for printing dassign : DimAssign dimLen -- only used for printing
dnames : NContext dimLen -- only used for printing dnames : BContext dimLen -- only used for printing
tctx : TContext 0 n tctx : TContext 0 n
tnames : NContext n tnames : BContext n
qtys : QContext n -- only used for printing qtys : QContext n -- only used for printing
%name EqContext ctx %name EqContext ctx
@ -54,8 +54,8 @@ record EqContext n where
public export public export
record WhnfContext d n where record WhnfContext d n where
constructor MkWhnfContext constructor MkWhnfContext
dnames : NContext d dnames : BContext d
tnames : NContext n tnames : BContext n
tctx : TContext d n tctx : TContext d n
%name WhnfContext ctx %name WhnfContext ctx
@ -76,7 +76,7 @@ extendLen (tel :< _) x = [|S $ extendLen tel x|]
public export public export
CtxExtension : Nat -> Nat -> Nat -> Type CtxExtension : Nat -> Nat -> Nat -> Type
CtxExtension d = Telescope ((Qty, BaseName,) . Term d) CtxExtension d = Telescope ((Qty, BindName,) . Term d)
namespace TyContext namespace TyContext
public export %inline public export %inline
@ -101,11 +101,11 @@ namespace TyContext
} }
export %inline export %inline
extendTy : Qty -> BaseName -> Term d n -> TyContext d n -> TyContext d (S n) extendTy : Qty -> BindName -> Term d n -> TyContext d n -> TyContext d (S n)
extendTy q x s = extendTyN [< (q, x, s)] extendTy q x s = extendTyN [< (q, x, s)]
export %inline export %inline
extendDim : BaseName -> TyContext d n -> TyContext (S d) n extendDim : BindName -> TyContext d n -> TyContext (S d) n
extendDim x (MkTyContext {dimLen, dctx, dnames, tctx, tnames, qtys}) = extendDim x (MkTyContext {dimLen, dctx, dnames, tctx, tnames, qtys}) =
MkTyContext { MkTyContext {
dctx = dctx :<? Nothing, dctx = dctx :<? Nothing,
@ -141,8 +141,8 @@ namespace QOutput
export export
makeDAssign : DSubst d 0 -> DimAssign d makeDAssign : DSubst d 0 -> DimAssign d
makeDAssign (Shift SZ) = [<] makeDAssign (Shift SZ) = [<]
makeDAssign (K e ::: th) = makeDAssign th :< e makeDAssign (K e _ ::: th) = makeDAssign th :< e
export export
makeEqContext' : {d : Nat} -> TyContext d n -> DSubst d 0 -> EqContext n makeEqContext' : {d : Nat} -> TyContext d n -> DSubst d 0 -> EqContext n
@ -172,8 +172,7 @@ namespace EqContext
null ctx = null ctx.dnames && null ctx.tnames null ctx = null ctx.dnames && null ctx.tnames
export %inline export %inline
extendTyN : Telescope (\n => (Qty, BaseName, Term 0 n)) from to -> extendTyN : CtxExtension 0 from to -> EqContext from -> EqContext to
EqContext from -> EqContext to
extendTyN xss (MkEqContext {termLen, dassign, dnames, tctx, tnames, qtys}) = extendTyN xss (MkEqContext {termLen, dassign, dnames, tctx, tnames, qtys}) =
let (qs, xs, ss) = unzip3 xss in let (qs, xs, ss) = unzip3 xss in
MkEqContext { MkEqContext {
@ -185,11 +184,11 @@ namespace EqContext
} }
export %inline export %inline
extendTy : Qty -> BaseName -> Term 0 n -> EqContext n -> EqContext (S n) extendTy : Qty -> BindName -> Term 0 n -> EqContext n -> EqContext (S n)
extendTy q x s = extendTyN [< (q, x, s)] extendTy q x s = extendTyN [< (q, x, s)]
export %inline export %inline
extendDim : BaseName -> DimConst -> EqContext n -> EqContext n extendDim : BindName -> DimConst -> EqContext n -> EqContext n
extendDim x e (MkEqContext {dassign, dnames, tctx, tnames, qtys}) = extendDim x e (MkEqContext {dassign, dnames, tctx, tnames, qtys}) =
MkEqContext {dassign = dassign :< e, dnames = dnames :< x, MkEqContext {dassign = dassign :< e, dnames = dnames :< x,
tctx, tnames, qtys} tctx, tnames, qtys}
@ -214,7 +213,7 @@ namespace WhnfContext
empty = MkWhnfContext [<] [<] [<] empty = MkWhnfContext [<] [<] [<]
export export
extendDimN : {s : Nat} -> NContext s -> WhnfContext d n -> extendDimN : {s : Nat} -> BContext s -> WhnfContext d n ->
WhnfContext (s + d) n WhnfContext (s + d) n
extendDimN ns (MkWhnfContext {dnames, tnames, tctx}) = extendDimN ns (MkWhnfContext {dnames, tnames, tctx}) =
MkWhnfContext { MkWhnfContext {
@ -224,16 +223,16 @@ namespace WhnfContext
} }
export export
extendDim : BaseName -> WhnfContext d n -> WhnfContext (S d) n extendDim : BindName -> WhnfContext d n -> WhnfContext (S d) n
extendDim i = extendDimN [< i] extendDim i = extendDimN [< i]
private private
data CtxBinder a = MkCtxBinder BaseName a data CtxBinder a = MkCtxBinder BindName a
PrettyHL a => PrettyHL (CtxBinder a) where PrettyHL a => PrettyHL (CtxBinder a) where
prettyM (MkCtxBinder x t) = pure $ prettyM (MkCtxBinder x t) = pure $
sep [hsep [!(pretty0M $ TV x), colonD], !(pretty0M t)] sep [hsep [!(pretty0M $ TV x.name), colonD], !(pretty0M t)]
parameters (unicode : Bool) parameters (unicode : Bool)
private private
@ -241,16 +240,16 @@ parameters (unicode : Bool)
pipeD = hl Syntax "|" pipeD = hl Syntax "|"
export covering export covering
prettyTContext : NContext d -> prettyTContext : BContext d ->
QContext n -> NContext n -> QContext n -> BContext n ->
TContext d n -> Doc HL TContext d n -> Doc HL
prettyTContext ds qs xs ctx = separate comma $ toList $ go qs xs ctx where prettyTContext ds qs xs ctx = separate comma $ toList $ go qs xs ctx where
go : QContext m -> NContext m -> TContext d m -> SnocList (Doc HL) go : QContext m -> BContext m -> TContext d m -> SnocList (Doc HL)
go [<] [<] [<] = [<] go [<] [<] [<] = [<]
go (qs :< q) (xs :< x) (ctx :< t) = go (qs :< q) (xs :< x) (ctx :< t) =
let bind = MkWithQty q $ MkCtxBinder x t in let bind = MkWithQty q $ MkCtxBinder x t in
go qs xs ctx :< go qs xs ctx :<
runPrettyWith unicode (toSnocList' ds) (toSnocList' xs) (pretty0M bind) runPrettyWith unicode (toNames ds) (toNames xs) (pretty0M bind)
export covering export covering
prettyTyContext : TyContext d n -> Doc HL prettyTyContext : TyContext d n -> Doc HL

View File

@ -1,5 +1,6 @@
module Quox.Typing.Error module Quox.Typing.Error
import Quox.Loc
import Quox.Syntax import Quox.Syntax
import Quox.Typing.Context import Quox.Typing.Context
import Quox.Typing.EqMode import Quox.Typing.EqMode
@ -12,8 +13,8 @@ import Control.Eff
public export public export
record NameContexts d n where record NameContexts d n where
constructor MkNameContexts constructor MkNameContexts
dnames : NContext d dnames : BContext d
tnames : NContext n tnames : BContext n
namespace NameContexts namespace NameContexts
export export
@ -21,11 +22,11 @@ namespace NameContexts
empty = MkNameContexts [<] [<] empty = MkNameContexts [<] [<]
export export
extendDimN : NContext s -> NameContexts d n -> NameContexts (s + d) n extendDimN : BContext s -> NameContexts d n -> NameContexts (s + d) n
extendDimN xs = {dnames $= (++ toSnocVect' xs)} extendDimN xs = {dnames $= (++ toSnocVect' xs)}
export export
extendDim : BaseName -> NameContexts d n -> NameContexts (S d) n extendDim : BindName -> NameContexts d n -> NameContexts (S d) n
extendDim i = extendDimN [< i] extendDim i = extendDimN [< i]
namespace TyContext namespace TyContext
@ -54,30 +55,30 @@ namespace WhnfContext
public export public export
data Error data Error
= ExpectedTYPE (NameContexts d n) (Term d n) = ExpectedTYPE Loc (NameContexts d n) (Term d n)
| ExpectedPi (NameContexts d n) (Term d n) | ExpectedPi Loc (NameContexts d n) (Term d n)
| ExpectedSig (NameContexts d n) (Term d n) | ExpectedSig Loc (NameContexts d n) (Term d n)
| ExpectedEnum (NameContexts d n) (Term d n) | ExpectedEnum Loc (NameContexts d n) (Term d n)
| ExpectedEq (NameContexts d n) (Term d n) | ExpectedEq Loc (NameContexts d n) (Term d n)
| ExpectedNat (NameContexts d n) (Term d n) | ExpectedNat Loc (NameContexts d n) (Term d n)
| ExpectedBOX (NameContexts d n) (Term d n) | ExpectedBOX Loc (NameContexts d n) (Term d n)
| BadUniverse Universe Universe | BadUniverse Loc Universe Universe
| TagNotIn TagVal (SortedSet TagVal) | TagNotIn Loc TagVal (SortedSet TagVal)
| BadCaseEnum (SortedSet TagVal) (SortedSet TagVal) | BadCaseEnum Loc (SortedSet TagVal) (SortedSet TagVal)
| BadQtys String (TyContext d n) (List (QOutput n, Term d n)) | BadQtys Loc String (TyContext d n) (List (QOutput n, Term d n))
-- first term arg of ClashT is the type -- first term arg of ClashT is the type
| ClashT (EqContext n) EqMode (Term 0 n) (Term 0 n) (Term 0 n) | ClashT Loc (EqContext n) EqMode (Term 0 n) (Term 0 n) (Term 0 n)
| ClashTy (EqContext n) EqMode (Term 0 n) (Term 0 n) | ClashTy Loc (EqContext n) EqMode (Term 0 n) (Term 0 n)
| ClashE (EqContext n) EqMode (Elim 0 n) (Elim 0 n) | ClashE Loc (EqContext n) EqMode (Elim 0 n) (Elim 0 n)
| ClashU EqMode Universe Universe | ClashU Loc EqMode Universe Universe
| ClashQ Qty Qty | ClashQ Loc Qty Qty
| NotInScope Name | NotInScope Loc Name
| NotType (TyContext d n) (Term d n) | NotType Loc (TyContext d n) (Term d n)
| WrongType (EqContext n) (Term 0 n) (Term 0 n) | WrongType Loc (EqContext n) (Term 0 n) (Term 0 n)
| MissingEnumArm TagVal (List TagVal) | MissingEnumArm Loc TagVal (List TagVal)
-- extra context -- extra context
| WhileChecking | WhileChecking
@ -166,18 +167,18 @@ expect : Has (Except e) fs =>
(a -> a -> e) -> (a -> a -> Bool) -> a -> a -> Eff fs () (a -> a -> e) -> (a -> a -> Bool) -> a -> a -> Eff fs ()
expect err cmp x y = unless (x `cmp` y) $ throw $ err x y expect err cmp x y = unless (x `cmp` y) $ throw $ err x y
parameters {auto _ : Has ErrorEff fs} parameters {auto _ : Has ErrorEff fs} (loc : Loc)
export %inline export %inline
expectEqualQ : Qty -> Qty -> Eff fs () expectEqualQ : Qty -> Qty -> Eff fs ()
expectEqualQ = expect ClashQ (==) expectEqualQ = expect (ClashQ loc) (==)
export %inline export %inline
expectCompatQ : Qty -> Qty -> Eff fs () expectCompatQ : Qty -> Qty -> Eff fs ()
expectCompatQ = expect ClashQ compat expectCompatQ = expect (ClashQ loc) compat
export %inline export %inline
expectModeU : EqMode -> Universe -> Universe -> Eff fs () expectModeU : EqMode -> Universe -> Universe -> Eff fs ()
expectModeU mode = expect (ClashU mode) $ ucmp mode expectModeU mode = expect (ClashU loc mode) $ ucmp mode
private private
@ -207,8 +208,8 @@ parameters (unicode : Bool)
dstermn ctx (S i t) = termn (extendDimN i ctx) t.term dstermn ctx (S i t) = termn (extendDimN i ctx) t.term
private private
filterSameQtys : NContext n -> List (QOutput n, z) -> filterSameQtys : BContext n -> List (QOutput n, z) ->
Exists $ \n' => (NContext n', List (QOutput n', z)) Exists $ \n' => (BContext n', List (QOutput n', z))
filterSameQtys [<] qts = Evidence 0 ([<], qts) filterSameQtys [<] qts = Evidence 0 ([<], qts)
filterSameQtys (ns :< n) qts = filterSameQtys (ns :< n) qts =
let (qs, qts) = unzip $ map (\(qs :< q, t) => (q, qs, t)) qts let (qs, qts) = unzip $ map (\(qs :< q, t) => (q, qs, t)) qts
@ -224,7 +225,7 @@ parameters (unicode : Bool)
private private
printCaseQtys : TyContext d n -> printCaseQtys : TyContext d n ->
NContext n' -> List (QOutput n', Term d n) -> BContext n' -> List (QOutput n', Term d n) ->
List (Doc HL) List (Doc HL)
printCaseQtys ctx ns qts = printCaseQtys ctx ns qts =
let Evidence l (ns, qts) = filterSameQtys ns qts in let Evidence l (ns, qts) = filterSameQtys ns qts in
@ -233,94 +234,97 @@ parameters (unicode : Bool)
commaList : PrettyHL a => Context' a l -> Doc HL commaList : PrettyHL a => Context' a l -> Doc HL
commaList = hseparate comma . map (pretty0 unicode) . toList' commaList = hseparate comma . map (pretty0 unicode) . toList'
line : NContext l -> (QOutput l, Term d n) -> Doc HL line : BContext l -> (QOutput l, Term d n) -> Doc HL
line ns (qs, t) = line ns (qs, t) =
"-" <++> asep ["the term", termn ctx.names t, "-" <++> asep ["the term", termn ctx.names t,
"uses variables", commaList $ TV <$> ns, "uses variables", commaList $ (TV . name) <$> ns,
"with quantities", commaList qs] "with quantities", commaList qs]
-- [todo] only show some contexts, probably -- [todo] only show some contexts, probably
export export
prettyError : (showContext : Bool) -> Error -> Doc HL prettyError : (showContext : Bool) -> Error -> Doc HL
prettyError showContext = \case prettyError showContext = \case
ExpectedTYPE ctx s => ExpectedTYPE loc ctx s =>
sep ["expected a type universe, but got", termn ctx s] sep [prettyLoc loc <++> "expected a type universe, but got", termn ctx s]
ExpectedPi ctx s => ExpectedPi loc ctx s =>
sep ["expected a function type, but got", termn ctx s] sep [prettyLoc loc <++> "expected a function type, but got", termn ctx s]
ExpectedSig ctx s => ExpectedSig loc ctx s =>
sep ["expected a pair type, but got", termn ctx s] sep [prettyLoc loc <++> "expected a pair type, but got", termn ctx s]
ExpectedEnum ctx s => ExpectedEnum loc ctx s =>
sep ["expected an enumeration type, but got", termn ctx s] sep [prettyLoc loc <++> "expected an enumeration type, but got",
termn ctx s]
ExpectedEq ctx s => ExpectedEq loc ctx s =>
sep ["expected an equality type, but got", termn ctx s] sep [prettyLoc loc <++> "expected an equality type, but got", termn ctx s]
ExpectedNat ctx s {d, n} => ExpectedNat loc ctx s {d, n} =>
sep ["expected the type", pretty0 unicode $ Nat {d, n}, sep [prettyLoc loc <++> "expected the type",
"but got", termn ctx s] pretty0 unicode $ Nat noLoc {d, n}, "but got", termn ctx s]
ExpectedBOX ctx s => ExpectedBOX loc ctx s =>
sep ["expected a box type, but got", termn ctx s] sep [prettyLoc loc <++> "expected a box type, but got", termn ctx s]
BadUniverse k l => BadUniverse loc k l =>
sep ["the universe level", prettyUniverse k, sep [prettyLoc loc <++> "the universe level", prettyUniverse k,
"is not strictly less than", prettyUniverse l] "is not strictly less than", prettyUniverse l]
TagNotIn tag set => TagNotIn loc tag set =>
sep [sep ["tag", prettyTag tag, "is not contained in"], sep [hsep [prettyLoc loc, "tag", prettyTag tag, "is not contained in"],
termn empty (Enum set)] termn empty (Enum set noLoc)]
BadCaseEnum type arms => BadCaseEnum loc type arms =>
sep ["case expression has head of type", termn empty (Enum type), sep [prettyLoc loc <++> "case expression has head of type",
"but cases for", termn empty (Enum arms)] termn empty (Enum type noLoc),
"but cases for", termn empty (Enum arms noLoc)]
BadQtys what ctx arms => BadQtys loc what ctx arms =>
hang 4 $ sep $ hang 4 $ sep $
("inconsistent variable usage in" <++> fromString what) :: hsep [prettyLoc loc, "inconsistent variable usage in", fromString what]
printCaseQtys ctx ctx.tnames arms :: printCaseQtys ctx ctx.tnames arms
ClashT ctx mode ty s t => ClashT loc ctx mode ty s t =>
inEContext ctx $ inEContext ctx $
sep ["the term", termn ctx.names0 s, sep [prettyLoc loc <++> "the term", termn ctx.names0 s,
hsep ["is not", prettyMode mode], termn ctx.names0 t, hsep ["is not", prettyMode mode], termn ctx.names0 t,
"at type", termn ctx.names0 ty] "at type", termn ctx.names0 ty]
ClashTy ctx mode a b => ClashTy loc ctx mode a b =>
inEContext ctx $ inEContext ctx $
sep ["the type", termn ctx.names0 a, sep [prettyLoc loc <++> "the type", termn ctx.names0 a,
hsep ["is not", prettyMode mode], termn ctx.names0 b] hsep ["is not", prettyMode mode], termn ctx.names0 b]
ClashE ctx mode e f => ClashE loc ctx mode e f =>
inEContext ctx $ inEContext ctx $
sep ["the term", termn ctx.names0 $ E e, sep [prettyLoc loc <++> "the term", termn ctx.names0 $ E e,
hsep ["is not", prettyMode mode], termn ctx.names0 $ E f] hsep ["is not", prettyMode mode], termn ctx.names0 $ E f]
ClashU mode k l => ClashU loc mode k l =>
sep ["the universe level", prettyUniverse k, sep [prettyLoc loc <++> "the universe level", prettyUniverse k,
hsep ["is not", prettyMode mode], prettyUniverse l] hsep ["is not", prettyMode mode], prettyUniverse l]
ClashQ pi rh => ClashQ loc pi rh =>
sep ["the quantity", pretty0 unicode pi, sep [prettyLoc loc <++> "the quantity", pretty0 unicode pi,
"is not equal to", pretty0 unicode rh] "is not equal to", pretty0 unicode rh]
NotInScope x => NotInScope loc x =>
hsep [hl' Free $ pretty0 unicode x, "is not in scope"] hsep [prettyLoc loc, hl' Free $ pretty0 unicode x, "is not in scope"]
NotType ctx s => NotType loc ctx s =>
inTContext ctx $ inTContext ctx $
sep ["the term", termn ctx.names s, "is not a type"] sep [prettyLoc loc <++> "the term", termn ctx.names s, "is not a type"]
WrongType ctx ty s => WrongType loc ctx ty s =>
inEContext ctx $ inEContext ctx $
sep ["the term", termn ctx.names0 s, sep [prettyLoc loc <++> "the term", termn ctx.names0 s,
"cannot have type", termn ctx.names0 ty] "cannot have type", termn ctx.names0 ty]
MissingEnumArm tag tags => MissingEnumArm loc tag tags =>
sep [hsep ["the tag", hl Tag $ pretty tag, "is not contained in"], sep [hsep [prettyLoc loc, "the tag", hl Tag $ pretty tag,
termn empty $ Enum $ fromList tags] "is not contained in"],
termn empty $ Enum (fromList tags) noLoc]
WhileChecking ctx pi s a err => WhileChecking ctx pi s a err =>
vsep [inTContext ctx $ vsep [inTContext ctx $

33
tests/AstExtra.idr Normal file
View File

@ -0,0 +1,33 @@
module AstExtra
import Quox.Syntax
import Quox.Parser.Syntax
import Quox.Typing.Context
prefix 9 ^
public export
(^) : (Loc -> a) -> a
(^) a = a noLoc
public export
FromString BindName where fromString str = BN (fromString str) noLoc
public export
FromString PatVar where fromString x = PV x noLoc
public export
empty01 : TyContext 0 0
empty01 = eqDim (^K Zero) (^K One) empty
anys : {n : Nat} -> QContext n
anys {n = 0} = [<]
anys {n = S n} = anys :< Any
public export
ctx, ctx01 : {n : Nat} -> Context (\n => (BindName, Term 0 n)) n ->
TyContext 0 n
ctx tel = let (ns, ts) = unzip tel in
MkTyContext new [<] ts ns anys
ctx01 tel = let (ns, ts) = unzip tel in
MkTyContext ZeroIsOne [<] ts ns anys

View File

@ -2,6 +2,7 @@ module Tests.DimEq
import Quox.Syntax.DimEq import Quox.Syntax.DimEq
import PrettyExtra import PrettyExtra
import AstExtra
import TAP import TAP
import Data.Maybe import Data.Maybe
@ -95,9 +96,9 @@ tests = "dimension constraints" :- [
testPrettyD ii new "𝑖", testPrettyD ii new "𝑖",
testPrettyD iijj (fromGround [< Zero, One]) testPrettyD iijj (fromGround [< Zero, One])
"𝑖, 𝑗, 𝑖 = 0, 𝑗 = 1", "𝑖, 𝑗, 𝑖 = 0, 𝑗 = 1",
testPrettyD iijj (C [< Just (K Zero), Nothing]) testPrettyD iijj (C [< Just (^K Zero), Nothing])
"𝑖, 𝑗, 𝑖 = 0", "𝑖, 𝑗, 𝑖 = 0",
testPrettyD iijjkk (C [< Nothing, Just (BV 0), Just (BV 1)]) testPrettyD iijjkk (C [< Nothing, Just (^BV 0), Just (^BV 1)])
"𝑖, 𝑗, 𝑘, 𝑗 = 𝑖, 𝑘 = 𝑖" "𝑖, 𝑗, 𝑘, 𝑗 = 𝑖, 𝑘 = 𝑖"
], ],
@ -108,56 +109,56 @@ tests = "dimension constraints" :- [
testNeq [<] new ZeroIsOne, testNeq [<] new ZeroIsOne,
testNeq iijj new ZeroIsOne, testNeq iijj new ZeroIsOne,
testSet iijj testSet iijj
(C [< Nothing, Just (BV 0)]) (C [< Nothing, Just (^BV 0)])
new [(BV 1, BV 0)], new [(^BV 1, ^BV 0)],
testSet iijj testSet iijj
(C [< Nothing, Just (BV 0)]) (C [< Nothing, Just (^BV 0)])
new [(BV 0, BV 1)], new [(^BV 0, ^BV 1)],
testNeq iijj testNeq iijj
new new
(C [< Nothing, Just (BV 0)]), (C [< Nothing, Just (^BV 0)]),
testSet [<] testSet [<]
ZeroIsOne ZeroIsOne
new [(K Zero, K One)], new [(^K Zero, ^K One)],
testSet iijjkk testSet iijjkk
(C [< Nothing, Just (BV 0), Just (BV 1)]) (C [< Nothing, Just (^BV 0), Just (^BV 1)])
new [(BV 0, BV 1), (BV 1, BV 2)], new [(^BV 0, ^BV 1), (^BV 1, ^BV 2)],
testSet iijjkk testSet iijjkk
(C [< Nothing, Just (BV 0), Just (BV 1)]) (C [< Nothing, Just (^BV 0), Just (^BV 1)])
new [(BV 0, BV 1), (BV 0, BV 2)], new [(^BV 0, ^BV 1), (^BV 0, ^BV 2)],
testSet iijjkk testSet iijjkk
(C [< Nothing, Nothing, Just (BV 0)]) (C [< Nothing, Nothing, Just (^BV 0)])
new [(BV 0, BV 1), (BV 0, BV 1)], new [(^BV 0, ^BV 1), (^BV 0, ^BV 1)],
testSet iijj testSet iijj
(C [< Just (K Zero), Just (K Zero)]) (C [< Just (^K Zero), Just (^K Zero)])
new [(BV 1, K Zero), (BV 0, BV 1)], new [(^BV 1, ^K Zero), (^BV 0, ^BV 1)],
testSet iijjkk testSet iijjkk
(C [< Just (K Zero), Just (K Zero), Just (K Zero)]) (C [< Just (^K Zero), Just (^K Zero), Just (^K Zero)])
new [(BV 2, K Zero), (BV 1, BV 2), (BV 0, BV 1)], new [(^BV 2, ^K Zero), (^BV 1, ^BV 2), (^BV 0, ^BV 1)],
testSet iijjkk testSet iijjkk
(C [< Just (K Zero), Just (K Zero), Just (K Zero)]) (C [< Just (^K Zero), Just (^K Zero), Just (^K Zero)])
new [(BV 2, K Zero), (BV 0, BV 1), (BV 1, BV 2)], new [(^BV 2, ^K Zero), (^BV 0, ^BV 1), (^BV 1, ^BV 2)],
testSet iijjkk testSet iijjkk
(C [< Just (K Zero), Just (K Zero), Just (K Zero)]) (C [< Just (^K Zero), Just (^K Zero), Just (^K Zero)])
new [(BV 0, BV 2), (BV 1, K Zero), (BV 2, BV 1)], new [(^BV 0, ^BV 2), (^BV 1, ^K Zero), (^BV 2, ^BV 1)],
testSet iijjkk testSet iijjkk
(C [< Nothing, Just (BV 0), Just (BV 1)]) (C [< Nothing, Just (^BV 0), Just (^BV 1)])
new [(BV 0, BV 2), (BV 2, BV 1)], new [(^BV 0, ^BV 2), (^BV 2, ^BV 1)],
testSet iijjkkll testSet iijjkkll
(C [< Nothing, Just (BV 0), Just (BV 1), Just (BV 2)]) (C [< Nothing, Just (^BV 0), Just (^BV 1), Just (^BV 2)])
new [(BV 2, BV 1), (BV 3, BV 0), (BV 2, BV 3)], new [(^BV 2, ^BV 1), (^BV 3, ^BV 0), (^BV 2, ^BV 3)],
testSet iijjkk testSet iijjkk
(C [< Just (K One), Just (K One), Just (K One)]) (C [< Just (^K One), Just (^K One), Just (^K One)])
(C [< Just (K One), Nothing, Just (BV 0)]) (C [< Just (^K One), Nothing, Just (^BV 0)])
[(BV 1, BV 2)], [(^BV 1, ^BV 2)],
testSet iijj testSet iijj
ZeroIsOne ZeroIsOne
(C [< Just (K One), Just (K Zero)]) (C [< Just (^K One), Just (^K Zero)])
[(BV 1, BV 0)], [(^BV 1, ^BV 0)],
testSet iijj testSet iijj
ZeroIsOne ZeroIsOne
(C [< Nothing, Just (BV 0)]) (C [< Nothing, Just (^BV 0)])
[(BV 1, K Zero), (BV 0, K One)] [(^BV 1, ^K Zero), (^BV 0, ^K One)]
], ],
"wf" :- [ "wf" :- [
@ -165,9 +166,9 @@ tests = "dimension constraints" :- [
testWf ii ZeroIsOne, testWf ii ZeroIsOne,
testWf [<] new, testWf [<] new,
testWf iijjkk new, testWf iijjkk new,
testWf iijjkk (C [< Nothing, Just (BV 0), Just (BV 1)]), testWf iijjkk (C [< Nothing, Just (^BV 0), Just (^BV 1)]),
testNwf iijjkk (C [< Nothing, Just (BV 0), Just (BV 0)]), testNwf iijjkk (C [< Nothing, Just (^BV 0), Just (^BV 0)]),
testWf iijj (C [< Just (K Zero), Just (K Zero)]), testWf iijj (C [< Just (^K Zero), Just (^K Zero)]),
testNwf iijj (C [< Just (K Zero), Just (BV 0)]) testNwf iijj (C [< Just (^K Zero), Just (^BV 0)])
] ]
] ]

View File

@ -5,52 +5,40 @@ import Quox.Typechecker
import public TypingImpls import public TypingImpls
import TAP import TAP
import Quox.EffExtra import Quox.EffExtra
import AstExtra
defGlobals : Definitions defGlobals : Definitions
defGlobals = fromList defGlobals = fromList
[("A", mkPostulate gzero $ TYPE 0), [("A", ^mkPostulate gzero (^TYPE 0)),
("B", mkPostulate gzero $ TYPE 0), ("B", ^mkPostulate gzero (^TYPE 0)),
("a", mkPostulate gany $ FT "A"), ("a", ^mkPostulate gany (^FT "A")),
("a'", mkPostulate gany $ FT "A"), ("a'", ^mkPostulate gany (^FT "A")),
("b", mkPostulate gany $ FT "B"), ("b", ^mkPostulate gany (^FT "B")),
("f", mkPostulate gany $ Arr One (FT "A") (FT "A")), ("f", ^mkPostulate gany (^Arr One (^FT "A") (^FT "A"))),
("id", mkDef gany (Arr One (FT "A") (FT "A")) ([< "x"] :\\ BVT 0)), ("id", ^mkDef gany (^Arr One (^FT "A") (^FT "A")) (^LamY "x" (^BVT 0))),
("eq-AB", mkPostulate gzero $ Eq0 (TYPE 0) (FT "A") (FT "B")), ("eq-AB", ^mkPostulate gzero (^Eq0 (^TYPE 0) (^FT "A") (^FT "B"))),
("two", mkDef gany Nat (Succ (Succ Zero)))] ("two", ^mkDef gany (^Nat) (^Succ (^Succ (^Zero))))]
parameters (label : String) (act : Lazy (TC ())) parameters (label : String) (act : Equal ())
{default defGlobals globals : Definitions} {default defGlobals globals : Definitions}
testEq : Test testEq : Test
testEq = test label $ runTC globals act testEq = test label $ runEqual globals act
testNeq : Test testNeq : Test
testNeq = testThrows label (const True) $ runTC globals act $> "()" testNeq = testThrows label (const True) $ runTC globals act $> "()"
parameters (0 d : Nat) (ctx : TyContext d n) parameters (ctx : TyContext d n)
subTD, equalTD : Term d n -> Term d n -> Term d n -> TC () subT, equalT : Term d n -> Term d n -> Term d n -> TC ()
subTD ty s t = Term.sub ctx ty s t subT ty s t = lift $ Term.sub noLoc ctx ty s t
equalTD ty s t = Term.equal ctx ty s t equalT ty s t = lift $ Term.equal noLoc ctx ty s t
equalTyD : Term d n -> Term d n -> TC () equalTy : Term d n -> Term d n -> TC ()
equalTyD s t = Term.equalType ctx s t equalTy s t = lift $ Term.equalType noLoc ctx s t
subED, equalED : Elim d n -> Elim d n -> TC () subE, equalE : Elim d n -> Elim d n -> TC ()
subED e f = Elim.sub ctx e f subE e f = lift $ Elim.sub noLoc ctx e f
equalED e f = Elim.equal ctx e f equalE e f = lift $ Elim.equal noLoc ctx e f
parameters (ctx : TyContext 0 n)
subT, equalT : Term 0 n -> Term 0 n -> Term 0 n -> TC ()
subT = subTD 0 ctx
equalT = equalTD 0 ctx
equalTy : Term 0 n -> Term 0 n -> TC ()
equalTy = equalTyD 0 ctx
subE, equalE : Elim 0 n -> Elim 0 n -> TC ()
subE = subED 0 ctx
equalE = equalED 0 ctx
empty01 : TyContext 0 0
empty01 = eqDim (K Zero) (K One) empty
export export
@ -61,410 +49,434 @@ tests = "equality & subtyping" :- [
"universes" :- [ "universes" :- [
testEq "★₀ = ★₀" $ testEq "★₀ = ★₀" $
equalT empty (TYPE 1) (TYPE 0) (TYPE 0), equalT empty (^TYPE 1) (^TYPE 0) (^TYPE 0),
testNeq "★₀ ≠ ★₁" $ testNeq "★₀ ≠ ★₁" $
equalT empty (TYPE 2) (TYPE 0) (TYPE 1), equalT empty (^TYPE 2) (^TYPE 0) (^TYPE 1),
testNeq "★₁ ≠ ★₀" $ testNeq "★₁ ≠ ★₀" $
equalT empty (TYPE 2) (TYPE 1) (TYPE 0), equalT empty (^TYPE 2) (^TYPE 1) (^TYPE 0),
testEq "★₀ <: ★₀" $ testEq "★₀ <: ★₀" $
subT empty (TYPE 1) (TYPE 0) (TYPE 0), subT empty (^TYPE 1) (^TYPE 0) (^TYPE 0),
testEq "★₀ <: ★₁" $ testEq "★₀ <: ★₁" $
subT empty (TYPE 2) (TYPE 0) (TYPE 1), subT empty (^TYPE 2) (^TYPE 0) (^TYPE 1),
testNeq "★₁ ≮: ★₀" $ testNeq "★₁ ≮: ★₀" $
subT empty (TYPE 2) (TYPE 1) (TYPE 0) subT empty (^TYPE 2) (^TYPE 1) (^TYPE 0)
], ],
"function types" :- [ "function types" :- [
note #""𝐴𝐵" for (1·𝐴)𝐵"#, note "cumulativity",
note #""𝐴𝐵" for (0·𝐴)𝐵"#, testEq "0.★₀ → ★₀ = 0.★₀ → ★₀" $
testEq "★₀ ⇾ ★₀ = ★₀ ⇾ ★₀" $ let tm = ^Arr Zero (^TYPE 0) (^TYPE 0) in
let tm = Arr Zero (TYPE 0) (TYPE 0) in equalT empty (^TYPE 1) tm tm,
equalT empty (TYPE 1) tm tm, testEq "0.★₀ → ★₀ <: 0.★₀ → ★₀" $
testEq "★₀ ⇾ ★₀ <: ★₀ ⇾ ★₀" $ let tm = ^Arr Zero (^TYPE 0) (^TYPE 0) in
let tm = Arr Zero (TYPE 0) (TYPE 0) in subT empty (^TYPE 1) tm tm,
subT empty (TYPE 1) tm tm, testNeq "0.★₁ → ★₀ ≠ 0.★₀ → ★₀" $
testNeq "★₁ ⊸ ★₀ ≠ ★₀ ⇾ ★₀" $ let tm1 = ^Arr Zero (^TYPE 1) (^TYPE 0)
let tm1 = Arr Zero (TYPE 1) (TYPE 0) tm2 = ^Arr Zero (^TYPE 0) (^TYPE 0) in
tm2 = Arr Zero (TYPE 0) (TYPE 0) in equalT empty (^TYPE 2) tm1 tm2,
equalT empty (TYPE 2) tm1 tm2, testEq "1.★₁ → ★₀ <: 1.★₀ → ★₀" $
testEq "★₁ ⊸ ★₀ <: ★₀ ⊸ ★₀" $ let tm1 = ^Arr One (^TYPE 1) (^TYPE 0)
let tm1 = Arr One (TYPE 1) (TYPE 0) tm2 = ^Arr One (^TYPE 0) (^TYPE 0) in
tm2 = Arr One (TYPE 0) (TYPE 0) in subT empty (^TYPE 2) tm1 tm2,
subT empty (TYPE 2) tm1 tm2, testEq "1.★₀ → ★₀ <: 1.★₀ → ★₁" $
testEq "★₀ ⊸ ★₀ <: ★₀ ⊸ ★₁" $ let tm1 = ^Arr One (^TYPE 0) (^TYPE 0)
let tm1 = Arr One (TYPE 0) (TYPE 0) tm2 = ^Arr One (^TYPE 0) (^TYPE 1) in
tm2 = Arr One (TYPE 0) (TYPE 1) in subT empty (^TYPE 2) tm1 tm2,
subT empty (TYPE 2) tm1 tm2, testEq "1.★₀ → ★₀ <: 1.★₀ → ★₁" $
testEq "★₀ ⊸ ★₀ <: ★₀ ⊸ ★₁" $ let tm1 = ^Arr One (^TYPE 0) (^TYPE 0)
let tm1 = Arr One (TYPE 0) (TYPE 0) tm2 = ^Arr One (^TYPE 0) (^TYPE 1) in
tm2 = Arr One (TYPE 0) (TYPE 1) in subT empty (^TYPE 2) tm1 tm2,
subT empty (TYPE 2) tm1 tm2, testEq "1.A → B = 1.A → B" $
testEq "A ⊸ B = A ⊸ B" $ let tm = ^Arr One (^FT "A") (^FT "B") in
let tm = Arr One (FT "A") (FT "B") in equalT empty (^TYPE 0) tm tm,
equalT empty (TYPE 0) tm tm, testEq "1.A → B <: 1.A → B" $
testEq "A ⊸ B <: A ⊸ B" $ let tm = ^Arr One (^FT "A") (^FT "B") in
let tm = Arr One (FT "A") (FT "B") in subT empty (^TYPE 0) tm tm,
subT empty (TYPE 0) tm tm,
note "incompatible quantities", note "incompatible quantities",
testNeq "★₀ ⊸ ★₀ ≠ ★₀ ⇾ ★₁" $ testNeq "1.★₀ → ★₀ ≠ 0.★₀ → ★₁" $
let tm1 = Arr Zero (TYPE 0) (TYPE 0) let tm1 = ^Arr Zero (^TYPE 0) (^TYPE 0)
tm2 = Arr Zero (TYPE 0) (TYPE 1) in tm2 = ^Arr Zero (^TYPE 0) (^TYPE 1) in
equalT empty (TYPE 2) tm1 tm2, equalT empty (^TYPE 2) tm1 tm2,
testNeq "A ⇾ B ≠ A ⊸ B" $ testNeq "0.A → B ≠ 1.A → B" $
let tm1 = Arr Zero (FT "A") (FT "B") let tm1 = ^Arr Zero (^FT "A") (^FT "B")
tm2 = Arr One (FT "A") (FT "B") in tm2 = ^Arr One (^FT "A") (^FT "B") in
equalT empty (TYPE 0) tm1 tm2, equalT empty (^TYPE 0) tm1 tm2,
testNeq "A ⇾ B ≮: A ⊸ B" $ testNeq "0.A → B ≮: 1.A → B" $
let tm1 = Arr Zero (FT "A") (FT "B") let tm1 = ^Arr Zero (^FT "A") (^FT "B")
tm2 = Arr One (FT "A") (FT "B") in tm2 = ^Arr One (^FT "A") (^FT "B") in
subT empty (TYPE 0) tm1 tm2, subT empty (^TYPE 0) tm1 tm2,
testEq "0=1 ⊢ A ⇾ B = A ⊸ B" $ testEq "0=1 ⊢ 0.A → B = 1.A → B" $
let tm1 = Arr Zero (FT "A") (FT "B") let tm1 = ^Arr Zero (^FT "A") (^FT "B")
tm2 = Arr One (FT "A") (FT "B") in tm2 = ^Arr One (^FT "A") (^FT "B") in
equalT empty01 (TYPE 0) tm1 tm2, equalT empty01 (^TYPE 0) tm1 tm2,
todo "dependent function types", todo "dependent function types"
note "[todo] should π ≤ ρ ⊢ (ρ·A) → B <: (π·A) → B?"
], ],
"lambda" :- [ "lambda" :- [
testEq "λ x ⇒ [x] = λ x ⇒ [x]" $ testEq "λ x ⇒ x = λ x ⇒ x" $
equalT empty (Arr One (FT "A") (FT "A")) equalT empty (^Arr One (^FT "A") (^FT "A"))
([< "x"] :\\ BVT 0) (^LamY "x" (^BVT 0))
([< "x"] :\\ BVT 0), (^LamY "x" (^BVT 0)),
testEq "λ x ⇒ [x] <: λ x ⇒ [x]" $ testEq "λ x ⇒ x <: λ x ⇒ x" $
subT empty (Arr One (FT "A") (FT "A")) subT empty (^Arr One (^FT "A") (^FT "A"))
([< "x"] :\\ BVT 0) (^LamY "x" (^BVT 0))
([< "x"] :\\ BVT 0), (^LamY "x" (^BVT 0)),
testEq "λ x ⇒ [x] = λ y ⇒ [y]" $ testEq "λ x ⇒ x = λ y ⇒ y" $
equalT empty (Arr One (FT "A") (FT "A")) equalT empty (^Arr One (^FT "A") (^FT "A"))
([< "x"] :\\ BVT 0) (^LamY "x" (^BVT 0))
([< "y"] :\\ BVT 0), (^LamY "y" (^BVT 0)),
testEq "λ x ⇒ [x] <: λ y ⇒ [y]" $ testEq "λ x ⇒ x <: λ y ⇒ y" $
equalT empty (Arr One (FT "A") (FT "A")) subT empty (^Arr One (^FT "A") (^FT "A"))
([< "x"] :\\ BVT 0) (^LamY "x" (^BVT 0))
([< "y"] :\\ BVT 0), (^LamY "y" (^BVT 0)),
testNeq "λ x y ⇒ [x] ≠ λ x y ⇒ [y]" $ testNeq "λ x y ⇒ x ≠ λ x y ⇒ y" $
equalT empty (Arr One (FT "A") $ Arr One (FT "A") (FT "A")) equalT empty
([< "x", "y"] :\\ BVT 1) (^Arr One (^FT "A") (^Arr One (^FT "A") (^FT "A")))
([< "x", "y"] :\\ BVT 0), (^LamY "x" (^LamY "y" (^BVT 1)))
testEq "λ x ⇒ [a] = λ x ⇒ [a] (Y vs N)" $ (^LamY "x" (^LamY "y" (^BVT 0))),
equalT empty (Arr Zero (FT "B") (FT "A")) testEq "λ x ⇒ a = λ x ⇒ a (Y vs N)" $
(Lam $ SY [< "x"] $ FT "a") equalT empty
(Lam $ SN $ FT "a"), (^Arr Zero (^FT "B") (^FT "A"))
testEq "λ x ⇒ [f [x]] = [f] (η)" $ (^LamY "x" (^FT "a"))
equalT empty (Arr One (FT "A") (FT "A")) (^LamN (^FT "a")),
([< "x"] :\\ E (F "f" :@ BVT 0)) testEq "λ x ⇒ f x = f (η)" $
(FT "f") equalT empty
(^Arr One (^FT "A") (^FT "A"))
(^LamY "x" (E $ ^App (^F "f") (^BVT 0)))
(^FT "f")
], ],
"eq type" :- [ "eq type" :- [
testEq "(★₀ ≡ ★₀ : ★₁) = (★₀ ≡ ★₀ : ★₁)" $ testEq "(★₀ ≡ ★₀ : ★₁) = (★₀ ≡ ★₀ : ★₁)" $
let tm = Eq0 (TYPE 1) (TYPE 0) (TYPE 0) in let tm = ^Eq0 (^TYPE 1) (^TYPE 0) (^TYPE 0) in
equalT empty (TYPE 2) tm tm, equalT empty (^TYPE 2) tm tm,
testEq "A ≔ ★₁ ⊢ (★₀ ≡ ★₀ : ★₁) = (★₀ ≡ ★₀ : A)" testEq "A ≔ ★₁ ⊢ (★₀ ≡ ★₀ : ★₁) = (★₀ ≡ ★₀ : A)"
{globals = fromList [("A", mkDef gzero (TYPE 2) (TYPE 1))]} $ {globals = fromList [("A", ^mkDef gzero (^TYPE 2) (^TYPE 1))]} $
equalT empty (TYPE 2) equalT empty (^TYPE 2)
(Eq0 (TYPE 1) (TYPE 0) (TYPE 0)) (^Eq0 (^TYPE 1) (^TYPE 0) (^TYPE 0))
(Eq0 (FT "A") (TYPE 0) (TYPE 0)), (^Eq0 (^FT "A") (^TYPE 0) (^TYPE 0)),
todo "dependent equality types" todo "dependent equality types"
], ],
"equalities and uip" :- "equalities and uip" :-
let refl : Term d n -> Term d n -> Elim d n let refl : Term d n -> Term d n -> Elim d n
refl a x = (DLam $ S [< "_"] $ N x) :# (Eq0 a x x) refl a x = ^Ann (^DLam (SN x)) (^Eq0 a x x)
in in
[ [
note #""refl [A] x" is an abbreviation for "(δ i ⇒ x)(x ≡ x : A)""#,
note "binds before ∥ are globals, after it are BVs", note "binds before ∥ are globals, after it are BVs",
testEq "refl [A] a = refl [A] a" $ note #"refl A x is an abbreviation for "(δ i ⇒ x)(x ≡ x : A)""#,
equalE empty (refl (FT "A") (FT "a")) (refl (FT "A") (FT "a")), testEq "refl A a = refl A a" $
equalE empty (refl (^FT "A") (^FT "a")) (refl (^FT "A") (^FT "a")),
testEq "p : (a ≡ a' : A), q : (a ≡ a' : A) ∥ ⊢ p = q (free)" testEq "p : (a ≡ a' : A), q : (a ≡ a' : A) ∥ ⊢ p = q (free)"
{globals = {globals =
let def = mkPostulate gzero $ Eq0 (FT "A") (FT "a") (FT "a'") in let def = ^mkPostulate gzero (^Eq0 (^FT "A") (^FT "a") (^FT "a'"))
defGlobals `mergeLeft` fromList [("p", def), ("q", def)]} $ in defGlobals `mergeLeft` fromList [("p", def), ("q", def)]} $
equalE empty (F "p") (F "q"), equalE empty (^F "p") (^F "q"),
testEq "∥ x : (a ≡ a' : A), y : (a ≡ a' : A) ⊢ x = y (bound)" $ testEq "∥ x : (a ≡ a' : A), y : (a ≡ a' : A) ⊢ x = y (bound)" $
let ty : forall n. Term 0 n := Eq0 (FT "A") (FT "a") (FT "a'") in let ty : forall n. Term 0 n := ^Eq0 (^FT "A") (^FT "a") (^FT "a'") in
equalE (extendTyN [< (Any, "x", ty), (Any, "y", ty)] empty) equalE (extendTyN [< (Any, "x", ty), (Any, "y", ty)] empty)
(BV 0) (BV 1), (^BV 0) (^BV 1),
testEq "∥ x : [(a ≡ a' : A) ∷ Type 0], y : [ditto] ⊢ x = y" $ testEq "∥ x : (a ≡ a' : A) ∷ Type 0, y : [ditto] ⊢ x = y" $
let ty : forall n. Term 0 n := let ty : forall n. Term 0 n :=
E (Eq0 (FT "A") (FT "a") (FT "a'") :# TYPE 0) in E $ ^Ann (^Eq0 (^FT "A") (^FT "a") (^FT "a'")) (^TYPE 0) in
equalE (extendTyN [< (Any, "x", ty), (Any, "y", ty)] empty) equalE (extendTyN [< (Any, "x", ty), (Any, "y", ty)] empty)
(BV 0) (BV 1), (^BV 0) (^BV 1),
testEq "E ≔ a ≡ a' : A, EE ≔ E ∥ x : EE, y : EE ⊢ x = y" testEq "E ≔ a ≡ a' : A, EE ≔ E ∥ x : EE, y : EE ⊢ x = y"
{globals = defGlobals `mergeLeft` fromList {globals = defGlobals `mergeLeft` fromList
[("E", mkDef gzero (TYPE 0) (Eq0 (FT "A") (FT "a") (FT "a'"))), [("E", ^mkDef gzero (^TYPE 0)
("EE", mkDef gzero (TYPE 0) (FT "E"))]} $ (^Eq0 (^FT "A") (^FT "a") (^FT "a'"))),
equalE (extendTyN [< (Any, "x", FT "EE"), (Any, "y", FT "EE")] empty) ("EE", ^mkDef gzero (^TYPE 0) (^FT "E"))]} $
(BV 0) (BV 1), equalE (extendTyN [< (Any, "x", ^FT "EE"), (Any, "y", ^FT "EE")] empty)
(^BV 0) (^BV 1),
testEq "E ≔ a ≡ a' : A, EE ≔ E ∥ x : EE, y : E ⊢ x = y" testEq "E ≔ a ≡ a' : A, EE ≔ E ∥ x : EE, y : E ⊢ x = y"
{globals = defGlobals `mergeLeft` fromList {globals = defGlobals `mergeLeft` fromList
[("E", mkDef gzero (TYPE 0) (Eq0 (FT "A") (FT "a") (FT "a'"))), [("E", ^mkDef gzero (^TYPE 0)
("EE", mkDef gzero (TYPE 0) (FT "E"))]} $ (^Eq0 (^FT "A") (^FT "a") (^FT "a'"))),
equalE (extendTyN [< (Any, "x", FT "EE"), (Any, "y", FT "E")] empty) ("EE", ^mkDef gzero (^TYPE 0) (^FT "E"))]} $
(BV 0) (BV 1), equalE (extendTyN [< (Any, "x", ^FT "EE"), (Any, "y", ^FT "E")] empty)
(^BV 0) (^BV 1),
testEq "E ≔ a ≡ a' : A ∥ x : E, y : E ⊢ x = y" testEq "E ≔ a ≡ a' : A ∥ x : E, y : E ⊢ x = y"
{globals = defGlobals `mergeLeft` fromList {globals = defGlobals `mergeLeft` fromList
[("E", mkDef gzero (TYPE 0) (Eq0 (FT "A") (FT "a") (FT "a'")))]} $ [("E", ^mkDef gzero (^TYPE 0)
equalE (extendTyN [< (Any, "x", FT "E"), (Any, "y", FT "E")] empty) (^Eq0 (^FT "A") (^FT "a") (^FT "a'")))]} $
(BV 0) (BV 1), equalE (extendTyN [< (Any, "x", ^FT "E"), (Any, "y", ^FT "E")] empty)
(^BV 0) (^BV 1),
testEq "E ≔ a ≡ a' : A ∥ x : (E×E), y : (E×E) ⊢ x = y" testEq "E ≔ a ≡ a' : A ∥ x : (E×E), y : (E×E) ⊢ x = y"
{globals = defGlobals `mergeLeft` fromList {globals = defGlobals `mergeLeft` fromList
[("E", mkDef gzero (TYPE 0) (Eq0 (FT "A") (FT "a") (FT "a'")))]} $ [("E", ^mkDef gzero (^TYPE 0)
let ty : forall n. Term 0 n := (^Eq0 (^FT "A") (^FT "a") (^FT "a'")))]} $
Sig (FT "E") $ S [< "_"] $ N $ FT "E" in let ty : forall n. Term 0 n := ^Sig (^FT "E") (SN $ ^FT "E") in
equalE (extendTyN [< (Any, "x", ty), (Any, "y", ty)] empty) equalE (extendTyN [< (Any, "x", ty), (Any, "y", ty)] empty)
(BV 0) (BV 1), (^BV 0) (^BV 1),
testEq "E ≔ a ≡ a' : A, W ≔ E × E ∥ x : W, y : W ⊢ x = y" testEq "E ≔ a ≡ a' : A, W ≔ E × E ∥ x : W, y : E×E ⊢ x = y"
{globals = defGlobals `mergeLeft` fromList {globals = defGlobals `mergeLeft` fromList
[("E", mkDef gzero (TYPE 0) (Eq0 (FT "A") (FT "a") (FT "a'"))), [("E", ^mkDef gzero (^TYPE 0)
("W", mkDef gzero (TYPE 0) (FT "E" `And` FT "E"))]} $ (^Eq0 (^FT "A") (^FT "a") (^FT "a'"))),
("W", ^mkDef gzero (^TYPE 0) (^And (^FT "E") (^FT "E")))]} $
equalE equalE
(extendTyN [< (Any, "x", FT "W"), (Any, "y", FT "W")] empty) (extendTyN [< (Any, "x", ^FT "W"),
(BV 0) (BV 1) (Any, "y", ^And (^FT "E") (^FT "E"))] empty)
(^BV 0) (^BV 1)
], ],
"term closure" :- [ "term closure" :- [
testEq "[#0]{} = [#0] : A" $ note "bold numbers for de bruijn indices",
equalT (extendTy Any "x" (FT "A") empty) testEq "𝟎{} = 𝟎 : A" $
(FT "A") equalT (extendTy Any "x" (^FT "A") empty)
(CloT (Sub (BVT 0) id)) (^FT "A")
(BVT 0), (CloT (Sub (^BVT 0) id))
testEq "[#0]{a} = [a] : A" $ (^BVT 0),
equalT empty (FT "A") testEq "𝟎{a} = a : A" $
(CloT (Sub (BVT 0) (F "a" ::: id))) equalT empty (^FT "A")
(FT "a"), (CloT (Sub (^BVT 0) (^F "a" ::: id)))
testEq "[#0]{a,b} = [a] : A" $ (^FT "a"),
equalT empty (FT "A") testEq "𝟎{a,b} = a : A" $
(CloT (Sub (BVT 0) (F "a" ::: F "b" ::: id))) equalT empty (^FT "A")
(FT "a"), (CloT (Sub (^BVT 0) (^F "a" ::: ^F "b" ::: id)))
testEq "[#1]{a,b} = [b] : A" $ (^FT "a"),
equalT empty (FT "A") testEq "𝟏{a,b} = b : A" $
(CloT (Sub (BVT 1) (F "a" ::: F "b" ::: id))) equalT empty (^FT "A")
(FT "b"), (CloT (Sub (^BVT 1) (^F "a" ::: ^F "b" ::: id)))
testEq "(λy ⇒ [#1]){a} = λy ⇒ [a] : B ⇾ A (N)" $ (^FT "b"),
equalT empty (Arr Zero (FT "B") (FT "A")) testEq "(λy ⇒ 𝟏){a} = λy ⇒ a : 0.B → A (N)" $
(CloT (Sub (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id))) equalT empty (^Arr Zero (^FT "B") (^FT "A"))
(Lam $ S [< "y"] $ N $ FT "a"), (CloT (Sub (^LamN (^BVT 0)) (^F "a" ::: id)))
testEq "(λy ⇒ [#1]){a} = λy ⇒ [a] : B ⇾ A (Y)" $ (^LamN (^FT "a")),
equalT empty (Arr Zero (FT "B") (FT "A")) testEq "(λy ⇒ 𝟏){a} = λy ⇒ a : 0.B → A (Y)" $
(CloT (Sub ([< "y"] :\\ BVT 1) (F "a" ::: id))) equalT empty (^Arr Zero (^FT "B") (^FT "A"))
([< "y"] :\\ FT "a") (CloT (Sub (^LamY "y" (^BVT 1)) (^F "a" ::: id)))
(^LamY "y" (^FT "a"))
], ],
"term d-closure" :- [ "term d-closure" :- [
testEq "★₀‹𝟎› = ★₀ : ★₁" $ testEq "★₀0 = ★₀ : ★₁" $
equalTD 1 equalT (extendDim "𝑗" empty)
(extendDim "𝑗" empty) (^TYPE 1) (DCloT (Sub (^TYPE 0) (^K Zero ::: id))) (^TYPE 0),
(TYPE 1) (DCloT (Sub (TYPE 0) (K Zero ::: id))) (TYPE 0), testEq "(δ i ⇒ a)0 = (δ i ⇒ a) : (a ≡ a : A)" $
testEq "(δ i ⇒ a)𝟎 = (δ i ⇒ a) : (a ≡ a : A)" $ equalT (extendDim "𝑗" empty)
equalTD 1 (^Eq0 (^FT "A") (^FT "a") (^FT "a"))
(extendDim "𝑗" empty) (DCloT (Sub (^DLamN (^FT "a")) (^K Zero ::: id)))
(Eq0 (FT "A") (FT "a") (FT "a")) (^DLamN (^FT "a")),
(DCloT (Sub ([< "i"] :\\% FT "a") (K Zero ::: id)))
([< "i"] :\\% FT "a"),
note "it is hard to think of well-typed terms with big dctxs" note "it is hard to think of well-typed terms with big dctxs"
], ],
"free var" :- "free var" :-
let au_bu = fromList let au_bu = fromList
[("A", mkDef gany (TYPE 1) (TYPE 0)), [("A", ^mkDef gany (^TYPE 1) (^TYPE 0)),
("B", mkDef gany (TYPE 1) (TYPE 0))] ("B", ^mkDef gany (^TYPE 1) (^TYPE 0))]
au_ba = fromList au_ba = fromList
[("A", mkDef gany (TYPE 1) (TYPE 0)), [("A", ^mkDef gany (^TYPE 1) (^TYPE 0)),
("B", mkDef gany (TYPE 1) (FT "A"))] ("B", ^mkDef gany (^TYPE 1) (^FT "A"))]
in [ in [
testEq "A = A" $ testEq "A = A" $
equalE empty (F "A") (F "A"), equalE empty (^F "A") (^F "A"),
testNeq "A ≠ B" $ testNeq "A ≠ B" $
equalE empty (F "A") (F "B"), equalE empty (^F "A") (^F "B"),
testEq "0=1 ⊢ A = B" $ testEq "0=1 ⊢ A = B" $
equalE empty01 (F "A") (F "B"), equalE empty01 (^F "A") (^F "B"),
testEq "A : ★₁ ≔ ★₀ ⊢ A = (★₀ ∷ ★₁)" {globals = au_bu} $ testEq "A : ★₁ ≔ ★₀ ⊢ A = (★₀ ∷ ★₁)" {globals = au_bu} $
equalE empty (F "A") (TYPE 0 :# TYPE 1), equalE empty (^F "A") (^Ann (^TYPE 0) (^TYPE 1)),
testEq "A : ★₁ ≔ ★₀ ⊢ [A] = ★₀" {globals = au_bu} $ testEq "A : ★₁ ≔ ★₀ ⊢ A = ★₀" {globals = au_bu} $
equalT empty (TYPE 1) (FT "A") (TYPE 0), equalT empty (^TYPE 1) (^FT "A") (^TYPE 0),
testEq "A ≔ ★₀, B ≔ ★₀ ⊢ A = B" {globals = au_bu} $ testEq "A ≔ ★₀, B ≔ ★₀ ⊢ A = B" {globals = au_bu} $
equalE empty (F "A") (F "B"), equalE empty (^F "A") (^F "B"),
testEq "A ≔ ★₀, B ≔ A ⊢ A = B" {globals = au_ba} $ testEq "A ≔ ★₀, B ≔ A ⊢ A = B" {globals = au_ba} $
equalE empty (F "A") (F "B"), equalE empty (^F "A") (^F "B"),
testEq "A <: A" $ testEq "A <: A" $
subE empty (F "A") (F "A"), subE empty (^F "A") (^F "A"),
testNeq "A ≮: B" $ testNeq "A ≮: B" $
subE empty (F "A") (F "B"), subE empty (^F "A") (^F "B"),
testEq "A : ★₃ ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B" testEq "A : ★₃ ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B"
{globals = fromList [("A", mkDef gany (TYPE 3) (TYPE 0)), {globals = fromList [("A", ^mkDef gany (^TYPE 3) (^TYPE 0)),
("B", mkDef gany (TYPE 3) (TYPE 2))]} $ ("B", ^mkDef gany (^TYPE 3) (^TYPE 2))]} $
subE empty (F "A") (F "B"), subE empty (^F "A") (^F "B"),
note "(A and B in different universes)", note "(A and B in different universes)",
testEq "A : ★₁ ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B" testEq "A : ★₁ ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B"
{globals = fromList [("A", mkDef gany (TYPE 1) (TYPE 0)), {globals = fromList [("A", ^mkDef gany (^TYPE 1) (^TYPE 0)),
("B", mkDef gany (TYPE 3) (TYPE 2))]} $ ("B", ^mkDef gany (^TYPE 3) (^TYPE 2))]} $
subE empty (F "A") (F "B"), subE empty (^F "A") (^F "B"),
testEq "0=1 ⊢ A <: B" $ testEq "0=1 ⊢ A <: B" $
subE empty01 (F "A") (F "B") subE empty01 (^F "A") (^F "B")
], ],
"bound var" :- [ "bound var" :- [
testEq "#0 = #0" $ note "bold numbers for de bruijn indices",
equalE (extendTy Any "A" (TYPE 0) empty) (BV 0) (BV 0), testEq "𝟎 = 𝟎" $
testEq "#0 <: #0" $ equalE (extendTy Any "A" (^TYPE 0) empty) (^BV 0) (^BV 0),
subE (extendTy Any "A" (TYPE 0) empty) (BV 0) (BV 0), testEq "𝟎 <: 𝟎" $
testNeq "#0 ≠ #1" $ subE (extendTy Any "A" (^TYPE 0) empty) (^BV 0) (^BV 0),
equalE (extendTyN [< (Any, "A", TYPE 0), (Any, "B", TYPE 0)] empty) testNeq "𝟎𝟏" $
(BV 0) (BV 1), equalE (extendTyN [< (Any, "A", ^TYPE 0), (Any, "B", ^TYPE 0)] empty)
testNeq "#0 ≮: #1" $ (^BV 0) (^BV 1),
subE (extendTyN [< (Any, "A", TYPE 0), (Any, "B", TYPE 0)] empty) testNeq "𝟎 ≮: 𝟏" $
(BV 0) (BV 1), subE (extendTyN [< (Any, "A", ^TYPE 0), (Any, "B", ^TYPE 0)] empty)
testEq "0=1 ⊢ #0 = #1" $ (^BV 0) (^BV 1),
equalE (extendTyN [< (Any, "A", TYPE 0), (Any, "B", TYPE 0)] empty01) testEq "0=1 ⊢ 𝟎 = 𝟏" $
(BV 0) (BV 1) equalE (extendTyN [< (Any, "A", ^TYPE 0), (Any, "B", ^TYPE 0)] empty01)
(^BV 0) (^BV 1)
], ],
"application" :- [ "application" :- [
testEq "f [a] = f [a]" $ testEq "f a = f a" $
equalE empty (F "f" :@ FT "a") (F "f" :@ FT "a"), equalE empty (^App (^F "f") (^FT "a")) (^App (^F "f") (^FT "a")),
testEq "f [a] <: f [a]" $ testEq "f a <: f a" $
subE empty (F "f" :@ FT "a") (F "f" :@ FT "a"), subE empty (^App (^F "f") (^FT "a")) (^App (^F "f") (^FT "a")),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a = ([a ∷ A] ∷ A) (β)" $ testEq "(λ x ⇒ x ∷ 1.A → A) a = ((a ∷ A) ∷ A) (β)" $
equalE empty equalE empty
((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a") (^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
(E (FT "a" :# FT "A") :# FT "A"), (^FT "a"))
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a = a (βυ)" $ (^Ann (E $ ^Ann (^FT "a") (^FT "A")) (^FT "A")),
testEq "(λ x ⇒ x ∷ A ⊸ A) a = a (βυ)" $
equalE empty equalE empty
((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a") (^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
(F "a"), (^FT "a"))
testEq "(λ g ⇒ [g [a]] ∷ ⋯)) [f] = (λ y ⇒ [f [y]] ∷ ⋯) [a] (β↘↙)" $ (^F "a"),
let a = FT "A"; a2a = (Arr One a a) in testEq "((λ g ⇒ g a) ∷ 1.(1.A → A) → A) f = ((λ y ⇒ f y) ∷ 1.A → A) a # β↘↙" $
let a = ^FT "A"; a2a = ^Arr One a a; aa2a = ^Arr One a2a a in
equalE empty equalE empty
((([< "g"] :\\ E (BV 0 :@ FT "a")) :# Arr One a2a a) :@ FT "f") (^App (^Ann (^LamY "g" (E $ ^App (^BV 0) (^FT "a"))) aa2a) (^FT "f"))
((([< "y"] :\\ E (F "f" :@ BVT 0)) :# a2a) :@ FT "a"), (^App (^Ann (^LamY "y" (E $ ^App (^F "f") (^BVT 0))) a2a) (^FT "a")),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a <: a" $ testEq "((λ x ⇒ x) ∷ 1.A → A) a <: a" $
subE empty subE empty
((([< "x"] :\\ BVT 0) :# (Arr One (FT "A") (FT "A"))) :@ FT "a") (^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
(F "a"), (^FT "a"))
note "id : A ⊸ A ≔ λ x ⇒ [x]", (^F "a"),
testEq "id [a] = a" $ equalE empty (F "id" :@ FT "a") (F "a"), note "id : A ⊸ A ≔ λ x ⇒ x",
testEq "id [a] <: a" $ subE empty (F "id" :@ FT "a") (F "a") testEq "id a = a" $ equalE empty (^App (^F "id") (^FT "a")) (^F "a"),
testEq "id a <: a" $ subE empty (^App (^F "id") (^FT "a")) (^F "a")
], ],
"dim application" :- [ "dim application" :- [
testEq "eq-AB @0 = eq-AB @0" $ testEq "eq-AB @0 = eq-AB @0" $
equalE empty (F "eq-AB" :% K Zero) (F "eq-AB" :% K Zero), equalE empty
(^DApp (^F "eq-AB") (^K Zero))
(^DApp (^F "eq-AB") (^K Zero)),
testNeq "eq-AB @0 ≠ eq-AB @1" $ testNeq "eq-AB @0 ≠ eq-AB @1" $
equalE empty (F "eq-AB" :% K Zero) (F "eq-AB" :% K One), equalE empty
(^DApp (^F "eq-AB") (^K Zero))
(^DApp (^F "eq-AB") (^K One)),
testEq "𝑖 | ⊢ eq-AB @𝑖 = eq-AB @𝑖" $ testEq "𝑖 | ⊢ eq-AB @𝑖 = eq-AB @𝑖" $
equalED 1 equalE
(extendDim "𝑖" empty) (extendDim "𝑖" empty)
(F "eq-AB" :% BV 0) (F "eq-AB" :% BV 0), (^DApp (^F "eq-AB") (^BV 0))
(^DApp (^F "eq-AB") (^BV 0)),
testNeq "𝑖 | ⊢ eq-AB @𝑖 ≠ eq-AB @0" $ testNeq "𝑖 | ⊢ eq-AB @𝑖 ≠ eq-AB @0" $
equalED 1 equalE (extendDim "𝑖" empty)
(extendDim "𝑖" empty) (^DApp (^F "eq-AB") (^BV 0))
(F "eq-AB" :% BV 0) (F "eq-AB" :% K Zero), (^DApp (^F "eq-AB") (^K Zero)),
testEq "𝑖, 𝑖=0 | ⊢ eq-AB @𝑖 = eq-AB @0" $ testEq "𝑖, 𝑖=0 | ⊢ eq-AB @𝑖 = eq-AB @0" $
equalED 1 equalE (eqDim (^BV 0) (^K Zero) $ extendDim "𝑖" empty)
(eqDim (BV 0) (K Zero) $ extendDim "𝑖" empty) (^DApp (^F "eq-AB") (^BV 0))
(F "eq-AB" :% BV 0) (F "eq-AB" :% K Zero), (^DApp (^F "eq-AB") (^K Zero)),
testNeq "𝑖, 𝑖=1 | ⊢ eq-AB @𝑖 ≠ eq-AB @0" $ testNeq "𝑖, 𝑖=1 | ⊢ eq-AB @𝑖 ≠ eq-AB @0" $
equalED 1 equalE (eqDim (^BV 0) (^K One) $ extendDim "𝑖" empty)
(eqDim (BV 0) (K One) $ extendDim "𝑖" empty) (^DApp (^F "eq-AB") (^BV 0))
(F "eq-AB" :% BV 0) (F "eq-AB" :% K Zero), (^DApp (^F "eq-AB") (^K Zero)),
testNeq "𝑖, 𝑗 | ⊢ eq-AB @𝑖 ≠ eq-AB @𝑗" $ testNeq "𝑖, 𝑗 | ⊢ eq-AB @𝑖 ≠ eq-AB @𝑗" $
equalED 2 equalE (extendDim "𝑗" $ extendDim "𝑖" empty)
(extendDim "𝑗" $ extendDim "𝑖" empty) (^DApp (^F "eq-AB") (^BV 1))
(F "eq-AB" :% BV 1) (F "eq-AB" :% BV 0), (^DApp (^F "eq-AB") (^BV 0)),
testEq "𝑖, 𝑗, 𝑖=𝑗 | ⊢ eq-AB @𝑖 = eq-AB @𝑗" $ testEq "𝑖, 𝑗, 𝑖=𝑗 | ⊢ eq-AB @𝑖 = eq-AB @𝑗" $
equalED 2 equalE (eqDim (^BV 0) (^BV 1) $ extendDim "𝑗" $ extendDim "𝑖" empty)
(eqDim (BV 0) (BV 1) $ extendDim "𝑗" $ extendDim "𝑖" empty) (^DApp (^F "eq-AB") (^BV 1))
(F "eq-AB" :% BV 1) (F "eq-AB" :% BV 0), (^DApp (^F "eq-AB") (^BV 0)),
testEq "𝑖, 𝑗, 𝑖=0, 𝑗=0 | ⊢ eq-AB @𝑖 = eq-AB @𝑗" $ testEq "𝑖, 𝑗, 𝑖=0, 𝑗=0 | ⊢ eq-AB @𝑖 = eq-AB @𝑗" $
equalED 2 equalE
(eqDim (BV 0) (K Zero) $ eqDim (BV 1) (K Zero) $ (eqDim (^BV 0) (^K Zero) $ eqDim (^BV 1) (^K Zero) $
extendDim "𝑗" $ extendDim "𝑖" empty) extendDim "𝑗" $ extendDim "𝑖" empty)
(F "eq-AB" :% BV 1) (F "eq-AB" :% BV 0), (^DApp (^F "eq-AB") (^BV 1))
(^DApp (^F "eq-AB") (^BV 0)),
testEq "0=1 | ⊢ eq-AB @𝑖 = eq-AB @𝑗" $ testEq "0=1 | ⊢ eq-AB @𝑖 = eq-AB @𝑗" $
equalED 2 equalE (extendDim "𝑗" $ extendDim "𝑖" empty01)
(extendDim "𝑗" $ extendDim "𝑖" empty01) (^DApp (^F "eq-AB") (^BV 1))
(F "eq-AB" :% BV 1) (F "eq-AB" :% BV 0), (^DApp (^F "eq-AB") (^BV 0)),
testEq "eq-AB @0 = A" $ equalE empty (F "eq-AB" :% K Zero) (F "A"), testEq "eq-AB @0 = A" $
testEq "eq-AB @1 = B" $ equalE empty (F "eq-AB" :% K One) (F "B"), equalE empty (^DApp (^F "eq-AB") (^K Zero)) (^F "A"),
testEq "((δ i ⇒ a) ∷ a ≡ a) @0 = a" $ testEq "eq-AB @1 = B" $
equalE empty (^DApp (^F "eq-AB") (^K One)) (^F "B"),
testEq "((δ i ⇒ a) ∷ a ≡ a : A) @0 = a" $
equalE empty equalE empty
(((DLam $ SN $ FT "a") :# Eq0 (FT "A") (FT "a") (FT "a")) :% K Zero) (^DApp (^Ann (^DLamN (^FT "a")) (^Eq0 (^FT "A") (^FT "a") (^FT "a")))
(F "a"), (^K Zero))
testEq "((δ i ⇒ a) ∷ a ≡ a) @0 = ((δ i ⇒ a) ∷ a ≡ a) @1" $ (^F "a"),
testEq "((δ i ⇒ a) ∷ a ≡ a : A) @0 = ((δ i ⇒ a) ∷ a ≡ a : A) @1" $
equalE empty equalE empty
(((DLam $ SN $ FT "a") :# Eq0 (FT "A") (FT "a") (FT "a")) :% K Zero) (^DApp (^Ann (^DLamN (^FT "a")) (^Eq0 (^FT "A") (^FT "a") (^FT "a")))
(((DLam $ SN $ FT "a") :# Eq0 (FT "A") (FT "a") (FT "a")) :% K One) (^K Zero))
(^DApp (^Ann (^DLamN (^FT "a")) (^Eq0 (^FT "A") (^FT "a") (^FT "a")))
(^K One))
], ],
"annotation" :- [ "annotation" :- [
testEq "(λ x ⇒ f [x]) ∷ A ⊸ A = [f] ∷ A ⊸ A" $ testEq "(λ x ⇒ f x) ∷ 1.A → A = f ∷ 1.A → A" $
equalE empty equalE empty
(([< "x"] :\\ E (F "f" :@ BVT 0)) :# Arr One (FT "A") (FT "A")) (^Ann (^LamY "x" (E $ ^App (^F "f") (^BVT 0)))
(FT "f" :# Arr One (FT "A") (FT "A")), (^Arr One (^FT "A") (^FT "A")))
testEq "[f] ∷ A ⊸ A = f" $ (^Ann (^FT "f") (^Arr One (^FT "A") (^FT "A"))),
equalE empty (FT "f" :# Arr One (FT "A") (FT "A")) (F "f"), testEq "f ∷ 1.A → A = f" $
testEq "(λ x ⇒ f [x]) ∷ A ⊸ A = f" $
equalE empty equalE empty
(([< "x"] :\\ E (F "f" :@ BVT 0)) :# Arr One (FT "A") (FT "A")) (^Ann (^FT "f") (^Arr One (^FT "A") (^FT "A")))
(F "f") (^F "f"),
testEq "(λ x ⇒ f x) ∷ 1.A → A = f" $
equalE empty
(^Ann (^LamY "x" (E $ ^App (^F "f") (^BVT 0)))
(^Arr One (^FT "A") (^FT "A")))
(^F "f")
], ],
"natural type" :- [ "natural type" :- [
testEq " = " $ equalTy empty Nat Nat, testEq " = " $ equalTy empty (^Nat) (^Nat),
testEq " = : ★₀" $ equalT empty (TYPE 0) Nat Nat, testEq " = : ★₀" $ equalT empty (^TYPE 0) (^Nat) (^Nat),
testEq " = : ★₆₉" $ equalT empty (TYPE 69) Nat Nat, testEq " = : ★₆₉" $ equalT empty (^TYPE 69) (^Nat) (^Nat),
testNeq " ≠ {}" $ equalTy empty Nat (enum []), testNeq " ≠ {}" $ equalTy empty (^Nat) (^enum []),
testEq "0=1 ⊢ = {}" $ equalTy empty01 Nat (enum []) testEq "0=1 ⊢ = {}" $ equalTy empty01 (^Nat) (^enum [])
], ],
"natural numbers" :- [ "natural numbers" :- [
testEq "zero = zero" $ equalT empty Nat Zero Zero, testEq "0 = 0" $ equalT empty (^Nat) (^Zero) (^Zero),
testEq "succ two = succ two" $ testEq "succ two = succ two" $
equalT empty Nat (Succ (FT "two")) (Succ (FT "two")), equalT empty (^Nat) (^Succ (^FT "two")) (^Succ (^FT "two")),
testNeq "succ two ≠ two" $ testNeq "succ two ≠ two" $
equalT empty Nat (Succ (FT "two")) (FT "two"), equalT empty (^Nat) (^Succ (^FT "two")) (^FT "two"),
testNeq "zero ≠ succ zero" $ testNeq "0 ≠ 1" $
equalT empty Nat Zero (Succ Zero), equalT empty (^Nat) (^Zero) (^Succ (^Zero)),
testEq "0=1 ⊢ zero = succ zero" $ testEq "0=1 ⊢ 0 = 1" $
equalT empty01 Nat Zero (Succ Zero) equalT empty01 (^Nat) (^Zero) (^Succ (^Zero))
], ],
"natural elim" :- [ "natural elim" :- [
testEq "caseω 0 return {a,b} of {zero ⇒ 'a; succ _ ⇒ 'b} = 'a" $ testEq "caseω 0 return {a,b} of {zero ⇒ 'a; succ _ ⇒ 'b} = 'a" $
equalT empty equalT empty
(enum ["a", "b"]) (^enum ["a", "b"])
(E $ CaseNat Any Zero (Zero :# Nat) (E $ ^CaseNat Any Zero (^Ann (^Zero) (^Nat))
(SN $ enum ["a", "b"]) (SN $ ^enum ["a", "b"])
(Tag "a") (^Tag "a")
(SN $ Tag "b")) (SN $ ^Tag "b"))
(Tag "a"), (^Tag "a"),
testEq "caseω 1 return {a,b} of {zero ⇒ 'a; succ _ ⇒ 'b} = 'b" $ testEq "caseω 1 return {a,b} of {zero ⇒ 'a; succ _ ⇒ 'b} = 'b" $
equalT empty equalT empty
(enum ["a", "b"]) (^enum ["a", "b"])
(E $ CaseNat Any Zero (Succ Zero :# Nat) (E $ ^CaseNat Any Zero (^Ann (^Succ (^Zero)) (^Nat))
(SN $ enum ["a", "b"]) (SN $ ^enum ["a", "b"])
(Tag "a") (^Tag "a")
(SN $ Tag "b")) (SN $ ^Tag "b"))
(Tag "b"), (^Tag "b"),
testEq "caseω 4 return of {0 ⇒ 0; succ n ⇒ n} = 3" $ testEq "caseω 4 return of {0 ⇒ 0; succ n ⇒ n} = 3" $
equalT empty equalT empty
Nat (^Nat)
(E $ CaseNat Any Zero (makeNat 4 :# Nat) (E $ ^CaseNat Any Zero (^Ann (^makeNat 4) (^Nat))
(SN Nat) (SN $ ^Nat)
Zero (^Zero)
(SY [< "n", Unused] $ BVT 1)) (SY [< "n", ^BN Unused] $ ^BVT 1))
(makeNat 3) (^makeNat 3)
], ],
todo "pair types", todo "pair types",
@ -472,24 +484,24 @@ tests = "equality & subtyping" :- [
"pairs" :- [ "pairs" :- [
testEq "('a, 'b) = ('a, 'b) : {a} × {b}" $ testEq "('a, 'b) = ('a, 'b) : {a} × {b}" $
equalT empty equalT empty
(enum ["a"] `And` enum ["b"]) (^And (^enum ["a"]) (^enum ["b"]))
(Tag "a" `Pair` Tag "b") (^Pair (^Tag "a") (^Tag "b"))
(Tag "a" `Pair` Tag "b"), (^Pair (^Tag "a") (^Tag "b")),
testNeq "('a, 'b) ≠ ('b, 'a) : {a,b} × {a,b}" $ testNeq "('a, 'b) ≠ ('b, 'a) : {a,b} × {a,b}" $
equalT empty equalT empty
(enum ["a", "b"] `And` enum ["a", "b"]) (^And (^enum ["a", "b"]) (^enum ["a", "b"]))
(Tag "a" `Pair` Tag "b") (^Pair (^Tag "a") (^Tag "b"))
(Tag "b" `Pair` Tag "a"), (^Pair (^Tag "b") (^Tag "a")),
testEq "0=1 ⊢ ('a, 'b) = ('b, 'a) : {a,b} × {a,b}" $ testEq "0=1 ⊢ ('a, 'b) = ('b, 'a) : {a,b} × {a,b}" $
equalT empty01 equalT empty01
(enum ["a", "b"] `And` enum ["a", "b"]) (^And (^enum ["a", "b"]) (^enum ["a", "b"]))
(Tag "a" `Pair` Tag "b") (^Pair (^Tag "a") (^Tag "b"))
(Tag "b" `Pair` Tag "a"), (^Pair (^Tag "b") (^Tag "a")),
testEq "0=1 ⊢ ('a, 'b) = ('b, 'a) : " $ testEq "0=1 ⊢ ('a, 'b) = ('b, 'a) : " $
equalT empty01 equalT empty01
Nat (^Nat)
(Tag "a" `Pair` Tag "b") (^Pair (^Tag "a") (^Tag "b"))
(Tag "b" `Pair` Tag "a") (^Pair (^Tag "b") (^Tag "a"))
], ],
todo "pair elim", todo "pair elim",
@ -503,61 +515,60 @@ tests = "equality & subtyping" :- [
todo "box elim", todo "box elim",
"elim closure" :- [ "elim closure" :- [
testEq "#0{a} = a" $ note "bold numbers for de bruijn indices",
equalE empty (CloE (Sub (BV 0) (F "a" ::: id))) (F "a"), testEq "𝟎{a} = a" $
testEq "#1{a} = #0" $ equalE empty (CloE (Sub (^BV 0) (^F "a" ::: id))) (^F "a"),
equalE (extendTy Any "x" (FT "A") empty) testEq "𝟏{a} = 𝟎" $
(CloE (Sub (BV 1) (F "a" ::: id))) (BV 0) equalE (extendTy Any "x" (^FT "A") empty)
(CloE (Sub (^BV 1) (^F "a" ::: id))) (^BV 0)
], ],
"elim d-closure" :- [ "elim d-closure" :- [
note "bold numbers for de bruijn indices",
note "0·eq-AB : (A ≡ B : ★₀)", note "0·eq-AB : (A ≡ B : ★₀)",
testEq "(eq-AB #0)𝟎 = eq-AB 𝟎" $ testEq "(eq-AB @𝟎)0 = eq-AB @0" $
equalED 1 equalE empty
(extendDim "𝑖" empty) (DCloE (Sub (^DApp (^F "eq-AB") (^BV 0)) (^K Zero ::: id)))
(DCloE (Sub (F "eq-AB" :% BV 0) (K Zero ::: id))) (^DApp (^F "eq-AB") (^K Zero)),
(F "eq-AB" :% K Zero), testEq "(eq-AB @𝟎)0 = A" $
testEq "(eq-AB #0)𝟎 = A" $ equalE empty
equalED 1 (DCloE (Sub (^DApp (^F "eq-AB") (^BV 0)) (^K Zero ::: id)))
(extendDim "𝑖" empty) (^F "A"),
(DCloE (Sub (F "eq-AB" :% BV 0) (K Zero ::: id))) (F "A"), testEq "(eq-AB @𝟎)1 = B" $
testEq "(eq-AB #0)𝟏 = B" $ equalE empty
equalED 1 (DCloE (Sub (^DApp (^F "eq-AB") (^BV 0)) (^K One ::: id)))
(extendDim "𝑖" empty) (^F "B"),
(DCloE (Sub (F "eq-AB" :% BV 0) (K One ::: id))) (F "B"), testNeq "(eq-AB @𝟎)1 ≠ A" $
testNeq "(eq-AB #0)𝟏 ≠ A" $ equalE empty
equalED 1 (DCloE (Sub (^DApp (^F "eq-AB") (^BV 0)) (^K One ::: id)))
(extendDim "𝑖" empty) (^F "A"),
(DCloE (Sub (F "eq-AB" :% BV 0) (K One ::: id))) (F "A"), testEq "(eq-AB @𝟎)𝟎,0 = (eq-AB 𝟎)" $
testEq "(eq-AB #0)#0,𝟎 = (eq-AB #0)" $ equalE (extendDim "𝑖" empty)
equalED 2 (DCloE (Sub (^DApp (^F "eq-AB") (^BV 0)) (^BV 0 ::: ^K Zero ::: id)))
(extendDim "𝑗" $ extendDim "𝑖" empty) (^DApp (^F "eq-AB") (^BV 0)),
(DCloE (Sub (F "eq-AB" :% BV 0) (BV 0 ::: K Zero ::: id))) testNeq "(eq-AB 𝟎)0 ≠ (eq-AB 0)" $
(F "eq-AB" :% BV 0), equalE (extendDim "𝑖" empty)
testNeq "(eq-AB #0)𝟎 ≠ (eq-AB 𝟎)" $ (DCloE (Sub (^DApp (^F "eq-AB") (^BV 0)) (^BV 0 ::: ^K Zero ::: id)))
equalED 2 (^DApp (^F "eq-AB") (^K Zero)),
(extendDim "𝑗" $ extendDim "𝑖" empty) testEq "𝟎0 = 𝟎 # term and dim vars distinct" $
(DCloE (Sub (F "eq-AB" :% BV 0) (BV 0 ::: K Zero ::: id))) equalE
(F "eq-AB" :% K Zero), (extendTy Any "x" (^FT "A") empty)
testEq "#0𝟎 = #0 # term and dim vars distinct" $ (DCloE (Sub (^BV 0) (^K Zero ::: id))) (^BV 0),
equalED 1 testEq "a0 = a" $
(extendTy Any "x" (FT "A") $ extendDim "𝑖" empty) equalE empty
(DCloE (Sub (BV 0) (K Zero ::: id))) (BV 0), (DCloE (Sub (^F "a") (^K Zero ::: id))) (^F "a"),
testEq "a𝟎 = a" $ testEq "(f a)0 = f0 a0" $
equalED 1 (extendDim "𝑖" empty) let th = ^K Zero ::: id in
(DCloE (Sub (F "a") (K Zero ::: id))) (F "a"), equalE empty
testEq "(f [a])𝟎 = f𝟎 [a]𝟎" $ (DCloE (Sub (^App (^F "f") (^FT "a")) th))
let th = K Zero ::: id in (^App (DCloE (Sub (^F "f") th)) (DCloT (Sub (^FT "a") th)))
equalED 1 (extendDim "𝑖" empty)
(DCloE (Sub (F "f" :@ FT "a") th))
(DCloE (Sub (F "f") th) :@ DCloT (Sub (FT "a") th))
], ],
"clashes" :- [ "clashes" :- [
testNeq "★₀ ≠ ★₀ ⇾ ★₀" $ testNeq "★₀ ≠ 0.★₀ → ★₀" $
equalT empty (TYPE 1) (TYPE 0) (Arr Zero (TYPE 0) (TYPE 0)), equalT empty (^TYPE 1) (^TYPE 0) (^Arr Zero (^TYPE 0) (^TYPE 0)),
testEq "0=1 ⊢ ★₀ = ★₀ ⇾ ★₀" $ testEq "0=1 ⊢ ★₀ = 0.★₀ → ★₀" $
equalT empty01 (TYPE 1) (TYPE 0) (Arr Zero (TYPE 0) (TYPE 0)), equalT empty01 (^TYPE 1) (^TYPE 0) (^Arr Zero (^TYPE 0) (^TYPE 0)),
todo "others" todo "others"
] ]
] ]

View File

@ -4,8 +4,9 @@ import Quox.Parser.FromParser
import Quox.Parser import Quox.Parser
import TypingImpls import TypingImpls
import Tests.Parser as TParser import Tests.Parser as TParser
import TAP
import Quox.EffExtra import Quox.EffExtra
import TAP
import AstExtra
import System.File import System.File
import Derive.Prelude import Derive.Prelude
@ -49,6 +50,11 @@ parameters {c : Bool} {auto _ : Show b}
parses : Test parses : Test
parses = parsesWith $ const True parses = parsesWith $ const True
%macro
parseMatch : TTImp -> Elab Test
parseMatch pat =
parsesWith <$> check `(\case ~(pat) => True; _ => False)
parsesAs : Eq b => b -> Test parsesAs : Eq b => b -> Test
parsesAs exp = parsesWith (== exp) parsesAs exp = parsesWith (== exp)
@ -59,11 +65,9 @@ parameters {c : Bool} {auto _ : Show b}
either (const $ Right ()) (Left . ExpectedFail . show) $ fromP pres either (const $ Right ()) (Left . ExpectedFail . show) $ fromP pres
FromString PatVar where fromString x = PV x Nothing
runFromParser : {default empty defs : Definitions} -> runFromParser : {default empty defs : Definitions} ->
Eff FromParserPure a -> Either FromParser.Error a Eff FromParserPure a -> Either FromParser.Error a
runFromParser = map fst . fromParserPure defs runFromParser = map fst . fst . fromParserPure 0 defs
export export
tests : Test tests : Test
@ -72,30 +76,35 @@ tests = "PTerm → Term" :- [
let fromPDim = runFromParser . fromPDimWith [< "𝑖", "𝑗"] let fromPDim = runFromParser . fromPDimWith [< "𝑖", "𝑗"]
in [ in [
note "dim ctx: [𝑖, 𝑗]", note "dim ctx: [𝑖, 𝑗]",
parsesAs dim fromPDim "𝑖" (BV 1), parseMatch dim fromPDim "𝑖" `(B (VS VZ) _),
parsesAs dim fromPDim "𝑗" (BV 0), parseMatch dim fromPDim "𝑗" `(B VZ _),
parseFails dim fromPDim "𝑘", parseFails dim fromPDim "𝑘",
parsesAs dim fromPDim "0" (K Zero), parseMatch dim fromPDim "0" `(K Zero _),
parsesAs dim fromPDim "1" (K One) parseMatch dim fromPDim "1" `(K One _)
], ],
"terms" :- "terms" :-
let defs = fromList [("f", mkDef gany Nat Zero)] let defs = fromList [("f", mkDef gany (Nat noLoc) (Zero noLoc) noLoc)]
-- doesn't have to be well typed yet, just well scoped -- doesn't have to be well typed yet, just well scoped
fromPTerm = runFromParser {defs} . fromPTerm = runFromParser {defs} .
fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"] fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"]
in [ in [
note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]", note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]",
parsesAs term fromPTerm "x" $ BVT 2, parseMatch term fromPTerm "x" `(E $ B (VS $ VS VZ) _),
parseFails term fromPTerm "𝑖", parseFails term fromPTerm "𝑖",
parsesAs term fromPTerm "f" $ FT "f", parseMatch term fromPTerm "f" `(E $ F "f" _),
parsesAs term fromPTerm "λ w ⇒ w" $ [< "w"] :\\ BVT 0, parseMatch term fromPTerm "λ w ⇒ w"
parsesAs term fromPTerm "λ w ⇒ x" $ [< "w"] :\\ BVT 3, `(Lam (S _ $ Y $ E $ B VZ _) _),
parsesAs term fromPTerm "λ x ⇒ x" $ [< "x"] :\\ BVT 0, parseMatch term fromPTerm "λ w ⇒ x"
parsesAs term fromPTerm "λ a b ⇒ f a b" $ `(Lam (S _ $ N $ E $ B (VS $ VS VZ) _) _),
[< "a", "b"] :\\ E (F "f" :@@ [BVT 1, BVT 0]), parseMatch term fromPTerm "λ x ⇒ x"
parsesAs term fromPTerm "f @𝑖" $ `(Lam (S _ $ Y $ E $ B VZ _) _),
E $ F "f" :% BV 1 parseMatch term fromPTerm "λ a b ⇒ f a b"
`(Lam (S _ $ Y $
Lam (S _ $ Y $
E $ App (App (F "f" _) (E $ B (VS VZ) _) _) (E $ B VZ _) _) _) _),
parseMatch term fromPTerm "f @𝑖" $
`(E $ DApp (F "f" _) (B (VS VZ) _) _)
], ],
todo "everything else" todo "everything else"

View File

@ -23,128 +23,179 @@ parameters (ds : NContext d) (ns : NContext n)
{default str label : String} -> Test {default str label : String} -> Test
testPrettyE1 e str {label} = testPrettyT1 (E e) str {label} testPrettyE1 e str {label} = testPrettyT1 (E e) str {label}
prefix 9 ^
(^) : (Loc -> a) -> a
(^) a = a noLoc
FromString BindName where fromString str = BN (fromString str) noLoc
export export
tests : Test tests : Test
tests = "pretty printing terms" :- [ tests = "pretty printing terms" :- [
"free vars" :- [ "free vars" :- [
testPrettyE1 [<] [<] (F "x") "x", testPrettyE1 [<] [<] (^F "x") "x",
testPrettyE1 [<] [<] (F $ MakeName [< "A", "B", "C"] "x") "A.B.C.x" testPrettyE1 [<] [<] (^F (MakeName [< "A", "B", "C"] "x")) "A.B.C.x"
], ],
"bound vars" :- [ "bound vars" :- [
testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"] (BV 0) "y", testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"] (^BV 0) "y",
testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"] (BV 1) "x", testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"] (^BV 1) "x",
testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"] (F "eq" :% BV 1) "eq @𝑖", testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"]
testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"] (F "eq" :% BV 1 :% BV 0) "eq @𝑖 @𝑗" (^DApp (^F "eq") (^BV 1))
"eq @𝑖",
testPrettyE1 [< "𝑖", "𝑗"] [< "x", "y"]
(^DApp (^DApp (^F "eq") (^BV 1)) (^BV 0))
"eq @𝑖 @𝑗"
], ],
"applications" :- [ "applications" :- [
testPrettyE1 [<] [<] (F "f" :@ FT "x") "f x", testPrettyE1 [<] [<]
testPrettyE1 [<] [<] (F "f" :@@ [FT "x", FT "y"]) "f x y", (^App (^F "f") (^FT "x"))
testPrettyE1 [<] [<] (F "f" :% K Zero) "f @0", "f x",
testPrettyE1 [<] [<] (F "f" :@ FT "x" :% K Zero) "f x @0", testPrettyE1 [<] [<]
testPrettyE1 [<] [<] (F "g" :% K One :@ FT "y") "g @1 y" (^App (^App (^F "f") (^FT "x")) (^FT "y"))
"f x y",
testPrettyE1 [<] [<]
(^DApp (^F "f") (^K Zero))
"f @0",
testPrettyE1 [<] [<]
(^DApp (^App (^F "f") (^FT "x")) (^K Zero))
"f x @0",
testPrettyE1 [<] [<]
(^App (^DApp (^F "g") (^K One)) (^FT "y"))
"g @1 y"
], ],
"lambda" :- [ "lambda" :- [
testPrettyT [<] [<] ([< "x"] :\\ BVT 0) "λ x ⇒ x" "fun x => x",
testPrettyT [<] [<] (Lam $ SN $ FT "a") "λ _ ⇒ a" "fun _ => a",
testPrettyT [<] [< "y"] ([< "x"] :\\ BVT 1) "λ x ⇒ y" "fun x => y",
testPrettyT [<] [<] testPrettyT [<] [<]
([< "x", "y", "f"] :\\ E (BV 0 :@@ [BVT 2, BVT 1])) (^LamY "x" (^BVT 0))
"λ x ⇒ x"
"fun x => x",
testPrettyT [<] [<]
(^LamN (^FT "a"))
"λ _ ⇒ a"
"fun _ => a",
testPrettyT [<] [< "y"]
(^LamY "x" (^BVT 1))
"λ x ⇒ y"
"fun x => y",
testPrettyT [<] [<]
(^LamY "x" (^LamY "y" (^LamY "f"
(E $ ^App (^App (^BV 0) (^BVT 2)) (^BVT 1)))))
"λ x y f ⇒ f x y" "λ x y f ⇒ f x y"
"fun x y f => f x y", "fun x y f => f x y",
testPrettyT [<] [<] (DLam $ SN $ FT "a") "δ _ ⇒ a" "dfun _ => a",
testPrettyT [<] [<] ([< "i"] :\\% FT "x") "δ i ⇒ x" "dfun i => x",
testPrettyT [<] [<] testPrettyT [<] [<]
([< "x"] :\\ [< "i"] :\\% E (BV 0 :% BV 0)) (^DLam (SN (^FT "a")))
"δ _ ⇒ a"
"dfun _ => a",
testPrettyT [<] [<]
(^DLamY "i" (^FT "x"))
"δ i ⇒ x"
"dfun i => x",
testPrettyT [<] [<]
(^LamY "x" (^DLamY "i" (E $ ^DApp (^BV 0) (^BV 0))))
"λ x ⇒ δ i ⇒ x @i" "λ x ⇒ δ i ⇒ x @i"
"fun x => dfun i => x @i" "fun x => dfun i => x @i"
], ],
"type universes" :- [ "type universes" :- [
testPrettyT [<] [<] (TYPE 0) "★₀" "Type0", testPrettyT [<] [<] (^TYPE 0) "★₀" "Type0",
testPrettyT [<] [<] (TYPE 100) "★₁₀₀" "Type100" testPrettyT [<] [<] (^TYPE 100) "★₁₀₀" "Type100"
], ],
"function types" :- [ "function types" :- [
testPrettyT [<] [<] (Arr One (FT "A") (FT "B")) "1.A → B" "1.A -> B",
testPrettyT [<] [<] testPrettyT [<] [<]
(PiY One "x" (FT "A") (E $ F "B" :@ BVT 0)) (^Arr One (^FT "A") (^FT "B"))
"1.A → B"
"1.A -> B",
testPrettyT [<] [<]
(^PiY One "x" (^FT "A") (E $ ^App (^F "B") (^BVT 0)))
"1.(x : A) → B x" "1.(x : A) → B x"
"1.(x : A) -> B x", "1.(x : A) -> B x",
testPrettyT [<] [<] testPrettyT [<] [<]
(PiY Zero "A" (TYPE 0) $ Arr Any (BVT 0) (BVT 0)) (^PiY Zero "A" (^TYPE 0) (^Arr Any (^BVT 0) (^BVT 0)))
"0.(A : ★₀) → ω.A → A" "0.(A : ★₀) → ω.A → A"
"0.(A : Type0) -> #.A -> A", "0.(A : Type0) -> #.A -> A",
testPrettyT [<] [<] testPrettyT [<] [<]
(Arr Any (Arr Any (FT "A") (FT "A")) (FT "A")) (^Arr Any (^Arr Any (^FT "A") (^FT "A")) (^FT "A"))
"ω.(ω.A → A) → A" "ω.(ω.A → A) → A"
"#.(#.A -> A) -> A", "#.(#.A -> A) -> A",
testPrettyT [<] [<] testPrettyT [<] [<]
(Arr Any (FT "A") (Arr Any (FT "A") (FT "A"))) (^Arr Any (^FT "A") (^Arr Any (^FT "A") (^FT "A")))
"ω.A → ω.A → A" "ω.A → ω.A → A"
"#.A -> #.A -> A", "#.A -> #.A -> A",
testPrettyT [<] [<] testPrettyT [<] [<]
(PiY Zero "P" (Arr Zero (FT "A") (TYPE 0)) (E $ BV 0 :@ FT "a")) (^PiY Zero "P" (^Arr Zero (^FT "A") (^TYPE 0))
(E $ ^App (^BV 0) (^FT "a")))
"0.(P : 0.A → ★₀) → P a" "0.(P : 0.A → ★₀) → P a"
"0.(P : 0.A -> Type0) -> P a" "0.(P : 0.A -> Type0) -> P a"
], ],
"pair types" :- [ "pair types" :- [
testPrettyT [<] [<] (FT "A" `And` FT "B") "A × B" "A ** B",
testPrettyT [<] [<] testPrettyT [<] [<]
(SigY "x" (FT "A") (E $ F "B" :@ BVT 0)) (^And (^FT "A") (^FT "B"))
"A × B"
"A ** B",
testPrettyT [<] [<]
(^SigY "x" (^FT "A") (E $ ^App (^F "B") (^BVT 0)))
"(x : A) × B x" "(x : A) × B x"
"(x : A) ** B x", "(x : A) ** B x",
testPrettyT [<] [<] testPrettyT [<] [<]
(SigY "x" (FT "A") $ (^SigY "x" (^FT "A")
SigY "y" (E $ F "B" :@ BVT 0) $ (^SigY "y" (E $ ^App (^F "B") (^BVT 0))
E $ F "C" :@@ [BVT 1, BVT 0]) (E $ ^App (^App (^F "C") (^BVT 1)) (^BVT 0))))
"(x : A) × (y : B x) × C x y" "(x : A) × (y : B x) × C x y"
"(x : A) ** (y : B x) ** C x y", "(x : A) ** (y : B x) ** C x y",
todo "non-dependent, left and right nested" todo "non-dependent, left and right nested"
], ],
"pairs" :- [ "pairs" :- [
testPrettyT1 [<] [<] (Pair (FT "A") (FT "B")) "(A, B)", testPrettyT1 [<] [<]
testPrettyT1 [<] [<] (Pair (FT "A") (Pair (FT "B") (FT "C"))) "(A, B, C)", (^Pair (^FT "A") (^FT "B"))
testPrettyT1 [<] [<] (Pair (Pair (FT "A") (FT "B")) (FT "C")) "((A, B), C)", "(A, B)",
testPrettyT1 [<] [<]
(^Pair (^FT "A") (^Pair (^FT "B") (^FT "C")))
"(A, B, C)",
testPrettyT1 [<] [<]
(^Pair (^Pair (^FT "A") (^FT "B")) (^FT "C"))
"((A, B), C)",
testPrettyT [<] [<] testPrettyT [<] [<]
(Pair ([< "x"] :\\ BVT 0) (Arr One (FT "B₁") (FT "B₂"))) (^Pair (^LamY "x" (^BVT 0)) (^Arr One (^FT "B₁") (^FT "B₂")))
"(λ x ⇒ x, 1.B₁ → B₂)" "(λ x ⇒ x, 1.B₁ → B₂)"
"(fun x => x, 1.B₁ -> B₂)" "(fun x => x, 1.B₁ -> B₂)"
], ],
"enum types" :- [ "enum types" :- [
testPrettyT1 [<] [<] (enum []) "{}", testPrettyT1 [<] [<] (^enum []) "{}",
testPrettyT1 [<] [<] (enum ["a"]) "{a}", testPrettyT1 [<] [<] (^enum ["a"]) "{a}",
testPrettyT1 [<] [<] (enum ["aa", "bb", "cc"]) "{aa, bb, cc}", testPrettyT1 [<] [<] (^enum ["aa", "bb", "cc"]) "{aa, bb, cc}",
testPrettyT1 [<] [<] (enum ["a b c"]) #"{"a b c"}"#, testPrettyT1 [<] [<] (^enum ["a b c"]) #"{"a b c"}"#,
testPrettyT1 [<] [<] (enum ["\"", ",", "\\"]) #" {"\"", ",", \} "# testPrettyT1 [<] [<] (^enum ["\"", ",", "\\"]) #" {"\"", ",", \} "#
{label = #"{"\"", ",", \} # 「\」 is an identifier"#} {label = #"{"\"", ",", \} # 「\」 is an identifier"#}
], ],
"tags" :- [ "tags" :- [
testPrettyT1 [<] [<] (Tag "a") "'a", testPrettyT1 [<] [<] (^Tag "a") "'a",
testPrettyT1 [<] [<] (Tag "hello") "'hello", testPrettyT1 [<] [<] (^Tag "hello") "'hello",
testPrettyT1 [<] [<] (Tag "qualified.tag") "'qualified.tag", testPrettyT1 [<] [<] (^Tag "qualified.tag") "'qualified.tag",
testPrettyT1 [<] [<] (Tag "non-identifier tag") #"'"non-identifier tag""#, testPrettyT1 [<] [<] (^Tag "non-identifier tag") #"'"non-identifier tag""#,
testPrettyT1 [<] [<] (Tag #"""#) #" '"\"" "# testPrettyT1 [<] [<] (^Tag #"""#) #" '"\"" "#
], ],
todo "equality types", todo "equality types",
"case" :- [ "case" :- [
testPrettyE [<] [<] testPrettyE [<] [<]
(CasePair One (F "a") (SN $ TYPE 1) (SN $ TYPE 0)) (^CasePair One (^F "a") (SN $ ^TYPE 1) (SN $ ^TYPE 0))
"case1 a return ★₁ of { (_, _) ⇒ ★₀ }" "case1 a return ★₁ of { (_, _) ⇒ ★₀ }"
"case1 a return Type1 of { (_, _) => Type0 }", "case1 a return Type1 of { (_, _) => Type0 }",
testPrettyT [<] [<] testPrettyT [<] [<]
([< "u"] :\\ (^LamY "u" (E $
E (CaseEnum One (F "u") ^CaseEnum One (^F "u")
(SY [< "x"] $ Eq0 (enum ["tt"]) (BVT 0) (Tag "tt")) (SY [< "x"] $ ^Eq0 (^enum ["tt"]) (^BVT 0) (^Tag "tt"))
(fromList [("tt", [< Unused] :\\% Tag "tt")]))) (fromList [("tt", ^DLamN (^Tag "tt"))])))
"λ u ⇒ case1 u return x ⇒ x ≡ 'tt : {tt} of { 'tt ⇒ δ _ ⇒ 'tt }" "λ u ⇒ case1 u return x ⇒ x ≡ 'tt : {tt} of { 'tt ⇒ δ _ ⇒ 'tt }"
""" """
fun u => fun u =>
@ -155,27 +206,30 @@ tests = "pretty printing terms" :- [
"type-case" :- [ "type-case" :- [
testPrettyE [<] [<] testPrettyE [<] [<]
{label = "type-case ∷ ★₀ return ★₀ of { ⋯ }"} {label = "type-case ∷ ★₀ return ★₀ of { ⋯ }"}
(TypeCase (Nat :# TYPE 0) (TYPE 0) empty Nat) (^TypeCase (^Ann (^Nat) (^TYPE 0)) (^TYPE 0) empty (^Nat))
"type-case ∷ ★₀ return ★₀ of { _ ⇒ }" "type-case ∷ ★₀ return ★₀ of { _ ⇒ }"
"type-case Nat :: Type0 return Type0 of { _ => Nat }" "type-case Nat :: Type0 return Type0 of { _ => Nat }"
], ],
"annotations" :- [ "annotations" :- [
testPrettyE [<] [<] (FT "a" :# FT "A") "a ∷ A" "a :: A",
testPrettyE [<] [<] testPrettyE [<] [<]
(FT "a" :# E (FT "A" :# FT "𝐀")) (^Ann (^FT "a") (^FT "A"))
"a ∷ A"
"a :: A",
testPrettyE [<] [<]
(^Ann (^FT "a") (E $ ^Ann (^FT "A") (^FT "𝐀")))
"a ∷ A ∷ 𝐀" "a ∷ A ∷ 𝐀"
"a :: A :: 𝐀", "a :: A :: 𝐀",
testPrettyE [<] [<] testPrettyE [<] [<]
(E (FT "α" :# FT "a") :# FT "A") (^Ann (E $ ^Ann (^FT "α") (^FT "a")) (^FT "A"))
"(α ∷ a) ∷ A" "(α ∷ a) ∷ A"
"(α :: a) :: A", "(α :: a) :: A",
testPrettyE [<] [<] testPrettyE [<] [<]
(([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
"(λ x ⇒ x) ∷ 1.A → A" "(λ x ⇒ x) ∷ 1.A → A"
"(fun x => x) :: 1.A -> A", "(fun x => x) :: 1.A -> A",
testPrettyE [<] [<] testPrettyE [<] [<]
(Arr One (FT "A") (FT "A") :# TYPE 7) (^Ann (^Arr One (^FT "A") (^FT "A")) (^TYPE 7))
"(1.A → A) ∷ ★₇" "(1.A → A) ∷ ★₇"
"(1.A -> A) :: Type7" "(1.A -> A) :: Type7"
] ]

View File

@ -3,7 +3,12 @@ module Tests.Reduce
import Quox.Syntax as Lib import Quox.Syntax as Lib
import Quox.Equal import Quox.Equal
import TypingImpls import TypingImpls
import AstExtra
import TAP import TAP
import Control.Eff
%hide Prelude.App
%hide Pretty.App
parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex} {d, n : Nat} parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex} {d, n : Nat}
@ -12,7 +17,7 @@ parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex} {d, n : Nat}
private private
testWhnf : String -> WhnfContext d n -> tm d n -> tm d n -> Test testWhnf : String -> WhnfContext d n -> tm d n -> tm d n -> Test
testWhnf label ctx from to = test "\{label} (whnf)" $ do testWhnf label ctx from to = test "\{label} (whnf)" $ do
result <- bimap toInfo fst $ whnf defs ctx from result <- mapFst toInfo $ runWhnf $ whnf0 defs ctx from
unless (result == to) $ Left [("exp", show to), ("got", show result)] unless (result == to) $ Left [("exp", show to), ("got", show result)]
private private
@ -20,7 +25,7 @@ parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex} {d, n : Nat}
testNoStep label ctx e = testWhnf label ctx e e testNoStep label ctx e = testWhnf label ctx e e
private private
ctx : Context (\n => (BaseName, Term 0 n)) n -> WhnfContext 0 n ctx : Context (\n => (BindName, Term 0 n)) n -> WhnfContext 0 n
ctx xs = let (ns, ts) = unzip xs in MkWhnfContext [<] ns ts ctx xs = let (ns, ts) = unzip xs in MkWhnfContext [<] ns ts
@ -28,91 +33,101 @@ export
tests : Test tests : Test
tests = "whnf" :- [ tests = "whnf" :- [
"head constructors" :- [ "head constructors" :- [
testNoStep "★₀" empty $ TYPE 0, testNoStep "★₀" empty $ ^TYPE 0,
testNoStep "[A] ⊸ [B]" empty $ testNoStep "1.A → B" empty $
Arr One (FT "A") (FT "B"), ^Arr One (^FT "A") (^FT "B"),
testNoStep "(x: [A]) ⊸ [B [x]]" empty $ testNoStep "(x: A) ⊸ B x" empty $
Pi One (FT "A") (S [< "x"] $ Y $ E $ F "B" :@ BVT 0), ^PiY One "x" (^FT "A") (E $ ^App (^F "B") (^BVT 0)),
testNoStep "λx. [x]" empty $ testNoStep "λ x ⇒ x" empty $
Lam $ S [< "x"] $ Y $ BVT 0, ^LamY "x" (^BVT 0),
testNoStep "[f [a]]" empty $ testNoStep "f a" empty $
E $ F "f" :@ FT "a" E $ ^App (^F "f") (^FT "a")
], ],
"neutrals" :- [ "neutrals" :- [
testNoStep "x" (ctx [< ("A", Nat)]) $ BV 0, testNoStep "x" (ctx [< ("A", ^Nat)]) $ ^BV 0,
testNoStep "a" empty $ F "a", testNoStep "a" empty $ ^F "a",
testNoStep "f [a]" empty $ F "f" :@ FT "a", testNoStep "f a" empty $ ^App (^F "f") (^FT "a"),
testNoStep "★₀ ∷ ★₁" empty $ TYPE 0 :# TYPE 1 testNoStep "★₀ ∷ ★₁" empty $ ^Ann (^TYPE 0) (^TYPE 1)
], ],
"redexes" :- [ "redexes" :- [
testWhnf "[a] ∷ [A]" empty testWhnf "a ∷ A" empty
(FT "a" :# FT "A") (^Ann (^FT "a") (^FT "A"))
(F "a"), (^F "a"),
testWhnf "[★₁ ∷ ★₃]" empty testWhnf "★₁ ∷ ★₃" empty
(E (TYPE 1 :# TYPE 3)) (E $ ^Ann (^TYPE 1) (^TYPE 3))
(TYPE 1), (^TYPE 1),
testWhnf "(λx. [x] ∷ [A] ⊸ [A]) [a]" empty testWhnf "(λ x ⇒ x ∷ 1.A → A) a" empty
((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a") (^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
(F "a") (^FT "a"))
(^F "a")
], ],
"definitions" :- [ "definitions" :- [
testWhnf "a (transparent)" empty testWhnf "a (transparent)" empty
{defs = fromList [("a", mkDef gzero (TYPE 1) (TYPE 0))]} {defs = fromList [("a", ^mkDef gzero (^TYPE 1) (^TYPE 0))]}
(F "a") (TYPE 0 :# TYPE 1) (^F "a") (^Ann (^TYPE 0) (^TYPE 1)),
testNoStep "a (opaque)" empty
{defs = fromList [("a", ^mkPostulate gzero (^TYPE 1))]}
(^F "a")
], ],
"elim closure" :- [ "elim closure" :- [
testWhnf "x{}" (ctx [< ("A", Nat)]) testWhnf "x{}" (ctx [< ("x", ^Nat)])
(CloE (Sub (BV 0) id)) (CloE (Sub (^BV 0) id))
(BV 0), (^BV 0),
testWhnf "x{a/x}" empty testWhnf "x{a/x}" empty
(CloE (Sub (BV 0) (F "a" ::: id))) (CloE (Sub (^BV 0) (^F "a" ::: id)))
(F "a"), (^F "a"),
testWhnf "x{x/x,a/y}" (ctx [< ("A", Nat)]) testWhnf "x{a/y}" (ctx [< ("x", ^Nat)])
(CloE (Sub (BV 0) (BV 0 ::: F "a" ::: id))) (CloE (Sub (^BV 0) (^BV 0 ::: ^F "a" ::: id)))
(BV 0), (^BV 0),
testWhnf "x{(y{a/y})/x}" empty testWhnf "x{(y{a/y})/x}" empty
(CloE (Sub (BV 0) ((CloE (Sub (BV 0) (F "a" ::: id))) ::: id))) (CloE (Sub (^BV 0) ((CloE (Sub (^BV 0) (^F "a" ::: id))) ::: id)))
(F "a"), (^F "a"),
testWhnf "(x y){f/x,a/y}" empty testWhnf "(x y){f/x,a/y}" empty
(CloE (Sub (BV 0 :@ BVT 1) (F "f" ::: F "a" ::: id))) (CloE (Sub (^App (^BV 0) (^BVT 1)) (^F "f" ::: ^F "a" ::: id)))
(F "f" :@ FT "a"), (^App (^F "f") (^FT "a")),
testWhnf "([y][x]){A/x}" (ctx [< ("A", Nat)]) testWhnf "(y ∷ x){A/x}" (ctx [< ("x", ^Nat)])
(CloE (Sub (BVT 1 :# BVT 0) (F "A" ::: id))) (CloE (Sub (^Ann (^BVT 1) (^BVT 0)) (^F "A" ::: id)))
(BV 0), (^BV 0),
testWhnf "([y][x]){A/x,a/y}" empty testWhnf "(y ∷ x){A/x,a/y}" empty
(CloE (Sub (BVT 1 :# BVT 0) (F "A" ::: F "a" ::: id))) (CloE (Sub (^Ann (^BVT 1) (^BVT 0)) (^F "A" ::: ^F "a" ::: id)))
(F "a") (^F "a")
], ],
"term closure" :- [ "term closure" :- [
testWhnf "y. x){a/x}" empty testWhnf " y ⇒ x){a/x}" empty
(CloT (Sub (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id))) (CloT (Sub (^LamN (^BVT 0)) (^F "a" ::: id)))
(Lam $ S [< "y"] $ N $ FT "a"), (^LamN (^FT "a")),
testWhnf "(λy. y){a/x}" empty testWhnf "(λy. y){a/x}" empty
(CloT (Sub ([< "y"] :\\ BVT 0) (F "a" ::: id))) (CloT (Sub (^LamY "y" (^BVT 0)) (^F "a" ::: id)))
([< "y"] :\\ BVT 0) (^LamY "y" (^BVT 0))
], ],
"looking inside […]" :- [ "looking inside `E`" :- [
testWhnf "[(λx. x ∷ A ⊸ A) [a]]" empty testWhnf "(λx. x ∷ A ⊸ A) a" empty
(E $ (([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a") (E $ ^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
(FT "a") (^FT "a"))
(^FT "a")
], ],
"nested redex" :- [ "nested redex" :- [
note "whnf only looks at top level redexes", testNoStep "λ y ⇒ ((λ x ⇒ x) ∷ 1.A → A) y" empty $
testNoStep "λy. [(λx. [x] ∷ [A] ⊸ [A]) [y]]" empty $ ^LamY "y" (E $
[< "y"] :\\ E ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ BVT 0), ^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
testNoStep "f [(λx. [x] ∷ [A] ⊸ [A]) [a]]" empty $ (^BVT 0)),
F "a" :@ testNoStep "f (((λ x ⇒ x) ∷ 1.A → A) a)" empty $
E ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a"), ^App (^F "f")
testNoStep "λx. [y [x]]{x/x,a/y}" (ctx [< ("A", Nat)]) $ (E $ ^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A") (^FT "A")))
[< "x"] :\\ CloT (Sub (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id)), (^FT "a")),
testNoStep "f ([y [x]]{x/x,a/y})" (ctx [< ("A", Nat)]) $ testNoStep "λx. (y x){x/x,a/y}" (ctx [< ("y", ^Nat)]) $
F "f" :@ CloT (Sub (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id)) ^LamY "x" (CloT $ Sub (E $ ^App (^BV 1) (^BVT 0))
(^BV 0 ::: ^F "a" ::: id)),
testNoStep "f (y x){x/x,a/y}" (ctx [< ("y", ^Nat)]) $
^App (^F "f")
(CloT (Sub (E $ ^App (^BV 1) (^BVT 0))
(^BV 0 ::: ^F "a" ::: id)))
] ]
] ]

View File

@ -5,6 +5,11 @@ import Quox.Typechecker as Lib
import public TypingImpls import public TypingImpls
import TAP import TAP
import Quox.EffExtra import Quox.EffExtra
import AstExtra
%hide Prelude.App
%hide Pretty.App
data Error' data Error'
@ -28,64 +33,75 @@ ToInfo Error' where
M = Eff [Except Error', DefsReader] M = Eff [Except Error', DefsReader]
inj : TC a -> M a inj : TC a -> M a
inj = rethrow . mapFst TCError <=< lift . runExcept inj act = rethrow $ mapFst TCError $ runTC !defs act
reflTy : Term d n reflTy : Term d n
reflTy = reflTy =
PiY Zero "A" (TYPE 0) $ ^PiY Zero "A" (^TYPE 0)
PiY One "x" (BVT 0) $ (^PiY One "x" (^BVT 0)
Eq0 (BVT 1) (BVT 0) (BVT 0) (^Eq0 (^BVT 1) (^BVT 0) (^BVT 0)))
reflDef : Term d n reflDef : Term d n
reflDef = [< "A","x"] :\\ [< "i"] :\\% BVT 0 reflDef = ^LamY "A" (^LamY "x" (^DLamY "i" (^BVT 0)))
fstTy : Term d n fstTy : Term d n
fstTy = fstTy =
(PiY Zero "A" (TYPE 1) $ ^PiY Zero "A" (^TYPE 1)
PiY Zero "B" (Arr Any (BVT 0) (TYPE 1)) $ (^PiY Zero "B" (^Arr Any (^BVT 0) (^TYPE 1))
Arr Any (SigY "x" (BVT 1) $ E $ BV 1 :@ BVT 0) (BVT 1)) (^Arr Any (^SigY "x" (^BVT 1) (E $ ^App (^BV 1) (^BVT 0))) (^BVT 1)))
fstDef : Term d n fstDef : Term d n
fstDef = fstDef =
([< "A","B","p"] :\\ ^LamY "A" (^LamY "B" (^LamY "p"
E (CasePair Any (BV 0) (SN $ BVT 2) (SY [< "x","y"] $ BVT 1))) (E $ ^CasePair Any (^BV 0) (SN $ ^BVT 2)
(SY [< "x", "y"] $ ^BVT 1))))
sndTy : Term d n sndTy : Term d n
sndTy = sndTy =
(PiY Zero "A" (TYPE 1) $ ^PiY Zero "A" (^TYPE 1)
PiY Zero "B" (Arr Any (BVT 0) (TYPE 1)) $ (^PiY Zero "B" (^Arr Any (^BVT 0) (^TYPE 1))
PiY Any "p" (SigY "x" (BVT 1) $ E $ BV 1 :@ BVT 0) $ (^PiY Any "p" (^SigY "x" (^BVT 1) (E $ ^App (^BV 1) (^BVT 0)))
E (BV 1 :@ E (F "fst" :@@ [BVT 2, BVT 1, BVT 0]))) (E $ ^App (^BV 1)
(E $ ^App (^App (^App (^F "fst") (^BVT 2)) (^BVT 1)) (^BVT 0)))))
sndDef : Term d n sndDef : Term d n
sndDef = sndDef =
([< "A","B","p"] :\\ -- λ A B p ⇒ caseω p return p' ⇒ B (fst A B p') of { (x, y) ⇒ y }
E (CasePair Any (BV 0) ^LamY "A" (^LamY "B" (^LamY "p"
(SY [< "p"] $ E $ BV 2 :@ E (F "fst" :@@ [BVT 3, BVT 2, BVT 0])) (E $ ^CasePair Any (^BV 0)
(SY [< "x","y"] $ BVT 0))) (SY [< "p"] $ E $
^App (^BV 2)
(E $ ^App (^App (^App (^F "fst") (^BVT 3)) (^BVT 2)) (^BVT 0)))
(SY [< "x", "y"] $ ^BVT 0))))
nat : Term d n
nat = ^Nat
defGlobals : Definitions defGlobals : Definitions
defGlobals = fromList defGlobals = fromList
[("A", mkPostulate gzero $ TYPE 0), [("A", ^mkPostulate gzero (^TYPE 0)),
("B", mkPostulate gzero $ TYPE 0), ("B", ^mkPostulate gzero (^TYPE 0)),
("C", mkPostulate gzero $ TYPE 1), ("C", ^mkPostulate gzero (^TYPE 1)),
("D", mkPostulate gzero $ TYPE 1), ("D", ^mkPostulate gzero (^TYPE 1)),
("P", mkPostulate gzero $ Arr Any (FT "A") (TYPE 0)), ("P", ^mkPostulate gzero (^Arr Any (^FT "A") (^TYPE 0))),
("a", mkPostulate gany $ FT "A"), ("a", ^mkPostulate gany (^FT "A")),
("a'", mkPostulate gany $ FT "A"), ("a'", ^mkPostulate gany (^FT "A")),
("b", mkPostulate gany $ FT "B"), ("b", ^mkPostulate gany (^FT "B")),
("f", mkPostulate gany $ Arr One (FT "A") (FT "A")), ("f", ^mkPostulate gany (^Arr One (^FT "A") (^FT "A"))),
("", mkPostulate gany $ Arr Any (FT "A") (FT "A")), ("", ^mkPostulate gany (^Arr Any (^FT "A") (^FT "A"))),
("g", mkPostulate gany $ Arr One (FT "A") (FT "B")), ("g", ^mkPostulate gany (^Arr One (^FT "A") (^FT "B"))),
("f2", mkPostulate gany $ Arr One (FT "A") $ Arr One (FT "A") (FT "B")), ("f2", ^mkPostulate gany
("p", mkPostulate gany $ PiY One "x" (FT "A") $ E $ F "P" :@ BVT 0), (^Arr One (^FT "A") (^Arr One (^FT "A") (^FT "B")))),
("q", mkPostulate gany $ PiY One "x" (FT "A") $ E $ F "P" :@ BVT 0), ("p", ^mkPostulate gany
("refl", mkDef gany reflTy reflDef), (^PiY One "x" (^FT "A") (E $ ^App (^F "P") (^BVT 0)))),
("fst", mkDef gany fstTy fstDef), ("q", ^mkPostulate gany
("snd", mkDef gany sndTy sndDef)] (^PiY One "x" (^FT "A") (E $ ^App (^F "P") (^BVT 0)))),
("refl", ^mkDef gany reflTy reflDef),
("fst", ^mkDef gany fstTy fstDef),
("snd", ^mkDef gany sndTy sndDef)]
parameters (label : String) (act : Lazy (M ())) parameters (label : String) (act : Lazy (M ()))
{default defGlobals globals : Definitions} {default defGlobals globals : Definitions}
@ -98,23 +114,10 @@ parameters (label : String) (act : Lazy (M ()))
(extract $ runExcept $ runReaderAt DEFS globals act) $> "()" (extract $ runExcept $ runReaderAt DEFS globals act) $> "()"
anys : {n : Nat} -> QContext n
anys {n = 0} = [<]
anys {n = S n} = anys :< Any
ctx, ctx01 : {n : Nat} -> Context (\n => (BaseName, Term 0 n)) n ->
TyContext 0 n
ctx tel = let (ns, ts) = unzip tel in
MkTyContext new [<] ts ns anys
ctx01 tel = let (ns, ts) = unzip tel in
MkTyContext ZeroIsOne [<] ts ns anys
empty01 : TyContext 0 0
empty01 = eqDim (K Zero) (K One) empty
inferredTypeEq : TyContext d n -> (exp, got : Term d n) -> M () inferredTypeEq : TyContext d n -> (exp, got : Term d n) -> M ()
inferredTypeEq ctx exp got = inferredTypeEq ctx exp got =
wrapErr (const $ WrongInfer exp got) $ inj $ equalType ctx exp got wrapErr (const $ WrongInfer exp got) $ inj $ lift $
equalType noLoc ctx exp got
qoutEq : (exp, got : QOutput n) -> M () qoutEq : (exp, got : QOutput n) -> M ()
qoutEq qout res = unless (qout == res) $ throw $ WrongQOut qout res qoutEq qout res = unless (qout == res) $ throw $ WrongQOut qout res
@ -156,153 +159,168 @@ tests : Test
tests = "typechecker" :- [ tests = "typechecker" :- [
"universes" :- [ "universes" :- [
testTC "0 · ★₀ ⇐ ★₁ # by checkType" $ testTC "0 · ★₀ ⇐ ★₁ # by checkType" $
checkType_ empty (TYPE 0) (Just 1), checkType_ empty (^TYPE 0) (Just 1),
testTC "0 · ★₀ ⇐ ★₁ # by check" $ testTC "0 · ★₀ ⇐ ★₁ # by check" $
check_ empty szero (TYPE 0) (TYPE 1), check_ empty szero (^TYPE 0) (^TYPE 1),
testTC "0 · ★₀ ⇐ ★₂" $ testTC "0 · ★₀ ⇐ ★₂" $
checkType_ empty (TYPE 0) (Just 2), checkType_ empty (^TYPE 0) (Just 2),
testTC "0 · ★₀ ⇐ ★_" $ testTC "0 · ★₀ ⇐ ★_" $
checkType_ empty (TYPE 0) Nothing, checkType_ empty (^TYPE 0) Nothing,
testTCFail "0 · ★₁ ⇍ ★₀" $ testTCFail "0 · ★₁ ⇍ ★₀" $
checkType_ empty (TYPE 1) (Just 0), checkType_ empty (^TYPE 1) (Just 0),
testTCFail "0 · ★₀ ⇍ ★₀" $ testTCFail "0 · ★₀ ⇍ ★₀" $
checkType_ empty (TYPE 0) (Just 0), checkType_ empty (^TYPE 0) (Just 0),
testTC "0=1 ⊢ 0 · ★₁ ⇐ ★₀" $ testTC "0=1 ⊢ 0 · ★₁ ⇐ ★₀" $
checkType_ empty01 (TYPE 1) (Just 0), checkType_ empty01 (^TYPE 1) (Just 0),
testTCFail "1 · ★₀ ⇍ ★₁ # by check" $ testTCFail "1 · ★₀ ⇍ ★₁ # by check" $
check_ empty sone (TYPE 0) (TYPE 1) check_ empty sone (^TYPE 0) (^TYPE 1)
], ],
"function types" :- [ "function types" :- [
note "A, B : ★₀; C, D : ★₁; P : A ⇾ ★₀", note "A, B : ★₀; C, D : ★₁; P : 0.A → ★₀",
testTC "0 · A ⊸ B ⇐ ★₀" $ testTC "0 · 1.A → B ⇐ ★₀" $
check_ empty szero (Arr One (FT "A") (FT "B")) (TYPE 0), check_ empty szero (^Arr One (^FT "A") (^FT "B")) (^TYPE 0),
note "subtyping", note "subtyping",
testTC "0 · A ⊸ B ⇐ ★₁" $ testTC "0 · 1.A → B ⇐ ★₁" $
check_ empty szero (Arr One (FT "A") (FT "B")) (TYPE 1), check_ empty szero (^Arr One (^FT "A") (^FT "B")) (^TYPE 1),
testTC "0 · C ⊸ D ⇐ ★₁" $ testTC "0 · 1.C → D ⇐ ★₁" $
check_ empty szero (Arr One (FT "C") (FT "D")) (TYPE 1), check_ empty szero (^Arr One (^FT "C") (^FT "D")) (^TYPE 1),
testTCFail "0 · C ⊸ D ⇍ ★₀" $ testTCFail "0 · 1.C → D ⇍ ★₀" $
check_ empty szero (Arr One (FT "C") (FT "D")) (TYPE 0), check_ empty szero (^Arr One (^FT "C") (^FT "D")) (^TYPE 0),
testTC "0 · (1·x : A) → P x ⇐ ★₀" $ testTC "0 · 1.(x : A) → P x ⇐ ★₀" $
check_ empty szero check_ empty szero
(PiY One "x" (FT "A") $ E $ F "P" :@ BVT 0) (^PiY One "x" (^FT "A") (E $ ^App (^F "P") (^BVT 0)))
(TYPE 0), (^TYPE 0),
testTCFail "0 · A ⊸ P ⇍ ★₀" $ testTCFail "0 · 1.A → P ⇍ ★₀" $
check_ empty szero (Arr One (FT "A") $ FT "P") (TYPE 0), check_ empty szero (^Arr One (^FT "A") (^FT "P")) (^TYPE 0),
testTC "0=1 ⊢ 0 · A ⊸ P ⇐ ★₀" $ testTC "0=1 ⊢ 0 · 1.A → P ⇐ ★₀" $
check_ empty01 szero (Arr One (FT "A") $ FT "P") (TYPE 0) check_ empty01 szero (^Arr One (^FT "A") (^FT "P")) (^TYPE 0)
], ],
"pair types" :- [ "pair types" :- [
note #""A × B" for "(_ : A) × B""#,
testTC "0 · A × A ⇐ ★₀" $ testTC "0 · A × A ⇐ ★₀" $
check_ empty szero (FT "A" `And` FT "A") (TYPE 0), check_ empty szero (^And (^FT "A") (^FT "A")) (^TYPE 0),
testTCFail "0 · A × P ⇍ ★₀" $ testTCFail "0 · A × P ⇍ ★₀" $
check_ empty szero (FT "A" `And` FT "P") (TYPE 0), check_ empty szero (^And (^FT "A") (^FT "P")) (^TYPE 0),
testTC "0 · (x : A) × P x ⇐ ★₀" $ testTC "0 · (x : A) × P x ⇐ ★₀" $
check_ empty szero check_ empty szero
(SigY "x" (FT "A") $ E $ F "P" :@ BVT 0) (TYPE 0), (^SigY "x" (^FT "A") (E $ ^App (^F "P") (^BVT 0)))
(^TYPE 0),
testTC "0 · (x : A) × P x ⇐ ★₁" $ testTC "0 · (x : A) × P x ⇐ ★₁" $
check_ empty szero check_ empty szero
(SigY "x" (FT "A") $ E $ F "P" :@ BVT 0) (TYPE 1), (^SigY "x" (^FT "A") (E $ ^App (^F "P") (^BVT 0)))
(^TYPE 1),
testTC "0 · (A : ★₀) × A ⇐ ★₁" $ testTC "0 · (A : ★₀) × A ⇐ ★₁" $
check_ empty szero (SigY "A" (TYPE 0) $ BVT 0) (TYPE 1), check_ empty szero
(^SigY "A" (^TYPE 0) (^BVT 0))
(^TYPE 1),
testTCFail "0 · (A : ★₀) × A ⇍ ★₀" $ testTCFail "0 · (A : ★₀) × A ⇍ ★₀" $
check_ empty szero (SigY "A" (TYPE 0) $ BVT 0) (TYPE 0), check_ empty szero
(^SigY "A" (^TYPE 0) (^BVT 0))
(^TYPE 0),
testTCFail "1 · A × A ⇍ ★₀" $ testTCFail "1 · A × A ⇍ ★₀" $
check_ empty sone (FT "A" `And` FT "A") (TYPE 0) check_ empty sone
(^And (^FT "A") (^FT "A"))
(^TYPE 0)
], ],
"enum types" :- [ "enum types" :- [
testTC "0 · {} ⇐ ★₀" $ check_ empty szero (enum []) (TYPE 0), testTC "0 · {} ⇐ ★₀" $ check_ empty szero (^enum []) (^TYPE 0),
testTC "0 · {} ⇐ ★₃" $ check_ empty szero (enum []) (TYPE 3), testTC "0 · {} ⇐ ★₃" $ check_ empty szero (^enum []) (^TYPE 3),
testTC "0 · {a,b,c} ⇐ ★₀" $ testTC "0 · {a,b,c} ⇐ ★₀" $
check_ empty szero (enum ["a", "b", "c"]) (TYPE 0), check_ empty szero (^enum ["a", "b", "c"]) (^TYPE 0),
testTC "0 · {a,b,c} ⇐ ★₃" $ testTC "0 · {a,b,c} ⇐ ★₃" $
check_ empty szero (enum ["a", "b", "c"]) (TYPE 3), check_ empty szero (^enum ["a", "b", "c"]) (^TYPE 3),
testTCFail "1 · {} ⇍ ★₀" $ check_ empty sone (enum []) (TYPE 0), testTCFail "1 · {} ⇍ ★₀" $ check_ empty sone (^enum []) (^TYPE 0),
testTC "0=1 ⊢ 1 · {} ⇐ ★₀" $ check_ empty01 sone (enum []) (TYPE 0) testTC "0=1 ⊢ 1 · {} ⇐ ★₀" $ check_ empty01 sone (^enum []) (^TYPE 0)
], ],
"free vars" :- [ "free vars" :- [
note "A : ★₀", note "A : ★₀",
testTC "0 · A ⇒ ★₀" $ testTC "0 · A ⇒ ★₀" $
inferAs empty szero (F "A") (TYPE 0), inferAs empty szero (^F "A") (^TYPE 0),
testTC "0 · [A] ⇐ ★₀" $ testTC "0 · [A] ⇐ ★₀" $
check_ empty szero (FT "A") (TYPE 0), check_ empty szero (^FT "A") (^TYPE 0),
note "subtyping", note "subtyping",
testTC "0 · [A] ⇐ ★₁" $ testTC "0 · [A] ⇐ ★₁" $
check_ empty szero (FT "A") (TYPE 1), check_ empty szero (^FT "A") (^TYPE 1),
note "(fail) runtime-relevant type", note "(fail) runtime-relevant type",
testTCFail "1 · A ⇏ ★₀" $ testTCFail "1 · A ⇏ ★₀" $
infer_ empty sone (F "A"), infer_ empty sone (^F "A"),
testTC "1 . f ⇒ 1.A → A" $ testTC "1 . f ⇒ 1.A → A" $
inferAs empty sone (F "f") (Arr One (FT "A") (FT "A")), inferAs empty sone (^F "f") (^Arr One (^FT "A") (^FT "A")),
testTC "1 . f ⇐ 1.A → A" $ testTC "1 . f ⇐ 1.A → A" $
check_ empty sone (FT "f") (Arr One (FT "A") (FT "A")), check_ empty sone (^FT "f") (^Arr One (^FT "A") (^FT "A")),
testTCFail "1 . f ⇍ 0.A → A" $ testTCFail "1 . f ⇍ 0.A → A" $
check_ empty sone (FT "f") (Arr Zero (FT "A") (FT "A")), check_ empty sone (^FT "f") (^Arr Zero (^FT "A") (^FT "A")),
testTCFail "1 . f ⇍ ω.A → A" $ testTCFail "1 . f ⇍ ω.A → A" $
check_ empty sone (FT "f") (Arr Any (FT "A") (FT "A")), check_ empty sone (^FT "f") (^Arr Any (^FT "A") (^FT "A")),
testTC "1 . (λ x ⇒ f x) ⇐ 1.A → A" $ testTC "1 . (λ x ⇒ f x) ⇐ 1.A → A" $
check_ empty sone check_ empty sone
([< "x"] :\\ E (F "f" :@ BVT 0)) (^LamY "x" (E $ ^App (^F "f") (^BVT 0)))
(Arr One (FT "A") (FT "A")), (^Arr One (^FT "A") (^FT "A")),
testTC "1 . (λ x ⇒ f x) ⇐ ω.A → A" $ testTC "1 . (λ x ⇒ f x) ⇐ ω.A → A" $
check_ empty sone check_ empty sone
([< "x"] :\\ E (F "f" :@ BVT 0)) (^LamY "x" (E $ ^App (^F "f") (^BVT 0)))
(Arr Any (FT "A") (FT "A")), (^Arr Any (^FT "A") (^FT "A")),
testTCFail "1 . (λ x ⇒ f x) ⇍ 0.A → A" $ testTCFail "1 . (λ x ⇒ f x) ⇍ 0.A → A" $
check_ empty sone check_ empty sone
([< "x"] :\\ E (F "f" :@ BVT 0)) (^LamY "x" (E $ ^App (^F "f") (^BVT 0)))
(Arr Zero (FT "A") (FT "A")), (^Arr Zero (^FT "A") (^FT "A")),
testTC "1 . fω ⇒ ω.A → A" $ testTC "1 . fω ⇒ ω.A → A" $
inferAs empty sone (F "") (Arr Any (FT "A") (FT "A")), inferAs empty sone (^F "") (^Arr Any (^FT "A") (^FT "A")),
testTC "1 . (λ x ⇒ fω x) ⇐ ω.A → A" $ testTC "1 . (λ x ⇒ fω x) ⇐ ω.A → A" $
check_ empty sone check_ empty sone
([< "x"] :\\ E (F "" :@ BVT 0)) (^LamY "x" (E $ ^App (^F "") (^BVT 0)))
(Arr Any (FT "A") (FT "A")), (^Arr Any (^FT "A") (^FT "A")),
testTCFail "1 . (λ x ⇒ fω x) ⇍ 0.A → A" $ testTCFail "1 . (λ x ⇒ fω x) ⇍ 0.A → A" $
check_ empty sone check_ empty sone
([< "x"] :\\ E (F "" :@ BVT 0)) (^LamY "x" (E $ ^App (^F "") (^BVT 0)))
(Arr Zero (FT "A") (FT "A")), (^Arr Zero (^FT "A") (^FT "A")),
testTCFail "1 . (λ x ⇒ fω x) ⇍ 1.A → A" $ testTCFail "1 . (λ x ⇒ fω x) ⇍ 1.A → A" $
check_ empty sone check_ empty sone
([< "x"] :\\ E (F "" :@ BVT 0)) (^LamY "x" (E $ ^App (^F "") (^BVT 0)))
(Arr One (FT "A") (FT "A")), (^Arr One (^FT "A") (^FT "A")),
note "refl : (0·A : ★₀) → (1·x : A) → (x ≡ x : A) ≔ (λ A x ⇒ δ _ ⇒ x)", note "refl : (0·A : ★₀) → (1·x : A) → (x ≡ x : A) ≔ (λ A x ⇒ δ _ ⇒ x)",
testTC "1 · refl ⇒ ⋯" $ inferAs empty sone (F "refl") reflTy, testTC "1 · refl ⇒ ⋯" $ inferAs empty sone (^F "refl") reflTy,
testTC "1 · [refl] ⇐ ⋯" $ check_ empty sone (FT "refl") reflTy testTC "1 · [refl] ⇐ ⋯" $ check_ empty sone (^FT "refl") reflTy
], ],
"bound vars" :- [ "bound vars" :- [
testTC "x : A ⊢ 1 · x ⇒ A ⊳ 1·x" $ testTC "x : A ⊢ 1 · x ⇒ A ⊳ 1·x" $
inferAsQ {n = 1} (ctx [< ("x", FT "A")]) sone inferAsQ (ctx [< ("x", ^FT "A")]) sone
(BV 0) (FT "A") [< One], (^BV 0) (^FT "A") [< One],
testTC "x : A ⊢ 1 · [x] ⇐ A ⊳ 1·x" $ testTC "x : A ⊢ 1 · x ⇐ A ⊳ 1·x" $
checkQ {n = 1} (ctx [< ("x", FT "A")]) sone (BVT 0) (FT "A") [< One], checkQ (ctx [< ("x", ^FT "A")]) sone (^BVT 0) (^FT "A") [< One],
note "f2 : A ⊸ A ⊸ B", note "f2 : 1.A → 1.A → B",
testTC "x : A ⊢ 1 · f2 [x] [x] ⇒ B ⊳ ω·x" $ testTC "x : A ⊢ 1 · f2 x x ⇒ B ⊳ ω·x" $
inferAsQ {n = 1} (ctx [< ("x", FT "A")]) sone inferAsQ (ctx [< ("x", ^FT "A")]) sone
(F "f2" :@@ [BVT 0, BVT 0]) (FT "B") [< Any] (^App (^App (^F "f2") (^BVT 0)) (^BVT 0)) (^FT "B") [< Any]
], ],
"lambda" :- [ "lambda" :- [
note "linear & unrestricted identity", note "linear & unrestricted identity",
testTC "1 · (λ x ⇒ x) ⇐ A ⊸ A" $
check_ empty sone ([< "x"] :\\ BVT 0) (Arr One (FT "A") (FT "A")),
testTC "1 · (λ x ⇒ x) ⇐ A → A" $ testTC "1 · (λ x ⇒ x) ⇐ A → A" $
check_ empty sone ([< "x"] :\\ BVT 0) (Arr Any (FT "A") (FT "A")), check_ empty sone
(^LamY "x" (^BVT 0))
(^Arr One (^FT "A") (^FT "A")),
testTC "1 · (λ x ⇒ x) ⇐ ω.A → A" $
check_ empty sone
(^LamY "x" (^BVT 0))
(^Arr Any (^FT "A") (^FT "A")),
note "(fail) zero binding used relevantly", note "(fail) zero binding used relevantly",
testTCFail "1 · (λ x ⇒ x) ⇍ A ⇾ A" $ testTCFail "1 · (λ x ⇒ x) ⇍ 0.A → A" $
check_ empty sone ([< "x"] :\\ BVT 0) (Arr Zero (FT "A") (FT "A")), check_ empty sone
(^LamY "x" (^BVT 0))
(^Arr Zero (^FT "A") (^FT "A")),
note "(but ok in overall erased context)", note "(but ok in overall erased context)",
testTC "0 · (λ x ⇒ x) ⇐ A ⇾ A" $ testTC "0 · (λ x ⇒ x) ⇐ A ⇾ A" $
check_ empty szero ([< "x"] :\\ BVT 0) (Arr Zero (FT "A") (FT "A")), check_ empty szero
(^LamY "x" (^BVT 0))
(^Arr Zero (^FT "A") (^FT "A")),
testTC "1 · (λ A x ⇒ refl A x) ⇐ ⋯ # (type of refl)" $ testTC "1 · (λ A x ⇒ refl A x) ⇐ ⋯ # (type of refl)" $
check_ empty sone check_ empty sone
([< "A", "x"] :\\ E (F "refl" :@@ [BVT 1, BVT 0])) (^LamY "A" (^LamY "x" (E $ ^App (^App (^F "refl") (^BVT 1)) (^BVT 0))))
reflTy, reflTy,
testTC "1 · (λ A x ⇒ δ i ⇒ x) ⇐ ⋯ # (def. and type of refl)" $ testTC "1 · (λ A x ⇒ δ i ⇒ x) ⇐ ⋯ # (def. and type of refl)" $
check_ empty sone reflDef reflTy check_ empty sone reflDef reflTy
@ -310,148 +328,153 @@ tests = "typechecker" :- [
"pairs" :- [ "pairs" :- [
testTC "1 · (a, a) ⇐ A × A" $ testTC "1 · (a, a) ⇐ A × A" $
check_ empty sone (Pair (FT "a") (FT "a")) (FT "A" `And` FT "A"), check_ empty sone
(^Pair (^FT "a") (^FT "a")) (^And (^FT "A") (^FT "A")),
testTC "x : A ⊢ 1 · (x, x) ⇐ A × A ⊳ ω·x" $ testTC "x : A ⊢ 1 · (x, x) ⇐ A × A ⊳ ω·x" $
checkQ (ctx [< ("x", FT "A")]) sone checkQ (ctx [< ("x", ^FT "A")]) sone
(Pair (BVT 0) (BVT 0)) (FT "A" `And` FT "A") [< Any], (^Pair (^BVT 0) (^BVT 0)) (^And (^FT "A") (^FT "A")) [< Any],
testTC "1 · (a, δ i ⇒ a) ⇐ (x : A) × (x ≡ a)" $ testTC "1 · (a, δ i ⇒ a) ⇐ (x : A) × (x ≡ a)" $
check_ empty sone check_ empty sone
(Pair (FT "a") ([< "i"] :\\% FT "a")) (^Pair (^FT "a") (^DLamN (^FT "a")))
(SigY "x" (FT "A") $ Eq0 (FT "A") (BVT 0) (FT "a")) (^SigY "x" (^FT "A") (^Eq0 (^FT "A") (^BVT 0) (^FT "a")))
], ],
"unpairing" :- [ "unpairing" :- [
testTC "x : A × A ⊢ 1 · (case1 x return B of (l,r) ⇒ f2 l r) ⇒ B ⊳ 1·x" $ testTC "x : A × A ⊢ 1 · (case1 x return B of (l,r) ⇒ f2 l r) ⇒ B ⊳ 1·x" $
inferAsQ (ctx [< ("x", FT "A" `And` FT "A")]) sone inferAsQ (ctx [< ("x", ^And (^FT "A") (^FT "A"))]) sone
(CasePair One (BV 0) (SN $ FT "B") (^CasePair One (^BV 0) (SN $ ^FT "B")
(SY [< "l", "r"] $ E $ F "f2" :@@ [BVT 1, BVT 0])) (SY [< "l", "r"] $ E $ ^App (^App (^F "f2") (^BVT 1)) (^BVT 0)))
(FT "B") [< One], (^FT "B") [< One],
testTC "x : A × A ⊢ 1 · (caseω x return B of (l,r) ⇒ f2 l r) ⇒ B ⊳ ω·x" $ testTC "x : A × A ⊢ 1 · (caseω x return B of (l,r) ⇒ f2 l r) ⇒ B ⊳ ω·x" $
inferAsQ (ctx [< ("x", FT "A" `And` FT "A")]) sone inferAsQ (ctx [< ("x", ^And (^FT "A") (^FT "A"))]) sone
(CasePair Any (BV 0) (SN $ FT "B") (^CasePair Any (^BV 0) (SN $ ^FT "B")
(SY [< "l", "r"] $ E $ F "f2" :@@ [BVT 1, BVT 0])) (SY [< "l", "r"] $ E $ ^App (^App (^F "f2") (^BVT 1)) (^BVT 0)))
(FT "B") [< Any], (^FT "B") [< Any],
testTC "x : A × A ⊢ 0 · (caseω x return B of (l,r) ⇒ f2 l r) ⇒ B ⊳ 0·x" $ testTC "x : A × A ⊢ 0 · (caseω x return B of (l,r) ⇒ f2 l r) ⇒ B ⊳ 0·x" $
inferAsQ (ctx [< ("x", FT "A" `And` FT "A")]) szero inferAsQ (ctx [< ("x", ^And (^FT "A") (^FT "A"))]) szero
(CasePair Any (BV 0) (SN $ FT "B") (^CasePair Any (^BV 0) (SN $ ^FT "B")
(SY [< "l", "r"] $ E $ F "f2" :@@ [BVT 1, BVT 0])) (SY [< "l", "r"] $ E $ ^App (^App (^F "f2") (^BVT 1)) (^BVT 0)))
(FT "B") [< Zero], (^FT "B") [< Zero],
testTCFail "x : A × A ⊢ 1 · (case0 x return B of (l,r) ⇒ f2 l r) ⇏" $ testTCFail "x : A × A ⊢ 1 · (case0 x return B of (l,r) ⇒ f2 l r) ⇏" $
infer_ (ctx [< ("x", FT "A" `And` FT "A")]) sone infer_ (ctx [< ("x", ^And (^FT "A") (^FT "A"))]) sone
(CasePair Zero (BV 0) (SN $ FT "B") (^CasePair Zero (^BV 0) (SN $ ^FT "B")
(SY [< "l", "r"] $ E $ F "f2" :@@ [BVT 1, BVT 0])), (SY [< "l", "r"] $ E $ ^App (^App (^F "f2") (^BVT 1)) (^BVT 0))),
testTC "x : A × B ⊢ 1 · (caseω x return A of (l,r) ⇒ l) ⇒ A ⊳ ω·x" $ testTC "x : A × B ⊢ 1 · (caseω x return A of (l,r) ⇒ l) ⇒ A ⊳ ω·x" $
inferAsQ (ctx [< ("x", FT "A" `And` FT "B")]) sone inferAsQ (ctx [< ("x", ^And (^FT "A") (^FT "B"))]) sone
(CasePair Any (BV 0) (SN $ FT "A") (^CasePair Any (^BV 0) (SN $ ^FT "A")
(SY [< "l", "r"] $ BVT 1)) (SY [< "l", "r"] $ ^BVT 1))
(FT "A") [< Any], (^FT "A") [< Any],
testTC "x : A × B ⊢ 0 · (case1 x return A of (l,r) ⇒ l) ⇒ A ⊳ 0·x" $ testTC "x : A × B ⊢ 0 · (case1 x return A of (l,r) ⇒ l) ⇒ A ⊳ 0·x" $
inferAsQ (ctx [< ("x", FT "A" `And` FT "B")]) szero inferAsQ (ctx [< ("x", ^And (^FT "A") (^FT "B"))]) szero
(CasePair One (BV 0) (SN $ FT "A") (^CasePair One (^BV 0) (SN $ ^FT "A")
(SY [< "l", "r"] $ BVT 1)) (SY [< "l", "r"] $ ^BVT 1))
(FT "A") [< Zero], (^FT "A") [< Zero],
testTCFail "x : A × B ⊢ 1 · (case1 x return A of (l,r) ⇒ l) ⇏" $ testTCFail "x : A × B ⊢ 1 · (case1 x return A of (l,r) ⇒ l) ⇏" $
infer_ (ctx [< ("x", FT "A" `And` FT "B")]) sone infer_ (ctx [< ("x", ^And (^FT "A") (^FT "B"))]) sone
(CasePair One (BV 0) (SN $ FT "A") (^CasePair One (^BV 0) (SN $ ^FT "A")
(SY [< "l", "r"] $ BVT 1)), (SY [< "l", "r"] $ ^BVT 1)),
note "fst : (0·A : ★₁) → (0·B : A ↠ ★₁) → ((x : A) × B x) ↠ A", note "fst : (0·A : ★₁) → (0·B : A ↠ ★₁) → ((x : A) × B x) ↠ A",
note " ≔ (λ A B p ⇒ caseω p return A of (x, y) ⇒ x)", note " ≔ (λ A B p ⇒ caseω p return A of (x, y) ⇒ x)",
testTC "0 · type of fst ⇐ ★₂" $ testTC "0 · type of fst ⇐ ★₂" $
check_ empty szero fstTy (TYPE 2), check_ empty szero fstTy (^TYPE 2),
testTC "1 · def of fsttype of fst" $ testTC "1 · def of fsttype of fst" $
check_ empty sone fstDef fstTy, check_ empty sone fstDef fstTy,
note "snd : (0·A : ★₁) → (0·B : A ↠ ★₁) → (ω·p : (x : A) × B x) → B (fst A B p)", note "snd : (0·A : ★₁) → (0·B : A ↠ ★₁) → (ω·p : (x : A) × B x) → B (fst A B p)",
note " ≔ (λ A B p ⇒ caseω p return p ⇒ B (fst A B p) of (x, y) ⇒ y)", note " ≔ (λ A B p ⇒ caseω p return p ⇒ B (fst A B p) of (x, y) ⇒ y)",
testTC "0 · type of snd ⇐ ★₂" $ testTC "0 · type of snd ⇐ ★₂" $
check_ empty szero sndTy (TYPE 2), check_ empty szero sndTy (^TYPE 2),
testTC "1 · def of sndtype of snd" $ testTC "1 · def of sndtype of snd" $
check_ empty sone sndDef sndTy, check_ empty sone sndDef sndTy,
testTC "0 · snd ★₀ (λ x ⇒ x) ⇒ (ω·p : (A : ★₀) × A) → fst ★₀ (λ x ⇒ x) p" $ testTC "0 · snd ★₀ (λ x ⇒ x) ⇒ (ω·p : (A : ★₀) × A) → fst ★₀ (λ x ⇒ x) p" $
inferAs empty szero inferAs empty szero
(F "snd" :@@ [TYPE 0, [< "x"] :\\ BVT 0]) (^App (^App (^F "snd") (^TYPE 0)) (^LamY "x" (^BVT 0)))
(PiY Any "A" (SigY "A" (TYPE 0) $ BVT 0) $ (^PiY Any "p" (^SigY "A" (^TYPE 0) (^BVT 0))
(E $ F "fst" :@@ [TYPE 0, [< "x"] :\\ BVT 0, BVT 0])) (E $ ^App (^App (^App (^F "fst") (^TYPE 0)) (^LamY "x" (^BVT 0)))
(^BVT 0)))
], ],
"enums" :- [ "enums" :- [
testTC "1 · 'a ⇐ {a}" $ testTC "1 · 'a ⇐ {a}" $
check_ empty sone (Tag "a") (enum ["a"]), check_ empty sone (^Tag "a") (^enum ["a"]),
testTC "1 · 'a ⇐ {a, b, c}" $ testTC "1 · 'a ⇐ {a, b, c}" $
check_ empty sone (Tag "a") (enum ["a", "b", "c"]), check_ empty sone (^Tag "a") (^enum ["a", "b", "c"]),
testTCFail "1 · 'a ⇍ {b, c}" $ testTCFail "1 · 'a ⇍ {b, c}" $
check_ empty sone (Tag "a") (enum ["b", "c"]), check_ empty sone (^Tag "a") (^enum ["b", "c"]),
testTC "0=1 ⊢ 1 · 'a ⇐ {b, c}" $ testTC "0=1 ⊢ 1 · 'a ⇐ {b, c}" $
check_ empty01 sone (Tag "a") (enum ["b", "c"]) check_ empty01 sone (^Tag "a") (^enum ["b", "c"])
], ],
"enum matching" :- [ "enum matching" :- [
testTC "ω.x : {tt} ⊢ 1 · case1 x return {tt} of { 'tt ⇒ 'tt } ⇒ {tt}" $ testTC "ω.x : {tt} ⊢ 1 · case1 x return {tt} of { 'tt ⇒ 'tt } ⇒ {tt}" $
inferAs (ctx [< ("x", enum ["tt"])]) sone inferAs (ctx [< ("x", ^enum ["tt"])]) sone
(CaseEnum One (BV 0) (SN (enum ["tt"])) $ (^CaseEnum One (^BV 0) (SN (^enum ["tt"]))
singleton "tt" (Tag "tt")) (singleton "tt" (^Tag "tt")))
(enum ["tt"]), (^enum ["tt"]),
testTCFail "ω.x : {tt} ⊢ 1 · case1 x return {tt} of { 'ff ⇒ 'tt } ⇏" $ testTCFail "ω.x : {tt} ⊢ 1 · case1 x return {tt} of { 'ff ⇒ 'tt } ⇏" $
infer_ (ctx [< ("x", enum ["tt"])]) sone infer_ (ctx [< ("x", ^enum ["tt"])]) sone
(CaseEnum One (BV 0) (SN (enum ["tt"])) $ (^CaseEnum One (^BV 0) (SN (^enum ["tt"]))
singleton "ff" (Tag "tt")) (singleton "ff" (^Tag "tt")))
], ],
"equality types" :- [ "equality types" :- [
testTC "0 · : ★₀ ⇐ Type" $ testTC "0 · : ★₀ ⇐ Type" $
checkType_ empty (Eq0 (TYPE 0) Nat Nat) Nothing, checkType_ empty (^Eq0 (^TYPE 0) nat nat) Nothing,
testTC "0 · : ★₀ ⇐ ★₁" $ testTC "0 · : ★₀ ⇐ ★₁" $
check_ empty szero (Eq0 (TYPE 0) Nat Nat) (TYPE 1), check_ empty szero (^Eq0 (^TYPE 0) nat nat) (^TYPE 1),
testTCFail "1 · : ★₀ ⇍ ★₁" $ testTCFail "1 · : ★₀ ⇍ ★₁" $
check_ empty sone (Eq0 (TYPE 0) Nat Nat) (TYPE 1), check_ empty sone (^Eq0 (^TYPE 0) nat nat) (^TYPE 1),
testTC "0 · : ★₀ ⇐ ★₂" $ testTC "0 · : ★₀ ⇐ ★₂" $
check_ empty szero (Eq0 (TYPE 0) Nat Nat) (TYPE 2), check_ empty szero (^Eq0 (^TYPE 0) nat nat) (^TYPE 2),
testTC "0 · : ★₁ ⇐ ★₂" $ testTC "0 · : ★₁ ⇐ ★₂" $
check_ empty szero (Eq0 (TYPE 1) Nat Nat) (TYPE 2), check_ empty szero (^Eq0 (^TYPE 1) nat nat) (^TYPE 2),
testTCFail "0 · : ★₁ ⇍ ★₁" $ testTCFail "0 · : ★₁ ⇍ ★₁" $
check_ empty szero (Eq0 (TYPE 1) Nat Nat) (TYPE 1), check_ empty szero (^Eq0 (^TYPE 1) nat nat) (^TYPE 1),
testTCFail "0 ≡ 'beep : {beep} ⇍ Type" $ testTCFail "0 ≡ 'beep : {beep} ⇍ Type" $
checkType_ empty (Eq0 (enum ["beep"]) Zero (Tag "beep")) Nothing, checkType_ empty
(^Eq0 (^enum ["beep"]) (^Zero) (^Tag "beep"))
Nothing,
testTC "ab : A ≡ B : ★₀, x : A, y : B ⊢ 0 · Eq [i ⇒ ab i] x y ⇐ ★₀" $ testTC "ab : A ≡ B : ★₀, x : A, y : B ⊢ 0 · Eq [i ⇒ ab i] x y ⇐ ★₀" $
check_ (ctx [< ("ab", Eq0 (TYPE 0) (FT "A") (FT "B")), check_ (ctx [< ("ab", ^Eq0 (^TYPE 0) (^FT "A") (^FT "B")),
("x", FT "A"), ("y", FT "B")]) szero ("x", ^FT "A"), ("y", ^FT "B")]) szero
(Eq (SY [< "i"] $ E $ BV 2 :% BV 0) (BVT 1) (BVT 0)) (^EqY "i" (E $ ^DApp (^BV 2) (^BV 0)) (^BVT 1) (^BVT 0))
(TYPE 0), (^TYPE 0),
testTCFail "ab : A ≡ B : ★₀, x : A, y : B ⊢ Eq [i ⇒ ab i] y x ⇍ Type" $ testTCFail "ab : A ≡ B : ★₀, x : A, y : B ⊢ Eq [i ⇒ ab i] y x ⇍ Type" $
checkType_ (ctx [< ("ab", Eq0 (TYPE 0) (FT "A") (FT "B")), check_ (ctx [< ("ab", ^Eq0 (^TYPE 0) (^FT "A") (^FT "B")),
("x", FT "A"), ("y", FT "B")]) ("x", ^FT "A"), ("y", ^FT "B")]) szero
(Eq (SY [< "i"] $ E $ BV 2 :% BV 0) (BVT 0) (BVT 1)) (^EqY "i" (E $ ^DApp (^BV 2) (^BV 0)) (^BVT 0) (^BVT 1))
Nothing (^TYPE 0)
], ],
"equalities" :- [ "equalities" :- [
testTC "1 · (δ i ⇒ a) ⇐ a ≡ a" $ testTC "1 · (δ i ⇒ a) ⇐ a ≡ a" $
check_ empty sone (DLam $ SN $ FT "a") check_ empty sone (^DLamN (^FT "a"))
(Eq0 (FT "A") (FT "a") (FT "a")), (^Eq0 (^FT "A") (^FT "a") (^FT "a")),
testTC "0 · (λ p q ⇒ δ i ⇒ p) ⇐ (ω·p q : a ≡ a') → p ≡ q" $ testTC "0 · (λ p q ⇒ δ i ⇒ p) ⇐ (ω·p q : a ≡ a') → p ≡ q # uip" $
check_ empty szero check_ empty szero
([< "p","q"] :\\ [< "i"] :\\% BVT 1) (^LamY "p" (^LamY "q" (^DLamN (^BVT 1))))
(PiY Any "p" (Eq0 (FT "A") (FT "a") (FT "a")) $ (^PiY Any "p" (^Eq0 (^FT "A") (^FT "a") (^FT "a"))
PiY Any "q" (Eq0 (FT "A") (FT "a") (FT "a")) $ (^PiY Any "q" (^Eq0 (^FT "A") (^FT "a") (^FT "a"))
Eq0 (Eq0 (FT "A") (FT "a") (FT "a")) (BVT 1) (BVT 0)), (^Eq0 (^Eq0 (^FT "A") (^FT "a") (^FT "a")) (^BVT 1) (^BVT 0)))),
testTC "0 · (λ p q ⇒ δ i ⇒ q) ⇐ (ω·p q : a ≡ a') → p ≡ q" $ testTC "0 · (λ p q ⇒ δ i ⇒ q) ⇐ (ω·p q : a ≡ a') → p ≡ q # uip(2)" $
check_ empty szero check_ empty szero
([< "p","q"] :\\ [< "i"] :\\% BVT 0) (^LamY "p" (^LamY "q" (^DLamN (^BVT 0))))
(PiY Any "p" (Eq0 (FT "A") (FT "a") (FT "a")) $ (^PiY Any "p" (^Eq0 (^FT "A") (^FT "a") (^FT "a"))
PiY Any "q" (Eq0 (FT "A") (FT "a") (FT "a")) $ (^PiY Any "q" (^Eq0 (^FT "A") (^FT "a") (^FT "a"))
Eq0 (Eq0 (FT "A") (FT "a") (FT "a")) (BVT 1) (BVT 0)) (^Eq0 (^Eq0 (^FT "A") (^FT "a") (^FT "a")) (^BVT 1) (^BVT 0))))
], ],
"natural numbers" :- [ "natural numbers" :- [
testTC "0 · ⇐ ★₀" $ check_ empty szero Nat (TYPE 0), testTC "0 · ⇐ ★₀" $ check_ empty szero nat (^TYPE 0),
testTC "0 · ⇐ ★₇" $ check_ empty szero Nat (TYPE 7), testTC "0 · ⇐ ★₇" $ check_ empty szero nat (^TYPE 7),
testTCFail "1 · ⇍ ★₀" $ check_ empty sone Nat (TYPE 0), testTCFail "1 · ⇍ ★₀" $ check_ empty sone nat (^TYPE 0),
testTC "1 · zero ⇐ " $ check_ empty sone Zero Nat, testTC "1 · zero ⇐ " $ check_ empty sone (^Zero) nat,
testTCFail "1 · zero ⇍ ×" $ check_ empty sone Zero (Nat `And` Nat), testTCFail "1 · zero ⇍ ×" $ check_ empty sone (^Zero) (^And nat nat),
testTC "ω·n : ⊢ 1 · succ n ⇐ " $ testTC "ω·n : ⊢ 1 · succ n ⇐ " $
check_ (ctx [< ("n", Nat)]) sone (Succ (BVT 0)) Nat, check_ (ctx [< ("n", nat)]) sone (^Succ (^BVT 0)) nat,
testTC "1 · λ n ⇒ succ n ⇐ 1." $ testTC "1 · λ n ⇒ succ n ⇐ 1." $
check_ empty sone ([< "n"] :\\ Succ (BVT 0)) (Arr One Nat Nat), check_ empty sone
todo "nat elim" (^LamY "n" (^Succ (^BVT 0)))
(^Arr One nat nat)
], ],
"natural elim" :- [ "natural elim" :- [
@ -459,25 +482,28 @@ tests = "typechecker" :- [
note " ⇐ 1.", note " ⇐ 1.",
testTC "pred" $ testTC "pred" $
check_ empty sone check_ empty sone
([< "n"] :\\ E (CaseNat One Zero (BV 0) (SN Nat) (^LamY "n" (E $
Zero (SY [< "n", Unused] $ BVT 1))) ^CaseNat One Zero (^BV 0) (SN nat)
(Arr One Nat Nat), (^Zero) (SY [< "n", ^BN Unused] $ ^BVT 1)))
(^Arr One nat nat),
note "1 · λ m n ⇒ case1 m return of { zero ⇒ n; succ _, 1.p ⇒ succ p }", note "1 · λ m n ⇒ case1 m return of { zero ⇒ n; succ _, 1.p ⇒ succ p }",
note " ⇐ 1. → 1. → 1.", note " ⇐ 1. → 1. → 1.",
testTC "plus" $ testTC "plus" $
check_ empty sone check_ empty sone
([< "m", "n"] :\\ E (CaseNat One One (BV 1) (SN Nat) (^LamY "m" (^LamY "n" (E $
(BVT 0) (SY [< Unused, "p"] $ Succ $ BVT 0))) ^CaseNat One One (^BV 1) (SN nat)
(Arr One Nat $ Arr One Nat Nat) (^BVT 0)
(SY [< ^BN Unused, "p"] $ ^Succ (^BVT 0)))))
(^Arr One nat (^Arr One nat nat))
], ],
"box types" :- [ "box types" :- [
testTC "0 · [0.] ⇐ ★₀" $ testTC "0 · [0.] ⇐ ★₀" $
check_ empty szero (BOX Zero Nat) (TYPE 0), check_ empty szero (^BOX Zero nat) (^TYPE 0),
testTC "0 · [0.★₀] ⇐ ★₁" $ testTC "0 · [0.★₀] ⇐ ★₁" $
check_ empty szero (BOX Zero (TYPE 0)) (TYPE 1), check_ empty szero (^BOX Zero (^TYPE 0)) (^TYPE 1),
testTCFail "0 · [0.★₀] ⇍ ★₀" $ testTCFail "0 · [0.★₀] ⇍ ★₀" $
check_ empty szero (BOX Zero (TYPE 0)) (TYPE 0) check_ empty szero (^BOX Zero (^TYPE 0)) (^TYPE 0)
], ],
todo "box values", todo "box values",
@ -486,10 +512,14 @@ tests = "typechecker" :- [
"type-case" :- [ "type-case" :- [
testTC "0 · type-case ∷ ★₀ return ★₀ of { _ ⇒ } ⇒ ★₀" $ testTC "0 · type-case ∷ ★₀ return ★₀ of { _ ⇒ } ⇒ ★₀" $
inferAs empty szero inferAs empty szero
(TypeCase (Nat :# TYPE 0) (TYPE 0) empty Nat) (^TypeCase (^Ann nat (^TYPE 0)) (^TYPE 0) empty nat)
(TYPE 0) (^TYPE 0)
], ],
todo "add the examples dir to the tests"
]
{-
"misc" :- [ "misc" :- [
note "0·A : Type, 0·P : A → Type, ω·p : (1·x : A) → P x", note "0·A : Type, 0·P : A → Type, ω·p : (1·x : A) → P x",
note "", note "",
@ -524,4 +554,4 @@ tests = "typechecker" :- [
-- return A -- return A
-- of { } -- of { }
] ]
] -}

View File

@ -5,6 +5,7 @@ depends = base, contrib, elab-util, snocvect, quox-lib, tap, eff
executable = quox-tests executable = quox-tests
main = Tests main = Tests
modules = modules =
AstExtra,
TypingImpls, TypingImpls,
PrettyExtra, PrettyExtra,
Tests.DimEq, Tests.DimEq,