Compare commits

...

15 Commits

31 changed files with 1622 additions and 410 deletions

View File

@ -1,6 +1,7 @@
load "misc.quox"
load "bool.quox"
load "either.quox"
load "maybe.quox"
load "nat.quox"
load "pair.quox"
load "list.quox"

View File

@ -10,7 +10,6 @@ def boolω : 1.Bool → [ω.Bool] =
def if : 0.(A : ★) → 1.Bool → ω.A → ω.A → A =
λ A b t f ⇒ case1 b return A of { 'true ⇒ t; 'false ⇒ f };
-- [todo]: universe lifting
def0 If : 1.Bool → 0.★ → 0.★ → ★ =
λ b T F ⇒ case1 b return ★ of { 'true ⇒ T; 'false ⇒ F };

69
examples/maybe.quox Normal file
View File

@ -0,0 +1,69 @@
load "misc.quox"
load "either.quox"
namespace maybe {
def0 Tag : ★ = {nothing, just}
def0 Payload : ω.Tag → ω.★ → ★ =
λ tag A ⇒ caseω tag return ★ of { 'nothing ⇒ True; 'just ⇒ A }
def0 Maybe : ω.★ → ★ =
λ A ⇒ (t : Tag) × Payload t A
def tag : 0.(A : ★) → ω.(Maybe A) → Tag =
λ _ x ⇒ caseω x return Tag of { (tag, _) ⇒ tag }
def Nothing : 0.(A : ★) → Maybe A =
λ _ ⇒ ('nothing, 'true)
def Just : 0.(A : ★) → 1.A → Maybe A =
λ _ x ⇒ ('just, x)
def0 IsJustTag : ω.Tag → ★ =
λ t ⇒ caseω t return ★ of { 'just ⇒ True; 'nothing ⇒ False }
def0 IsJust : 0.(A : ★) → ω.(Maybe A) → ★ =
λ A x ⇒ IsJustTag (tag A x)
def is-just? : 0.(A : ★) → ω.(x : Maybe A) → Dec (IsJust A x) =
λ A x ⇒
caseω tag A x return t ⇒ Dec (IsJustTag t) of {
'just ⇒ Yes True 'true;
'nothing ⇒ No False (λ x ⇒ x)
}
def0 nothing-unique :
0.(A : ★) → ω.(x : True) → ('nothing, x) ≡ Nothing A : Maybe A =
λ A x ⇒
caseω x return x' ⇒ ('nothing, x') ≡ Nothing A : Maybe A of {
'true ⇒ δ _ ⇒ ('nothing, 'true)
}
def elim :
0.(A : ★) →
0.(P : 0.(Maybe A) → ★) →
ω.(P (Nothing A)) →
ω.(ω.(x : A) → P (Just A x)) →
1.(x : Maybe A) → P x =
λ A P n j x ⇒
caseω x return x' ⇒ P x' of {
(tag, payload) ⇒
(caseω tag
return t ⇒
0.(eq : tag ≡ t : Tag) → P (t, coe (i ⇒ Payload (eq @i) A) payload)
of {
'nothing ⇒
λ eq ⇒
caseω coe (i ⇒ Payload (eq @i) A) payload
return p ⇒ P ('nothing, p)
of { 'true ⇒ n };
'just ⇒ λ eq ⇒ j (coe (i ⇒ Payload (eq @i) A) payload)
}) (δ _ ⇒ tag)
}
}
def0 Maybe = maybe.Maybe
def0 Just = maybe.Just
def0 Nothing = maybe.Nothing

View File

@ -74,10 +74,6 @@ public export
DefsState : Type -> Type
DefsState = StateL DEFS Definitions
export
defs : Has DefsReader fs => Eff fs Definitions
defs = askAt DEFS
public export %inline
lookupElim : {d, n : Nat} -> Name -> Definitions -> Maybe (Elim d n)

View File

@ -347,9 +347,10 @@ parameters (defs : Definitions)
-- Γ ⊢ 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
local_ Equal $ do
Term.compare0 ctx sTy.zero sl tl
Term.compare0 ctx sTy.one sr tr
Term.compare0 ctx ty.zero sl tl
Term.compare0 ctx ty.one sr tr
compareType' ctx s@(Enum tags1 {}) t@(Enum tags2 {}) = do
-- ------------------
@ -636,7 +637,7 @@ parameters (loc : Loc) (ctx : TyContext d n)
private
runCompare : (Definitions -> EqContext n -> DSubst d 0 -> Equal_ ()) ->
Equal ()
runCompare act = fromEqual_ $ eachFace $ act !defs
runCompare act = fromEqual_ $ eachFace $ act !(askAt DEFS)
namespace Term
export covering

View File

@ -1,12 +1,19 @@
module Quox.NatExtra
import public Data.Nat
import public Data.Nat.Views
import Data.Nat.Division
import Data.SnocList
import Data.Vect
import Syntax.PreorderReasoning
%default total
infixl 8 `shiftL`, `shiftR`
infixl 7 .&.
infixl 6 `xor`
infixl 5 .|.
public export
data LTE' : Nat -> Nat -> Type where
@ -55,3 +62,148 @@ parameters {base : Nat} {auto 0 _ : base `GTE` 2} (chars : Vect base Char)
export
showHex : Nat -> String
showHex = showAtBase $ fromList $ unpack "0123456789ABCDEF"
export
0 notEvenOdd : (a, b : Nat) -> Not (a + a = S (b + b))
notEvenOdd 0 b prf = absurd prf
notEvenOdd (S a) b prf =
notEvenOdd b a $ Calc $
|~ b + b
~~ a + S a ..<(inj S prf)
~~ S (a + a) ..<(plusSuccRightSucc {})
export
0 doubleInj : (m, n : Nat) -> m + m = n + n -> m = n
doubleInj 0 0 _ = Refl
doubleInj (S m) (S n) prf =
cong S $ doubleInj m n $
inj S $ Calc $
|~ S (m + m)
~~ m + S m ...(plusSuccRightSucc {})
~~ n + S n ...(inj S prf)
~~ S (n + n) ..<(plusSuccRightSucc {})
export
0 halfDouble : (n : Nat) -> half (n + n) = HalfEven n
halfDouble n with (half (n + n)) | (n + n) proof nn
_ | HalfOdd k | S (k + k) = void $ notEvenOdd n k nn
_ | HalfEven k | k + k = rewrite doubleInj n k nn in Refl
export
floorHalf : Nat -> Nat
floorHalf k = case half k of
HalfOdd n => n
HalfEven n => n
||| like in intercal ☺
|||
||| take all the bits of `subj` that are set in `mask`, and squish them down
||| towards the lsb
export
select : (mask, subj : Nat) -> Nat
select mask subj = go 1 (halfRec mask) subj 0 where
go : forall mask. Nat -> HalfRec mask -> Nat -> Nat -> Nat
go bit HalfRecZ subj res = res
go bit (HalfRecEven _ rec) subj res = go bit rec (floorHalf subj) res
go bit (HalfRecOdd _ rec) subj res = case half subj of
HalfOdd subj => go (bit + bit) rec subj (res + bit)
HalfEven subj => go (bit + bit) rec subj res
||| take the i least significant bits of subj (where i = popCount mask),
||| and place them where mask's set bits are
|||
||| left inverse of select if mask .|. subj = mask
export
spread : (mask, subj : Nat) -> Nat
spread mask subj = go 1 (halfRec mask) subj 0 where
go : forall mask. Nat -> HalfRec mask -> Nat -> Nat -> Nat
go bit HalfRecZ subj res = res
go bit (HalfRecEven _ rec) subj res = go (bit + bit) rec subj res
go bit (HalfRecOdd _ rec) subj res = case half subj of
HalfOdd subj => go (bit + bit) rec subj (res + bit)
HalfEven subj => go (bit + bit) rec subj res
public export
data BitwiseRec : Nat -> Nat -> Type where
BwDone : BitwiseRec 0 0
Bw00 : (m, n : Nat) -> Lazy (BitwiseRec m n) ->
BitwiseRec (m + m) (n + n)
Bw01 : (m, n : Nat) -> Lazy (BitwiseRec m n) ->
BitwiseRec (m + m) (S (n + n))
Bw10 : (m, n : Nat) -> Lazy (BitwiseRec m n) ->
BitwiseRec (S (m + m)) (n + n)
Bw11 : (m, n : Nat) -> Lazy (BitwiseRec m n) ->
BitwiseRec (S (m + m)) (S (n + n))
export
bitwiseRec : (m, n : Nat) -> BitwiseRec m n
bitwiseRec m n = go (halfRec m) (halfRec n) where
go : forall m, n. HalfRec m -> HalfRec n -> BitwiseRec m n
go HalfRecZ HalfRecZ = BwDone
go HalfRecZ (HalfRecEven n nr) = Bw00 0 n $ go HalfRecZ nr
go HalfRecZ (HalfRecOdd n nr) = Bw01 0 n $ go HalfRecZ nr
go (HalfRecEven m mr) HalfRecZ = Bw00 m 0 $ go mr HalfRecZ
go (HalfRecEven m mr) (HalfRecEven n nr) = Bw00 m n $ go mr nr
go (HalfRecEven m mr) (HalfRecOdd n nr) = Bw01 m n $ go mr nr
go (HalfRecOdd m mr) HalfRecZ = Bw10 m 0 $ go mr HalfRecZ
go (HalfRecOdd m mr) (HalfRecEven n nr) = Bw10 m n $ go mr nr
go (HalfRecOdd m mr) (HalfRecOdd n nr) = Bw11 m n $ go mr nr
export
bitwise : (Bool -> Bool -> Bool) -> Nat -> Nat -> Nat
bitwise f m n = go 1 (bitwiseRec m n) 0 where
one : Bool -> Bool -> Nat -> Nat -> Nat
one p q bit res = if f p q then bit + res else res
go : forall m, n. Nat -> BitwiseRec m n -> Nat -> Nat
go bit BwDone res = res
go bit (Bw00 m n rec) res = go (bit + bit) rec $ one False False bit res
go bit (Bw01 m n rec) res = go (bit + bit) rec $ one False True bit res
go bit (Bw10 m n rec) res = go (bit + bit) rec $ one True False bit res
go bit (Bw11 m n rec) res = go (bit + bit) rec $ one True True bit res
export
(.&.) : Nat -> Nat -> Nat
(.&.) = bitwise $ \p, q => p && q
private %foreign "scheme:blodwen-and"
primAnd : Nat -> Nat -> Nat
%transform "NatExtra.(.&.)" NatExtra.(.&.) m n = primAnd m n
export
(.|.) : Nat -> Nat -> Nat
(.|.) = bitwise $ \p, q => p || q
private %foreign "scheme:blodwen-or"
primOr : Nat -> Nat -> Nat
%transform "NatExtra.(.|.)" NatExtra.(.|.) m n = primOr m n
export
xor : Nat -> Nat -> Nat
xor = bitwise (/=)
private %foreign "scheme:blodwen-xor"
primXor : Nat -> Nat -> Nat
%transform "NatExtra.xor" NatExtra.xor m n = primXor m n
export
shiftL : Nat -> Nat -> Nat
shiftL n 0 = n
shiftL n (S i) = shiftL (n + n) i
private %foreign "scheme:blodwen-shl"
primShiftL : Nat -> Nat -> Nat
%transform "NatExtra.shiftL" NatExtra.shiftL n i = primShiftL n i
export
shiftR : Nat -> Nat -> Nat
shiftR n 0 = n
shiftR n (S i) = shiftL (floorHalf n) i
private %foreign "scheme:blodwen-shr"
primShiftR : Nat -> Nat -> Nat
%transform "NatExtra.shiftR" NatExtra.shiftR n i = primShiftR n i

View File

@ -255,7 +255,6 @@ mutual
t => let ctx = MkNameContexts (map fromPatVar ds) (map fromPatVar ns) in
throw $ AnnotationNeeded t.loc ctx t
-- [todo] use SN if the var is named but still unused
private
fromPTermTScope : {s : Nat} -> Context' PatVar d -> Context' PatVar n ->
SnocVect s PatVar -> PTerm ->

View File

@ -6,6 +6,8 @@ import System.File
import Quox.Pretty
%hide Text.PrettyPrint.Prettyprinter.Doc.infixr.(<++>)
public export
TypeError, LexError, ParseError : Type

View File

@ -14,7 +14,6 @@ import Control.Eff
%default total
-- [fixme] rename this to Whnf and the interface to CanWhnf
public export
Whnf : List (Type -> Type)
Whnf = [NameGen, Except Error]

View File

@ -1,5 +1,6 @@
module Quox.Syntax.Term.Base
import public Quox.Thin
import public Quox.Syntax.Var
import public Quox.Syntax.Shift
import public Quox.Syntax.Subst
@ -18,9 +19,6 @@ import Data.Maybe
import Data.Nat
import public Data.So
import Data.String
import public Data.SortedMap
import public Data.SortedMap.Dependent
import public Data.SortedSet
import Derive.Prelude
%default total
@ -46,345 +44,300 @@ TagVal : Type
TagVal = String
||| type-checkable terms, which consists of types and constructor forms.
|||
||| first argument `d` is dimension scope size; second `n` is term scope size
public export
data ScopedBody : Nat -> (Nat -> Type) -> Nat -> Type where
Y : (body : f (s + n)) -> ScopedBody s f n
N : (body : f n) -> ScopedBody s f n
%name ScopedBody body
data Term : (d, n : Nat) -> Type
%name Term s, t, r
export %inline %hint
EqScopedBody : (forall n. Eq (f n)) => Eq (ScopedBody s f n)
EqScopedBody = deriveEq
export %inline %hint
ShowScopedBody : (forall n. Show (f n)) => Show (ScopedBody s f n)
ShowScopedBody = deriveShow
||| a scoped term with names
||| inferrable terms, which consists of elimination forms like application and
||| `case` (as well as other terms with an annotation)
|||
||| first argument `d` is dimension scope size; second `n` is term scope size
public export
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
constructor S
names : BContext s
body : ScopedBody s f n
%name Scoped body
export %inline
(forall n. Eq (f n)) => Eq (Scoped s f n) where
s == t = s.body == t.body
export %inline %hint
ShowScoped : (forall n. Show (f n)) => Show (Scoped s f n)
ShowScoped = deriveShow
data Elim : (d, n : Nat) -> Type
%name Elim e, f
infixl 8 :#
infixl 9 :@, :%
mutual
public export
TSubst : TSubstLike
TSubst d = Subst $ \n => Elim d n
public export
ScopeTermN : Nat -> TermLike
ScopeTermN s d n = ScopedN s (\n => Term d n) n
||| first argument `d` is dimension scope size;
||| second `n` is term scope size
public export
data Term : (d, n : Nat) -> Type where
||| type of types
TYPE : (l : Universe) -> (loc : Loc) -> Term d n
public export
DScopeTermN : Nat -> TermLike
DScopeTermN s d n = ScopedN s (\d => Term d n) d
||| function type
Pi : (qty : Qty) -> (arg : Term d n) ->
(res : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| function term
Lam : (body : ScopeTerm d n) -> (loc : Loc) -> Term d n
public export
ScopeTerm : TermLike
ScopeTerm = ScopeTermN 1
||| pair type
Sig : (fst : Term d n) -> (snd : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| pair value
Pair : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
||| enumeration type
Enum : (cases : SortedSet TagVal) -> (loc : Loc) -> Term d n
||| enumeration value
Tag : (tag : TagVal) -> (loc : Loc) -> Term d n
||| equality type
Eq : (ty : DScopeTerm d n) -> (l, r : Term d n) -> (loc : Loc) -> Term d n
||| equality term
DLam : (body : DScopeTerm d n) -> (loc : Loc) -> Term d n
||| natural numbers (temporary until 𝐖 gets added)
Nat : (loc : Loc) -> Term d n
-- [todo] can these be elims?
Zero : (loc : Loc) -> Term d n
Succ : (p : Term d n) -> (loc : Loc) -> Term d n
||| "box" (package a value up with a certain quantity)
BOX : (qty : Qty) -> (ty : Term d n) -> (loc : Loc) -> Term d n
Box : (val : Term d n) -> (loc : Loc) -> Term d n
||| elimination
E : (e : Elim d n) -> Term d n
||| term closure/suspended substitution
CloT : WithSubst (Term d) (Elim d) n -> Term d n
||| dimension closure/suspended substitution
DCloT : WithSubst (\d => Term d n) Dim d -> Term d n
%name Term s, t, r
||| first argument `d` is dimension scope size, second `n` is term scope size
public export
data Elim : (d, n : Nat) -> Type where
||| free variable, possibly with a displacement (see @crude, or @mugen for a
||| more abstract and formalised take)
|||
||| e.g. if f : ★₀ → ★₁, then f¹ : ★₁ → ★₂
F : (x : Name) -> (u : Universe) -> (loc : Loc) -> Elim d n
||| bound variable
B : (i : Var n) -> (loc : Loc) -> Elim d n
||| term application
App : (fun : Elim d n) -> (arg : Term d n) -> (loc : Loc) -> Elim d n
||| pair destruction
|||
||| `CasePair 𝜋 𝑒 ([𝑟], 𝐴) ([𝑥, 𝑦], 𝑡)` is
||| `𝐜𝐚𝐬𝐞 𝜋 · 𝑒 𝐫𝐞𝐭𝐮𝐫𝐧 𝑟𝐴 𝐨𝐟 { (𝑥, 𝑦) ⇒ 𝑡 }`
CasePair : (qty : Qty) -> (pair : Elim d n) ->
(ret : ScopeTerm d n) ->
(body : ScopeTermN 2 d n) ->
(loc : Loc) ->
Elim d n
||| enum matching
CaseEnum : (qty : Qty) -> (tag : Elim d n) ->
(ret : ScopeTerm d n) ->
(arms : CaseEnumArms d n) ->
(loc : Loc) ->
Elim d n
||| nat matching
CaseNat : (qty, qtyIH : Qty) -> (nat : Elim d n) ->
(ret : ScopeTerm d n) ->
(zero : Term d n) ->
(succ : ScopeTermN 2 d n) ->
(loc : Loc) ->
Elim d n
||| unboxing
CaseBox : (qty : Qty) -> (box : Elim d n) ->
(ret : ScopeTerm d n) ->
(body : ScopeTerm d n) ->
(loc : Loc) ->
Elim d n
||| dim application
DApp : (fun : Elim d n) -> (arg : Dim d) -> (loc : Loc) -> Elim d n
||| type-annotated term
Ann : (tm, ty : Term d n) -> (loc : Loc) -> Elim d n
||| coerce a value along a type equality, or show its coherence
||| [@xtt; §2.1.1]
Coe : (ty : DScopeTerm d n) -> (p, q : Dim d) ->
(val : Term d n) -> (loc : Loc) -> Elim d n
||| "generalised composition" [@xtt; §2.1.2]
Comp : (ty : Term d n) -> (p, q : Dim d) ->
(val : Term d n) -> (r : Dim d) ->
(zero, one : DScopeTerm d n) -> (loc : Loc) -> Elim d n
||| match on types. needed for b.s. of coercions [@xtt; §2.2]
TypeCase : (ty : Elim d n) -> (ret : Term d n) ->
(arms : TypeCaseArms d n) -> (def : Term d n) ->
(loc : Loc) ->
Elim d n
||| term closure/suspended substitution
CloE : WithSubst (Elim d) (Elim d) n -> Elim d n
||| dimension closure/suspended substitution
DCloE : WithSubst (\d => Elim d n) Dim d -> Elim d n
%name Elim e, f
public export
CaseEnumArms : TermLike
CaseEnumArms d n = SortedMap TagVal (Term d n)
public export
TypeCaseArms : TermLike
TypeCaseArms d n = SortedDMap TyConKind (\k => TypeCaseArmBody k d n)
public export
TypeCaseArm : TermLike
TypeCaseArm d n = (k ** TypeCaseArmBody k d n)
public export
TypeCaseArmBody : TyConKind -> TermLike
TypeCaseArmBody k = ScopeTermN (arity k)
public export
DScopeTerm : TermLike
DScopeTerm = DScopeTermN 1
public export
ScopeTermN, DScopeTermN : Nat -> TermLike
ScopeTermN s d n = Scoped s (Term d) n
DScopeTermN s d n = Scoped s (\d => Term d n) d
public export
TermT : TermLike
TermT = Thinned2 (\d, n => Term d n)
public export
ScopeTerm, DScopeTerm : TermLike
ScopeTerm = ScopeTermN 1
DScopeTerm = DScopeTermN 1
public export
ElimT : TermLike
ElimT = Thinned2 (\d, n => Elim d n)
mutual
export %hint
EqTerm : Eq (Term d n)
EqTerm = assert_total {a = Eq (Term d n)} deriveEq
export %hint
EqElim : Eq (Elim d n)
EqElim = assert_total {a = Eq (Elim d n)} deriveEq
public export
DimArg : TermLike
DimArg d n = Dim d
mutual
export %hint
ShowTerm : Show (Term d n)
ShowTerm = assert_total {a = Show (Term d n)} deriveShow
export %hint
ShowElim : Show (Elim d n)
ShowElim = assert_total {a = Show (Elim d n)} deriveShow
data Term where
||| type of types
TYPE : (l : Universe) -> (loc : Loc) -> Term 0 0
||| scope which ignores all its binders
public export %inline
SN : {s : Nat} -> f n -> Scoped s f n
SN = S (replicate s $ BN Unused noLoc) . N
||| function type
Pi : Qty -> Subterms [Term, ScopeTerm] d n -> Loc -> Term d n
||| function value
Lam : ScopeTerm d n -> Loc -> Term d n
||| scope which uses its binders
public export %inline
SY : BContext s -> f (s + n) -> Scoped s f n
SY ns = S ns . Y
||| pair type
Sig : Subterms [Term, ScopeTerm] d n -> Loc -> Term d n
||| pair value
Pair : Subterms [Term, Term] d n -> Loc -> Term d n
public export %inline
name : Scoped 1 f n -> BindName
name (S [< x] _) = x
||| enum type
Enum : List TagVal -> Loc -> Term 0 0
||| enum value
Tag : TagVal -> Loc -> Term 0 0
public export %inline
(.name) : Scoped 1 f n -> BindName
s.name = name s
||| equality type
Eq : Subterms [DScopeTerm, Term, Term] d n -> Loc -> Term d n
||| equality value
DLam : DScopeTerm d n -> Loc -> Term d n
||| more convenient Pi
public export %inline
PiY : (qty : Qty) -> (x : BindName) ->
(arg : Term d n) -> (res : Term d (S n)) -> (loc : Loc) -> Term d n
PiY {qty, x, arg, res, loc} = Pi {qty, arg, res = SY [< x] res, loc}
||| natural numbers (temporary until 𝐖 gets added)
Nat : Loc -> Term 0 0
Zero : Loc -> Term 0 0
Succ : Term d n -> Loc -> Term 0 0
||| more convenient Lam
public export %inline
LamY : (x : BindName) -> (body : Term d (S n)) -> (loc : Loc) -> Term d n
LamY {x, body, loc} = Lam {body = SY [< x] body, loc}
||| package a value with a quantity
||| e.g. a value of [ω. A], when unpacked, can be used ω times,
||| even if the box itself is linear
BOX : Qty -> Term d n -> Loc -> Term d n
Box : Term d n -> Loc -> Term d n
public export %inline
LamN : (body : Term d n) -> (loc : Loc) -> Term d n
LamN {body, loc} = Lam {body = SN body, loc}
E : Elim d n -> Term d n
||| non dependent function type
public export %inline
Arr : (qty : Qty) -> (arg, res : Term d n) -> (loc : Loc) -> Term d n
Arr {qty, arg, res, loc} = Pi {qty, arg, res = SN res, loc}
||| term closure/suspended substitution
CloT : WithSubst (Term d) (Elim d) n -> Term d n
||| dimension closure/suspended substitution
DCloT : WithSubst (\d => Term d n) Dim d -> Term d n
||| more convenient Sig
public export %inline
SigY : (x : BindName) -> (fst : Term d n) ->
(snd : Term d (S n)) -> (loc : Loc) -> Term d n
SigY {x, fst, snd, loc} = Sig {fst, snd = SY [< x] snd, loc}
public export
data Elim where
||| free variable, possibly with a displacement (see @crude, or @mugen for a
||| more abstract and formalised take)
|||
||| e.g. if f : ★₀ → ★₁, then f¹ : ★₁ → ★₂
F : Name -> Universe -> Loc -> Elim 0 0
||| bound variable
B : Loc -> Elim 0 1
||| non dependent pair type
public export %inline
And : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
And {fst, snd, loc} = Sig {fst, snd = SN snd, loc}
||| term application
App : Subterms [Elim, Term] d n -> Loc -> Elim d n
||| more convenient Eq
public export %inline
EqY : (i : BindName) -> (ty : Term (S d) n) ->
(l, r : Term d n) -> (loc : Loc) -> Term d n
EqY {i, ty, l, r, loc} = Eq {ty = SY [< i] ty, l, r, loc}
||| pair match
||| - the subterms are, in order: [head, return type, body]
||| - the quantity is that of the head, and since pairs only have one
||| constructor, can be 0
CasePair : Qty -> Subterms [Elim, ScopeTerm, ScopeTermN 2] d n ->
Loc -> Elim d n
||| more convenient DLam
public export %inline
DLamY : (i : BindName) -> (body : Term (S d) n) -> (loc : Loc) -> Term d n
DLamY {i, body, loc} = DLam {body = SY [< i] body, loc}
||| enum match
CaseEnum : Qty -> (arms : List TagVal) ->
Subterms (Elim :: ScopeTerm :: (Term <$ arms)) d n ->
Loc -> Elim d n
public export %inline
DLamN : (body : Term d n) -> (loc : Loc) -> Term d n
DLamN {body, loc} = DLam {body = SN body, loc}
||| nat match
CaseNat : Qty -> Qty ->
Subterms [Elim, ScopeTerm, Term, ScopeTermN 2] d n ->
Loc -> Elim d n
||| box match
CaseBox : Qty -> Subterms [Elim, ScopeTerm, ScopeTerm] d n -> Loc -> Elim d n
||| dim application
DApp : Subterms [Elim, DimArg] d n -> Loc -> Elim d n
||| type-annotated term
Ann : Subterms [Term, Term] d n -> Loc -> Elim d n
||| coerce a value along a type equality, or show its coherence
||| [@xtt; §2.1.1]
Coe : Subterms [DScopeTerm, DimArg, DimArg, Term] d n ->
Loc -> Elim d n
||| "generalised composition" [@xtt; §2.1.2]
Comp : Subterms [Term, DimArg, DimArg, Term,
DimArg, DScopeTerm, DScopeTerm] d n ->
Loc -> Elim d n
||| match on types. needed for b.s. of coercions [@xtt; §2.2]
TypeCase : Subterms [Elim, Term, -- head, type
Term, -- ★
ScopeTermN 2, -- pi
ScopeTermN 2, -- sig
Term, -- enum
ScopeTermN 5, -- eq
Term, -- nat
ScopeTerm -- box
] d n -> Loc -> Elim d n
||| term closure/suspended substitution
CloE : WithSubst (Elim d) (Elim d) n -> Elim d n
||| dimension closure/suspended substitution
DCloE : WithSubst (\d => Elim d n) Dim d -> Elim d n
-- this kills the idris ☹
-- export %hint
-- EqTerm : Eq (Term d n)
-- export %hint
-- EqElim : Eq (Elim d n)
-- EqTerm = deriveEq
-- EqElim = deriveEq
-- mutual
-- export %hint
-- ShowTerm : Show (Term d n)
-- ShowTerm = assert_total {a = Show (Term d n)} deriveShow
-- export %hint
-- ShowElim : Show (Elim d n)
-- ShowElim = assert_total {a = Show (Elim d n)} deriveShow
-- ||| scope which ignores all its binders
-- public export %inline
-- SN : {s : Nat} -> f n -> Scoped s f n
-- SN = S (replicate s $ BN Unused noLoc) . N
-- ||| scope which uses its binders
-- public export %inline
-- SY : BContext s -> f (s + n) -> Scoped s f n
-- SY ns = S ns . Y
-- public export %inline
-- name : Scoped 1 f n -> BindName
-- name (S [< x] _) = x
-- public export %inline
-- (.name) : Scoped 1 f n -> BindName
-- s.name = name s
-- ||| more convenient Pi
-- public export %inline
-- PiY : (qty : Qty) -> (x : BindName) ->
-- (arg : Term d n) -> (res : Term d (S n)) -> (loc : Loc) -> Term d n
-- PiY {qty, x, arg, res, loc} = Pi {qty, arg, res = SY [< x] res, loc}
-- ||| more convenient Lam
-- public export %inline
-- LamY : (x : BindName) -> (body : Term d (S n)) -> (loc : Loc) -> Term d n
-- LamY {x, body, loc} = Lam {body = SY [< x] body, loc}
-- public export %inline
-- LamN : (body : Term d n) -> (loc : Loc) -> Term d n
-- LamN {body, loc} = Lam {body = SN body, loc}
-- ||| non dependent function type
-- public export %inline
-- Arr : (qty : Qty) -> (arg, res : Term d n) -> (loc : Loc) -> Term d n
-- Arr {qty, arg, res, loc} = Pi {qty, arg, res = SN res, loc}
-- ||| more convenient Sig
-- public export %inline
-- SigY : (x : BindName) -> (fst : Term d n) ->
-- (snd : Term d (S n)) -> (loc : Loc) -> Term d n
-- SigY {x, fst, snd, loc} = Sig {fst, snd = SY [< x] snd, loc}
-- ||| non dependent pair type
-- public export %inline
-- And : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
-- And {fst, snd, loc} = Sig {fst, snd = SN snd, loc}
-- ||| more convenient Eq
-- public export %inline
-- EqY : (i : BindName) -> (ty : Term (S d) n) ->
-- (l, r : Term d n) -> (loc : Loc) -> Term d n
-- EqY {i, ty, l, r, loc} = Eq {ty = SY [< i] ty, l, r, loc}
-- ||| more convenient DLam
-- public export %inline
-- DLamY : (i : BindName) -> (body : Term (S d) n) -> (loc : Loc) -> Term d n
-- DLamY {i, body, loc} = DLam {body = SY [< i] body, loc}
-- public export %inline
-- DLamN : (body : Term d n) -> (loc : Loc) -> Term d n
-- DLamN {body, loc} = DLam {body = SN body, loc}
-- ||| non dependent equality type
-- public export %inline
-- Eq0 : (ty, l, r : Term d n) -> (loc : Loc) -> Term d n
-- Eq0 {ty, l, r, loc} = Eq {ty = SN ty, l, r, loc}
||| non dependent equality type
public export %inline
Eq0 : (ty, l, r : Term d n) -> (loc : Loc) -> Term d n
Eq0 {ty, l, r, loc} = Eq {ty = SN ty, l, r, loc}
||| same as `F` but as a term
public export %inline
FT : Name -> Universe -> Loc -> Term d n
FT : Name -> Universe -> Loc -> Term 0 0
FT x u loc = E $ F x u loc
||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))`
public export %inline
BV : (i : Nat) -> (0 _ : LT i n) => (loc : Loc) -> Elim d n
BV i loc = B (V i) loc
BV : (i : Fin n) -> (loc : Loc) -> ElimT d n
BV i loc = Th2 zero (one' i) $ B loc
||| same as `BV` but as a term
public export %inline
BVT : (i : Nat) -> (0 _ : LT i n) => (loc : Loc) -> Term d n
BVT i loc = E $ BV i loc
BVT : (i : Fin n) -> (loc : Loc) -> TermT d n
BVT i loc = Th2 zero (one' i) $ E $ B loc
public export
makeNat : Nat -> Loc -> Term d n
makeNat 0 loc = Zero loc
makeNat : Nat -> Loc -> Term 0 0
makeNat 0 loc = Zero loc
makeNat (S k) loc = Succ (makeNat k loc) loc
public export %inline
enum : List TagVal -> Loc -> Term d n
enum ts loc = Enum (SortedSet.fromList ts) loc
public export %inline
typeCase : Elim d n -> Term d n ->
List (TypeCaseArm d n) -> Term d n -> Loc -> Elim d n
typeCase ty ret arms def loc = TypeCase ty ret (fromList arms) def loc
public export %inline
typeCase1Y : Elim d n -> Term d n ->
(k : TyConKind) -> BContext (arity k) -> Term d (arity k + n) ->
(loc : Loc) ->
{default (Nat loc) def : Term d n} ->
Elim d n
typeCase1Y ty ret k ns body loc = typeCase ty ret [(k ** SY ns body)] def loc
export
Located (Elim d n) where
(F _ _ loc).loc = loc
(B _ loc).loc = loc
(App _ _ loc).loc = loc
(CasePair _ _ _ _ loc).loc = loc
(CaseEnum _ _ _ _ loc).loc = loc
(CaseNat _ _ _ _ _ _ loc).loc = loc
(CaseBox _ _ _ _ loc).loc = loc
(DApp _ _ loc).loc = loc
(Ann _ _ loc).loc = loc
(Coe _ _ _ _ loc).loc = loc
(Comp _ _ _ _ _ _ _ loc).loc = loc
(TypeCase _ _ _ _ loc).loc = loc
(CloE (Sub e _)).loc = e.loc
(DCloE (Sub e _)).loc = e.loc
(F _ _ loc).loc = loc
(B loc).loc = loc
(App _ loc).loc = loc
(CasePair _ _ loc).loc = loc
(CaseEnum _ _ _ loc).loc = loc
(CaseNat _ _ _ loc).loc = loc
(CaseBox _ _ loc).loc = loc
(DApp _ loc).loc = loc
(Ann _ loc).loc = loc
(Coe _ loc).loc = loc
(Comp _ loc).loc = loc
(TypeCase _ loc).loc = loc
(CloE (Sub e _)).loc = e.loc
(DCloE (Sub e _)).loc = e.loc
export
Located (Term d n) where
(TYPE _ loc).loc = loc
(Pi _ _ _ loc).loc = loc
(Pi _ _ loc).loc = loc
(Lam _ loc).loc = loc
(Sig _ _ loc).loc = loc
(Pair _ _ loc).loc = loc
(Sig _ loc).loc = loc
(Pair _ loc).loc = loc
(Enum _ loc).loc = loc
(Tag _ loc).loc = loc
(Eq _ _ _ loc).loc = loc
(Eq _ loc).loc = loc
(DLam _ loc).loc = loc
(Nat loc).loc = loc
(Zero loc).loc = loc
@ -395,69 +348,40 @@ Located (Term d n) where
(CloT (Sub t _)).loc = t.loc
(DCloT (Sub t _)).loc = t.loc
export
Located1 f => Located (ScopedBody s f n) where
(Y t).loc = t.loc
(N t).loc = t.loc
export
Located1 f => Located (Scoped s f n) where
t.loc = t.body.loc
export
Relocatable (Elim d n) where
setLoc loc (F x u _) = F x u loc
setLoc loc (B i _) = B i loc
setLoc loc (App fun arg _) = App fun arg loc
setLoc loc (CasePair qty pair ret body _) =
CasePair qty pair ret body loc
setLoc loc (CaseEnum qty tag ret arms _) =
CaseEnum qty tag ret arms loc
setLoc loc (CaseNat qty qtyIH nat ret zero succ _) =
CaseNat qty qtyIH nat ret zero succ loc
setLoc loc (CaseBox qty box ret body _) =
CaseBox qty box ret body loc
setLoc loc (DApp fun arg _) =
DApp fun arg loc
setLoc loc (Ann tm ty _) =
Ann tm ty loc
setLoc loc (Coe ty p q val _) =
Coe ty p q val loc
setLoc loc (Comp ty p q val r zero one _) =
Comp ty p q val r zero one loc
setLoc loc (TypeCase ty ret arms def _) =
TypeCase ty ret arms def loc
setLoc loc (CloE (Sub term subst)) =
CloE $ Sub (setLoc loc term) subst
setLoc loc (DCloE (Sub term subst)) =
DCloE $ Sub (setLoc loc term) subst
setLoc loc (F x u _) = F x u loc
setLoc loc (B _) = B loc
setLoc loc (App ts _) = App ts loc
setLoc loc (CasePair qty ts _) = CasePair qty ts loc
setLoc loc (CaseEnum qty arms ts _) = CaseEnum qty arms ts loc
setLoc loc (CaseNat qty qtyIH ts _) = CaseNat qty qtyIH ts loc
setLoc loc (CaseBox qty ts _) = CaseBox qty ts loc
setLoc loc (DApp ts _) = DApp ts loc
setLoc loc (Ann ts _) = Ann ts loc
setLoc loc (Coe ts _) = Coe ts loc
setLoc loc (Comp ts _) = Comp ts loc
setLoc loc (TypeCase ts _) = TypeCase ts loc
setLoc loc (CloE (Sub term subst)) = CloE $ Sub (setLoc loc term) subst
setLoc loc (DCloE (Sub term subst)) = DCloE $ Sub (setLoc loc term) subst
export
Relocatable (Term d n) where
setLoc loc (TYPE l _) = TYPE l loc
setLoc loc (Pi qty arg res _) = Pi qty arg res loc
setLoc loc (Lam body _) = Lam body loc
setLoc loc (Sig fst snd _) = Sig fst snd loc
setLoc loc (Pair fst snd _) = Pair fst snd loc
setLoc loc (Enum cases _) = Enum cases loc
setLoc loc (Tag tag _) = Tag tag loc
setLoc loc (Eq ty l r _) = Eq ty l r loc
setLoc loc (DLam body _) = DLam body loc
setLoc loc (Nat _) = Nat loc
setLoc loc (Zero _) = Zero loc
setLoc loc (Succ p _) = Succ p loc
setLoc loc (BOX qty ty _) = BOX qty ty loc
setLoc loc (Box val _) = Box val loc
setLoc loc (E e) = E $ setLoc loc e
setLoc loc (CloT (Sub term subst)) = CloT $ Sub (setLoc loc term) subst
setLoc loc (TYPE l _) = TYPE l loc
setLoc loc (Pi qty ts _) = Pi qty ts loc
setLoc loc (Lam body _) = Lam body loc
setLoc loc (Sig ts _) = Sig ts loc
setLoc loc (Pair ts _) = Pair ts loc
setLoc loc (Enum cases _) = Enum cases loc
setLoc loc (Tag tag _) = Tag tag loc
setLoc loc (Eq ts _) = Eq ts loc
setLoc loc (DLam body _) = DLam body loc
setLoc loc (Nat _) = Nat loc
setLoc loc (Zero _) = Zero loc
setLoc loc (Succ p _) = Succ p loc
setLoc loc (BOX qty ty _) = BOX qty ty loc
setLoc loc (Box val _) = Box val loc
setLoc loc (E e) = E $ setLoc loc e
setLoc loc (CloT (Sub term subst)) = CloT $ Sub (setLoc loc term) subst
setLoc loc (DCloT (Sub term subst)) = DCloT $ Sub (setLoc loc term) subst
export
Relocatable1 f => Relocatable (ScopedBody s f n) where
setLoc loc (Y body) = Y $ setLoc loc body
setLoc loc (N body) = N $ setLoc loc body
export
Relocatable1 f => Relocatable (Scoped s f n) where
setLoc loc (S names body) = S (setLoc loc <$> names) (setLoc loc body)

View File

@ -58,11 +58,12 @@ pbname (_, x, _) = x
private
record SplitPi d n where
constructor MkSplitPi
{0 inner : Nat}
binds : Telescope (PiBind d) n inner
cod : Term d inner
private
splitPi : Telescope (PiBind d) n inner -> Term d inner -> SplitPi d n
splitPi : Telescope (PiBind d) n n' -> Term d n' -> SplitPi d n
splitPi binds (Pi qty arg res _) =
splitPi (binds :< (qty, res.name, arg)) $
assert_smaller res $ pushSubsts' res.term
@ -87,7 +88,7 @@ prettyPiBind1 pi x dnames tnames s = hcat <$> sequence
private
prettyPiBinds : {opts : _} ->
BContext d -> BContext n ->
Telescope (PiBind d) n inner ->
Telescope (PiBind d) n n' ->
Eff Pretty (SnocList (Doc opts))
prettyPiBinds _ _ [<] = pure [<]
prettyPiBinds dnames tnames (binds :< (q, x, t)) =
@ -103,11 +104,12 @@ SigBind d n = (BindName, Term d n)
private
record SplitSig d n where
constructor MkSplitSig
{0 inner : Nat}
binds : Telescope (SigBind d) n inner
last : Term d inner
private
splitSig : Telescope (SigBind d) n inner -> Term d inner -> SplitSig d n
splitSig : Telescope (SigBind d) n n' -> Term d n' -> SplitSig d n
splitSig binds (Sig fst snd _) =
splitSig (binds :< (snd.name, fst)) $
assert_smaller snd $ pushSubsts' snd.term
@ -129,7 +131,7 @@ prettySigBind1 x dnames tnames s = hcat <$> sequence
private
prettySigBinds : {opts : _} ->
BContext d -> BContext n ->
Telescope (SigBind d) n inner ->
Telescope (SigBind d) n n' ->
Eff Pretty (SnocList (Doc opts))
prettySigBinds _ _ [<] = pure [<]
prettySigBinds dnames tnames (binds :< (x, t)) =
@ -163,6 +165,7 @@ NameChunks = SnocList (NameSort, SnocList BindName)
private
record SplitLams d n where
constructor MkSplitLams
{0 dinner, ninner : Nat}
dnames : BTelescope d dinner
tnames : BTelescope n ninner
chunks : NameChunks
@ -178,9 +181,9 @@ where
if s == s' then xss :< (s', xs :< y)
else xss :< (s', xs) :< (s, [< y])
go : BTelescope d dinner -> BTelescope n ninner ->
go : BTelescope d d' -> BTelescope n n' ->
SnocList (NameSort, SnocList BindName) ->
Term dinner ninner -> SplitLams d n
Term d' n' -> SplitLams d n
go dnames tnames chunks (Lam b _) =
go dnames (tnames :< b.name) (push T b.name chunks) $
assert_smaller b $ pushSubsts' b.term
@ -235,6 +238,7 @@ prettyDTApps dnames tnames f xs = do
private
record CaseArm opts d n where
constructor MkCaseArm
{0 dinner, ninner : Nat}
pat : Doc opts
dbinds : BTelescope d dinner -- 🍴
tbinds : BTelescope n ninner

View File

@ -10,7 +10,7 @@ import Data.SnocVect
namespace CanDSubst
public export
interface CanDSubst (0 tm : TermLike) where
(//) : tm dfrom n -> Lazy (DSubst dfrom dto) -> tm dto n
(//) : tm d1 n -> Lazy (DSubst d1 d2) -> tm d2 n
||| does the minimal reasonable work:
||| - deletes the closure around an atomic constant like `TYPE`
@ -25,7 +25,7 @@ CanDSubst Term where
s // th = DCloT $ Sub s th
private
subDArgs : Elim dfrom n -> DSubst dfrom dto -> Elim dto n
subDArgs : Elim d1 n -> DSubst d1 d2 -> Elim d2 n
subDArgs (DApp f d loc) th = DApp (subDArgs f th) (d // th) loc
subDArgs e th = DCloE $ Sub e th
@ -47,16 +47,16 @@ CanDSubst Elim where
namespace DSubst.ScopeTermN
export %inline
(//) : ScopeTermN s dfrom n -> Lazy (DSubst dfrom dto) ->
ScopeTermN s dto n
(//) : ScopeTermN s d1 n -> Lazy (DSubst d1 d2) ->
ScopeTermN s d2 n
S ns (Y body) // th = S ns $ Y $ body // th
S ns (N body) // th = S ns $ N $ body // th
namespace DSubst.DScopeTermN
export %inline
(//) : {s : Nat} ->
DScopeTermN s dfrom n -> Lazy (DSubst dfrom dto) ->
DScopeTermN s dto n
DScopeTermN s d1 n -> Lazy (DSubst d1 d2) ->
DScopeTermN s d2 n
S ns (Y body) // th = S ns $ Y $ body // pushN s th
S ns (N body) // th = S ns $ N $ body // th
@ -83,7 +83,7 @@ CanSubstSelf (Elim d) where
namespace CanTSubst
public export
interface CanTSubst (0 tm : TermLike) where
(//) : tm d from -> Lazy (TSubst d from to) -> tm d to
(//) : tm d n1 -> Lazy (TSubst d n1 n2) -> tm d n2
||| does the minimal reasonable work:
||| - deletes the closure around an atomic constant like `TYPE`
@ -103,16 +103,15 @@ CanTSubst Term where
namespace ScopeTermN
export %inline
(//) : {s : Nat} ->
ScopeTermN s d from -> Lazy (TSubst d from to) ->
ScopeTermN s d to
ScopeTermN s d n1 -> Lazy (TSubst d n1 n2) ->
ScopeTermN s d n2
S ns (Y body) // th = S ns $ Y $ body // pushN s th
S ns (N body) // th = S ns $ N $ body // th
namespace DScopeTermN
export %inline
(//) : {s : Nat} ->
DScopeTermN s d from -> Lazy (TSubst d from to) ->
DScopeTermN s d to
DScopeTermN s d n1 -> Lazy (TSubst d n1 n2) -> DScopeTermN s d n2
S ns (Y body) // th = S ns $ Y $ body // map (// shift s) th
S ns (N body) // th = S ns $ N $ body // th
@ -125,8 +124,7 @@ export %inline
export %inline
comp : DSubst dfrom dto -> TSubst dfrom from mid -> TSubst dto mid to ->
TSubst dto from to
comp : DSubst d1 d2 -> TSubst d1 n1 mid -> TSubst d2 mid n2 -> TSubst d2 n1 n2
comp th ps ph = map (// th) ps . ph
@ -205,8 +203,8 @@ public export
CloTest tm = forall d, n. tm d n -> Bool
interface PushSubsts (0 tm : TermLike) (0 isClo : CloTest tm) | tm where
pushSubstsWith : DSubst dfrom dto -> TSubst dto from to ->
tm dfrom from -> Subset (tm dto to) (No . isClo)
pushSubstsWith : DSubst d1 d2 -> TSubst d2 n1 n2 ->
tm d1 n1 -> Subset (tm d2 n2) (No . isClo)
public export
0 NotClo : {isClo : CloTest tm} -> PushSubsts tm isClo => Pred (tm d n)
@ -230,8 +228,7 @@ parameters {0 isClo : CloTest tm} {auto _ : PushSubsts tm isClo}
pushSubsts s = pushSubstsWith id id s
export %inline
pushSubstsWith' : DSubst dfrom dto -> TSubst dto from to ->
tm dfrom from -> tm dto to
pushSubstsWith' : DSubst d1 d2 -> TSubst d2 n1 n2 -> tm d1 n1 -> tm d2 n2
pushSubstsWith' th ph x = fst $ pushSubstsWith th ph x
export %inline
@ -334,8 +331,7 @@ private %inline
CompHY : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(r : Dim d) -> (zero, one : DScopeTerm d n) -> (loc : Loc) -> Elim d n
CompHY {ty, p, q, val, r, zero, one, loc} =
-- [fixme] maintain location of existing B VZ
let ty' = SY ty.names $ ty.term // (B VZ noLoc ::: shift 2) in
let ty' = SY ty.names $ ty.term // (B VZ ty.loc ::: shift 2) in
Comp {
ty = dsub1 ty q, p, q,
val = E $ Coe ty p q val val.loc, r,

View File

@ -8,10 +8,10 @@ import public Quox.OPE
export
Tighten (Shift from) where
Tighten (Shift f) where
-- `OPE m n` is a spicy `m ≤ n`,
-- and `Shift from n` is a (different) spicy `from ≤ n`
-- so the value is `from ≤ m` (as a `Shift`), if that is the case
-- and `Shift f n` is a (different) spicy `f ≤ n`
-- so the value is `f ≤ m` (as a `Shift`), if that is the case
tighten _ SZ = Nothing
tighten Id by = Just by
tighten (Drop p) (SS by) = tighten p by
@ -26,12 +26,12 @@ Tighten Dim where
export
tightenSub : (forall m, n. OPE m n -> env n -> Maybe (env m)) ->
OPE to1 to2 -> Subst env from to2 -> Maybe (Subst env from to1)
OPE t1 t2 -> Subst env f t2 -> Maybe (Subst env f t1)
tightenSub f p (Shift by) = [|Shift $ tighten p by|]
tightenSub f p (t ::: th) = [|f p t !::: tightenSub f p th|]
export
Tighten env => Tighten (Subst env from) where
Tighten env => Tighten (Subst env f) where
tighten p th = tightenSub tighten p th

13
lib/Quox/Thin.idr Normal file
View File

@ -0,0 +1,13 @@
module Quox.Thin
import public Quox.Thin.Base
import public Quox.Thin.View
import public Quox.Thin.Eqv
import public Quox.Thin.Cons
import public Quox.Thin.List
import public Quox.Thin.Append
import public Quox.Thin.Comp
import public Quox.Thin.Cover
import public Quox.Thin.Coprod
import public Quox.Thin.Split
import public Quox.Thin.Term

27
lib/Quox/Thin/Append.idr Normal file
View File

@ -0,0 +1,27 @@
module Quox.Thin.Append
import public Quox.Thin.Base
import public Quox.Thin.View
import Data.DPair
%default total
public export
app' : OPE m1 n1 mask1 -> OPE m2 n2 mask2 -> Exists (OPE (m1 + m2) (n1 + n2))
app' Stop ope2 = Evidence _ ope2
app' (Drop ope1 Refl) ope2 = Evidence _ $ Drop (app' ope1 ope2).snd Refl
app' (Keep ope1 Refl) ope2 = Evidence _ $ Keep (app' ope1 ope2).snd Refl
public export
(++) : {n1, n2, mask1, mask2 : Nat} ->
(0 ope1 : OPE m1 n1 mask1) -> (0 ope2 : OPE m2 n2 mask2) ->
Subset Nat (OPE (m1 + m2) (n1 + n2))
ope1 ++ ope2 with %syntactic (view ope1)
Stop ++ ope2 | StopV = Element _ ope2
Drop ope1 Refl ++ ope2 | DropV mask ope1 =
Element _ $ Drop (ope1 ++ ope2).snd Refl
Keep ope1 Refl ++ ope2 | KeepV mask ope1 =
Element _ $ Keep (ope1 ++ ope2).snd Refl
-- [todo] this mask is just (mask1 << n2) | mask2
-- prove it and add %transform

81
lib/Quox/Thin/Base.idr Normal file
View File

@ -0,0 +1,81 @@
module Quox.Thin.Base
import Data.Fin
import Data.DPair
%default total
||| "order preserving embeddings", for recording a correspondence between a
||| smaller scope and part of a larger one. the third argument is a bitmask
||| representing this OPE, unique for a given `n`.
public export
data OPE : (m, n, mask : Nat) -> Type where
[search m n]
Stop : OPE 0 0 0
Drop : OPE m n mask -> mask' = mask + mask -> OPE m (S n) mask'
Keep : OPE m n mask -> mask' = (S (mask + mask)) -> OPE (S m) (S n) mask'
%name OPE ope
export
Show (OPE m n mask) where
showPrec d Stop = "Stop"
showPrec d (Drop ope Refl) = showCon d "Drop" $ showArg ope ++ " Refl"
showPrec d (Keep ope Refl) = showCon d "Keep" $ showArg ope ++ " Refl"
public export %inline
drop : OPE m n mask -> OPE m (S n) (mask + mask)
drop ope = Drop ope Refl
public export %inline
keep : OPE m n mask -> OPE (S m) (S n) (S (mask + mask))
keep ope = Keep ope Refl
public export
data IsStop : OPE m n mask -> Type where ItIsStop : IsStop Stop
public export
data IsDrop : OPE m n mask -> Type where ItIsDrop : IsDrop (Drop ope eq)
public export
data IsKeep : OPE m n mask -> Type where ItIsKeep : IsKeep (Keep ope eq)
export
0 zeroIsStop : (ope : OPE m 0 mask) -> IsStop ope
zeroIsStop Stop = ItIsStop
||| everything selected
public export
id : {m : Nat} -> Subset Nat (OPE m m)
id {m = 0} = Element _ Stop
id {m = S m} = Element _ $ Keep id.snd Refl
public export %inline
0 id' : OPE m m Base.id.fst
id' = id.snd
||| nothing selected
public export
zero : {m : Nat} -> OPE 0 m 0
zero {m = 0} = Stop
zero {m = S m} = Drop zero Refl
||| a single slot selected
public export
one : Fin n -> Subset Nat (OPE 1 n)
one FZ = Element _ $ keep zero
one (FS i) = Element _ $ drop (one i).snd
public export %inline
0 one' : (i : Fin n) -> OPE 1 n (one i).fst
one' i = (one i).snd
public export
record SomeOPE n where
constructor MkOPE
{0 scope : Nat}
{mask : Nat}
0 ope : OPE scope n mask

59
lib/Quox/Thin/Comp.idr Normal file
View File

@ -0,0 +1,59 @@
module Quox.Thin.Comp
import public Quox.Thin.Base
import public Quox.Thin.View
import Data.Singleton
%default total
||| inductive definition of OPE composition
public export
data Comp : (l : OPE n p mask1) -> (r : OPE m n mask2) ->
(res : OPE m p mask3) -> Type where
[search l r]
StopZ : Comp Stop Stop Stop
DropZ : Comp a b c -> Comp (Drop a Refl) b (Drop c Refl)
KeepZ : Comp a b c -> Comp (Keep a Refl) (Keep b Refl) (Keep c Refl)
KDZ : Comp a b c -> Comp (Keep a Refl) (Drop b Refl) (Drop c Refl)
public export
record CompResult (ope1 : OPE n p mask1) (ope2 : OPE m n mask2) where
constructor MkComp
{mask : Nat}
{0 ope : OPE m p mask}
0 comp : Comp ope1 ope2 ope
%name CompResult comp
||| compose two OPEs, if the middle scope size is already known at runtime
export
comp' : {n, p, mask1, mask2 : Nat} ->
(0 ope1 : OPE n p mask1) -> (0 ope2 : OPE m n mask2) ->
CompResult ope1 ope2
comp' ope1 ope2 with %syntactic (view ope1) | (view ope2)
comp' Stop Stop | StopV | StopV =
MkComp StopZ
comp' (Drop ope1 Refl) ope2 | DropV _ ope1 | _ =
MkComp $ DropZ (comp' ope1 ope2).comp
comp' (Keep ope1 Refl) (Drop ope2 Refl) | KeepV _ ope1 | DropV _ ope2 =
MkComp $ KDZ (comp' ope1 ope2).comp
comp' (Keep ope1 Refl) (Keep ope2 Refl) | KeepV _ ope1 | KeepV _ ope2 =
MkComp $ KeepZ (comp' ope1 ope2).comp
||| compose two OPEs, after recomputing the middle scope size using `appOpe`
export
comp : {p, mask1, mask2 : Nat} ->
(0 ope1 : OPE n p mask1) -> (0 ope2 : OPE m n mask2) ->
CompResult ope1 ope2
comp ope1 ope2 = let Val n = appOpe p ope1 in comp' ope1 ope2
-- [todo] is there a quick way to compute the mask of comp?
export
0 (.) : (ope1 : OPE n p mask1) -> (ope2 : OPE m n mask2) ->
OPE m p (comp ope1 ope2).mask
ope1 . ope2 = (comp ope1 ope2).ope
-- export
-- 0 compMask : (ope1 : OPE n p mask1) -> (ope2 : OPE m n mask2) ->
-- (ope3 : OPE m p mask3) -> Comp ope1 ope2 ope3 ->
-- mask3 = ?aaa

74
lib/Quox/Thin/Cons.idr Normal file
View File

@ -0,0 +1,74 @@
module Quox.Thin.Cons
import public Quox.Thin.Base
import Quox.Thin.Eqv
import Quox.Thin.View
import Data.DPair
import Control.Relation
%default total
public export
data IsHead : (ope : OPE m (S n) mask) -> Bool -> Type where
[search ope]
DropH : IsHead (Drop ope eq) False
KeepH : IsHead (Keep ope eq) True
public export
data IsTail : (full : OPE m (S n) mask) -> OPE m' n mask' -> Type where
[search full]
DropT : IsTail (Drop ope eq) ope
KeepT : IsTail (Keep ope eq) ope
public export
record Uncons (ope : OPE m (S n) mask) where
constructor MkUncons
0 head : Bool
{tailMask : Nat}
0 tail : OPE scope n tailMask
{auto isHead : IsHead ope head}
{auto 0 isTail : IsTail ope tail}
public export
uncons : {n, mask : Nat} -> (0 ope : OPE m (S n) mask) -> Uncons ope
uncons ope with %syntactic (view ope)
uncons (Drop ope Refl) | DropV _ ope = MkUncons False ope
uncons (Keep ope Refl) | KeepV _ ope = MkUncons True ope
public export
head : {n, mask : Nat} -> (0 ope : OPE m (S n) mask) -> Exists $ IsHead ope
head ope = Evidence _ (uncons ope).isHead
public export
record Tail (ope : OPE m (S n) mask) where
constructor MkTail
{tailMask : Nat}
0 tail : OPE scope n tailMask
{auto 0 isTail : IsTail ope tail}
public export
tail : {n, mask : Nat} -> (0 ope : OPE m (S n) mask) -> Tail ope
tail ope = let u = uncons ope in MkTail u.tail @{u.isTail}
export
cons : {mask : Nat} -> (head : Bool) -> (0 tail : OPE m n mask) ->
Subset Nat (OPE (if head then S m else m) (S n))
cons False tail = Element _ $ drop tail
cons True tail = Element _ $ keep tail
export
0 consEquiv' : (self : OPE m' (S n) mask') ->
(head : Bool) -> (tail : OPE m n mask) ->
IsHead self head -> IsTail self tail ->
(cons head tail).snd `Eqv` self
consEquiv' (Drop tail _) False tail DropH DropT = EqvDrop reflexive
consEquiv' (Keep tail _) True tail KeepH KeepT = EqvKeep reflexive
export
0 consEquiv : (full : OPE m' (S n) mask') ->
(cons (uncons full).head (uncons full).tail).snd `Eqv` full
consEquiv full with %syntactic (uncons full)
_ | MkUncons head tail {isHead, isTail} =
consEquiv' full head tail isHead isTail

171
lib/Quox/Thin/Coprod.idr Normal file
View File

@ -0,0 +1,171 @@
module Quox.Thin.Coprod
import public Quox.Thin.Base
import public Quox.Thin.Comp
import public Quox.Thin.View
import public Quox.Thin.List
import public Quox.Thin.Cover
import Data.DPair
import Data.Nat
import Control.Function
%default total
namespace Coprod
public export
data Comps : OPE scope n scopeMask ->
OPEList scope -> OPEList n -> Type where
Nil : Comps sub [] []
(::) : Comp sub inner full ->
Comps sub inners fulls ->
Comps sub (inner :: inners) (full :: fulls)
%name Comps comps
public export
record Coprod (fulls : OPEList n) where
constructor MkCoprod
{scopeMask : Nat}
{0 sub : OPE scope n scopeMask}
inners : OPEList scope
0 comps : Comps sub inners fulls
0 cov : Cover inners
%name Coprod cop
export
0 compsLength : Comps s ts us -> length ts = length us
compsLength [] = Refl
compsLength (_ :: comps) = cong S $ compsLength comps
export
coprodNil : Coprod []
coprodNil = MkCoprod [] [] [] {sub = zero}
private
coprodHead : {n : Nat} -> (opes : OPEList (S n)) ->
Either (Cover1 opes) (All IsDrop opes)
coprodHead [] = Right []
coprodHead (ope :: opes) = case view ope of
DropV {} => case coprodHead opes of
Left cov1 => Left $ There cov1
Right drops => Right $ ItIsDrop :: drops
KeepV {} => Left Here
private
0 compsConsDrop : (opes : OPEList (S n)) ->
All IsDrop opes ->
All2 IsTail opes tails ->
Comps sub inners tails -> Comps (drop sub) inners opes
compsConsDrop [] [] [] [] = []
compsConsDrop (Drop ope Refl :: opes) (ItIsDrop :: ds) (DropT :: ts) (c :: cs) =
DropZ c :: compsConsDrop opes ds ts cs
compsConsDrop (_ :: _) [] _ _ impossible
private
coprodConsDrop : (0 opes : OPEList (S n)) ->
(0 ds : All IsDrop opes) ->
(0 ts : All2 IsTail opes tails) ->
Coprod tails -> Coprod opes
coprodConsDrop opes ds ts (MkCoprod inners comps cov) =
MkCoprod inners (compsConsDrop opes ds ts comps) cov
private
copyHeads : {m : Nat} ->
(src : OPEList (S m)) -> (tgt : OPEList n) ->
(0 eq : length src = length tgt) => OPEList (S n)
copyHeads [] [] = []
copyHeads (s :: ss) (t :: ts) =
case view s of
DropV mask ope => drop t :: copyHeads ss ts @{inj S eq}
KeepV mask ope => keep t :: copyHeads ss ts @{inj S eq}
private
0 copyHeadsComps : (eq : length outers = length inners) ->
All2 IsTail outers tails ->
Comps sub inners tails ->
Comps (keep sub) (copyHeads outers inners) outers
copyHeadsComps _ [] [] = []
copyHeadsComps eq (DropT {eq = eq2} :: ps) ((c :: cs) {full}) =
let (Refl) = eq2 in -- coverage checker weirdness
rewrite viewDrop full Refl in
KDZ c :: copyHeadsComps (inj S eq) ps cs
copyHeadsComps eq (KeepT {eq = eq2} :: ps) ((c :: cs) {full}) =
let (Refl) = eq2 in
rewrite viewKeep full Refl in
KeepZ c :: copyHeadsComps (inj S eq) ps cs
-- should be erased (coverage checker weirdness)
-- it is possibly https://github.com/idris-lang/Idris2/issues/1417 that keeps
-- happening. not 100% sure
private
cover1CopyHeads : {m : Nat} ->
(ss : OPEList (S m)) -> (ts : OPEList n) ->
(eq : length ss = length ts) ->
(cov1 : Cover1 ss) -> Cover1 (copyHeads ss ts)
cover1CopyHeads (Keep s Refl :: ss) (t :: ts) eq Here =
rewrite viewKeep s Refl in Here
cover1CopyHeads (s :: ss) (t :: ts) eq (There c) with (view s)
cover1CopyHeads (Drop {} :: ss) (t :: ts) eq (There c) | DropV {} =
There $ cover1CopyHeads ss ts (inj S eq) c
cover1CopyHeads (Keep {} :: ss) (t :: ts) eq (There c) | KeepV {} =
Here
private
copyHeadsTails : {m : Nat} ->
(ss : OPEList (S m)) -> (ts : OPEList n) ->
(eq : length ss = length ts) ->
All2 IsTail (copyHeads ss ts) ts
copyHeadsTails [] [] eq = []
copyHeadsTails (s :: ss) (t :: ts) eq with (view s)
copyHeadsTails (Drop ope Refl :: ss) (t :: ts) eq | DropV mask ope =
DropT :: copyHeadsTails ss ts (inj S eq)
copyHeadsTails (Keep ope Refl :: ss) (t :: ts) eq | KeepV mask ope =
KeepT :: copyHeadsTails ss ts (inj S eq)
private
coprodConsKeep : {n : Nat} ->
(opes : OPEList (S n)) ->
{0 tails : OPEList n} ->
(cov1 : Cover1 opes) ->
(0 ts : All2 IsTail opes tails) ->
Coprod tails -> Coprod opes
coprodConsKeep opes cov1 ts (MkCoprod inners comps cov) =
MkCoprod
(copyHeads opes inners @{all2Length ts `trans` sym (compsLength comps)})
(copyHeadsComps _ ts comps)
((cover1CopyHeads {cov1, _} :: cov) @{copyHeadsTails {}})
export
coprod : {n : Nat} -> (opes : OPEList n) -> Coprod opes
private
coprod0 : (opes : OPEList 0) -> Coprod opes
private
coprodS : {n : Nat} -> (opes : OPEList (S n)) -> Coprod opes
coprod {n = 0} opes = coprod0 opes
coprod {n = S n} opes = coprodS opes
coprod0 [] = coprodNil
coprod0 (ope :: opes) with %syntactic 0 (zeroIsStop ope) | (coprod opes)
coprod0 (Stop :: opes)
| ItIsStop | MkCoprod {sub} inners comps cov
with %syntactic 0 (zeroIsStop sub)
coprod0 (Stop :: opes)
| ItIsStop | MkCoprod {sub = Stop} inners comps cov | ItIsStop
= MkCoprod (Stop :: inners) (StopZ :: comps) []
coprodS [] = coprodNil
coprodS opes =
let hs = heads opes
Element ts tprf = tails_ opes
tcop = coprod $ assert_smaller opes ts
in
case coprodHead opes of
Left cov1 => coprodConsKeep opes cov1 tprf tcop
Right drops => coprodConsDrop opes drops tprf tcop

27
lib/Quox/Thin/Cover.idr Normal file
View File

@ -0,0 +1,27 @@
module Quox.Thin.Cover
import public Quox.Thin.Base
import public Quox.Thin.List
%default total
||| an OPE list is a cover if at least one of the OPEs has `Keep` as the head,
||| and the tails are also a cover
|||
||| in @egtbs it is a binary relation which is fine for ×ᵣ but i don't want to
||| write my AST in universe-of-syntaxes style. sorry
public export data Cover : OPEList n -> Type
||| the "`Keep` in the head" condition of a cover
public export data Cover1 : OPEList n -> Type
data Cover where
Nil : Cover opes {n = 0}
(::) : Cover1 opes -> All2 IsTail opes tails => Cover tails -> Cover opes
%name Cover cov
data Cover1 where
Here : Cover1 (Keep ope eq :: opes)
There : Cover1 opes -> Cover1 (ope :: opes)
%name Cover1 cov1
%builtin Natural Cover1

128
lib/Quox/Thin/Eqv.idr Normal file
View File

@ -0,0 +1,128 @@
module Quox.Thin.Eqv
import public Quox.Thin.Base
import public Quox.Thin.View
import Quox.NatExtra
import Syntax.PreorderReasoning
%default total
infix 6 `Eqv`
private
uip : (p, q : a = b) -> p = q
uip Refl Refl = Refl
public export
data Eqv : OPE m1 n1 mask1 -> OPE m2 n2 mask2 -> Type where
EqvStop : Eqv Stop Stop
EqvDrop : {0 p : OPE m1 n1 mask1} ->
{0 q : OPE m2 n2 mask2} ->
Eqv p q -> Eqv (Drop p eq1) (Drop q eq2)
EqvKeep : {0 p : OPE m1 n1 mask1} ->
{0 q : OPE m2 n2 mask2} ->
Eqv p q -> Eqv (Keep p eq1) (Keep q eq2)
%name Eqv eqv
export Uninhabited (Stop `Eqv` Drop p e) where uninhabited _ impossible
export Uninhabited (Stop `Eqv` Keep p e) where uninhabited _ impossible
export Uninhabited (Drop p e `Eqv` Stop) where uninhabited _ impossible
export Uninhabited (Drop p e `Eqv` Keep q f) where uninhabited _ impossible
export Uninhabited (Keep p e `Eqv` Stop) where uninhabited _ impossible
export Uninhabited (Keep p e `Eqv` Drop q f) where uninhabited _ impossible
export
Reflexive (OPE m n mask) Eqv where
reflexive {x = Stop} = EqvStop
reflexive {x = Drop {}} = EqvDrop reflexive
reflexive {x = Keep {}} = EqvKeep reflexive
export
symmetric : p `Eqv` q -> q `Eqv` p
symmetric EqvStop = EqvStop
symmetric (EqvDrop eqv) = EqvDrop $ symmetric eqv
symmetric (EqvKeep eqv) = EqvKeep $ symmetric eqv
export
transitive : p `Eqv` q -> q `Eqv` r -> p `Eqv` r
transitive EqvStop EqvStop = EqvStop
transitive (EqvDrop eqv1) (EqvDrop eqv2) = EqvDrop (transitive eqv1 eqv2)
transitive (EqvKeep eqv1) (EqvKeep eqv2) = EqvKeep (transitive eqv1 eqv2)
private
recompute' : {mask1, mask2, n1, n2 : Nat} ->
(0 p : OPE m1 n1 mask1) -> (0 q : OPE m2 n2 mask2) ->
(0 eqv : p `Eqv` q) -> p `Eqv` q
recompute' p q eqv with %syntactic (view p) | (view q)
recompute' Stop Stop _ | StopV | StopV = EqvStop
recompute' (Drop p _) (Drop q _) eqv | DropV _ p | DropV _ q =
EqvDrop $ recompute' {eqv = let EqvDrop e = eqv in e, _}
recompute' (Keep p _) (Keep q _) eqv | KeepV _ p | KeepV _ q =
EqvKeep $ recompute' {eqv = let EqvKeep e = eqv in e, _}
recompute' (Drop p _) (Keep q _) eqv | DropV _ p | KeepV _ q =
void $ absurd eqv
recompute' (Keep p _) (Drop q _) eqv | KeepV _ p | DropV _ q =
void $ absurd eqv
private
recompute : {mask1, mask2, n1, n2 : Nat} ->
{0 p : OPE m1 n1 mask1} -> {0 q : OPE m2 n2 mask2} ->
(0 _ : p `Eqv` q) -> p `Eqv` q
recompute eqv = recompute' {eqv, _}
export
eqvIndices : {0 p : OPE m1 n1 mask1} -> {0 q : OPE m2 n2 mask2} ->
p `Eqv` q -> (m1 = m2, n1 = n2, mask1 = mask2)
eqvIndices EqvStop = (Refl, Refl, Refl)
eqvIndices (EqvDrop eqv {eq1 = Refl, eq2 = Refl}) =
let (Refl, Refl, Refl) = eqvIndices eqv in (Refl, Refl, Refl)
eqvIndices (EqvKeep eqv {eq1 = Refl, eq2 = Refl}) =
let (Refl, Refl, Refl) = eqvIndices eqv in (Refl, Refl, Refl)
export
0 eqvMask : (p : OPE m1 n mask1) -> (q : OPE m2 n mask2) ->
mask1 = mask2 -> p `Eqv` q
eqvMask Stop Stop _ = EqvStop
eqvMask (Drop ope1 Refl) (Drop {mask = mm2} ope2 eq2) Refl =
EqvDrop $ eqvMask ope1 ope2 (doubleInj _ _ eq2)
eqvMask (Drop ope1 Refl) (Keep ope2 eq2) Refl =
void $ notEvenOdd _ _ eq2
eqvMask (Keep ope1 eq1) (Keep ope2 eq2) Refl =
EqvKeep $ eqvMask ope1 ope2 (doubleInj _ _ $ inj S $ trans (sym eq1) eq2)
eqvMask (Keep ope1 eq1) (Drop ope2 eq2) Refl =
void $ notEvenOdd _ _ $ trans (sym eq2) eq1
export
0 eqvEq : (p, q : OPE m n mask) -> p `Eqv` q -> p === q
eqvEq Stop Stop EqvStop = Refl
eqvEq (Drop p eq1) (Drop q eq2) (EqvDrop eqv)
with %syntactic (doubleInj _ _ $ trans (sym eq1) eq2)
_ | Refl = cong2 Drop (eqvEq p q eqv) (uip eq1 eq2)
eqvEq (Keep p eq1) (Keep q eq2) (EqvKeep eqv)
with %syntactic (doubleInj _ _ $ inj S $ trans (sym eq1) eq2)
_ | Refl = cong2 Keep (eqvEq p q eqv) (uip eq1 eq2)
export
0 eqvEq' : (p : OPE m1 n1 mask1) -> (q : OPE m2 n2 mask2) ->
p `Eqv` q -> p ~=~ q
eqvEq' p q eqv = let (Refl, Refl, Refl) = eqvIndices eqv in eqvEq p q eqv
export
0 maskEqInner : (0 ope1 : OPE m1 n mask1) -> (0 ope2 : OPE m2 n mask2) ->
mask1 = mask2 -> m1 = m2
maskEqInner Stop Stop _ = Refl
maskEqInner (Drop ope1 Refl) (Drop ope2 Refl) eq =
maskEqInner ope1 ope2 (doubleInj _ _ eq)
maskEqInner (Keep ope1 Refl) (Keep ope2 Refl) eq =
cong S $ maskEqInner ope1 ope2 $ doubleInj _ _ $ inj S eq
maskEqInner (Drop ope1 Refl) (Keep ope2 Refl) eq = void $ notEvenOdd _ _ eq
maskEqInner (Keep {mask = mask1'} ope1 eq1) (Drop {mask = mask2'} ope2 eq2) eq =
-- matching on eq1, eq2, or eq here triggers that weird coverage bug ☹
void $ notEvenOdd _ _ $ Calc $
|~ mask2' + mask2'
~~ mask2 ..<(eq2)
~~ mask1 ..<(eq)
~~ S (mask1' + mask1') ...(eq1)

127
lib/Quox/Thin/List.idr Normal file
View File

@ -0,0 +1,127 @@
module Quox.Thin.List
import public Quox.Thin.Base
import public Quox.Thin.Cons
import Data.DPair
import Data.Nat
import Control.Function
%default total
||| a list of OPEs of a given outer scope size
||| (at runtime just the masks)
public export
data OPEList : Nat -> Type where
Nil : OPEList n
(::) : {mask : Nat} -> (0 ope : OPE m n mask) -> OPEList n -> OPEList n
%name OPEList opes
public export
length : OPEList n -> Nat
length [] = 0
length (_ :: opes) = S $ length opes
public export
toList : OPEList n -> List (SomeOPE n)
toList [] = []
toList (ope :: opes) = MkOPE ope :: toList opes
public export
fromList : List (SomeOPE n) -> OPEList n
fromList [] = []
fromList (MkOPE ope :: xs) = ope :: fromList xs
public export
0 Pred : Nat -> Type
Pred n = forall m, mask. OPE m n mask -> Type
public export
0 Rel : Nat -> Nat -> Type
Rel n1 n2 = forall m1, m2, mask1, mask2.
OPE m1 n1 mask1 -> OPE m2 n2 mask2 -> Type
namespace All
public export
data All : Pred n -> OPEList n -> Type where
Nil : {0 p : Pred n} -> All p []
(::) : {0 p : Pred n} -> p ope -> All p opes -> All p (ope :: opes)
%name All.All ps, qs
namespace All2
public export
data All2 : Rel n1 n2 -> OPEList n1 -> OPEList n2 -> Type where
Nil : {0 p : Rel n1 n2} -> All2 p [] []
(::) : {0 p : Rel n1 n2} -> p a b -> All2 p as bs ->
All2 p (a :: as) (b :: bs)
%name All2.All2 ps, qs
export
0 all2Length : {p : Rel m n} -> All2 p ss ts -> length ss = length ts
all2Length [] = Refl
all2Length (p :: ps) = cong S $ all2Length ps
namespace Any
public export
data Any : Pred n -> OPEList n -> Type where
Here : {0 p : Pred n} -> p ope -> Any p (ope :: opes)
There : {0 p : Pred n} -> Any p opes -> Any p (ope :: opes)
%name Any.Any p, q
export
{0 p : Pred n} -> Uninhabited (Any p []) where uninhabited _ impossible
export
all : {0 p : Pred n} ->
(forall m. {mask : Nat} -> (0 ope : OPE m n mask) -> p ope) ->
(opes : OPEList n) -> All p opes
all f [] = []
all f (ope :: opes) = f ope :: all f opes
export
allDec : {0 p : Pred n} ->
(forall m. {mask : Nat} -> (0 ope : OPE m n mask) -> Dec (p ope)) ->
(opes : OPEList n) -> Dec (All p opes)
allDec f [] = Yes []
allDec f (ope :: opes) = case f ope of
Yes y => case allDec f opes of
Yes ys => Yes $ y :: ys
No k => No $ \(_ :: ps) => k ps
No k => No $ \(p :: _) => k p
export
anyDec : {0 p : Pred n} ->
(forall m. {mask : Nat} -> (0 ope : OPE m n mask) -> Dec (p ope)) ->
(opes : OPEList n) -> Dec (Any p opes)
anyDec f [] = No absurd
anyDec f (ope :: opes) = case f ope of
Yes y => Yes $ Here y
No nh => case anyDec f opes of
Yes y => Yes $ There y
No nt => No $ \case Here h => nh h; There t => nt t
export
unconses : {n : Nat} -> (opes : OPEList (S n)) -> All Uncons opes
unconses = all uncons
export
heads : {n : Nat} -> (opes : OPEList (S n)) -> All (Exists . IsHead) opes
heads = all head
export
tails : {n : Nat} -> (opes : OPEList (S n)) -> All Tail opes
tails = all tail
export
tails_ : {n : Nat} -> (opes : OPEList (S n)) ->
Subset (OPEList n) (All2 IsTail opes)
tails_ [] = Element [] []
tails_ (ope :: opes) = Element _ $ (tail ope).isTail :: (tails_ opes).snd
export
conses : (heads : List Bool) -> (tails : OPEList n) ->
(0 len : length heads = length tails) =>
OPEList (S n)
conses [] [] = []
conses (h :: hs) (t :: ts) = snd (cons h t) :: conses hs ts @{inj S len}

57
lib/Quox/Thin/Split.idr Normal file
View File

@ -0,0 +1,57 @@
module Quox.Thin.Split
import public Quox.Thin.Base
import public Quox.Thin.View
import public Quox.Thin.Eqv
import public Quox.Thin.Append
import public Quox.Thin.Cover
import Data.DPair
import Control.Relation
%default total
public export
record Chunks m n where
constructor MkChunks
{leftMask : Nat}
{rightMask : Nat}
0 left : OPE m (m + n) leftMask
0 right : OPE n (m + n) rightMask
{auto 0 isCover : Cover [left, right]}
%name Chunks chunks
export
chunks : (m, n : Nat) -> Chunks m n
chunks 0 0 = MkChunks Stop Stop
chunks 0 (S n) =
let MkChunks l r = chunks 0 n in
MkChunks (Drop l Refl) (Keep r Refl)
chunks (S m) n =
let MkChunks l r = chunks m n in
MkChunks (Keep l Refl) (Drop r Refl)
-- [todo] the masks here are just ((2 << m) - 1) << n and (2 << n) - 1
public export
record SplitAt m n1 n2 (ope : OPE m (n1 + n2) mask) where
constructor MkSplitAt
{leftMask, rightMask : Nat}
{0 leftScope, rightScope : Nat}
0 left : OPE leftScope n1 leftMask
0 right : OPE rightScope n2 rightMask
0 scopePrf : m = leftScope + rightScope
0 opePrf : ope `Eqv` (left `app'` right).snd
%name SplitAt split
export
splitAt : (n1 : Nat) -> {n2, mask : Nat} -> (0 ope : OPE m (n1 + n2) mask) ->
SplitAt m n1 n2 ope
splitAt 0 ope = MkSplitAt zero ope Refl reflexive
splitAt (S n1) ope with %syntactic (view ope)
splitAt (S n1) (Drop ope Refl) | DropV _ ope with %syntactic (splitAt n1 ope)
_ | MkSplitAt left right scopePrf opePrf =
MkSplitAt (Drop left Refl) right scopePrf (EqvDrop opePrf)
splitAt (S n1) (Keep ope Refl) | KeepV _ ope with %syntactic (splitAt n1 ope)
_ | MkSplitAt left right scopePrf opePrf =
MkSplitAt (Keep left Refl) right (cong S scopePrf) (EqvKeep opePrf)

179
lib/Quox/Thin/Term.idr Normal file
View File

@ -0,0 +1,179 @@
module Quox.Thin.Term
import public Quox.Thin.Base
import public Quox.Thin.Comp
import public Quox.Thin.List
import Quox.Thin.Eqv
import public Quox.Thin.Cover
import Quox.Name
import Quox.Loc
import Data.DPair
import public Data.List.Quantifiers
import Data.Vect
import Data.Singleton
import Decidable.Equality
%default total
public export
record Thinned f n where
constructor Th
{0 scope : Nat}
{scopeMask : Nat}
0 ope : OPE scope n scopeMask
term : f scope
%name Thinned s, t, u
export
(forall n. Eq (f n)) => Eq (Thinned f n) where
s == t = case decEq s.scopeMask t.scopeMask of
Yes eq => s.term == (rewrite maskEqInner s.ope t.ope eq in t.term)
No _ => False
export
(forall n. Located (f n)) => Located (Thinned f n) where
term.loc = term.term.loc
export
(forall n. Relocatable (f n)) => Relocatable (Thinned f n) where
setLoc loc = {term $= setLoc loc}
namespace Thinned
export
pure : {n : Nat} -> f n -> Thinned f n
pure term = Th id.snd term
export
join : {n : Nat} -> Thinned (Thinned f) n -> Thinned f n
join (Th ope1 (Th ope2 term)) = Th (ope1 . ope2) term
public export
record ScopedN (s : Nat) (f : Nat -> Type) (n : Nat) where
constructor S
names : Vect s BindName
{0 scope : Nat}
{mask : Nat}
0 ope : OPE scope s mask
body : f (scope + n)
export
(forall n. Eq (f n)) => Eq (ScopedN s f n) where
s1 == s2 = case decEq s1.mask s2.mask of
Yes eq =>
s1.names == s2.names &&
s1.body == (rewrite maskEqInner s1.ope s2.ope eq in s2.body)
No _ => False
export
{s : Nat} -> (forall n. Show (f n)) => Show (ScopedN s f n) where
showPrec d (S ns ope body) = showCon d "S" $
showArg ns ++ showArg (unVal $ maskToOpe ope) ++ showArg body
public export
Scoped : (Nat -> Type) -> Nat -> Type
Scoped d n = ScopedN 1 d n
(.name) : Scoped f n -> BindName
(S {names = [x], _}).name = x
export
(forall n. Located (f n)) => Located (ScopedN s f n) where
s.loc = s.body.loc
export
(forall n. Relocatable (f n)) => Relocatable (ScopedN s f n) where
setLoc loc = {body $= setLoc loc}
public export
record Thinned2 f d n where
constructor Th2
{0 dscope, tscope : Nat}
{dmask, tmask : Nat}
0 dope : OPE dscope d dmask
0 tope : OPE tscope n tmask
term : f dscope tscope
%name Thinned2 term
export
(forall d, n. Eq (f d n)) => Eq (Thinned2 f d n) where
s == t = case (decEq s.dmask t.dmask, decEq s.tmask t.tmask) of
(Yes deq, Yes teq) =>
s.term == (rewrite maskEqInner s.dope t.dope deq in
rewrite maskEqInner s.tope t.tope teq in t.term)
_ => False
export
(forall d, n. Located (f d n)) => Located (Thinned2 f d n) where
term.loc = term.term.loc
export
(forall d, n. Relocatable (f d n)) => Relocatable (Thinned2 f d n) where
setLoc loc = {term $= setLoc loc}
namespace Thinned2
export
pure : {d, n : Nat} -> f d n -> Thinned2 f d n
pure term = Th2 id.snd id.snd term
export
join : {d, n : Nat} -> Thinned2 (Thinned2 f) d n -> Thinned2 f d n
join (Th2 dope1 tope1 (Th2 dope2 tope2 term)) =
Th2 (dope1 . dope2) (tope1 . tope2) term
namespace TermList
public export
data Element : (Nat -> Nat -> Type) ->
OPE dscope d dmask -> OPE tscope n tmask -> Type where
T : f dscope tscope ->
{dmask : Nat} -> (0 dope : OPE dscope d dmask) ->
{tmask : Nat} -> (0 tope : OPE tscope n tmask) ->
Element f dope tope
%name TermList.Element s, t, u
export
elementEq : (forall d, n. Eq (f d n)) =>
Element {d, n} f dope1 tope1 -> Element {d, n} f dope2 tope2 ->
Bool
elementEq (T s dope1 tope1 {dmask = dm1, tmask = tm1})
(T t dope2 tope2 {dmask = dm2, tmask = tm2}) =
case (decEq dm1 dm2, decEq tm1 tm2) of
(Yes deq, Yes teq) =>
s == (rewrite maskEqInner dope1 dope2 deq in
rewrite maskEqInner tope1 tope2 teq in t)
_ => False
public export
data TermList : List (Nat -> Nat -> Type) ->
OPEList d -> OPEList n -> Type where
Nil : TermList [] [] []
(::) : Element f dope tope ->
TermList fs dopes topes ->
TermList (f :: fs) (dope :: dopes) (tope :: topes)
%name TermList ss, ts, us
export
termListEq : All (\f => forall d, n. Eq (f d n)) fs =>
TermList {d, n} fs dopes1 topes1 ->
TermList {d, n} fs dopes2 topes2 ->
Bool
termListEq [] [] = True
termListEq (s :: ss) (t :: ts) @{eq :: eqs} =
elementEq s t && termListEq ss ts
public export
record Subterms (fs : List (Nat -> Nat -> Type)) d n where
constructor Sub
{0 dopes : OPEList d}
{0 topes : OPEList n}
terms : TermList fs dopes topes
0 dcov : Cover dopes
0 tcov : Cover topes
%name Subterms ss, ts, us
export
All (\f => forall d, n. Eq (f d n)) fs => Eq (Subterms fs d n) where
ss == ts = ss.terms `termListEq` ts.terms

103
lib/Quox/Thin/View.idr Normal file
View File

@ -0,0 +1,103 @@
module Quox.Thin.View
import public Quox.Thin.Base
import Quox.NatExtra
import Data.Singleton
import Data.SnocVect
%default total
public export
data View : OPE m n mask -> Type where
StopV : View Stop
DropV : (mask : Nat) -> (0 ope : OPE m n mask) -> View (Drop ope Refl)
KeepV : (mask : Nat) -> (0 ope : OPE m n mask) -> View (Keep ope Refl)
%name View.View v
private
0 stopEqs : OPE m 0 mask -> (m = 0, mask = 0)
stopEqs Stop = (Refl, Refl)
private
0 fromStop : (ope : OPE 0 0 0) -> ope = Stop
fromStop Stop = Refl
private
0 fromDrop : (ope : OPE m (S n) (k + k)) ->
(inner : OPE m n k ** ope === Drop inner Refl)
fromDrop (Drop ope eq) with %syntactic (doubleInj _ _ eq)
fromDrop (Drop ope Refl) | Refl = (ope ** Refl)
fromDrop (Keep ope eq) = void $ notEvenOdd _ _ eq
private
0 fromKeep : (ope : OPE (S m) (S n) (S (k + k))) ->
(inner : OPE m n k ** ope === Keep inner Refl)
fromKeep (Drop ope eq) = void $ notEvenOdd _ _ $ sym eq
fromKeep (Keep ope eq) with %syntactic (doubleInj _ _ $ inj S eq)
fromKeep (Keep ope Refl) | Refl = (ope ** Refl)
private
0 keepIsSucc : (ope : OPE m n (S (k + k))) -> IsSucc m
keepIsSucc (Drop ope eq) = void $ notEvenOdd _ _ $ sym eq
keepIsSucc (Keep ope _) = ItIsSucc
export
view : {0 m : Nat} -> {n, mask : Nat} -> (0 ope : OPE m n mask) -> View ope
view {n = 0} ope with %syntactic 0 (fst $ stopEqs ope) | 0 (snd $ stopEqs ope)
_ | Refl | Refl = rewrite fromStop ope in StopV
view {n = S n} ope with %syntactic (half mask)
_ | HalfOdd mask' with %syntactic 0 (keepIsSucc ope)
_ | ItIsSucc with %syntactic 0 (fromKeep ope)
_ | (ope' ** eq) = rewrite eq in KeepV mask' ope'
_ | HalfEven mask' with %syntactic 0 (fromDrop ope)
_ | (ope' ** eq) = rewrite eq in DropV mask' ope'
export
appOpe : {0 m : Nat} -> (n : Nat) -> {mask : Nat} ->
(0 ope : OPE m n mask) -> Singleton m
appOpe n ope with %syntactic (view ope)
appOpe 0 Stop | StopV = Val 0
appOpe (S n) (Drop ope' _) | DropV _ ope' = appOpe n ope'
appOpe (S n) (Keep ope' _) | KeepV _ ope' = [|S $ appOpe n ope'|]
export
maskToOpe : {n, mask : Nat} -> (0 ope : OPE m n mask) -> Singleton ope
maskToOpe ope with %syntactic (view ope)
maskToOpe Stop | StopV = [|Stop|]
maskToOpe (Drop ope Refl) | DropV k ope = [|drop $ maskToOpe ope|]
maskToOpe (Keep ope Refl) | KeepV k ope = [|keep $ maskToOpe ope|]
export
0 outerInnerZero : OPE m 0 mask -> m = 0
outerInnerZero Stop = Refl
export
0 outerMaskZero : OPE m 0 mask -> mask = 0
outerMaskZero Stop = Refl
export
0 viewStop : view Stop = StopV
viewStop = Refl
export
0 viewDrop : (ope : OPE m n mask) -> (eq : mask2 = mask + mask) ->
view (Drop {mask} ope eq) = DropV mask ope
viewDrop ope eq with (view (Drop ope eq))
viewDrop ope Refl | DropV _ ope = Refl
export
0 viewKeep : (ope : OPE m n mask) -> (eq : mask2 = S (mask + mask)) ->
view (Keep {mask} ope eq) = KeepV mask ope
viewKeep ope eq with (view (Keep ope eq))
viewKeep ope Refl | KeepV _ ope = Refl
namespace SnocVect
export
select : {n, mask : Nat} -> (0 ope : OPE m n mask) ->
SnocVect n a -> SnocVect m a
select ope sx with (view ope)
select _ [<] | StopV = [<]
select _ (sx :< x) | DropV _ ope = select ope sx
select _ (sx :< x) | KeepV _ ope = select ope sx :< x

View File

@ -163,7 +163,7 @@ mutual
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
TC (CheckResult' n)
toCheckType ctx sg t ty = do
u <- expectTYPE !defs ctx ty.loc ty
u <- expectTYPE !(askAt DEFS) ctx ty.loc ty
expectEqualQ t.loc Zero sg.fst
checkTypeNoWrap ctx t (Just u)
pure $ zeroFor ctx
@ -178,7 +178,7 @@ mutual
check' ctx sg t@(Pi {}) ty = toCheckType ctx sg t ty
check' ctx sg (Lam body loc) ty = do
(qty, arg, res) <- expectPi !defs ctx ty.loc ty
(qty, arg, res) <- expectPi !(askAt DEFS) ctx ty.loc ty
-- if Ψ | Γ, x : A ⊢ σ · t ⇐ B ⊳ Σ, ρ·x
-- with ρ ≤ σπ
let qty' = sg.fst * qty
@ -189,7 +189,7 @@ mutual
check' ctx sg t@(Sig {}) ty = toCheckType ctx sg t ty
check' ctx sg (Pair fst snd loc) ty = do
(tfst, tsnd) <- expectSig !defs ctx ty.loc ty
(tfst, tsnd) <- expectSig !(askAt DEFS) ctx ty.loc ty
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁
qfst <- checkC ctx sg fst tfst
let tsnd = sub1 tsnd (Ann fst tfst fst.loc)
@ -201,7 +201,7 @@ mutual
check' ctx sg t@(Enum {}) ty = toCheckType ctx sg t ty
check' ctx sg (Tag t loc) ty = do
tags <- expectEnum !defs ctx ty.loc ty
tags <- expectEnum !(askAt DEFS) ctx ty.loc ty
-- if t ∈ ts
unless (t `elem` tags) $ throw $ TagNotIn loc t tags
-- then Ψ | Γ ⊢ σ · t ⇐ {ts} ⊳ 𝟎
@ -210,7 +210,7 @@ mutual
check' ctx sg t@(Eq {}) ty = toCheckType ctx sg t ty
check' ctx sg (DLam body loc) ty = do
(ty, l, r) <- expectEq !defs ctx ty.loc ty
(ty, l, r) <- expectEq !(askAt DEFS) ctx ty.loc ty
let ctx' = extendDim body.name ctx
ty = ty.term
body = body.term
@ -226,17 +226,17 @@ mutual
check' ctx sg t@(Nat {}) ty = toCheckType ctx sg t ty
check' ctx sg (Zero {}) ty = do
expectNat !defs ctx ty.loc ty
expectNat !(askAt DEFS) ctx ty.loc ty
pure $ zeroFor ctx
check' ctx sg (Succ n {}) ty = do
expectNat !defs ctx ty.loc ty
expectNat !(askAt DEFS) ctx ty.loc ty
checkC ctx sg n ty
check' ctx sg t@(BOX {}) ty = toCheckType ctx sg t ty
check' ctx sg (Box val loc) ty = do
(q, ty) <- expectBOX !defs ctx ty.loc ty
(q, ty) <- expectBOX !(askAt DEFS) ctx ty.loc ty
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ
valout <- checkC ctx sg val ty
-- then Ψ | Γ ⊢ σ · [s] ⇐ [π.A] ⊳ πΣ
@ -314,7 +314,7 @@ mutual
-- if Ψ | Γ ⊢ Type <: Type 𝓀
case u of
Just u => lift $ subtype e.loc ctx infres.type (TYPE u noLoc)
Nothing => ignore $ expectTYPE !defs ctx e.loc infres.type
Nothing => ignore $ expectTYPE !(askAt DEFS) ctx e.loc infres.type
-- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀
@ -333,7 +333,7 @@ mutual
infer' ctx sg (F x u loc) = do
-- if π·x : A {≔ s} in global context
g <- lookupFree x loc !defs
g <- lookupFree x loc !(askAt DEFS)
-- if σ ≤ π
expectCompatQ loc sg.fst g.qty.fst
-- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎
@ -355,7 +355,7 @@ mutual
infer' ctx sg (App fun arg loc) = do
-- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁
funres <- inferC ctx sg fun
(qty, argty, res) <- expectPi !defs ctx fun.loc funres.type
(qty, argty, res) <- expectPi !(askAt DEFS) ctx fun.loc funres.type
-- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂
argout <- checkC ctx (subjMult sg qty) arg argty
-- then Ψ | Γ ⊢ σ · f s ⇒ B[s] ⊳ Σ₁ + πΣ₂
@ -372,7 +372,7 @@ mutual
pairres <- inferC ctx sg pair
-- if Ψ | Γ, p : (x : A) × B ⊢₀ ret ⇐ Type
checkTypeC (extendTy Zero ret.name pairres.type ctx) ret.term Nothing
(tfst, tsnd) <- expectSig !defs ctx pair.loc pairres.type
(tfst, tsnd) <- expectSig !(askAt DEFS) ctx pair.loc pairres.type
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐
-- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y
-- with ρ₁, ρ₂ ≤ πσ
@ -391,7 +391,7 @@ mutual
infer' ctx sg (CaseEnum pi t ret arms loc) {d, n} = do
-- if Ψ | Γ ⊢ σ · t ⇒ {ts} ⊳ Σ₁
tres <- inferC ctx sg t
ttags <- expectEnum !defs ctx t.loc tres.type
ttags <- expectEnum !(askAt DEFS) ctx t.loc tres.type
-- if 1 ≤ π, OR there is only zero or one option
unless (length (SortedSet.toList ttags) <= 1) $ expectCompatQ loc One pi
-- if Ψ | Γ, x : {ts} ⊢₀ A ⇐ Type
@ -415,7 +415,7 @@ mutual
-- if Ψ | Γ ⊢ σ · n ⇒ ⊳ Σn
nres <- inferC ctx sg n
let nat = nres.type
expectNat !defs ctx n.loc nat
expectNat !(askAt DEFS) ctx n.loc nat
-- if Ψ | Γ, n : ⊢₀ A ⇐ Type
checkTypeC (extendTy Zero ret.name nat ctx) ret.term Nothing
-- if Ψ | Γ ⊢ σ · zer ⇐ A[0 ∷ /n] ⊳ Σz
@ -439,7 +439,7 @@ mutual
infer' ctx sg (CaseBox pi box ret body loc) = do
-- if Ψ | Γ ⊢ σ · b ⇒ [ρ.A] ⊳ Σ₁
boxres <- inferC ctx sg box
(q, ty) <- expectBOX !defs ctx box.loc boxres.type
(q, ty) <- expectBOX !(askAt DEFS) ctx box.loc boxres.type
-- if Ψ | Γ, x : [ρ.A] ⊢₀ R ⇐ Type
checkTypeC (extendTy Zero ret.name boxres.type ctx) ret.term Nothing
-- if Ψ | Γ, x : A ⊢ t ⇐ R[[x] ∷ [ρ.A/x]] ⊳ Σ₂, ς·x
@ -457,7 +457,7 @@ mutual
infer' ctx sg (DApp fun dim loc) = do
-- if Ψ | Γ ⊢ σ · f ⇒ Eq [𝑖 ⇒ A] l r ⊳ Σ
InfRes {type, qout} <- inferC ctx sg fun
ty <- fst <$> expectEq !defs ctx fun.loc type
ty <- fst <$> expectEq !(askAt DEFS) ctx fun.loc type
-- then Ψ | Γ ⊢ σ · f p ⇒ Ap/𝑖 ⊳ Σ
pure $ InfRes {type = dsub1 ty dim, qout}
@ -485,7 +485,7 @@ mutual
-- if σ = 0
expectEqualQ loc Zero sg.fst
-- if Ψ, Γ ⊢₀ e ⇒ Type u
u <- expectTYPE !defs ctx ty.loc . type =<< inferC ctx szero ty
u <- expectTYPE !(askAt DEFS) ctx ty.loc . type =<< inferC ctx szero ty
-- if Ψ, Γ ⊢₀ C ⇐ Type (non-dependent return type)
checkTypeC ctx ret Nothing
-- if Ψ, Γ' ⊢₀ A ⇐ C for each rhs A

View File

@ -69,7 +69,7 @@ namespace TContext
zeroFor ctx = Zero <$ ctx
private
extendLen : Telescope a from to -> Singleton from -> Singleton to
extendLen : Telescope a n1 n2 -> Singleton n1 -> Singleton n2
extendLen [<] x = x
extendLen (tel :< _) x = [|S $ extendLen tel x|]
@ -89,7 +89,7 @@ namespace TyContext
null ctx = null ctx.dnames && null ctx.tnames
export %inline
extendTyN : CtxExtension d from to -> TyContext d from -> TyContext d to
extendTyN : CtxExtension d n1 n2 -> TyContext d n1 -> TyContext d n2
extendTyN xss (MkTyContext {termLen, dctx, dnames, tctx, tnames, qtys}) =
let (qs, xs, ss) = unzip3 xss in
MkTyContext {
@ -172,7 +172,7 @@ namespace EqContext
null ctx = null ctx.dnames && null ctx.tnames
export %inline
extendTyN : CtxExtension 0 from to -> EqContext from -> EqContext to
extendTyN : CtxExtension 0 n1 n2 -> EqContext n1 -> EqContext n2
extendTyN xss (MkEqContext {termLen, dassign, dnames, tctx, tnames, qtys}) =
let (qs, xs, ss) = unzip3 xss in
MkEqContext {

View File

@ -16,8 +16,19 @@ modules =
Quox.Decidable,
Quox.No,
Quox.Loc,
Quox.OPE,
Quox.Pretty,
Quox.Thin.Base,
Quox.Thin.View,
Quox.Thin.Eqv,
Quox.Thin.Cons,
Quox.Thin.List,
Quox.Thin.Append,
Quox.Thin.Comp,
Quox.Thin.Cover,
Quox.Thin.Coprod,
Quox.Thin.Split,
Quox.Thin.Term,
Quox.Thin,
Quox.Syntax,
Quox.Syntax.Dim,
Quox.Syntax.DimEq,

View File

@ -1,4 +1,4 @@
collection = "nightly-230522"
collection = "nightly-230622"
[custom.all.tap]
type = "git"

View File

@ -366,3 +366,16 @@
url = {https://www.cs.bham.ac.uk/~pbl/papers/hosc05.pdf},
doi = {10.1007/s10990-006-0480-6},
}
@article{egtbs,
doi = {10.4204/eptcs.275.6},
url = {https://doi.org/10.4204%2Feptcs.275.6},
year = 2018,
month = {jul},
publisher = {Open Publishing Association},
volume = {275},
pages = {53--69},
author = {Conor McBride},
title = {Everybody's Got To Be Somewhere},
journal = {Electronic Proceedings in Theoretical Computer Science}
}

View File

@ -34,7 +34,7 @@ ToInfo Error' where
M = Eff [Except Error', DefsReader]
inj : TC a -> M a
inj act = rethrow $ mapFst TCError $ runTC !defs act
inj act = rethrow $ mapFst TCError $ runTC !(askAt DEFS) act
reflTy : Term d n