use NContext/SnocVect for scope name lists etc

This commit is contained in:
rhiannon morris 2023-03-16 18:18:49 +01:00
parent 32f38238ef
commit 6dc7177be5
12 changed files with 165 additions and 134 deletions

View file

@ -7,6 +7,7 @@ import Quox.Name
import Data.DPair
import Data.Nat
import Data.SnocList
import Data.SnocVect
import Data.Vect
import Control.Monad.Identity
@ -68,6 +69,21 @@ export %inline
toList' : Telescope' a _ _ -> List a
toList' = map snd . toList
export
fromSnocVect : SnocVect n a -> Context' a n
fromSnocVect [<] = [<]
fromSnocVect (sx :< x) = fromSnocVect sx :< x
public export
tabulate : ((n : Nat) -> tm n) -> (n : Nat) -> Context tm n
tabulate f 0 = [<]
tabulate f (S k) = tabulate f k :< f k
public export
replicate : (n : Nat) -> a -> Context' a n
replicate n x = tabulate (const x) n
public export
(.) : Telescope tm from mid -> Telescope tm mid to -> Telescope tm from to
@ -81,6 +97,10 @@ export
rewrite plusSuccRightSucc n to in
(tel :< x) <>< xs
export
(++) : Telescope' a from to -> SnocVect n a -> Telescope' a from (n + to)
tel ++ [<] = tel
tel ++ (sx :< x) = (tel ++ sx) :< x
public export
getShiftWith : (forall from, to. tm from -> Shift from to -> tm to) ->
@ -236,6 +256,12 @@ foldLazy : Monoid a => Telescope' (Lazy a) from to -> a
foldLazy = foldMap force
export %inline
all, any : (forall n. tm n -> Bool) -> Telescope tm from to -> Bool
all p = foldMap @{All} p
any p = foldMap @{Any} p
export %inline
(forall n. Eq (tm n)) => Eq (Telescope tm from to) where
(==) = foldLazy @{All} .: zipWithLazy (==)

View file

@ -368,7 +368,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, IsQty q)}
ety <- computeElimType ctx e (noOr1 ne)
compareType (extendTy zero eret.name ety ctx) eret.term fret.term
(fst, snd) <- expectSigE defs ctx ety
let [x, y] = ebody.names
let [< x, y] = ebody.names
Term.compare0 (extendTyN [< (epi, x, fst), (epi, y, snd.term)] ctx)
(substCasePairRet ety eret)
ebody.term fbody.term

View file

@ -4,8 +4,9 @@ module Quox.Parser.FromParser
import Quox.Parser.Syntax
import Quox.Parser.Parser
import Quox.Typechecker
import Data.List
import Data.List
import Data.SnocVect
import public Control.Monad.Either
import public Control.Monad.State
import public Control.Monad.Reader
@ -91,16 +92,16 @@ mutual
Pi pi x s t =>
Pi pi <$> fromPTermWith ds ns s
<*> fromPTermTScope ds ns [x] t
<*> fromPTermTScope ds ns [< x] t
Lam x s =>
Lam <$> fromPTermTScope ds ns [x] s
Lam <$> fromPTermTScope ds ns [< x] s
s :@ t =>
map E $ (:@) <$> fromPTermElim ds ns s <*> fromPTermWith ds ns t
Sig x s t =>
Sig <$> fromPTermWith ds ns s <*> fromPTermTScope ds ns [x] t
Sig <$> fromPTermWith ds ns s <*> fromPTermTScope ds ns [< x] t
Pair s t =>
Pair <$> fromPTermWith ds ns s <*> fromPTermWith ds ns t
@ -108,13 +109,13 @@ mutual
Case pi pair (r, ret) (CasePair (x, y) body) =>
map E $ Base.CasePair pi
<$> fromPTermElim ds ns pair
<*> fromPTermTScope ds ns [r] ret
<*> fromPTermTScope ds ns [x, y] body
<*> fromPTermTScope ds ns [< r] ret
<*> fromPTermTScope ds ns [< x, y] body
Case pi tag (r, ret) (CaseEnum arms) =>
map E $ Base.CaseEnum pi
<$> fromPTermElim ds ns tag
<*> fromPTermTScope ds ns [r] ret
<*> fromPTermTScope ds ns [< r] ret
<*> assert_total fromPTermEnumArms ds ns arms
Enum strs =>
@ -127,12 +128,12 @@ mutual
Tag str => pure $ Tag str
Eq (i, ty) s t =>
Eq <$> fromPTermDScope ds ns [i] ty
Eq <$> fromPTermDScope ds ns [< i] ty
<*> fromPTermWith ds ns s
<*> fromPTermWith ds ns t
DLam i s =>
DLam <$> fromPTermDScope ds ns [i] s
DLam <$> fromPTermDScope ds ns [< i] s
s :% p =>
map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p
@ -163,24 +164,26 @@ mutual
private
fromPTermTScope : {s : Nat} -> CanError m =>
Context' BName d -> Context' BName n ->
Vect s BName ->
SnocVect s BName ->
PTerm -> m (ScopeTermN s Three d n)
fromPTermTScope ds ns xs t =
if all isNothing xs then
SN <$> fromPTermWith ds ns t
else
SY (map (maybe Unused UN) xs) <$> fromPTermWith ds (ns <>< xs) t
SY (fromSnocVect $ map (maybe Unused UN) xs)
<$> fromPTermWith ds (ns ++ xs) t
private
fromPTermDScope : {s : Nat} -> CanError m =>
Context' BName d -> Context' BName n ->
Vect s BName ->
SnocVect s BName ->
PTerm -> m (DScopeTermN s Three d n)
fromPTermDScope ds ns xs t =
if all isNothing xs then
SN <$> fromPTermWith ds ns t
else
SY (map (maybe Unused UN) xs) <$> fromPTermWith (ds <>< xs) ns t
SY (fromSnocVect $ map (maybe Unused UN) xs)
<$> fromPTermWith (ds ++ xs) ns t
export %inline

View file

@ -121,14 +121,14 @@ mutual
toPTermWith' ds ns s = case s of
TYPE l =>
TYPE l
Pi qty arg (S [x] res) =>
Pi qty arg (S [< x] res) =>
Pi qty (Just $ show x)
(toPTermWith ds ns arg)
(toPTermWith ds (ns :< baseStr x) res.term)
Lam (S [x] body) =>
Lam (S [< x] body) =>
Lam (Just $ show x) $
toPTermWith ds (ns :< baseStr x) body.term
Sig fst (S [x] snd) =>
Sig fst (S [< x] snd) =>
Sig (Just $ show x)
(toPTermWith ds ns fst)
(toPTermWith ds (ns :< baseStr x) snd.term)
@ -138,10 +138,10 @@ mutual
Enum $ SortedSet.toList cases
Tag tag =>
Tag tag
Eq (S [i] ty) l r =>
Eq (S [< i] ty) l r =>
Eq (Just $ show i, toPTermWith (ds :< baseStr i) ns ty.term)
(toPTermWith ds ns l) (toPTermWith ds ns r)
DLam (S [i] body) =>
DLam (S [< i] body) =>
DLam (Just $ show i) $ toPTermWith (ds :< baseStr i) ns body.term
E e =>
toPTermWith ds ns e
@ -165,12 +165,12 @@ mutual
V $ MakePName [<] $ ns !!! i
fun :@ arg =>
toPTermWith ds ns fun :@ toPTermWith ds ns arg
CasePair qty pair (S [r] ret) (S [x, y] body) =>
CasePair qty pair (S [< r] ret) (S [< x, y] body) =>
Case qty (toPTermWith ds ns pair)
(Just $ show r, toPTermWith ds (ns :< baseStr r) ret.term)
(CasePair (Just $ show x, Just $ show y) $
toPTermWith ds (ns :< baseStr x :< baseStr y) body.term)
CaseEnum qty tag (S [r] ret) arms =>
CaseEnum qty tag (S [< r] ret) arms =>
Case qty (toPTermWith ds ns tag)
(Just $ show r, toPTermWith ds (ns :< baseStr r) ret.term)
(CaseEnum $ mapSnd (toPTermWith ds ns) <$> SortedMap.toList arms)

View file

@ -6,6 +6,7 @@ import public Quox.Syntax.Subst
import public Quox.Syntax.Qty
import public Quox.Syntax.Dim
import public Quox.Name
import public Quox.Context
-- import public Quox.OPE
import Quox.Pretty
@ -16,7 +17,6 @@ import Data.Maybe
import Data.Nat
import public Data.So
import Data.String
import Data.Vect
import public Data.SortedMap
import public Data.SortedSet
@ -132,7 +132,7 @@ mutual
public export
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
constructor S
names : Vect s BaseName
names : NContext s
body : ScopedBody s f n
public export
@ -158,16 +158,16 @@ mutual
||| scope which ignores all its binders
public export %inline
SN : {s : Nat} -> f n -> Scoped s f n
SN = S (replicate s "_") . N
SN = S (replicate s Unused) . N
||| scope which uses its binders
public export %inline
SY : Vect s BaseName -> f (s + n) -> Scoped s f n
SY : NContext s -> f (s + n) -> Scoped s f n
SY ns = S ns . Y
public export %inline
name : Scoped 1 f n -> BaseName
name (S [x] _) = x
name (S [< x] _) = x
public export %inline
(.name) : Scoped 1 f n -> BaseName
@ -177,7 +177,7 @@ s.name = name s
public export %inline
Pi_ : (qty : q) -> (x : BaseName) ->
(arg : Term q d n) -> (res : Term q d (S n)) -> Term q d n
Pi_ {qty, x, arg, res} = Pi {qty, arg, res = S [x] $ Y res}
Pi_ {qty, x, arg, res} = Pi {qty, arg, res = S [< x] $ Y res}
||| non dependent function type
public export %inline
@ -188,7 +188,7 @@ Arr {qty, arg, res} = Pi {qty, arg, res = SN res}
public export %inline
Sig_ : (x : BaseName) -> (fst : Term q d n) ->
(snd : Term q d (S n)) -> Term q d n
Sig_ {x, fst, snd} = Sig {fst, snd = S [x] $ Y snd}
Sig_ {x, fst, snd} = Sig {fst, snd = S [< x] $ Y snd}
||| non dependent pair type
public export %inline
@ -199,7 +199,7 @@ And {fst, snd} = Sig {fst, snd = SN snd}
public export %inline
Eq_ : (i : BaseName) -> (ty : Term q (S d) n) ->
(l, r : Term q d n) -> Term q d n
Eq_ {i, ty, l, r} = Eq {ty = S [i] $ Y ty, l, r}
Eq_ {i, ty, l, r} = Eq {ty = S [< i] $ Y ty, l, r}
||| non dependent equality type
public export %inline

View file

@ -73,12 +73,13 @@ prettyArm sort xs pat body = do
export
prettyLams : PrettyHL a => Pretty.HasEnv m =>
Maybe (Doc HL) -> BinderSort -> List BaseName -> a -> m (Doc HL)
Maybe (Doc HL) -> BinderSort -> SnocList BaseName -> a ->
m (Doc HL)
prettyLams lam sort names body = do
let var = case sort of T => TVar; D => DVar
header <- sequence $ [hlF var $ prettyM x | x <- names]
header <- sequence $ [hlF var $ prettyM x | x <- toList names]
let header = sep $ maybe header (:: header) lam
parensIfM Outer =<< prettyArm sort (cast names) header body
parensIfM Outer =<< prettyArm sort names header body
export
prettyApps : PrettyHL f => PrettyHL a => Pretty.HasEnv m =>
@ -110,7 +111,7 @@ prettyCase : PrettyHL a => PrettyHL b => PrettyHL c => PrettyHL q =>
m (Doc HL)
prettyCase pi elim r ret arms = do
elim <- prettyQtyBinds [pi] elim
ret <- prettyLams Nothing T [r] ret
ret <- prettyLams Nothing T [< r] ret
arms <- prettyArms arms
pure $ asep [caseD <++> elim, returnD <++> ret, ofD <++> arms]
@ -127,12 +128,12 @@ parameters (showSubsts : Bool)
where
prettyM (TYPE l) =
parensIfM App $ !typeD <+> hl Syntax !(prettyUnivSuffix l)
prettyM (Pi qty s (S [x] t)) =
prettyM (Pi qty s (S [< x] t)) =
prettyBindType [qty] x s !arrowD t.term
prettyM (Lam (S x t)) =
let GotLams {names, body, _} = getLams' x t.term Refl in
prettyLams (Just !lamD) T (toList names) body
prettyM (Sig s (S [x] t)) =
prettyLams (Just !lamD) T (toSnocList' names) body
prettyM (Sig s (S [< x] t)) =
prettyBindType {q} [] x s !timesD t.term
prettyM (Pair s t) =
let GotPairs {init, last, _} = getPairs' [< s] t in
@ -147,14 +148,14 @@ parameters (showSubsts : Bool)
r <- withPrec InEq $ prettyM r
ty <- withPrec InEq $ prettyM ty
parensIfM Eq $ asep [l <++> !eqndD, r <++> colonD, ty]
prettyM (Eq (S [i] (Y ty)) l r) = do
ty <- bracks <$> withPrec Outer (prettyLams Nothing D [i] ty)
prettyM (Eq (S [< i] (Y ty)) l r) = do
ty <- bracks <$> withPrec Outer (prettyLams Nothing D [< i] ty)
l <- withPrec Arg $ prettyM l
r <- withPrec Arg $ prettyM r
parensIfM App $ eqD <++> asep [ty, l, r]
prettyM (DLam (S i t)) =
let GotDLams {names, body, _} = getDLams' i t.term Refl in
prettyLams (Just !dlamD) D (toList names) body
prettyLams (Just !dlamD) D (toSnocList' names) body
prettyM (E e) = prettyM e
prettyM (CloT s th) =
if showSubsts then
@ -179,10 +180,10 @@ parameters (showSubsts : Bool)
prettyM (e :@ s) =
let GotArgs {fun, args, _} = getArgs' e [s] in
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]
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
[([<], prettyTag t, b) | (t, b) <- SortedMap.toList arms]
prettyM (e :% d) =

View file

@ -2,6 +2,7 @@ module Quox.Syntax.Term.Split
import Quox.Syntax.Term.Base
import Quox.Syntax.Term.Subst
import Quox.Context
import public Quox.No
import public Data.Vect
@ -130,25 +131,23 @@ getDArgs e = getDArgs' e []
infixr 1 :\\
public export
absN : Vect m BaseName -> Term q d (m + n) -> Term q d n
absN {m = 0} [] s = s
absN {m = S m} (x :: xs) s =
Lam $ S [x] $ Y $ absN xs $ rewrite sym $ plusSuccRightSucc m n in s
absN : NContext m -> Term q d (m + n) -> Term q d n
absN [<] s = s
absN (xs :< x) s = absN xs $ Lam $ SY [< x] s
public export %inline
(:\\) : Vect m BaseName -> Term q d (m + n) -> Term q d n
(:\\) : NContext m -> Term q d (m + n) -> Term q d n
(:\\) = absN
infixr 1 :\\%
public export
dabsN : Vect m BaseName -> Term q (m + d) n -> Term q d n
dabsN {m = 0} [] s = s
dabsN {m = S m} (x :: xs) s =
DLam $ S [x] $ Y $ dabsN xs $ rewrite sym $ plusSuccRightSucc m d in s
dabsN : NContext m -> Term q (m + d) n -> Term q d n
dabsN [<] s = s
dabsN (xs :< x) s = dabsN xs $ DLam $ SY [< x] s
public export %inline
(:\\%) : Vect m BaseName -> Term q (m + d) n -> Term q d n
(:\\%) : NContext m -> Term q (m + d) n -> Term q d n
(:\\%) = dabsN
@ -156,7 +155,7 @@ public export
record GetLams q d n where
constructor GotLams
{0 lams, rest : Nat}
names : Vect lams BaseName
names : NContext lams
body : Term q d rest
0 eq : lams + n = rest
0 notLam : NotLam body
@ -164,7 +163,7 @@ record GetLams q d n where
mutual
export %inline
getLams' : forall lams, rest.
Vect lams BaseName -> Term q d rest -> (0 eq : lams + n = rest) ->
NContext lams -> Term q d rest -> (0 eq : lams + n = rest) ->
GetLams q d n
getLams' xs s0 eq =
let Element s nc = pushSubsts s0 in
@ -172,25 +171,25 @@ mutual
private
getLamsNc : forall lams, rest.
Vect lams BaseName ->
NContext lams ->
(t : Term q d rest) -> (0 nc : NotClo t) =>
(0 eq : lams + n = rest) ->
GetLams q d n
getLamsNc xs s Refl = case nchoose $ isLam s of
Left y => let Lam (S [x] body) = s in
getLams' (x :: xs) (assert_smaller s body.term) Refl
Left y => let Lam (S [< x] body) = s in
getLams' (xs :< x) (assert_smaller s body.term) Refl
Right n => GotLams xs s Refl n
export %inline
getLams : Term q d n -> GetLams q d n
getLams s = getLams' [] s Refl
getLams s = getLams' [<] s Refl
public export
record GetDLams q d n where
constructor GotDLams
{0 lams, rest : Nat}
names : Vect lams BaseName
names : NContext lams
body : Term q rest n
0 eq : lams + d = rest
0 notDLam : NotDLam body
@ -198,7 +197,7 @@ record GetDLams q d n where
mutual
export %inline
getDLams' : forall lams, rest.
Vect lams BaseName -> Term q rest n -> (0 eq : lams + d = rest) ->
NContext lams -> Term q rest n -> (0 eq : lams + d = rest) ->
GetDLams q d n
getDLams' xs s0 eq =
let Element s nc = pushSubsts s0 in
@ -206,18 +205,18 @@ mutual
private
getDLamsNc : forall lams, rest.
Vect lams BaseName ->
NContext lams ->
(t : Term q rest n) -> (0 nc : NotClo t) =>
(0 eq : lams + d = rest) ->
GetDLams q d n
getDLamsNc is s Refl = case nchoose $ isDLam s of
Left y => let DLam (S [i] body) = s in
getDLams' (i :: is) (assert_smaller s body.term) Refl
Left y => let DLam (S [< i] body) = s in
getDLams' (is :< i) (assert_smaller s body.term) Refl
Right n => GotDLams is s Refl n
export %inline
getDLams : Term q d n -> GetDLams q d n
getDLams s = getDLams' [] s Refl
getDLams s = getDLams' [<] s Refl
public export

View file

@ -290,7 +290,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐
-- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y
-- with ρ₁, ρ₂ ≤ πσ
let [x, y] = body.names
let [< x, y] = body.names
pisg = pi * sg.fst
bodyctx = extendTyN [< (pisg, x, tfst), (pisg, y, tsnd.term)] ctx
bodyty = substCasePairRet pairres.type ret