add W to internal language
This commit is contained in:
parent
00d92d3f25
commit
1da902d13a
14 changed files with 544 additions and 94 deletions
|
@ -79,6 +79,8 @@ sameTyCon (Pi {}) (Pi {}) = True
|
|||
sameTyCon (Pi {}) _ = False
|
||||
sameTyCon (Sig {}) (Sig {}) = True
|
||||
sameTyCon (Sig {}) _ = False
|
||||
sameTyCon (W {}) (W {}) = True
|
||||
sameTyCon (W {}) _ = False
|
||||
sameTyCon (Enum {}) (Enum {}) = True
|
||||
sameTyCon (Enum {}) _ = False
|
||||
sameTyCon (Eq {}) (Eq {}) = True
|
||||
|
@ -111,6 +113,7 @@ isSubSing defs ctx ty0 = do
|
|||
Sig {fst, snd, _} =>
|
||||
isSubSing defs ctx fst `andM`
|
||||
isSubSing defs (extendTy Zero snd.name fst ctx) snd.term
|
||||
W {} => pure False
|
||||
Enum {cases, _} =>
|
||||
pure $ length (SortedSet.toList cases) <= 1
|
||||
Eq {} => pure True
|
||||
|
@ -120,6 +123,7 @@ isSubSing defs ctx ty0 = do
|
|||
E _ => pure False
|
||||
Lam {} => pure False
|
||||
Pair {} => pure False
|
||||
Sup {} => pure False
|
||||
Tag {} => pure False
|
||||
DLam {} => pure False
|
||||
Zero {} => pure False
|
||||
|
@ -209,9 +213,9 @@ parameters (defs : Definitions)
|
|||
case (s, t) of
|
||||
-- Γ ⊢ s₁ = t₁ : A Γ ⊢ s₂ = t₂ : B{s₁/x}
|
||||
-- --------------------------------------------
|
||||
-- Γ ⊢ (s₁, t₁) = (s₂,t₂) : (x : A) × B
|
||||
-- Γ ⊢ (s₁, t₁) = (s₂, t₂) : (x : A) × B
|
||||
--
|
||||
-- [todo] η for π ≥ 0 maybe
|
||||
-- [todo] η for π ≰ 1 maybe
|
||||
(Pair sFst sSnd {}, Pair tFst tSnd {}) => do
|
||||
compare0 ctx fst sFst tFst
|
||||
compare0 ctx (sub1 snd (Ann sFst fst fst.loc)) sSnd tSnd
|
||||
|
@ -225,6 +229,27 @@ parameters (defs : Definitions)
|
|||
(E _, t) => wrongType t.loc ctx ty t
|
||||
(s, _) => wrongType s.loc ctx ty s
|
||||
|
||||
compare0' ctx ty@(W {shape, body, _}) s t =
|
||||
case (s, t) of
|
||||
-- Γ ⊢ s₁ = t₁ : A
|
||||
-- Γ ⊢ s₂ = t₂ : 1.B[s₁∷A/x] → (x : A) ⊲ B
|
||||
-- -----------------------------------------
|
||||
-- Γ ⊢ s₁⋄s₂ = t₁⋄t₂ : (x : A) ⊲ B
|
||||
(Sup sRoot sSub {}, Sup tRoot tSub {}) => do
|
||||
compare0 ctx shape sRoot tRoot
|
||||
let arg = sub1 body (Ann sRoot shape sRoot.loc)
|
||||
subTy = Arr One arg ty ty.loc
|
||||
compare0 ctx subTy sSub tSub
|
||||
|
||||
(E e, E f) => Elim.compare0 ctx e f
|
||||
|
||||
(Sup {}, E {}) => clashT s.loc ctx ty s t
|
||||
(E {}, Sup {}) => clashT s.loc ctx ty s t
|
||||
|
||||
(Sup {}, t) => wrongType t.loc ctx ty t
|
||||
(E {}, t) => wrongType t.loc ctx ty t
|
||||
(s, _) => wrongType s.loc ctx ty s
|
||||
|
||||
compare0' ctx ty@(Enum {}) s t = local_ Equal $
|
||||
case (s, t) of
|
||||
-- --------------------
|
||||
|
@ -322,12 +347,16 @@ parameters (defs : Definitions)
|
|||
-- Γ ⊢ Type 𝓀 <: Type ℓ
|
||||
expectModeU a.loc !mode k l
|
||||
|
||||
compareType' ctx a@(Pi {qty = sQty, arg = sArg, res = sRes, _})
|
||||
(Pi {qty = tQty, arg = tArg, res = tRes, _}) = do
|
||||
compareType' ctx (Pi {qty = sQty, arg = sArg, res = sRes, loc})
|
||||
(Pi {qty = tQty, arg = tArg, res = tRes, _}) = do
|
||||
-- Γ ⊢ A₁ :> A₂ Γ, x : A₁ ⊢ B₁ <: B₂
|
||||
-- ----------------------------------------
|
||||
-- Γ ⊢ (π·x : A₁) → B₁ <: (π·x : A₂) → B₂
|
||||
expectEqualQ a.loc sQty tQty
|
||||
-- Γ ⊢ π.(x : A₁) → B₁ <: π.(x : A₂) → B₂
|
||||
--
|
||||
-- no quantity subtyping since that would need a runtime coercion
|
||||
-- e.g. if ⌊0.A → B⌋ ⇝ ⌊B⌋, then the promotion to ω.A → B would need
|
||||
-- to re-add the abstraction
|
||||
expectEqualQ loc sQty tQty
|
||||
local flip $ compareType ctx sArg tArg -- contra
|
||||
compareType (extendTy Zero sRes.name sArg ctx) sRes.term tRes.term
|
||||
|
||||
|
@ -339,12 +368,20 @@ parameters (defs : Definitions)
|
|||
compareType ctx sFst tFst
|
||||
compareType (extendTy Zero sSnd.name sFst ctx) sSnd.term tSnd.term
|
||||
|
||||
compareType' ctx (W {shape = sShape, body = sBody, loc})
|
||||
(W {shape = tShape, body = tBody, _}) = do
|
||||
-- Γ ⊢ A₁ <: A₂ Γ, x : A₁ ⊢ B₁ <: B₂
|
||||
-- --------------------------------------
|
||||
-- Γ ⊢ (x : A₁) ⊲ B₁ <: (x : A₂) ⊲ B₂
|
||||
compareType ctx sShape tShape
|
||||
compareType (extendTy Zero sBody.name sShape ctx) sBody.term tBody.term
|
||||
|
||||
compareType' ctx (Eq {ty = sTy, l = sl, r = sr, _})
|
||||
(Eq {ty = tTy, l = tl, r = tr, _}) = do
|
||||
-- Γ ⊢ A₁‹ε/i› <: A₂‹ε/i›
|
||||
-- Γ ⊢ l₁ = l₂ : A₁‹𝟎/i› Γ ⊢ r₁ = r₂ : A₁‹𝟏/i›
|
||||
-- Γ ⊢ l₁ = l₂ : A₁‹0/i› Γ ⊢ r₁ = r₂ : A₁‹1/i›
|
||||
-- ------------------------------------------------
|
||||
-- Γ ⊢ Eq [i ⇒ A₁] l₁ r₂ <: Eq [i ⇒ A₂] l₂ r₂
|
||||
-- Γ ⊢ 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
|
||||
let ty = case !mode of Super => sTy; _ => tTy
|
||||
|
@ -366,6 +403,9 @@ parameters (defs : Definitions)
|
|||
pure ()
|
||||
|
||||
compareType' ctx (BOX pi a loc) (BOX rh b {}) = do
|
||||
-- Γ ⊢ A <: B
|
||||
-- --------------------
|
||||
-- Γ ⊢ [π.A] <: [π.B]
|
||||
expectEqualQ loc pi rh
|
||||
compareType ctx a b
|
||||
|
||||
|
@ -439,6 +479,36 @@ parameters (defs : Definitions)
|
|||
expectEqualQ e.loc epi fpi
|
||||
compare0' ctx e@(CasePair {}) f _ _ = clashE e.loc ctx e f
|
||||
|
||||
-- Ψ | Γ ⊢ e = f ⇒ (x : A) ⊲ B
|
||||
-- Ψ | Γ, p : (x : A) ⊲ B ⊢ Q = R ⇐ Type
|
||||
-- Ψ | Γ, x : A, y : 1.B → (x : A) ⊲ B, ih : 1.(z : B) → Q[y z/p]
|
||||
-- ⊢ s = t ⇐ Q[(x⋄y ∷ (x : A) ⊲ B)/p]
|
||||
-- ----------------------------------------------------------------
|
||||
-- Ψ | Γ ⊢ caseπ e return Q of { x ⋄ y, ς.ih ⇒ s }
|
||||
-- = caseπ f return R of { x ⋄ y, ς.ih ⇒ t } ⇒ Q[e/p]
|
||||
compare0' ctx (CaseW epi epi' e eret ebody eloc)
|
||||
(CaseW fpi fpi' f fret fbody floc) ne nf =
|
||||
local_ Equal $ do
|
||||
compare0 ctx e f
|
||||
ety <- computeElimTypeE defs ctx e @{noOr1 ne}
|
||||
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
|
||||
(shape, tbody) <- expectW defs ctx eloc ety
|
||||
let [< x, y, ih] = ebody.names
|
||||
z <- mnb "z" ih.loc
|
||||
let xbind = (epi, x, shape)
|
||||
ybind = (epi, y, Arr One tbody.term (weakT 1 ety) y.loc)
|
||||
ihbind = (epi', ih,
|
||||
PiY One z
|
||||
(sub1 (weakS 2 tbody) (BV 1 x.loc))
|
||||
(sub1 (weakS 3 eret) (App (BV 1 y.loc) (BVT 0 z.loc) ih.loc))
|
||||
ih.loc)
|
||||
ctx' = extendTyN [< xbind, ybind, ihbind] ctx
|
||||
Term.compare0 ctx' (substCaseWRet ebody.names ety eret)
|
||||
ebody.term fbody.term
|
||||
expectEqualQ e.loc epi fpi
|
||||
expectEqualQ e.loc epi' fpi'
|
||||
compare0' ctx e@(CaseW {}) f _ _ = clashE e.loc ctx e f
|
||||
|
||||
-- Ψ | Γ ⊢ e = f ⇒ {𝐚s}
|
||||
-- Ψ | Γ, x : {𝐚s} ⊢ Q = R
|
||||
-- Ψ | Γ ⊢ sᵢ = tᵢ ⇐ Q[𝐚ᵢ∷{𝐚s}]
|
||||
|
@ -598,6 +668,13 @@ parameters (defs : Definitions)
|
|||
(Zero, b, Arr Zero (BVT 0 b.loc) (TYPE u b.loc) b.loc)] ctx
|
||||
compare0 ctx (weakT 2 ret) b1.term b2.term
|
||||
|
||||
compareArm_ ctx KW ret u b1 b2 = do
|
||||
let [< a, b] = b1.names
|
||||
ctx = extendTyN
|
||||
[< (Zero, a, TYPE u a.loc),
|
||||
(Zero, b, Arr Zero (BVT 0 b.loc) (TYPE u b.loc) b.loc)] ctx
|
||||
compare0 ctx (weakT 2 ret) b1.term b2.term
|
||||
|
||||
compareArm_ ctx KEnum ret u b1 b2 =
|
||||
compare0 ctx ret b1.term b2.term
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue