move Scoped to separate module
This commit is contained in:
parent
cf3ed604a4
commit
08fb686bf6
3 changed files with 42 additions and 31 deletions
40
lib/Quox/Scoped.idr
Normal file
40
lib/Quox/Scoped.idr
Normal file
|
@ -0,0 +1,40 @@
|
||||||
|
module Quox.Scoped
|
||||||
|
|
||||||
|
import public Quox.Var
|
||||||
|
import public Quox.Context
|
||||||
|
|
||||||
|
import Derive.Prelude
|
||||||
|
|
||||||
|
%language ElabReflection
|
||||||
|
%default total
|
||||||
|
|
||||||
|
|
||||||
|
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
|
||||||
|
|
||||||
|
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
|
||||||
|
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
|
|
@ -1,6 +1,7 @@
|
||||||
module Quox.Syntax.Term.Base
|
module Quox.Syntax.Term.Base
|
||||||
|
|
||||||
import public Quox.Var
|
import public Quox.Var
|
||||||
|
import public Quox.Scoped
|
||||||
import public Quox.Syntax.Shift
|
import public Quox.Syntax.Shift
|
||||||
import public Quox.Syntax.Subst
|
import public Quox.Syntax.Subst
|
||||||
import public Quox.Syntax.Qty
|
import public Quox.Syntax.Qty
|
||||||
|
@ -46,37 +47,6 @@ TagVal : Type
|
||||||
TagVal = String
|
TagVal = String
|
||||||
|
|
||||||
|
|
||||||
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
|
|
||||||
|
|
||||||
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
|
|
||||||
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
|
|
||||||
|
|
||||||
|
|
||||||
infixl 8 :#
|
infixl 8 :#
|
||||||
infixl 9 :@, :%
|
infixl 9 :@, :%
|
||||||
mutual
|
mutual
|
||||||
|
|
|
@ -18,6 +18,7 @@ modules =
|
||||||
Quox.No,
|
Quox.No,
|
||||||
Quox.Loc,
|
Quox.Loc,
|
||||||
Quox.Var,
|
Quox.Var,
|
||||||
|
Quox.Scoped,
|
||||||
Quox.OPE,
|
Quox.OPE,
|
||||||
Quox.Pretty,
|
Quox.Pretty,
|
||||||
Quox.Syntax,
|
Quox.Syntax,
|
||||||
|
|
Loading…
Reference in a new issue