quox/lib/Quox/Equal.idr

606 lines
22 KiB
Idris
Raw Normal View History

2022-02-26 20:18:16 -05:00
module Quox.Equal
import Quox.BoolExtra
import public Quox.Typing
import Data.Maybe
2023-03-31 13:23:30 -04:00
import Quox.EffExtra
2022-02-26 20:18:16 -05:00
2023-03-26 18:07:39 -04:00
%default total
2022-02-26 20:18:16 -05:00
public export
2023-03-31 13:23:30 -04:00
0 EqModeState : Type -> Type
EqModeState = State EqMode
2023-01-08 14:44:25 -05:00
public export
2023-04-01 13:16:43 -04:00
0 EqualEff : List (Type -> Type)
EqualEff = [ErrorEff, EqModeState]
2023-01-08 14:44:25 -05:00
public export
0 Equal : Type -> Type
Equal = Eff $ EqualEff
export
runEqual : EqMode -> Equal a -> Either Error a
runEqual mode = extract . runExcept . evalState mode
2023-01-08 14:44:25 -05:00
2023-03-31 13:23:30 -04:00
export %inline
mode : Has EqModeState fs => Eff fs EqMode
mode = get
2023-01-08 14:44:25 -05:00
2023-04-01 13:16:43 -04:00
parameters (ctx : EqContext n)
private %inline
clashT : Term 0 n -> Term 0 n -> Term 0 n -> Equal a
2023-03-31 13:23:30 -04:00
clashT ty s t = throw $ ClashT ctx !mode ty s t
2023-01-08 14:44:25 -05:00
private %inline
clashTy : Term 0 n -> Term 0 n -> Equal a
2023-03-31 13:23:30 -04:00
clashTy s t = throw $ ClashTy ctx !mode s t
private %inline
clashE : Elim 0 n -> Elim 0 n -> Equal a
2023-03-31 13:23:30 -04:00
clashE e f = throw $ ClashE ctx !mode e f
2023-01-08 14:44:25 -05:00
2023-03-26 08:38:37 -04:00
private %inline
wrongType : Term 0 n -> Term 0 n -> Equal a
2023-03-31 13:23:30 -04:00
wrongType ty s = throw $ WrongType ctx ty s
2023-03-26 08:38:37 -04:00
2023-01-08 14:44:25 -05:00
2023-02-12 15:30:08 -05:00
public export %inline
2023-04-01 13:16:43 -04:00
sameTyCon : (s, t : Term d n) ->
2023-04-03 11:46:23 -04:00
(0 ts : So (isTyConE s)) => (0 tt : So (isTyConE t)) =>
2023-02-12 15:30:08 -05:00
Bool
sameTyCon (TYPE {}) (TYPE {}) = True
sameTyCon (TYPE {}) _ = False
sameTyCon (Pi {}) (Pi {}) = True
sameTyCon (Pi {}) _ = False
sameTyCon (Sig {}) (Sig {}) = True
sameTyCon (Sig {}) _ = False
sameTyCon (Enum {}) (Enum {}) = True
sameTyCon (Enum {}) _ = False
2023-02-12 15:30:08 -05:00
sameTyCon (Eq {}) (Eq {}) = True
sameTyCon (Eq {}) _ = False
2023-03-26 08:40:54 -04:00
sameTyCon Nat Nat = True
sameTyCon Nat _ = False
2023-03-31 13:11:35 -04:00
sameTyCon (BOX {}) (BOX {}) = True
sameTyCon (BOX {}) _ = False
2023-02-12 15:30:08 -05:00
sameTyCon (E {}) (E {}) = True
sameTyCon (E {}) _ = False
2023-04-15 09:13:01 -04:00
||| true if a type is known to be a subsingleton purely by its form.
||| a subsingleton is a type with only zero or one possible values.
||| equality/subtyping accepts immediately on values of subsingleton types.
|||
||| * a function type is a subsingleton if its codomain is.
||| * a pair type is a subsingleton if both its elements are.
||| * equality types are subsingletons because of uip.
||| * an enum type is a subsingleton if it has zero or one tags.
||| * a box type is a subsingleton if its content is
public export covering
isSubSing : Has ErrorEff fs => {n : Nat} ->
2023-04-17 17:58:24 -04:00
Mods -> Definitions -> EqContext n -> Term 0 n -> Eff fs Bool
isSubSing ns defs ctx ty0 = do
Element ty0 nc <- whnf ns defs ctx ty0
2023-04-15 09:13:01 -04:00
case ty0 of
TYPE _ => pure False
2023-04-17 17:58:24 -04:00
Pi _ arg res => isSubSing ns defs (extendTy Zero res.name arg ctx) res.term
Sig fst snd => isSubSing ns defs ctx fst `andM`
isSubSing ns defs (extendTy Zero snd.name fst ctx) snd.term
2023-04-15 09:13:01 -04:00
Enum tags => pure $ length (SortedSet.toList tags) <= 1
Eq {} => pure True
Nat => pure False
2023-04-17 17:58:24 -04:00
BOX _ ty => isSubSing ns defs ctx ty
E (s :# _) => isSubSing ns defs ctx s
2023-04-15 09:13:01 -04:00
E _ => pure False
Lam _ => pure False
Pair _ _ => pure False
Tag _ => pure False
DLam _ => pure False
Zero => pure False
Succ _ => pure False
Box _ => pure False
2023-02-10 15:40:44 -05:00
2023-03-15 10:54:51 -04:00
export
2023-04-01 13:16:43 -04:00
ensureTyCon : Has ErrorEff fs =>
(ctx : EqContext n) -> (t : Term 0 n) ->
2023-04-03 11:46:23 -04:00
Eff fs (So (isTyConE t))
ensureTyCon ctx t = case nchoose $ isTyConE t of
2023-03-15 10:54:51 -04:00
Left y => pure y
2023-03-31 13:23:30 -04:00
Right n => throw $ NotType (toTyContext ctx) (t // shift0 ctx.dimLen)
2023-03-15 10:54:51 -04:00
2023-04-15 09:13:01 -04:00
||| performs the minimum work required to recompute the type of an elim.
|||
||| ⚠ **assumes the elim is already typechecked.** ⚠
private covering
2023-04-17 17:58:24 -04:00
computeElimTypeE : (ns : Mods) -> (defs : Definitions) -> EqContext n ->
(e : Elim 0 n) -> (0 ne : NotRedex ns defs e) =>
2023-04-15 09:13:01 -04:00
Equal (Term 0 n)
2023-04-17 17:58:24 -04:00
computeElimTypeE ns defs ectx e =
2023-04-15 09:13:01 -04:00
let Val n = ectx.termLen in
2023-04-17 17:58:24 -04:00
rethrow $ computeElimType ns defs (toWhnfContext ectx) e
2023-04-15 09:13:01 -04:00
2023-04-17 17:58:24 -04:00
parameters (ns : Mods) (defs : Definitions)
mutual
namespace Term
2023-02-19 11:04:57 -05:00
||| `compare0 ctx ty s t` compares `s` and `t` at type `ty`, according to
||| the current variance `mode`.
|||
||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠
2023-02-10 15:40:44 -05:00
export covering %inline
compare0 : EqContext n -> (ty, s, t : Term 0 n) -> Equal ()
2023-02-19 11:54:39 -05:00
compare0 ctx ty s t =
wrapErr (WhileComparingT ctx !mode ty s t) $ do
let Val n = ctx.termLen
2023-04-17 17:58:24 -04:00
Element ty _ <- whnf ns defs ctx ty
Element s _ <- whnf ns defs ctx s
Element t _ <- whnf ns defs ctx t
tty <- ensureTyCon ctx ty
2023-02-19 11:54:39 -05:00
compare0' ctx ty s t
2023-02-10 15:40:44 -05:00
2023-02-19 11:04:57 -05:00
||| converts an elim "Γ ⊢ e" to "Γ, x ⊢ e x", for comparing with
||| a lambda "Γ ⊢ λx ⇒ t" that has been converted to "Γ, x ⊢ t".
2023-02-10 15:40:44 -05:00
private %inline
2023-04-01 13:16:43 -04:00
toLamBody : Elim d n -> Term d (S n)
2023-04-15 09:13:01 -04:00
toLamBody e = E $ weakE 1 e :@ BVT 0
2023-02-10 15:40:44 -05:00
private covering
2023-04-01 13:16:43 -04:00
compare0' : EqContext n ->
(ty, s, t : Term 0 n) ->
2023-04-17 17:58:24 -04:00
(0 _ : NotRedex ns defs ty) => (0 _ : So (isTyConE ty)) =>
(0 _ : NotRedex ns defs s) => (0 _ : NotRedex ns defs t) =>
Equal ()
2023-02-10 15:40:44 -05:00
compare0' ctx (TYPE _) s t = compareType ctx s t
2023-03-31 13:23:30 -04:00
compare0' ctx ty@(Pi {qty, arg, res}) s t {n} = local_ Equal $
2023-02-10 15:40:44 -05:00
case (s, t) of
2023-03-15 10:54:51 -04:00
-- Γ, x : A ⊢ s = t : B
-- -------------------------------------------
-- Γ ⊢ (λ x ⇒ s) = (λ x ⇒ t) : (π·x : A) → B
2023-02-22 01:40:19 -05:00
(Lam b1, Lam b2) => compare0 ctx' res.term b1.term b2.term
2023-02-19 11:04:57 -05:00
-- Γ, x : A ⊢ s = e x : B
2023-03-15 10:54:51 -04:00
-- -----------------------------------
-- Γ ⊢ (λ x ⇒ s) = e : (π·x : A) → B
2023-02-22 01:40:19 -05:00
(E e, Lam b) => eta e b
(Lam b, E e) => eta e b
2023-02-19 11:04:57 -05:00
(E e, E f) => Elim.compare0 ctx e f
2023-03-15 10:54:51 -04:00
2023-03-26 08:38:37 -04:00
(Lam _, t) => wrongType ctx ty t
(E _, t) => wrongType ctx ty t
(s, _) => wrongType ctx ty s
2023-02-19 11:04:57 -05:00
where
2023-04-01 13:16:43 -04:00
ctx' : EqContext (S n)
2023-03-15 10:54:51 -04:00
ctx' = extendTy qty res.name arg ctx
2023-02-19 11:04:57 -05:00
eta : Elim 0 n -> ScopeTerm 0 n -> Equal ()
2023-02-22 01:40:19 -05:00
eta e (S _ (Y b)) = compare0 ctx' res.term (toLamBody e) b
eta e (S _ (N _)) = clashT ctx ty s t
2023-02-10 15:40:44 -05:00
2023-03-31 13:23:30 -04:00
compare0' ctx ty@(Sig {fst, snd, _}) s t = local_ Equal $
2023-02-10 15:40:44 -05:00
case (s, t) of
2023-02-19 11:04:57 -05:00
-- Γ ⊢ s₁ = t₁ : A Γ ⊢ s₂ = t₂ : B{s₁/x}
2023-03-15 10:54:51 -04:00
-- --------------------------------------------
-- Γ ⊢ (s₁, t₁) = (s₂,t₂) : (x : A) × B
2023-02-19 11:04:57 -05:00
--
-- [todo] η for π ≥ 0 maybe
2023-02-10 15:40:44 -05:00
(Pair sFst sSnd, Pair tFst tSnd) => do
compare0 ctx fst sFst tFst
compare0 ctx (sub1 snd (sFst :# fst)) sSnd tSnd
2023-02-19 11:04:57 -05:00
2023-03-15 10:54:51 -04:00
(E e, E f) => Elim.compare0 ctx e f
2023-03-26 08:38:37 -04:00
(Pair {}, E _) => clashT ctx ty s t
(E _, Pair {}) => clashT ctx ty s t
(Pair {}, t) => wrongType ctx ty t
(E _, t) => wrongType ctx ty t
(s, _) => wrongType ctx ty s
2023-02-10 15:40:44 -05:00
2023-03-31 13:23:30 -04:00
compare0' ctx ty@(Enum tags) s t = local_ Equal $
case (s, t) of
-- --------------------
-- Γ ⊢ `t = `t : {ts}
--
-- t ∈ ts is in the typechecker, not here, ofc
(Tag t1, Tag t2) => unless (t1 == t2) $ clashT ctx ty s t
2023-03-15 10:54:51 -04:00
(E e, E f) => Elim.compare0 ctx e f
2023-03-26 08:38:37 -04:00
(Tag _, E _) => clashT ctx ty s t
(E _, Tag _) => clashT ctx ty s t
(Tag _, t) => wrongType ctx ty t
(E _, t) => wrongType ctx ty t
(s, _) => wrongType ctx ty s
2023-02-19 11:04:57 -05:00
compare0' _ (Eq {}) _ _ =
-- ✨ uip ✨
--
2023-04-15 09:13:01 -04:00
-- ----------------------------
-- Γ ⊢ e = f : Eq [i ⇒ A] s t
2023-02-19 11:04:57 -05:00
pure ()
2023-02-10 15:40:44 -05:00
2023-03-31 13:23:30 -04:00
compare0' ctx Nat s t = local_ Equal $
2023-03-26 08:40:54 -04:00
case (s, t) of
-- ---------------
-- Γ ⊢ 0 = 0 :
(Zero, Zero) => pure ()
-- Γ ⊢ m = n :
-- -------------------------
-- Γ ⊢ succ m = succ n :
(Succ m, Succ n) => compare0 ctx Nat m n
(E e, E f) => Elim.compare0 ctx e f
(Zero, Succ _) => clashT ctx Nat s t
(Zero, E _) => clashT ctx Nat s t
(Succ _, Zero) => clashT ctx Nat s t
(Succ _, E _) => clashT ctx Nat s t
(E _, Zero) => clashT ctx Nat s t
(E _, Succ _) => clashT ctx Nat s t
(Zero, t) => wrongType ctx Nat t
(Succ _, t) => wrongType ctx Nat t
(E _, t) => wrongType ctx Nat t
(s, _) => wrongType ctx Nat s
2023-03-31 13:23:30 -04:00
compare0' ctx ty@(BOX q ty') s t = local_ Equal $
2023-03-31 13:11:35 -04:00
case (s, t) of
-- Γ ⊢ s = t : A
-- -----------------------
-- Γ ⊢ [s] = [t] : [π.A]
(Box s, Box t) => compare0 ctx ty' s t
(E e, E f) => Elim.compare0 ctx e f
(Box _, t) => wrongType ctx ty t
(E _, t) => wrongType ctx ty t
(s, _) => wrongType ctx ty s
2023-02-10 15:40:44 -05:00
compare0' ctx ty@(E _) s t = do
-- a neutral type can only be inhabited by neutral values
-- e.g. an abstract value in an abstract type, bound variables, …
2023-03-26 08:38:37 -04:00
E e <- pure s | _ => wrongType ctx ty s
E f <- pure t | _ => wrongType ctx ty t
2023-02-12 15:30:08 -05:00
Elim.compare0 ctx e f
2023-02-10 15:40:44 -05:00
2023-02-19 11:04:57 -05:00
||| compares two types, using the current variance `mode` for universes.
||| fails if they are not types, even if they would happen to be equal.
2023-02-19 11:43:14 -05:00
export covering %inline
compareType : EqContext n -> (s, t : Term 0 n) -> Equal ()
2023-02-10 15:40:44 -05:00
compareType ctx s t = do
let Val n = ctx.termLen
2023-04-17 17:58:24 -04:00
Element s _ <- whnf ns defs ctx s
Element t _ <- whnf ns defs ctx t
ts <- ensureTyCon ctx s
tt <- ensureTyCon ctx t
2023-04-17 17:58:24 -04:00
st <- either pure (const $ clashTy ctx s t) $ nchoose $ sameTyCon s t
2023-02-10 15:40:44 -05:00
compareType' ctx s t
private covering
2023-04-01 13:16:43 -04:00
compareType' : EqContext n -> (s, t : Term 0 n) ->
2023-04-17 17:58:24 -04:00
(0 _ : NotRedex ns defs s) => (0 _ : So (isTyConE s)) =>
(0 _ : NotRedex ns defs t) => (0 _ : So (isTyConE t)) =>
(0 _ : So (sameTyCon s t)) =>
Equal ()
2023-02-19 11:04:57 -05:00
-- equality is the same as subtyping, except with the
-- "≤" in the TYPE rule being replaced with "="
2023-02-12 15:30:08 -05:00
compareType' ctx (TYPE k) (TYPE l) =
2023-02-19 11:04:57 -05:00
-- 𝓀
-- ----------------------
-- Γ ⊢ Type 𝓀 <: Type
2023-02-12 15:30:08 -05:00
expectModeU !mode k l
compareType' ctx (Pi {qty = sQty, arg = sArg, res = sRes, _})
(Pi {qty = tQty, arg = tArg, res = tRes, _}) = do
2023-02-19 11:04:57 -05:00
-- Γ ⊢ A₁ :> A₂ Γ, x : A₁ ⊢ B₁ <: B₂
-- ----------------------------------------
-- Γ ⊢ (π·x : A₁) → B₁ <: (π·x : A₂) → B₂
2023-02-12 15:30:08 -05:00
expectEqualQ sQty tQty
2023-03-31 13:23:30 -04:00
local flip $ compareType ctx sArg tArg -- contra
2023-04-01 13:16:43 -04:00
compareType (extendTy Zero sRes.name sArg ctx) sRes.term tRes.term
2023-02-12 15:30:08 -05:00
compareType' ctx (Sig {fst = sFst, snd = sSnd, _})
(Sig {fst = tFst, snd = tSnd, _}) = do
2023-02-19 11:04:57 -05:00
-- Γ ⊢ A₁ <: A₂ Γ, x : A₁ ⊢ B₁ <: B₂
-- --------------------------------------
-- Γ ⊢ (x : A₁) × B₁ <: (x : A₂) × B₂
2023-02-12 15:30:08 -05:00
compareType ctx sFst tFst
2023-04-01 13:16:43 -04:00
compareType (extendTy Zero sSnd.name sFst ctx) sSnd.term tSnd.term
2023-02-12 15:30:08 -05:00
compareType' ctx (Eq {ty = sTy, l = sl, r = sr, _})
(Eq {ty = tTy, l = tl, r = tr, _}) = do
2023-02-19 11:04:57 -05:00
-- Γ ⊢ A₁ε/i <: A₂ε/i
-- Γ ⊢ l₁ = l₂ : A₁𝟎/i Γ ⊢ r₁ = r₂ : A₁𝟏/i
-- ------------------------------------------------
-- Γ ⊢ Eq [i ⇒ A₁] l₁ r₂ <: Eq [i ⇒ A₂] l₂ r₂
compareType (extendDim sTy.name Zero ctx) sTy.zero tTy.zero
compareType (extendDim sTy.name One ctx) sTy.one tTy.one
2023-03-31 13:23:30 -04:00
local_ Equal $ do
2023-02-12 15:30:08 -05:00
Term.compare0 ctx sTy.zero sl tl
Term.compare0 ctx sTy.one sr tr
compareType' ctx s@(Enum tags1) t@(Enum tags2) = do
-- ------------------
-- Γ ⊢ {ts} <: {ts}
--
-- no subtyping based on tag subsets, since that would need
-- a runtime coercion
unless (tags1 == tags2) $ clashTy ctx s t
2023-03-26 08:40:54 -04:00
compareType' ctx Nat Nat =
-- ------------
-- Γ ⊢ <:
pure ()
2023-03-31 13:11:35 -04:00
compareType' ctx (BOX pi a) (BOX rh b) = do
expectEqualQ pi rh
compareType ctx a b
2023-02-12 15:30:08 -05:00
compareType' ctx (E e) (E f) = do
-- no fanciness needed here cos anything other than a neutral
-- has been inlined by whnf
2023-02-12 15:30:08 -05:00
Elim.compare0 ctx e f
2023-02-11 12:15:50 -05:00
private covering
2023-04-01 13:16:43 -04:00
replaceEnd : EqContext n ->
2023-04-17 17:58:24 -04:00
(e : Elim 0 n) -> DimConst -> (0 ne : NotRedex ns defs e) ->
Equal (Elim 0 n)
2023-02-11 12:15:50 -05:00
replaceEnd ctx e p ne = do
2023-04-17 17:58:24 -04:00
(ty, l, r) <- expectEq ns defs ctx !(computeElimTypeE ns defs ctx e)
2023-02-11 12:15:50 -05:00
pure $ ends l r p :# dsub1 ty (K p)
namespace Elim
2023-02-11 12:15:50 -05:00
-- [fixme] the following code ends up repeating a lot of work in the
-- computeElimType calls. the results should be shared better
2023-02-19 11:04:57 -05:00
||| compare two eliminations according to the given variance `mode`.
|||
||| ⚠ **assumes that they have both been typechecked, and have
||| equal types.** ⚠
2023-02-10 15:40:44 -05:00
export covering %inline
compare0 : EqContext n -> (e, f : Elim 0 n) -> Equal ()
2023-02-10 15:40:44 -05:00
compare0 ctx e f =
wrapErr (WhileComparingE ctx !mode e f) $ do
let Val n = ctx.termLen
2023-04-17 17:58:24 -04:00
Element e ne <- whnf ns defs ctx e
Element f nf <- whnf ns defs ctx f
2023-02-19 11:54:39 -05:00
-- [fixme] there is a better way to do this "isSubSing" stuff for sure
2023-04-17 17:58:24 -04:00
unless !(isSubSing ns defs ctx !(computeElimTypeE ns defs ctx e)) $
2023-02-19 11:54:39 -05:00
compare0' ctx e f ne nf
2023-02-10 15:40:44 -05:00
private covering
2023-04-01 13:16:43 -04:00
compare0' : EqContext n ->
(e, f : Elim 0 n) ->
2023-04-17 17:58:24 -04:00
(0 ne : NotRedex ns defs e) -> (0 nf : NotRedex ns defs f) ->
Equal ()
2023-02-11 12:15:50 -05:00
-- replace applied equalities with the appropriate end first
2023-02-12 15:30:08 -05:00
-- e.g. e : Eq [i ⇒ A] s t ⊢ e 𝟎 = s : A𝟎/i
2023-02-19 11:04:57 -05:00
--
-- [todo] maybe have typed whnf and do this (and η???) there instead
2023-02-12 15:30:08 -05:00
compare0' ctx (e :% K p) f ne nf =
2023-02-11 12:15:50 -05:00
compare0 ctx !(replaceEnd ctx e p $ noOr1 ne) f
2023-02-12 15:30:08 -05:00
compare0' ctx e (f :% K q) ne nf =
2023-02-11 12:15:50 -05:00
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 _) f _ _ = clashE ctx e f
2023-02-10 15:40:44 -05:00
compare0' ctx e@(B i) f@(B j) _ _ = unless (i == j) $ clashE ctx e f
compare0' ctx e@(B _) f _ _ = clashE ctx e f
2023-02-10 15:40:44 -05:00
2023-02-12 15:30:08 -05:00
compare0' ctx (e :@ s) (f :@ t) ne nf =
2023-03-31 13:23:30 -04:00
local_ Equal $ do
2023-02-12 15:30:08 -05:00
compare0 ctx e f
2023-04-17 17:58:24 -04:00
(_, arg, _) <- expectPi ns defs ctx =<<
computeElimTypeE ns defs ctx e @{noOr1 ne}
2023-02-12 15:30:08 -05:00
Term.compare0 ctx arg s t
compare0' ctx e@(_ :@ _) f _ _ = clashE ctx e f
2023-02-10 15:40:44 -05:00
2023-02-22 01:40:19 -05:00
compare0' ctx (CasePair epi e eret ebody)
(CasePair fpi f fret fbody) ne nf =
2023-03-31 13:23:30 -04:00
local_ Equal $ do
2023-02-11 12:15:50 -05:00
compare0 ctx e f
2023-04-17 17:58:24 -04:00
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne}
2023-04-01 13:16:43 -04:00
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
2023-04-17 17:58:24 -04:00
(fst, snd) <- expectSig ns defs ctx ety
let [< x, y] = ebody.names
2023-03-15 10:54:51 -04:00
Term.compare0 (extendTyN [< (epi, x, fst), (epi, y, snd.term)] ctx)
(substCasePairRet ety eret)
ebody.term fbody.term
expectEqualQ epi fpi
compare0' ctx e@(CasePair {}) f _ _ = clashE ctx e f
2023-02-10 15:40:44 -05:00
compare0' ctx (CaseEnum epi e eret earms)
(CaseEnum fpi f fret farms) ne nf =
2023-03-31 13:23:30 -04:00
local_ Equal $ do
compare0 ctx e f
2023-04-17 17:58:24 -04:00
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne}
2023-04-01 13:16:43 -04:00
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
2023-04-17 17:58:24 -04:00
for_ !(expectEnum ns defs ctx ety) $ \t =>
compare0 ctx (sub1 eret $ Tag t :# ety)
!(lookupArm t earms) !(lookupArm t farms)
expectEqualQ epi fpi
where
lookupArm : TagVal -> CaseEnumArms d n -> Equal (Term d n)
lookupArm t arms = case lookup t arms of
Just arm => pure arm
2023-03-31 13:23:30 -04:00
Nothing => throw $ TagNotIn t (fromList $ keys arms)
compare0' ctx e@(CaseEnum {}) f _ _ = clashE ctx e f
2023-03-26 08:40:54 -04:00
compare0' ctx (CaseNat epi epi' e eret ezer esuc)
(CaseNat fpi fpi' f fret fzer fsuc) ne nf =
2023-03-31 13:23:30 -04:00
local_ Equal $ do
2023-03-26 08:40:54 -04:00
compare0 ctx e f
2023-04-17 17:58:24 -04:00
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne}
2023-04-01 13:16:43 -04:00
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
2023-03-26 08:40:54 -04:00
compare0 ctx (sub1 eret (Zero :# Nat)) ezer fzer
let [< p, ih] = esuc.names
compare0 (extendTyN [< (epi, p, Nat), (epi', ih, eret.term)] ctx)
2023-03-31 13:26:55 -04:00
(substCaseSuccRet eret)
2023-03-26 08:40:54 -04:00
esuc.term fsuc.term
expectEqualQ epi fpi
expectEqualQ epi' fpi'
compare0' ctx e@(CaseNat {}) f _ _ = clashE ctx e f
2023-03-31 13:11:35 -04:00
compare0' ctx (CaseBox epi e eret ebody)
(CaseBox fpi f fret fbody) ne nf =
2023-03-31 13:23:30 -04:00
local_ Equal $ do
2023-03-31 13:11:35 -04:00
compare0 ctx e f
2023-04-17 17:58:24 -04:00
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne}
2023-04-01 13:16:43 -04:00
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
2023-04-17 17:58:24 -04:00
(q, ty) <- expectBOX ns defs ctx ety
2023-03-31 13:11:35 -04:00
compare0 (extendTy (epi * q) ebody.name ty ctx)
(substCaseBoxRet ety eret)
ebody.term fbody.term
expectEqualQ epi fpi
compare0' ctx e@(CaseBox {}) f _ _ = clashE ctx e f
2023-02-22 01:40:19 -05:00
compare0' ctx (s :# a) (t :# b) _ _ =
2023-04-15 09:13:01 -04:00
let ty = case !mode of Super => a; _ => b in
Term.compare0 ctx ty s t
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 q1) (dsub1 ty2 q2)
compare0 ctx val1 val2
compare0' ctx e@(Coe {}) f _ _ = clashE ctx e f
2023-02-22 01:40:19 -05:00
2023-04-15 09:13:01 -04:00
compare0' ctx (Comp _ _ _ _ (K _) _ _) _ ne _ = void $ absurd $ noOr2 ne
compare0' ctx (Comp _ _ _ _ (B i) _ _) _ _ _ = absurd i
compare0' ctx _ (Comp _ _ _ _ (K _) _ _) _ nf = void $ absurd $ noOr2 nf
compare0' ctx (TypeCase ty1 ret1 arms1 def1)
(TypeCase ty2 ret2 arms2 def2)
2023-04-03 11:46:23 -04:00
ne _ =
local_ Equal $ do
compare0 ctx ty1 ty2
2023-04-17 17:58:24 -04:00
u <- expectTYPE ns defs ctx =<<
computeElimTypeE ns defs ctx ty1 @{noOr1 ne}
2023-04-03 11:46:23 -04:00
compareType ctx ret1 ret2
2023-04-15 09:13:01 -04:00
compareType ctx def1 def2
for_ allKinds $ \k =>
compareArm ctx k ret1 u
(lookupPrecise k arms1) (lookupPrecise k arms2) def1
2023-04-03 11:46:23 -04:00
compare0' ctx e@(TypeCase {}) f _ _ = clashE ctx e f
2023-02-12 15:30:08 -05:00
compare0' ctx (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@(_ :# _) f _ _ = clashE ctx e f
2023-02-10 15:40:44 -05:00
2023-04-15 09:13:01 -04:00
||| 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
||| universe the head is in.
private covering
compareArm : EqContext n -> (k : TyConKind) ->
(ret : Term 0 n) -> (u : Universe) ->
(b1, b2 : Maybe (TypeCaseArmBody k 0 n)) ->
(def : Term 0 n) ->
Equal ()
compareArm {b1 = Nothing, b2 = Nothing, _} = pure ()
compareArm ctx k ret u b1 b2 def =
let def = SN def in
compareArm_ ctx k ret u (fromMaybe def b1) (fromMaybe def b2)
private covering
compareArm_ : EqContext n -> (k : TyConKind) ->
(ret : Term 0 n) -> (u : Universe) ->
(b1, b2 : TypeCaseArmBody k 0 n) ->
Equal ()
compareArm_ ctx KTYPE ret u b1 b2 =
compare0 ctx ret b1.term b2.term
compareArm_ ctx KPi ret u b1 b2 = do
let [< a, b] = b1.names
ctx = extendTyN
[< (Zero, a, TYPE u),
(Zero, b, Arr Zero (BVT 0) (TYPE u))] ctx
compare0 ctx (weakT 2 ret) b1.term b2.term
compareArm_ ctx KSig ret u b1 b2 = do
let [< a, b] = b1.names
ctx = extendTyN
[< (Zero, a, TYPE u),
(Zero, b, Arr Zero (BVT 0) (TYPE u))] ctx
compare0 ctx (weakT 2 ret) b1.term b2.term
compareArm_ ctx KEnum ret u b1 b2 =
compare0 ctx ret b1.term b2.term
compareArm_ ctx KEq ret u b1 b2 = do
let [< a0, a1, a, l, r] = b1.names
ctx = extendTyN
[< (Zero, a0, TYPE u),
(Zero, a1, TYPE u),
(Zero, a, Eq0 (TYPE u) (BVT 1) (BVT 0)),
(Zero, l, BVT 2),
(Zero, r, BVT 2)] ctx
compare0 ctx (weakT 5 ret) b1.term b2.term
compareArm_ ctx KNat ret u b1 b2 =
compare0 ctx ret b1.term b2.term
compareArm_ ctx KBOX ret u b1 b2 = do
let ctx = extendTy Zero b1.name (TYPE u) ctx
compare0 ctx (weakT 1 ret) b1.term b1.term
2023-02-10 15:40:44 -05:00
2023-04-17 17:58:24 -04:00
parameters {auto _ : (Has DefsReader fs, Has CurNSReader fs, Has ErrorEff fs)}
(ctx : TyContext d n)
-- [todo] only split on the dvars that are actually used anywhere in
-- the calls to `splits`
2023-02-10 15:40:44 -05:00
parameters (mode : EqMode)
namespace Term
2023-02-10 15:40:44 -05:00
export covering
2023-04-01 13:16:43 -04:00
compare : (ty, s, t : Term d n) -> Eff fs ()
2023-04-17 17:58:24 -04:00
compare ty s t =
2023-03-31 13:23:30 -04:00
map fst $ runState @{Z} mode $
for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in
2023-04-17 17:58:24 -04:00
lift $ compare0 !ns !defs ectx (ty // th) (s // th) (t // th)
2023-02-10 15:40:44 -05:00
export covering
2023-04-01 13:16:43 -04:00
compareType : (s, t : Term d n) -> Eff fs ()
2023-04-17 17:58:24 -04:00
compareType s t =
2023-03-31 13:23:30 -04:00
map fst $ runState @{Z} mode $
for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in
2023-04-17 17:58:24 -04:00
lift $ compareType !ns !defs ectx (s // th) (t // th)
namespace Elim
2023-02-11 12:15:50 -05:00
||| you don't have to pass the type in but the arguments must still be
||| of the same type!!
export covering %inline
2023-04-01 13:16:43 -04:00
compare : (e, f : Elim d n) -> Eff fs ()
2023-04-17 17:58:24 -04:00
compare e f =
2023-03-31 13:23:30 -04:00
map fst $ runState @{Z} mode $
for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in
2023-04-17 17:58:24 -04:00
lift $ compare0 !ns !defs ectx (e // th) (f // th)
namespace Term
export covering %inline
2023-04-01 13:16:43 -04:00
equal, sub, super : (ty, s, t : Term d n) -> Eff fs ()
2023-02-10 15:40:44 -05:00
equal = compare Equal
sub = compare Sub
2023-02-14 16:29:06 -05:00
super = compare Super
export covering %inline
2023-04-01 13:16:43 -04:00
equalType, subtype, supertype : (s, t : Term d n) -> Eff fs ()
2023-02-10 15:40:44 -05:00
equalType = compareType Equal
subtype = compareType Sub
2023-02-14 16:29:06 -05:00
supertype = compareType Super
namespace Elim
export covering %inline
2023-04-01 13:16:43 -04:00
equal, sub, super : (e, f : Elim d n) -> Eff fs ()
2023-02-10 15:40:44 -05:00
equal = compare Equal
sub = compare Sub
2023-02-14 16:29:06 -05:00
super = compare Super