check for 0=1 in typechecker

This commit is contained in:
rhiannon morris 2023-02-19 17:51:44 +01:00
parent 195791e158
commit 858b5db530
5 changed files with 95 additions and 55 deletions

View file

@ -28,6 +28,12 @@ data DimEq : Nat -> Type where
%name DimEq eqs %name DimEq eqs
public export
data IfConsistent : DimEq d -> a -> Type where
Nothing : IfConsistent ZeroIsOne a
Just : a -> IfConsistent (C eqs) a
export %inline export %inline
zeroEq : DimEq 0 zeroEq : DimEq 0
zeroEq = C [<] zeroEq = C [<]

View file

@ -1,11 +1,7 @@
module Quox.Typechecker module Quox.Typechecker
import public Quox.Syntax
import public Quox.Typing import public Quox.Typing
import public Quox.Equal import public Quox.Equal
import public Control.Monad.Either
import Decidable.Decidable
import Data.SnocVect
%default total %default total
@ -22,7 +18,7 @@ CanTC q = CanTC' q IsGlobal
private private
popQs : HasErr q m => IsQty q => popQs : HasErr q m => IsQty q =>
SnocVect s q -> QOutput q (s + n) -> m (QOutput q n) QOutput q s -> QOutput q (s + n) -> m (QOutput q n)
popQs [<] qout = pure qout popQs [<] qout = pure qout
popQs (pis :< pi) (qout :< rh) = do expectCompatQ rh pi; popQs pis qout popQs (pis :< pi) (qout :< rh) = do expectCompatQ rh pi; popQs pis qout
@ -33,20 +29,15 @@ popQ pi = popQs [< pi]
private %inline private %inline
tail : TyContext q d (S n) -> TyContext q d n weakI : IsQty q => InferResult' q d n -> InferResult' q d (S n)
tail = {tctx $= tail}
private %inline
weakI : IsQty q => InferResult q d n -> InferResult q d (S n)
weakI = {type $= weakT, qout $= (:< zero)} weakI = {type $= weakT, qout $= (:< zero)}
private private
lookupBound : IsQty q => q -> Var n -> TyContext q d n -> InferResult q d n lookupBound : IsQty q => q -> Var n -> TContext q d n -> InferResult' q d n
lookupBound pi VZ (MkTyContext {tctx = tctx :< ty, _}) = lookupBound pi VZ (ctx :< ty) =
InfRes {type = weakT ty, qout = (zero <$ tctx) :< pi} InfRes {type = weakT ty, qout = (zero <$ ctx) :< pi}
lookupBound pi (VS i) ctx = lookupBound pi (VS i) (ctx :< _) =
weakI $ lookupBound pi i (tail ctx) weakI $ lookupBound pi i ctx
private %inline private %inline
lookupFree : CanTC' q g m => Name -> m (Definition' q g) lookupFree : CanTC' q g m => Name -> m (Definition' q g)
@ -65,12 +56,16 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
||| `check ctx sg subj ty` checks that in the context `ctx`, the term ||| `check ctx sg subj ty` checks that in the context `ctx`, the term
||| `subj` has the type `ty`, with quantity `sg`. if so, returns the ||| `subj` has the type `ty`, with quantity `sg`. if so, returns the
||| quantities of all bound variables that it used. ||| quantities of all bound variables that it used.
|||
||| if the dimension context is inconsistent, then return `Nothing`, without
||| doing any further work.
export covering %inline export covering %inline
check : TyContext q d n -> SQty q -> Term q d n -> Term q d n -> check : (ctx : TyContext q d n) -> SQty q -> Term q d n -> Term q d n ->
m (CheckResult q n) m (CheckResult ctx.dctx q n)
check ctx sg subj ty = check ctx@(MkTyContext {dctx, _}) sg subj ty =
let Element subj nc = pushSubsts subj in case dctx of
check' ctx sg subj nc ty ZeroIsOne => pure Nothing
C _ => Just <$> checkC ctx sg subj ty
||| "Ψ | Γ ⊢₀ s ⇐ A" ||| "Ψ | Γ ⊢₀ s ⇐ A"
||| |||
@ -80,21 +75,43 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
check0 ctx tm ty = ignore $ check ctx szero tm ty check0 ctx tm ty = ignore $ check ctx szero tm ty
-- the output will always be 𝟎 because the subject quantity is 0 -- the output will always be 𝟎 because the subject quantity is 0
||| `check`, assuming the dimension context is consistent
export covering %inline
checkC : (ctx : TyContext q d n) -> SQty q -> Term q d n -> Term q d n ->
m (CheckResult' q n)
checkC ctx sg subj ty =
let Element subj nc = pushSubsts subj in
check' ctx sg subj nc ty
||| "Ψ | Γ ⊢ σ · e ⇒ A ⊳ Σ" ||| "Ψ | Γ ⊢ σ · e ⇒ A ⊳ Σ"
||| |||
||| `infer ctx sg subj` infers the type of `subj` in the context `ctx`, ||| `infer ctx sg subj` infers the type of `subj` in the context `ctx`,
||| and returns its type and the bound variables it used. ||| and returns its type and the bound variables it used.
|||
||| if the dimension context is inconsistent, then return `Nothing`, without
||| doing any further work.
export covering %inline export covering %inline
infer : TyContext q d n -> SQty q -> Elim q d n -> m (InferResult q d n) infer : (ctx : TyContext q d n) -> SQty q -> Elim q d n ->
infer ctx sg subj = m (InferResult ctx.dctx q d n)
infer ctx@(MkTyContext {dctx, _}) sg subj =
case dctx of
ZeroIsOne => pure Nothing
C _ => Just <$> inferC ctx sg subj
||| `infer`, assuming the dimension context is consistent
export covering %inline
inferC : (ctx : TyContext q d n) -> SQty q -> Elim q d n ->
m (InferResult' q d n)
inferC ctx sg subj =
let Element subj nc = pushSubsts subj in let Element subj nc = pushSubsts subj in
infer' ctx sg subj nc infer' ctx sg subj nc
export covering private covering
check' : TyContext q d n -> SQty q -> check' : TyContext q d n -> SQty q ->
(subj : Term q d n) -> (0 nc : NotClo subj) -> Term q d n -> (subj : Term q d n) -> (0 nc : NotClo subj) -> Term q d n ->
m (CheckResult q n) m (CheckResult' q n)
check' ctx sg (TYPE k) _ ty = do check' ctx sg (TYPE k) _ ty = do
-- if 𝓀 < then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type -- if 𝓀 < then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type
@ -120,7 +137,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
-- 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 <- check (extendTy arg ctx) sg body.term res.term qout <- checkC (extendTy arg ctx) sg body.term res.term
-- then Ψ | Γ ⊢ σ · (λx ⇒ t) ⇐ (π·x : A) → B ⊳ Σ -- then Ψ | Γ ⊢ σ · (λx ⇒ t) ⇐ (π·x : A) → B ⊳ Σ
popQ qty' qout popQ qty' qout
@ -139,10 +156,10 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
check' ctx sg (Pair fst snd) _ ty = do check' ctx sg (Pair fst snd) _ ty = do
(tfst, tsnd) <- expectSig !ask ty (tfst, tsnd) <- expectSig !ask ty
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁
qfst <- check ctx sg fst tfst qfst <- checkC ctx sg fst tfst
let tsnd = sub1 tsnd (fst :# tfst) let tsnd = sub1 tsnd (fst :# tfst)
-- if Ψ | Γ ⊢ σ · t ⇐ B[s] ⊳ Σ₂ -- if Ψ | Γ ⊢ σ · t ⇐ B[s] ⊳ Σ₂
qsnd <- check 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
@ -163,7 +180,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
check' ctx sg (DLam i body) _ ty = do check' ctx sg (DLam i body) _ ty = do
(ty, l, r) <- expectEq !ask ty (ty, l, r) <- expectEq !ask ty
-- if Ψ, i | Γ ⊢ σ · t ⇐ A ⊳ Σ -- if Ψ, i | Γ ⊢ σ · t ⇐ A ⊳ Σ
qout <- check (extendDim ctx) sg body.term ty.term qout <- checkC (extendDim ctx) sg body.term ty.term
-- if Ψ | Γ ⊢ t0 = l : A0 -- if Ψ | Γ ⊢ t0 = l : A0
equal ctx ty.zero body.zero l equal ctx ty.zero body.zero l
-- if Ψ | Γ ⊢ t1 = r : A1 -- if Ψ | Γ ⊢ t1 = r : A1
@ -173,16 +190,16 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
check' ctx sg (E e) _ ty = do check' ctx sg (E e) _ ty = do
-- if Ψ | Γ ⊢ σ · e ⇒ A' ⊳ Σ -- if Ψ | Γ ⊢ σ · e ⇒ A' ⊳ Σ
infres <- infer ctx sg e infres <- inferC ctx sg e
-- if Ψ | Γ ⊢ A' <: A -- if Ψ | Γ ⊢ A' <: A
subtype ctx infres.type ty subtype ctx infres.type ty
-- then Ψ | Γ ⊢ σ · e ⇐ A ⊳ Σ -- then Ψ | Γ ⊢ σ · e ⇐ A ⊳ Σ
pure infres.qout pure infres.qout
export covering private covering
infer' : TyContext q d n -> SQty q -> infer' : TyContext q d n -> SQty q ->
(subj : Elim q d n) -> (0 nc : NotClo subj) -> (subj : Elim q d n) -> (0 nc : NotClo subj) ->
m (InferResult q d n) m (InferResult' q d n)
infer' ctx sg (F x) _ = do infer' ctx sg (F x) _ = do
-- if π·x : A {≔ s} in global context -- if π·x : A {≔ s} in global context
@ -195,14 +212,14 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
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 pure $ lookupBound sg.fst i ctx.tctx
infer' ctx sg (fun :@ arg) _ = do infer' ctx sg (fun :@ arg) _ = do
-- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁
funres <- infer ctx sg fun funres <- inferC ctx sg fun
(qty, argty, res) <- expectPi !ask funres.type (qty, argty, res) <- expectPi !ask funres.type
-- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂ -- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂
argout <- check 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 $ arg :# argty,
@ -213,14 +230,14 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
-- if 1 ≤ π -- if 1 ≤ π
expectCompatQ one pi expectCompatQ one pi
-- if Ψ | Γ ⊢ 1 · pair ⇒ (x : A) × B ⊳ Σ₁ -- if Ψ | Γ ⊢ 1 · pair ⇒ (x : A) × B ⊳ Σ₁
pairres <- infer ctx sone pair pairres <- inferC ctx sone pair
check0 (extendTy pairres.type ctx) ret.term (TYPE UAny) check0 (extendTy pairres.type ctx) ret.term (TYPE UAny)
(tfst, tsnd) <- expectSig !ask pairres.type (tfst, tsnd) <- expectSig !ask pairres.type
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐ ret[(x, y)] ⊳ Σ₂, ρ₁·x, ρ₂·y -- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐ ret[(x, y)] ⊳ Σ₂, ρ₁·x, ρ₂·y
-- with ρ₁, ρ₂ ≤ π -- with ρ₁, ρ₂ ≤ π
let bodyctx = extendTyN [< tfst, tsnd.term] ctx let bodyctx = extendTyN [< tfst, tsnd.term] ctx
bodyty = substCasePairRet pairres.type ret bodyty = substCasePairRet pairres.type ret
bodyout <- check bodyctx sg body.term bodyty >>= popQs [< pi, pi] bodyout <- checkC bodyctx sg body.term bodyty >>= popQs [< pi, pi]
-- then Ψ | Γ ⊢ σ · case ⋯ ⇒ ret[pair] ⊳ πΣ₁ + Σ₂ -- then Ψ | Γ ⊢ σ · case ⋯ ⇒ ret[pair] ⊳ πΣ₁ + Σ₂
pure $ InfRes { pure $ InfRes {
type = sub1 ret pair, type = sub1 ret pair,
@ -229,7 +246,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
infer' ctx sg (fun :% dim) _ = do infer' ctx sg (fun :% dim) _ = do
-- if Ψ | Γ ⊢ σ · f ⇒ Eq [i ⇒ A] l r ⊳ Σ -- if Ψ | Γ ⊢ σ · f ⇒ Eq [i ⇒ A] l r ⊳ Σ
InfRes {type, qout} <- infer ctx sg fun InfRes {type, qout} <- inferC ctx sg fun
(ty, _, _) <- expectEq !ask type (ty, _, _) <- expectEq !ask type
-- then Ψ | Γ ⊢ σ · f p ⇒ Ap ⊳ Σ -- then Ψ | Γ ⊢ σ · f p ⇒ Ap ⊳ Σ
pure $ InfRes {type = dsub1 ty dim, qout} pure $ InfRes {type = dsub1 ty dim, qout}
@ -238,6 +255,6 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
-- if Ψ | Γ ⊢₀ A ⇐ Type -- if Ψ | Γ ⊢₀ A ⇐ Type
check0 ctx type (TYPE UAny) check0 ctx type (TYPE UAny)
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ -- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ
qout <- check ctx sg term type qout <- checkC ctx sg term type
-- then Ψ | Γ ⊢ σ · (s ∷ A) ⇒ A ⊳ Σ -- then Ψ | Γ ⊢ σ · (s ∷ A) ⇒ A ⊳ Σ
pure $ InfRes {type, qout} pure $ InfRes {type, qout}

View file

@ -82,15 +82,23 @@ namespace QOutput
public export public export
CheckResult : Type -> Nat -> Type CheckResult' : Type -> Nat -> Type
CheckResult = QOutput CheckResult' = QOutput
public export public export
record InferResult q d n where CheckResult : DimEq d -> Type -> Nat -> Type
CheckResult eqs q n = IfConsistent eqs $ CheckResult' q n
public export
record InferResult' q d n where
constructor InfRes constructor InfRes
type : Term q d n type : Term q d n
qout : QOutput q n qout : QOutput q n
public export
InferResult : DimEq d -> TermLike
InferResult eqs q d n = IfConsistent eqs $ InferResult' q d n
public export public export
data EqMode = Equal | Sub | Super data EqMode = Equal | Sub | Super

View file

@ -384,7 +384,7 @@ tests = "equality & subtyping" :- [
testEq "a𝟎 = a" $ testEq "a𝟎 = a" $
equalED 1 empty (DCloE (F "a") (K Zero ::: id)) (F "a"), equalED 1 empty (DCloE (F "a") (K Zero ::: id)) (F "a"),
testEq "(f [a])𝟎 = f𝟎 [a]𝟎" $ testEq "(f [a])𝟎 = f𝟎 [a]𝟎" $
let th = (K Zero ::: id) in let th = K Zero ::: id in
equalED 1 empty equalED 1 empty
(DCloE (F "f" :@ FT "a") th) (DCloE (F "f" :@ FT "a") th)
(DCloE (F "f") th :@ DCloT (FT "a") th) (DCloE (F "f") th :@ DCloT (FT "a") th)

View file

@ -83,25 +83,29 @@ qoutEq qout res = unless (qout == res) $ throwError $ WrongQOut qout res
inferAs : TyContext Three d n -> (sg : SQty Three) -> inferAs : TyContext Three d n -> (sg : SQty Three) ->
Elim Three d n -> Term Three d n -> M () Elim Three d n -> Term Three d n -> M ()
inferAs ctx sg e ty = do inferAs ctx@(MkTyContext {dctx, _}) sg e ty = do
res <- inj $ infer ctx sg e case !(inj $ infer ctx sg e) of
inferredTypeEq ctx ty res.type Just res => inferredTypeEq ctx ty res.type
Nothing => pure ()
inferAsQ : TyContext Three d n -> (sg : SQty Three) -> inferAsQ : TyContext Three d n -> (sg : SQty Three) ->
Elim Three d n -> Term Three d n -> QOutput Three n -> M () Elim Three d n -> Term Three d n -> QOutput Three n -> M ()
inferAsQ ctx sg e ty qout = do inferAsQ ctx@(MkTyContext {dctx, _}) sg e ty qout = do
res <- inj $ infer ctx sg e case !(inj $ infer ctx sg e) of
Just res => do
inferredTypeEq ctx ty res.type inferredTypeEq ctx ty res.type
qoutEq qout res.qout qoutEq qout res.qout
Nothing => pure ()
infer_ : TyContext Three d n -> (sg : SQty Three) -> Elim Three d n -> M () infer_ : TyContext Three d n -> (sg : SQty Three) -> Elim Three d n -> M ()
infer_ ctx sg e = ignore $ inj $ infer ctx sg e infer_ ctx sg e = ignore $ inj $ infer ctx sg e
checkQ : TyContext Three d n -> SQty Three -> checkQ : TyContext Three d n -> SQty Three ->
Term Three d n -> Term Three d n -> QOutput Three n -> M () Term Three d n -> Term Three d n -> QOutput Three n -> M ()
checkQ ctx sg s ty qout = do checkQ ctx@(MkTyContext {dctx, _}) sg s ty qout = do
res <- inj $ check ctx sg s ty case !(inj $ check ctx sg s ty) of
qoutEq qout res Just res => qoutEq qout res
Nothing => pure ()
check_ : TyContext Three d n -> SQty Three -> check_ : TyContext Three d n -> SQty Three ->
Term Three d n -> Term Three d n -> M () Term Three d n -> Term Three d n -> M ()
@ -117,7 +121,9 @@ tests = "typechecker" :- [
testTCFail "0 · ★₁ ⇍ ★₀" $ check_ (ctx [<]) szero (TYPE 1) (TYPE 0), testTCFail "0 · ★₁ ⇍ ★₀" $ check_ (ctx [<]) szero (TYPE 1) (TYPE 0),
testTCFail "0 · ★₀ ⇍ ★₀" $ check_ (ctx [<]) szero (TYPE 0) (TYPE 0), testTCFail "0 · ★₀ ⇍ ★₀" $ check_ (ctx [<]) szero (TYPE 0) (TYPE 0),
testTCFail "0 · ★_ ⇍ ★_" $ check_ (ctx [<]) szero (TYPE UAny) (TYPE UAny), testTCFail "0 · ★_ ⇍ ★_" $ check_ (ctx [<]) szero (TYPE UAny) (TYPE UAny),
testTCFail "1 · ★₀ ⇍ ★₁" $ check_ (ctx [<]) sone (TYPE 0) (TYPE 1) testTCFail "1 · ★₀ ⇍ ★₁" $ check_ (ctx [<]) sone (TYPE 0) (TYPE 1),
testTC "0=1 ⊢ 0 · ★₁ ⇐ ★₀" $
check_ (MkTyContext (ZeroIsOne {d = 0}) [<]) szero (TYPE 1) (TYPE 0)
], ],
"function types" :- [ "function types" :- [
@ -136,7 +142,10 @@ tests = "typechecker" :- [
(Pi One "x" (FT "A") $ TUsed $ E $ F "P" :@ BVT 0) (Pi One "x" (FT "A") $ TUsed $ E $ F "P" :@ BVT 0)
(TYPE 0), (TYPE 0),
testTCFail "0 · A ⊸ P ⇍ ★₀" $ testTCFail "0 · A ⊸ P ⇍ ★₀" $
check_ (ctx [<]) szero (Arr One (FT "A") $ FT "P") (TYPE 0) check_ (ctx [<]) szero (Arr One (FT "A") $ FT "P") (TYPE 0),
testTC "0=1 ⊢ 0 · A ⊸ P ⇐ ★₀" $
check_ (MkTyContext (ZeroIsOne {d = 0}) [<]) szero
(Arr One (FT "A") $ FT "P") (TYPE 0)
], ],
"pair types" :- [ "pair types" :- [