WIP: 𝕎 #25

Closed
rhi wants to merge 8 commits from 𝕎 into 🐉
11 changed files with 36 additions and 36 deletions
Showing only changes of commit ca1adcce47 - Show all commits

View file

@ -113,6 +113,10 @@ public export
0 Located1 : (a -> Type) -> Type 0 Located1 : (a -> Type) -> Type
Located1 f = forall x. Located (f x) Located1 f = forall x. Located (f x)
public export
0 Located2 : (a -> b -> Type) -> Type
Located2 f = forall x, y. Located (f x y)
public export public export
interface Located a => Relocatable a where setLoc : Loc -> a -> a interface Located a => Relocatable a where setLoc : Loc -> a -> a

View file

@ -215,14 +215,14 @@ tycaseRhsDef0 def k arms = fromMaybe def $ tycaseRhs0 k arms
export export
weakAtT : CanTSubst f => (by, at : Nat) -> weakAtT : (CanTSubst f, Located2 f) => (by, at : Nat) ->
f d (at + n) -> f d (at + (by + n)) f d (at + n) -> f d (at + (by + n))
weakAtT by at t = t `CanTSubst.(//)` pushN at (shift by) weakAtT by at t = t `CanTSubst.(//)` pushN at (shift by) t.loc
export export
weakAtD : CanDSubst f => (by, at : Nat) -> weakAtD : (CanDSubst f, Located2 f) => (by, at : Nat) ->
f (at + d) n -> f (at + (by + d)) n f (at + d) n -> f (at + (by + d)) n
weakAtD by at t = t `CanDSubst.(//)` pushN at (shift by) weakAtD by at t = t `CanDSubst.(//)` pushN at (shift by) t.loc
parameters {s : Nat} parameters {s : Nat}
export export

View file

@ -83,7 +83,7 @@ DSubst : Nat -> Nat -> Type
DSubst = Subst Dim DSubst = Subst Dim
public export FromVar Dim where fromVarLoc = B public export FromVar Dim where fromVar = B
export export

View file

@ -70,13 +70,13 @@ toMaybe (Just x) = Just x
export export
fromGround' : Context' DimConst d -> DimEq' d fromGround' : BContext d -> Context' DimConst d -> DimEq' d
fromGround' [<] = [<] fromGround' [<] [<] = [<]
fromGround' (ctx :< e) = fromGround' ctx :< Just (K e noLoc) fromGround' (xs :< x) (ctx :< e) = fromGround' xs ctx :< Just (K e x.loc)
export export
fromGround : Context' DimConst d -> DimEq d fromGround : BContext d -> Context' DimConst d -> DimEq d
fromGround = C . fromGround' fromGround = C .: fromGround'
public export %inline public export %inline
@ -243,7 +243,7 @@ prettyCsts : {opts : _} -> BContext d -> DimEq' d ->
prettyCsts [<] [<] = pure [<] prettyCsts [<] [<] = pure [<]
prettyCsts dnames (eqs :< Nothing) = prettyCsts (tail dnames) eqs prettyCsts dnames (eqs :< Nothing) = prettyCsts (tail dnames) eqs
prettyCsts dnames (eqs :< Just q) = prettyCsts dnames (eqs :< Just q) =
[|prettyCsts (tail dnames) eqs :< prettyCst dnames (BV 0 noLoc) (weakD 1 q)|] [|prettyCsts (tail dnames) eqs :< prettyCst dnames (BV 0 q.loc) (weakD 1 q)|]
export export
prettyDimEq' : {opts : _} -> BContext d -> DimEq' d -> Eff Pretty (Doc opts) prettyDimEq' : {opts : _} -> BContext d -> DimEq' d -> Eff Pretty (Doc opts)

View file

@ -49,7 +49,7 @@ interface FromVar term => CanSubstSelf term where
public export public export
getLoc : FromVar term => Subst term from to -> Var from -> Loc -> term to getLoc : FromVar term => Subst term from to -> Var from -> Loc -> term to
getLoc (Shift by) i loc = fromVarLoc (shift by i) loc getLoc (Shift by) i loc = fromVar (shift by i) loc
getLoc (t ::: th) VZ _ = t getLoc (t ::: th) VZ _ = t
getLoc (t ::: th) (VS i) loc = getLoc th i loc getLoc (t ::: th) (VS i) loc = getLoc th i loc
@ -95,18 +95,18 @@ map f (t ::: th) = f t ::: map f th
public export %inline public export %inline
push : CanSubstSelf f => Subst f from to -> Subst f (S from) (S to) push : CanSubstSelf f => Subst f from to -> Loc -> Subst f (S from) (S to)
push th = fromVar VZ ::: (th . shift 1) push th loc = fromVar VZ loc ::: (th . shift 1)
-- [fixme] a better way to do this? -- [fixme] a better way to do this?
public export public export
pushN : CanSubstSelf f => (s : Nat) -> pushN : CanSubstSelf f => (s : Nat) ->
Subst f from to -> Subst f (s + from) (s + to) Subst f from to -> Loc -> Subst f (s + from) (s + to)
pushN 0 th = th pushN 0 th _ = th
pushN (S s) th = pushN (S s) th loc =
rewrite plusSuccRightSucc s from in rewrite plusSuccRightSucc s from in
rewrite plusSuccRightSucc s to in rewrite plusSuccRightSucc s to in
pushN s $ fromVar VZ ::: (th . shift 1) pushN s (fromVar VZ loc ::: (th . shift 1)) loc
public export public export
drop1 : Subst f (S from) to -> Subst f from to drop1 : Subst f (S from) to -> Subst f from to

View file

@ -56,12 +56,12 @@ namespace DSubst.DScopeTermN
(//) : {s : Nat} -> (//) : {s : Nat} ->
DScopeTermN s d1 n -> Lazy (DSubst d1 d2) -> DScopeTermN s d1 n -> Lazy (DSubst d1 d2) ->
DScopeTermN s d2 n DScopeTermN s d2 n
S ns (Y body) // th = S ns $ Y $ body // pushN s th S ns (Y body) // th = S ns $ Y $ body // pushN s th body.loc
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 fromVarLoc = B export %inline FromVar (Elim d) where fromVar = B
export %inline FromVar (Term d) where fromVarLoc = E .: fromVar export %inline FromVar (Term d) where fromVar = E .: fromVar
||| does the minimal reasonable work: ||| does the minimal reasonable work:
@ -104,7 +104,7 @@ namespace ScopeTermN
(//) : {s : Nat} -> (//) : {s : Nat} ->
ScopeTermN s d n1 -> Lazy (TSubst d n1 n2) -> ScopeTermN s d n1 -> Lazy (TSubst d n1 n2) ->
ScopeTermN s d n2 ScopeTermN s d n2
S ns (Y body) // th = S ns $ Y $ body // pushN s th S ns (Y body) // th = S ns $ Y $ body // pushN s th body.loc
S ns (N body) // th = S ns $ N $ body // th S ns (N body) // th = S ns $ N $ body // th
namespace DScopeTermN namespace DScopeTermN
@ -189,12 +189,12 @@ dsub1 t p = dsubN t [< p]
public export %inline public export %inline
(.zero) : DScopeTerm d n -> {default noLoc loc : Loc} -> Term d n (.zero) : DScopeTerm d n -> Term d n
body.zero = dsub1 body $ K Zero loc body.zero = dsub1 body $ K Zero body.loc
public export %inline public export %inline
(.one) : DScopeTerm d n -> {default noLoc loc : Loc} -> Term d n (.one) : DScopeTerm d n -> Term d n
body.one = dsub1 body $ K One loc body.one = dsub1 body $ K One body.loc
public export public export

View file

@ -139,13 +139,9 @@ weakIsSpec p i = toNatInj $ trans (weakCorrect p i) (sym $ weakSpecCorrect p i)
public export public export
interface FromVar f where %inline fromVarLoc : Var n -> Loc -> f n interface FromVar f where %inline fromVar : Var n -> Loc -> f n
public export %inline public export FromVar Var where fromVar x _ = x
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

@ -347,7 +347,7 @@ mutual
infres <- inferC ctx szero e infres <- inferC ctx szero e
-- if Ψ | Γ ⊢ Type <: Type 𝓀 -- if Ψ | Γ ⊢ Type <: Type 𝓀
case u of case u of
Just u => lift $ subtype e.loc ctx infres.type (TYPE u noLoc) Just u => lift $ subtype e.loc ctx infres.type (TYPE u e.loc)
Nothing => ignore $ expectTYPE !(askAt DEFS) ctx e.loc infres.type Nothing => ignore $ expectTYPE !(askAt DEFS) ctx e.loc infres.type
-- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀 -- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀

View file

@ -62,7 +62,7 @@ substCaseWRet [< x, y, ih] dty retty =
public export public export
substCaseSuccRet : BContext 2 -> ScopeTerm d n -> Term d (2 + n) substCaseSuccRet : BContext 2 -> ScopeTerm d n -> Term d (2 + n)
substCaseSuccRet [< p, ih] retty = substCaseSuccRet [< p, ih] retty =
let arg = Ann (Succ (BVT 1 p.loc) p.loc) (Nat noLoc) $ p.loc `extendL` ih.loc let arg = Ann (Succ (BVT 1 p.loc) p.loc) (Nat p.loc) $ p.loc `extendL` ih.loc
in in
retty.term // (arg ::: shift 2) retty.term // (arg ::: shift 2)

View file

@ -197,7 +197,7 @@ namespace EqContext
toTyContext : (e : EqContext n) -> TyContext e.dimLen n toTyContext : (e : EqContext n) -> TyContext e.dimLen n
toTyContext (MkEqContext {dimLen, dassign, dnames, tctx, tnames, qtys}) = toTyContext (MkEqContext {dimLen, dassign, dnames, tctx, tnames, qtys}) =
MkTyContext { MkTyContext {
dctx = fromGround dassign, dctx = fromGround dnames dassign,
tctx = map (// shift0 dimLen) tctx, tctx = map (// shift0 dimLen) tctx,
dnames, tnames, qtys dnames, tnames, qtys
} }

View file

@ -97,7 +97,7 @@ tests = "dimension constraints" :- [
testPrettyD iijj ZeroIsOne "𝑖, 𝑗, 0 = 1", testPrettyD iijj ZeroIsOne "𝑖, 𝑗, 0 = 1",
testPrettyD [<] new "" {label = "[empty output from empty context]"}, testPrettyD [<] new "" {label = "[empty output from empty context]"},
testPrettyD ii new "𝑖", testPrettyD ii new "𝑖",
testPrettyD iijj (fromGround [< Zero, One]) testPrettyD iijj (fromGround iijj [< Zero, One])
"𝑖, 𝑗, 𝑖 = 0, 𝑗 = 1", "𝑖, 𝑗, 𝑖 = 0, 𝑗 = 1",
testPrettyD iijj (C [< Just (^K Zero), Nothing]) testPrettyD iijj (C [< Just (^K Zero), Nothing])
"𝑖, 𝑗, 𝑖 = 0", "𝑖, 𝑗, 𝑖 = 0",