mirror of
https://github.com/edwinb/Idris2-boot.git
synced 2024-12-28 23:33:27 +03:00
Merge remote-tracking branch 'upstream/master' into misc-fixes
This commit is contained in:
commit
1652a6be7d
@ -110,6 +110,7 @@ modules =
|
||||
TTImp.ProcessParams,
|
||||
TTImp.ProcessRecord,
|
||||
TTImp.ProcessType,
|
||||
TTImp.Reflect,
|
||||
TTImp.TTImp,
|
||||
TTImp.Unelab,
|
||||
TTImp.Utils,
|
||||
|
@ -1,6 +1,6 @@
|
||||
module Control.Monad.Identity
|
||||
|
||||
public export
|
||||
public export
|
||||
record Identity (a : Type) where
|
||||
constructor Id
|
||||
runIdentity : a
|
||||
|
@ -27,16 +27,16 @@ public export
|
||||
implementation Monad f => Applicative (StateT stateType f) where
|
||||
pure x = ST (\st => pure (x, st))
|
||||
|
||||
(ST f) <*> (ST a)
|
||||
= ST (\st =>
|
||||
(ST f) <*> (ST a)
|
||||
= ST (\st =>
|
||||
do (g, r) <- f st
|
||||
(b, t) <- a r
|
||||
pure (g b, t))
|
||||
|
||||
public export
|
||||
implementation Monad m => Monad (StateT stateType m) where
|
||||
(ST f) >>= k
|
||||
= ST (\st =>
|
||||
(ST f) >>= k
|
||||
= ST (\st =>
|
||||
do (v, st') <- f st
|
||||
let ST kv = k v
|
||||
kv st')
|
||||
@ -48,8 +48,8 @@ implementation Monad m => MonadState stateType (StateT stateType m) where
|
||||
|
||||
public export
|
||||
implementation MonadTrans (StateT stateType) where
|
||||
lift x
|
||||
= ST (\st =>
|
||||
lift x
|
||||
= ST (\st =>
|
||||
do r <- x
|
||||
pure (r, st))
|
||||
|
||||
@ -61,14 +61,14 @@ implementation (Monad f, Alternative f) => Alternative (StateT st f) where
|
||||
||| Apply a function to modify the context of this computation
|
||||
public export
|
||||
modify : MonadState stateType m => (stateType -> stateType) -> m ()
|
||||
modify f
|
||||
modify f
|
||||
= do s <- get
|
||||
put (f s)
|
||||
|
||||
||| Evaluate a function in the context held by this computation
|
||||
public export
|
||||
gets : MonadState stateType m => (stateType -> a) -> m a
|
||||
gets f
|
||||
gets f
|
||||
= do s <- get
|
||||
pure (f s)
|
||||
|
||||
|
@ -66,7 +66,7 @@ sizeInd : Sized a => {0 P : a -> Type} ->
|
||||
sizeInd step z = accInd step z (sizeAccessible z)
|
||||
|
||||
export
|
||||
sizeRec : Sized a =>
|
||||
sizeRec : Sized a =>
|
||||
(step : (x : a) -> ((y : a) -> Smaller y x -> b) -> b) ->
|
||||
(z : a) -> b
|
||||
sizeRec step z = accRec step z (sizeAccessible z)
|
||||
|
@ -75,7 +75,7 @@ bufferData buf
|
||||
unpackTo (val :: acc) (loc - 1)
|
||||
|
||||
export
|
||||
readBufferFromFile : BinaryFile -> Buffer -> (maxbytes : Int) ->
|
||||
readBufferFromFile : BinaryFile -> Buffer -> (maxbytes : Int) ->
|
||||
IO (Either FileError Buffer)
|
||||
readBufferFromFile (FHandle h) (MkBuffer buf size loc) max
|
||||
= do read <- schemeCall Int "blodwen-readbuffer" [h, buf, loc, max]
|
||||
@ -84,7 +84,7 @@ readBufferFromFile (FHandle h) (MkBuffer buf size loc) max
|
||||
else pure (Left FileReadError)
|
||||
|
||||
export
|
||||
writeBufferToFile : BinaryFile -> Buffer -> (maxbytes : Int) ->
|
||||
writeBufferToFile : BinaryFile -> Buffer -> (maxbytes : Int) ->
|
||||
IO (Either FileError Buffer)
|
||||
writeBufferToFile (FHandle h) (MkBuffer buf size loc) max
|
||||
= do let maxwrite = size - loc
|
||||
|
@ -103,7 +103,7 @@ last : {n : _} -> Fin (S n)
|
||||
last {n=Z} = FZ
|
||||
last {n=S _} = FS last
|
||||
|
||||
export total
|
||||
export total
|
||||
FSinjective : {f : Fin n} -> {f' : Fin n} -> (FS f = FS f') -> f = f'
|
||||
FSinjective Refl = Refl
|
||||
|
||||
@ -119,7 +119,7 @@ implementation Ord (Fin n) where
|
||||
public export
|
||||
natToFin : Nat -> (n : Nat) -> Maybe (Fin n)
|
||||
natToFin Z (S j) = Just FZ
|
||||
natToFin (S k) (S j)
|
||||
natToFin (S k) (S j)
|
||||
= case natToFin k j of
|
||||
Just k' => Just (FS k')
|
||||
Nothing => Nothing
|
||||
@ -157,7 +157,7 @@ restrict n val = let val' = assert_total (abs (mod val (cast (S n)))) in
|
||||
-- DecEq
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
export total
|
||||
export total
|
||||
FZNotFS : {f : Fin n} -> FZ {k = n} = FS f -> Void
|
||||
FZNotFS Refl impossible
|
||||
|
||||
|
@ -16,7 +16,7 @@ data IORef : Type -> Type where
|
||||
|
||||
export
|
||||
newIORef : a -> IO (IORef a)
|
||||
newIORef val
|
||||
newIORef val
|
||||
= do m <- primIO (prim__newIORef val)
|
||||
pure (MkRef m)
|
||||
|
||||
|
@ -9,19 +9,19 @@ lengthSuc : (xs : List a) -> (y : a) -> (ys : List a) ->
|
||||
lengthSuc [] _ _ = Refl
|
||||
lengthSuc (x :: xs) y ys = cong S (lengthSuc xs y ys)
|
||||
|
||||
lengthLT : (xs : List a) -> (ys : List a) ->
|
||||
lengthLT : (xs : List a) -> (ys : List a) ->
|
||||
LTE (length xs) (length (ys ++ xs))
|
||||
lengthLT xs [] = lteRefl
|
||||
lengthLT xs (x :: ys) = lteSuccRight (lengthLT _ _)
|
||||
|
||||
smallerLeft : (ys : List a) -> (y : a) -> (zs : List a) ->
|
||||
smallerLeft : (ys : List a) -> (y : a) -> (zs : List a) ->
|
||||
LTE (S (S (length ys))) (S (length (ys ++ (y :: zs))))
|
||||
smallerLeft [] y zs = LTESucc (LTESucc LTEZero)
|
||||
smallerLeft (z :: ys) y zs = LTESucc (smallerLeft ys _ _)
|
||||
|
||||
smallerRight : (ys : List a) -> (zs : List a) ->
|
||||
smallerRight : (ys : List a) -> (zs : List a) ->
|
||||
LTE (S (S (length zs))) (S (length (ys ++ (y :: zs))))
|
||||
smallerRight {y} ys zs = rewrite lengthSuc ys y zs in
|
||||
smallerRight {y} ys zs = rewrite lengthSuc ys y zs in
|
||||
(LTESucc (LTESucc (lengthLT _ _)))
|
||||
|
||||
||| View for splitting a list in half, non-recursively
|
||||
@ -29,17 +29,17 @@ public export
|
||||
data Split : List a -> Type where
|
||||
SplitNil : Split []
|
||||
SplitOne : (x : a) -> Split [x]
|
||||
SplitPair : (x : a) -> (xs : List a) ->
|
||||
(y : a) -> (ys : List a) ->
|
||||
SplitPair : (x : a) -> (xs : List a) ->
|
||||
(y : a) -> (ys : List a) ->
|
||||
Split (x :: xs ++ y :: ys)
|
||||
|
||||
splitHelp : (head : a) ->
|
||||
(xs : List a) ->
|
||||
(xs : List a) ->
|
||||
(counter : List a) -> Split (head :: xs)
|
||||
splitHelp head [] counter = SplitOne _
|
||||
splitHelp head (x :: xs) [] = SplitPair head [] x xs
|
||||
splitHelp head (x :: xs) [y] = SplitPair head [] x xs
|
||||
splitHelp head (x :: xs) (_ :: _ :: ys)
|
||||
splitHelp head (x :: xs) (_ :: _ :: ys)
|
||||
= case splitHelp head xs ys of
|
||||
SplitOne x => SplitPair x [] _ []
|
||||
SplitPair x' xs y' ys => SplitPair x' (x :: xs) y' ys
|
||||
@ -79,11 +79,11 @@ data SnocList : List a -> Type where
|
||||
Snoc : (x : a) -> (xs : List a) ->
|
||||
(rec : SnocList xs) -> SnocList (xs ++ [x])
|
||||
|
||||
snocListHelp : {input : _} ->
|
||||
snocListHelp : {input : _} ->
|
||||
SnocList input -> (rest : List a) -> SnocList (input ++ rest)
|
||||
snocListHelp snoc [] = rewrite appendNilRightNeutral input in snoc
|
||||
snocListHelp snoc (x :: xs)
|
||||
= rewrite appendAssociative input [x] xs in
|
||||
snocListHelp snoc (x :: xs)
|
||||
= rewrite appendAssociative input [x] xs in
|
||||
snocListHelp (Snoc x input snoc) xs
|
||||
|
||||
||| Covering function for the `SnocList` view
|
||||
|
@ -89,7 +89,7 @@ export
|
||||
isLTE : (m, n : Nat) -> Dec (LTE m n)
|
||||
isLTE Z n = Yes LTEZero
|
||||
isLTE (S k) Z = No succNotLTEzero
|
||||
isLTE (S k) (S j)
|
||||
isLTE (S k) (S j)
|
||||
= case isLTE k j of
|
||||
No contra => No (contra . fromLteSucc)
|
||||
Yes prf => Yes (LTESucc prf)
|
||||
@ -245,7 +245,7 @@ cmp (S x) (S y) with (cmp x y)
|
||||
cmp (S x) (S x) | CmpEQ = CmpEQ
|
||||
cmp (S (y + (S k))) (S y) | CmpGT k = CmpGT k
|
||||
|
||||
-- Proofs on +
|
||||
-- Proofs on
|
||||
|
||||
export
|
||||
plusZeroLeftNeutral : (right : Nat) -> 0 + right = right
|
||||
|
@ -10,7 +10,7 @@ module Data.So
|
||||
||| it may be appropriate to define a type of evidence for the property that you
|
||||
||| care about instead.
|
||||
public export
|
||||
data So : Bool -> Type where
|
||||
data So : Bool -> Type where
|
||||
Oh : So True
|
||||
|
||||
export
|
||||
|
@ -23,7 +23,7 @@ length [] = 0
|
||||
length (x::xs) = 1 + length xs
|
||||
|
||||
||| Show that the length function on vectors in fact calculates the length
|
||||
private
|
||||
private
|
||||
lengthCorrect : (len : Nat) -> (xs : Vect len elem) -> length xs = len
|
||||
lengthCorrect Z [] = Refl
|
||||
lengthCorrect (S n) (x :: xs) = rewrite lengthCorrect n xs in Refl
|
||||
@ -295,7 +295,7 @@ public export
|
||||
unzip3 : (xs : Vect n (a, b, c)) -> (Vect n a, Vect n b, Vect n c)
|
||||
unzip3 [] = ([], [], [])
|
||||
unzip3 ((l,c,r)::xs) with (unzip3 xs)
|
||||
unzip3 ((l,c,r)::xs) | (lefts, centers, rights)
|
||||
unzip3 ((l,c,r)::xs) | (lefts, centers, rights)
|
||||
= (l::lefts, c::centers, r::rights)
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
@ -315,7 +315,7 @@ implementation (Eq elem) => Eq (Vect len elem) where
|
||||
public export
|
||||
implementation Ord elem => Ord (Vect len elem) where
|
||||
compare [] [] = EQ
|
||||
compare (x::xs) (y::ys)
|
||||
compare (x::xs) (y::ys)
|
||||
= case compare x y of
|
||||
EQ => compare xs ys
|
||||
x => x
|
||||
@ -528,7 +528,7 @@ findIndex p (x :: xs) = if p x then Just FZ else map FS (findIndex p xs)
|
||||
public export
|
||||
findIndices : (elem -> Bool) -> Vect m elem -> List (Fin m)
|
||||
findIndices p [] = []
|
||||
findIndices p (x :: xs)
|
||||
findIndices p (x :: xs)
|
||||
= let is = map FS $ findIndices p xs in
|
||||
if p x then FZ :: is else is
|
||||
|
||||
@ -787,7 +787,7 @@ transpose (x :: xs) = zipWith (::) x (transpose xs) -- = [| x :: xs |]
|
||||
--------------------------------------------------------------------------------
|
||||
-- Applicative/Monad/Traversable
|
||||
--------------------------------------------------------------------------------
|
||||
-- These only work if the length is known at run time!
|
||||
-- These only work if the length is known at run time!
|
||||
|
||||
implementation {k : Nat} -> Applicative (Vect k) where
|
||||
pure = replicate _
|
||||
@ -849,7 +849,7 @@ replaceByElem (x::xs) Here y = y :: xs
|
||||
replaceByElem (x::xs) (There xinxs) y = x :: replaceByElem xs xinxs y
|
||||
|
||||
public export
|
||||
mapElem : {0 xs : Vect k t} -> {0 f : t -> u} ->
|
||||
mapElem : {0 xs : Vect k t} -> {0 f : t -> u} ->
|
||||
Elem x xs -> Elem (f x) (map f xs)
|
||||
mapElem Here = Here
|
||||
mapElem (There e) = There (mapElem e)
|
||||
|
@ -18,12 +18,12 @@ interface DecEq t where
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
||| The negation of equality is symmetric (follows from symmetry of equality)
|
||||
export total
|
||||
export total
|
||||
negEqSym : forall a, b . (a = b -> Void) -> (b = a -> Void)
|
||||
negEqSym p h = p (sym h)
|
||||
|
||||
||| Everything is decidably equal to itself
|
||||
export total
|
||||
export total
|
||||
decEqSelfIsYes : DecEq a => {x : a} -> decEq x x = Yes Refl
|
||||
decEqSelfIsYes {x} with (decEq x x)
|
||||
decEqSelfIsYes {x} | Yes Refl = Refl
|
||||
@ -78,7 +78,7 @@ implementation (DecEq t) => DecEq (Maybe t) where
|
||||
decEq Nothing Nothing = Yes Refl
|
||||
decEq (Just x') (Just y') with (decEq x' y')
|
||||
decEq (Just x') (Just y') | Yes p = Yes $ cong Just p
|
||||
decEq (Just x') (Just y') | No p
|
||||
decEq (Just x') (Just y') | No p
|
||||
= No $ \h : Just x' = Just y' => p $ justInjective h
|
||||
decEq Nothing (Just _) = No nothingNotJust
|
||||
decEq (Just _) Nothing = No (negEqSym nothingNotJust)
|
||||
@ -86,7 +86,7 @@ implementation (DecEq t) => DecEq (Maybe t) where
|
||||
-- TODO: Other prelude data types
|
||||
|
||||
-- For the primitives, we have to cheat because we don't have access to their
|
||||
-- internal implementations. We use believe_me for the inequality proofs
|
||||
-- internal implementations. We use believe_me for the inequality proofs
|
||||
-- because we don't them to reduce (and they should never be needed anyway...)
|
||||
-- A postulate would be better, but erasure analysis may think they're needed
|
||||
-- for computation in a higher order setting.
|
||||
@ -97,7 +97,7 @@ implementation (DecEq t) => DecEq (Maybe t) where
|
||||
export
|
||||
implementation DecEq Int where
|
||||
decEq x y = case x == y of -- Blocks if x or y not concrete
|
||||
True => Yes primitiveEq
|
||||
True => Yes primitiveEq
|
||||
False => No primitiveNotEq
|
||||
where primitiveEq : forall x, y . x = y
|
||||
primitiveEq = believe_me (Refl {x})
|
||||
@ -110,7 +110,7 @@ implementation DecEq Int where
|
||||
export
|
||||
implementation DecEq Char where
|
||||
decEq x y = case x == y of -- Blocks if x or y not concrete
|
||||
True => Yes primitiveEq
|
||||
True => Yes primitiveEq
|
||||
False => No primitiveNotEq
|
||||
where primitiveEq : forall x, y . x = y
|
||||
primitiveEq = believe_me (Refl {x})
|
||||
@ -123,7 +123,7 @@ implementation DecEq Char where
|
||||
export
|
||||
implementation DecEq Integer where
|
||||
decEq x y = case x == y of -- Blocks if x or y not concrete
|
||||
True => Yes primitiveEq
|
||||
True => Yes primitiveEq
|
||||
False => No primitiveNotEq
|
||||
where primitiveEq : forall x, y . x = y
|
||||
primitiveEq = believe_me (Refl {x})
|
||||
@ -136,7 +136,7 @@ implementation DecEq Integer where
|
||||
export
|
||||
implementation DecEq String where
|
||||
decEq x y = case x == y of -- Blocks if x or y not concrete
|
||||
True => Yes primitiveEq
|
||||
True => Yes primitiveEq
|
||||
False => No primitiveNotEq
|
||||
where primitiveEq : forall x, y . x = y
|
||||
primitiveEq = believe_me (Refl {x})
|
||||
|
@ -21,7 +21,7 @@ data NameType : Type where
|
||||
TyCon : (tag : Int) -> (arity : Nat) -> NameType
|
||||
|
||||
public export
|
||||
data Constant
|
||||
data Constant
|
||||
= I Int
|
||||
| BI Integer
|
||||
| Str String
|
||||
@ -58,14 +58,14 @@ data LazyReason = LInf | LLazy | LUnknown
|
||||
-- Type checked terms in the core TT
|
||||
public export
|
||||
data TT : List Name -> Type where
|
||||
Local : FC -> (idx : Nat) -> (n : Name) ->
|
||||
Local : FC -> (idx : Nat) -> (n : Name) ->
|
||||
(0 prf : IsVar name idx vars) -> TT vars
|
||||
Ref : FC -> NameType -> Name -> TT vars
|
||||
Pi : FC -> Count -> PiInfo ->
|
||||
(x : Name) -> (argTy : TT vars) -> (retTy : TT (x :: vars)) ->
|
||||
Pi : FC -> Count -> PiInfo ->
|
||||
(x : Name) -> (argTy : TT vars) -> (retTy : TT (x :: vars)) ->
|
||||
TT vars
|
||||
Lam : FC -> Count -> PiInfo ->
|
||||
(x : Name) -> (argTy : TT vars) -> (scope : TT (x :: vars)) ->
|
||||
Lam : FC -> Count -> PiInfo ->
|
||||
(x : Name) -> (argTy : TT vars) -> (scope : TT (x :: vars)) ->
|
||||
TT vars
|
||||
App : FC -> TT vars -> TT vars -> TT vars
|
||||
TDelayed : FC -> LazyReason -> TT vars -> TT vars
|
||||
@ -75,3 +75,143 @@ data TT : List Name -> Type where
|
||||
Erased : FC -> TT vars
|
||||
TType : FC -> TT vars
|
||||
|
||||
public export
|
||||
data Visibility = Private | Export | Public
|
||||
|
||||
-- Unchecked terms and declarations in the intermediate language
|
||||
mutual
|
||||
public export
|
||||
data BindMode = PI Count | PATTERN | NONE
|
||||
|
||||
-- For as patterns matching linear arguments, select which side is
|
||||
-- consumed
|
||||
public export
|
||||
data UseSide = UseLeft | UseRight
|
||||
|
||||
data TTImp : Type where
|
||||
IVar : FC -> Name -> TTImp
|
||||
IPi : FC -> Count -> PiInfo -> Maybe Name ->
|
||||
(argTy : TTImp) -> (retTy : TTImp) -> TTImp
|
||||
ILam : FC -> Count -> PiInfo -> Maybe Name ->
|
||||
(argTy : TTImp) -> (lamTy : TTImp) -> TTImp
|
||||
ILet : FC -> Count -> Name ->
|
||||
(nTy : TTImp) -> (nVal : TTImp) ->
|
||||
(scope : TTImp) -> TTImp
|
||||
ICase : FC -> TTImp -> (ty : TTImp) ->
|
||||
List Clause -> TTImp
|
||||
ILocal : FC -> List Decl -> TTImp -> TTImp
|
||||
IUpdate : FC -> List IFieldUpdate -> TTImp -> TTImp
|
||||
|
||||
IApp : FC -> TTImp -> TTImp -> TTImp
|
||||
IImplicitApp : FC -> TTImp -> Maybe Name -> TTImp -> TTImp
|
||||
IWithApp : FC -> TTImp -> TTImp -> TTImp
|
||||
|
||||
ISearch : FC -> (depth : Nat) -> TTImp
|
||||
IAlternative : FC -> AltType -> List TTImp -> TTImp
|
||||
IRewrite : FC -> TTImp -> TTImp -> TTImp
|
||||
|
||||
-- Any implicit bindings in the scope should be bound here, using
|
||||
-- the given binder
|
||||
IBindHere : FC -> BindMode -> TTImp -> TTImp
|
||||
-- A name which should be implicitly bound
|
||||
IBindVar : FC -> String -> TTImp
|
||||
-- An 'as' pattern, valid on the LHS of a clause only
|
||||
IAs : FC -> UseSide -> Name -> TTImp -> TTImp
|
||||
-- A 'dot' pattern, i.e. one which must also have the given value
|
||||
-- by unification
|
||||
IMustUnify : FC -> (reason : String) -> TTImp -> TTImp
|
||||
|
||||
-- Laziness annotations
|
||||
IDelayed : FC -> LazyReason -> TTImp -> TTImp -- the type
|
||||
IDelay : FC -> TTImp -> TTImp -- delay constructor
|
||||
IForce : FC -> TTImp -> TTImp
|
||||
|
||||
IPrimVal : FC -> (c : Constant) -> TTImp
|
||||
IType : FC -> TTImp
|
||||
IHole : FC -> String -> TTImp
|
||||
|
||||
-- An implicit value, solved by unification, but which will also be
|
||||
-- bound (either as a pattern variable or a type variable) if unsolved
|
||||
-- at the end of elaborator
|
||||
Implicit : FC -> (bindIfUnsolved : Bool) -> TTImp
|
||||
|
||||
public export
|
||||
data IFieldUpdate : Type where
|
||||
ISetField : (path : List String) -> TTImp -> IFieldUpdate
|
||||
ISetFieldApp : (path : List String) -> TTImp -> IFieldUpdate
|
||||
|
||||
public export
|
||||
data AltType : Type where
|
||||
FirstSuccess : AltType
|
||||
Unique : AltType
|
||||
UniqueDefault : TTImp -> AltType
|
||||
|
||||
public export
|
||||
data FnOpt : Type where
|
||||
Inline : FnOpt
|
||||
-- Flag means the hint is a direct hint, not a function which might
|
||||
-- find the result (e.g. chasing parent interface dictionaries)
|
||||
Hint : Bool -> FnOpt
|
||||
-- Flag means to use as a default if all else fails
|
||||
GlobalHint : Bool -> FnOpt
|
||||
ExternFn : FnOpt
|
||||
-- Defined externally, list calling conventions
|
||||
ForeignFn : List String -> FnOpt
|
||||
-- assume safe to cancel arguments in unification
|
||||
Invertible : FnOpt
|
||||
Total : FnOpt
|
||||
Covering : FnOpt
|
||||
PartialOK : FnOpt
|
||||
|
||||
public export
|
||||
data ITy : Type where
|
||||
MkTy : FC -> (n : Name) -> (ty : TTImp) -> ITy
|
||||
|
||||
public export
|
||||
data DataOpt : Type where
|
||||
SearchBy : List Name -> DataOpt -- determining arguments
|
||||
NoHints : DataOpt -- Don't generate search hints for constructors
|
||||
|
||||
public export
|
||||
data Data : Type where
|
||||
MkData : FC -> (n : Name) -> (tycon : TTImp) ->
|
||||
(opts : List DataOpt) ->
|
||||
(datacons : List ITy) -> Data
|
||||
MkLater : FC -> (n : Name) -> (tycon : TTImp) -> Data
|
||||
|
||||
public export
|
||||
data IField : Type where
|
||||
MkIField : FC -> Count -> PiInfo -> Name -> TTImp ->
|
||||
IField
|
||||
|
||||
public export
|
||||
data Record : Type where
|
||||
MkRecord : FC -> (n : Name) ->
|
||||
(params : List (Name, TTImp)) ->
|
||||
(conName : Maybe Name) ->
|
||||
(fields : List IField) ->
|
||||
Record
|
||||
|
||||
public export
|
||||
data Clause : Type where
|
||||
PatClause : FC -> (lhs : TTImp) -> (rhs : TTImp) -> Clause
|
||||
WithClause : FC -> (lhs : TTImp) -> (wval : TTImp) ->
|
||||
List Clause -> Clause
|
||||
ImpossibleClause : FC -> (lhs : TTImp) -> Clause
|
||||
|
||||
public export
|
||||
data Decl : Type where
|
||||
IClaim : FC -> Count -> Visibility -> List FnOpt ->
|
||||
ITy -> Decl
|
||||
IData : FC -> Visibility -> Data -> Decl
|
||||
IDef : FC -> Name -> List Clause -> Decl
|
||||
IParameters : FC -> List (Name, TTImp) ->
|
||||
List Decl -> Decl
|
||||
IRecord : FC -> Visibility -> Record -> Decl
|
||||
INamespace : FC ->
|
||||
(nested : Bool) ->
|
||||
-- ^ if True, parent namespaces in the same file can also
|
||||
-- look inside and see private/export names in full
|
||||
List String -> List Decl -> Decl
|
||||
ITransform : FC -> TTImp -> TTImp -> Decl
|
||||
ILog : Nat -> Decl
|
||||
|
@ -82,13 +82,13 @@ stderr = FHandle prim__stderr
|
||||
|
||||
export
|
||||
openFile : String -> Mode -> IO (Either FileError File)
|
||||
openFile f m
|
||||
openFile f m
|
||||
= do res <- primIO (prim__open f (modeStr m) 0)
|
||||
fpure (map FHandle res)
|
||||
|
||||
export
|
||||
openBinaryFile : String -> Mode -> IO (Either FileError BinaryFile)
|
||||
openBinaryFile f m
|
||||
openBinaryFile f m
|
||||
= do res <- primIO (prim__open f (modeStr m) 1)
|
||||
fpure (map FHandle res)
|
||||
|
||||
@ -98,7 +98,7 @@ closeFile (FHandle f) = primIO (prim__close f)
|
||||
|
||||
export
|
||||
fGetLine : (h : File) -> IO (Either FileError String)
|
||||
fGetLine (FHandle f)
|
||||
fGetLine (FHandle f)
|
||||
= do res <- primIO (prim__readLine f)
|
||||
fpure res
|
||||
|
||||
@ -110,14 +110,14 @@ fPutStr (FHandle f) str
|
||||
|
||||
export
|
||||
fPutStrLn : (h : File) -> String -> IO (Either FileError ())
|
||||
fPutStrLn f str = fPutStr f (str ++ "\n")
|
||||
fPutStrLn f str = fPutStr f (str ++ "\n")
|
||||
|
||||
export
|
||||
fEOF : (h : File) -> IO Bool
|
||||
fEOF (FHandle f)
|
||||
= do res <- primIO (prim__eof f)
|
||||
pure (res /= 0)
|
||||
|
||||
|
||||
export
|
||||
readFile : String -> IO (Either FileError String)
|
||||
readFile file
|
||||
|
@ -17,7 +17,7 @@ replWith acc prompt fn
|
||||
then pure ()
|
||||
else do x <- getLine
|
||||
case fn acc x of
|
||||
Just (out, acc') =>
|
||||
Just (out, acc') =>
|
||||
do putStr out
|
||||
replWith acc' prompt fn
|
||||
Nothing => pure ()
|
||||
|
@ -69,7 +69,7 @@ public export
|
||||
|
||||
%inline
|
||||
public export
|
||||
rewrite__impl : {0 x, y : a} -> (0 p : _) ->
|
||||
rewrite__impl : {0 x, y : a} -> (0 p : _) ->
|
||||
(0 rule : x = y) -> (1 val : p y) -> p x
|
||||
rewrite__impl p Refl prf = prf
|
||||
|
||||
|
@ -10,21 +10,30 @@ export
|
||||
data IO : Type -> Type where
|
||||
MkIO : (1 fn : (1 x : %World) -> IORes a) -> IO a
|
||||
|
||||
public export
|
||||
PrimIO : Type -> Type
|
||||
PrimIO a = (1 x : %World) -> IORes a
|
||||
|
||||
export
|
||||
prim_io_pure : a -> PrimIO a
|
||||
prim_io_pure x = \w => MkIORes x w
|
||||
|
||||
export
|
||||
io_pure : a -> IO a
|
||||
io_pure x = MkIO (\w => MkIORes x w)
|
||||
|
||||
export
|
||||
prim_io_bind : (1 act : PrimIO a) -> (1 k : a -> PrimIO b) -> PrimIO b
|
||||
prim_io_bind fn k w
|
||||
= let MkIORes x' w' = fn w in k x' w'
|
||||
|
||||
export
|
||||
io_bind : (1 act : IO a) -> (1 k : a -> IO b) -> IO b
|
||||
io_bind (MkIO fn)
|
||||
= \k => MkIO (\w => let MkIORes x' w' = fn w
|
||||
= \k => MkIO (\w => let MkIORes x' w' = fn w
|
||||
MkIO res = k x' in
|
||||
res w')
|
||||
|
||||
public export
|
||||
PrimIO : Type -> Type
|
||||
PrimIO a = (1 x : %World) -> IORes a
|
||||
|
||||
%extern prim__putStr : String -> (1 x : %World) -> IORes ()
|
||||
%extern prim__getStr : (1 x : %World) -> IORes String
|
||||
|
||||
@ -46,9 +55,9 @@ data FArgList : Type where
|
||||
Nil : FArgList
|
||||
(::) : {a : Type} -> (1 arg : a) -> (1 args : FArgList) -> FArgList
|
||||
|
||||
%extern prim__cCall : (ret : Type) -> String -> (1 args : FArgList) ->
|
||||
%extern prim__cCall : (ret : Type) -> String -> (1 args : FArgList) ->
|
||||
(1 x : %World) -> IORes ret
|
||||
%extern prim__schemeCall : (ret : Type) -> String -> (1 args : FArgList) ->
|
||||
%extern prim__schemeCall : (ret : Type) -> String -> (1 args : FArgList) ->
|
||||
(1 x : %World) -> IORes ret
|
||||
|
||||
export %inline
|
||||
@ -56,17 +65,17 @@ primIO : (1 fn : (1 x : %World) -> IORes a) -> IO a
|
||||
primIO op = MkIO op
|
||||
|
||||
export %inline
|
||||
toPrim : IO a -> PrimIO a
|
||||
toPrim : (1 act : IO a) -> PrimIO a
|
||||
toPrim (MkIO fn) = fn
|
||||
|
||||
export %inline
|
||||
schemeCall : (ret : Type) -> String -> (1 args : FArgList) -> IO ret
|
||||
schemeCall ret fn args = primIO (prim__schemeCall ret fn args)
|
||||
|
||||
|
||||
export %inline
|
||||
cCall : (ret : Type) -> String -> FArgList -> IO ret
|
||||
cCall ret fn args = primIO (prim__cCall ret fn args)
|
||||
|
||||
|
||||
export
|
||||
putStr : String -> IO ()
|
||||
putStr str = primIO (prim__putStr str)
|
||||
@ -83,6 +92,10 @@ export
|
||||
fork : (1 prog : IO ()) -> IO ThreadID
|
||||
fork (MkIO act) = schemeCall ThreadID "blodwen-thread" [act]
|
||||
|
||||
export
|
||||
prim_fork : (1 prog : PrimIO ()) -> PrimIO ThreadID
|
||||
prim_fork act w = prim__schemeCall ThreadID "blodwen-thread" [act] w
|
||||
|
||||
unsafeCreateWorld : (1 f : (1 x : %World) -> a) -> a
|
||||
unsafeCreateWorld f = f %MkWorld
|
||||
|
||||
|
@ -50,7 +50,7 @@ getAllDesc (n :: rest) ns defs
|
||||
Nothing =>
|
||||
case !(lookupCtxtExact n (gamma defs)) of
|
||||
Nothing => getAllDesc rest ns defs
|
||||
Just def =>
|
||||
Just def =>
|
||||
let refs = refersTo def in
|
||||
getAllDesc (rest ++ keys refs) (insert n () ns) defs
|
||||
|
||||
@ -76,7 +76,7 @@ natHackNames
|
||||
-- Find all the names which need compiling, from a given expression, and compile
|
||||
-- them to CExp form (and update that in the Defs)
|
||||
export
|
||||
findUsedNames : {auto c : Ref Ctxt Defs} -> Term vars ->
|
||||
findUsedNames : {auto c : Ref Ctxt Defs} -> Term vars ->
|
||||
Core (List Name, NameTags)
|
||||
findUsedNames tm
|
||||
= do defs <- get Ctxt
|
||||
@ -89,7 +89,7 @@ findUsedNames tm
|
||||
-- Use '1' for '->' constructor
|
||||
let tyconInit = insert (UN "->") 1 $
|
||||
insert (UN "Type") 2 $
|
||||
primTags 3 empty
|
||||
primTags 3 empty
|
||||
[IntType, IntegerType, StringType,
|
||||
CharType, DoubleType, WorldType]
|
||||
tycontags <- mkNameTags defs tyconInit 100 cns
|
||||
@ -136,8 +136,8 @@ parseCC "" = Nothing
|
||||
parseCC str
|
||||
= case span (/= ':') str of
|
||||
(target, "") => Just (trim target, [])
|
||||
(target, opts) => Just (trim target,
|
||||
map trim (getOpts
|
||||
(target, opts) => Just (trim target,
|
||||
map trim (getOpts
|
||||
(assert_total (strTail opts))))
|
||||
where
|
||||
getOpts : String -> List String
|
||||
|
@ -175,7 +175,7 @@ mutual
|
||||
toCExpTm tags n (TDelayed fc _ _) = pure $ CErased fc
|
||||
toCExpTm tags n (TDelay fc _ _ arg)
|
||||
= pure (CDelay fc !(toCExp tags n arg))
|
||||
toCExpTm tags n (TForce fc arg)
|
||||
toCExpTm tags n (TForce fc _ arg)
|
||||
= pure (CForce fc !(toCExp tags n arg))
|
||||
toCExpTm tags n (PrimVal fc c)
|
||||
= let t = constTag c in
|
||||
|
@ -30,19 +30,19 @@ checkLengthMatch : (xs : List a) -> (ys : List b) -> Maybe (LengthMatch xs ys)
|
||||
checkLengthMatch [] [] = Just NilMatch
|
||||
checkLengthMatch [] (x :: xs) = Nothing
|
||||
checkLengthMatch (x :: xs) [] = Nothing
|
||||
checkLengthMatch (x :: xs) (y :: ys)
|
||||
checkLengthMatch (x :: xs) (y :: ys)
|
||||
= Just (ConsMatch !(checkLengthMatch xs ys))
|
||||
|
||||
extend : EEnv free vars -> (args : List (CExp free)) -> (args' : List Name) ->
|
||||
LengthMatch args args' -> EEnv free (args' ++ vars)
|
||||
extend env [] [] NilMatch = env
|
||||
extend env (a :: xs) (n :: ns) (ConsMatch w)
|
||||
extend env (a :: xs) (n :: ns) (ConsMatch w)
|
||||
= a :: extend env xs ns w
|
||||
|
||||
extendLoc : FC -> EEnv free vars -> (args' : List Name) ->
|
||||
extendLoc : FC -> EEnv free vars -> (args' : List Name) ->
|
||||
EEnv (args' ++ free) (args' ++ vars)
|
||||
extendLoc fc env [] = env
|
||||
extendLoc fc env (n :: ns)
|
||||
extendLoc fc env (n :: ns)
|
||||
= CLocal fc First :: weakenEnv (extendLoc fc env ns)
|
||||
|
||||
Stack : List Name -> Type
|
||||
@ -61,9 +61,9 @@ getArity (MkCon _ arity) = arity
|
||||
getArity (MkForeign _ args _) = length args
|
||||
getArity (MkError _) = 0
|
||||
|
||||
takeFromStack : EEnv free vars -> Stack free -> (args : List Name) ->
|
||||
takeFromStack : EEnv free vars -> Stack free -> (args : List Name) ->
|
||||
Maybe (EEnv free (args ++ vars), Stack free)
|
||||
takeFromStack env (e :: es) (a :: as)
|
||||
takeFromStack env (e :: es) (a :: as)
|
||||
= do (env', stk') <- takeFromStack env es as
|
||||
pure (e :: env', stk')
|
||||
takeFromStack env stk [] = pure (env, stk)
|
||||
@ -71,97 +71,97 @@ takeFromStack env [] args = Nothing
|
||||
|
||||
thinAll : (ns : List Name) -> CExp (outer ++ inner) -> CExp (outer ++ ns ++ inner)
|
||||
thinAll [] exp = exp
|
||||
thinAll {outer} {inner} (n :: ns) exp
|
||||
thinAll {outer} {inner} (n :: ns) exp
|
||||
= thin {outer} {inner = ns ++ inner} n (thinAll ns exp)
|
||||
|
||||
mutual
|
||||
evalLocal : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> List Name -> Stack free ->
|
||||
EEnv free vars ->
|
||||
EEnv free vars ->
|
||||
{idx : Nat} -> .(IsVar x idx (vars ++ free)) ->
|
||||
Core (CExp free)
|
||||
evalLocal {vars = []} fc rec stk env p
|
||||
evalLocal {vars = []} fc rec stk env p
|
||||
= pure $ unload stk (CLocal fc p)
|
||||
evalLocal {vars = x :: xs} fc rec stk (v :: env) First
|
||||
= eval rec env stk (weakenNs xs v)
|
||||
evalLocal {vars = x :: xs} fc rec stk (_ :: env) (Later p)
|
||||
evalLocal {vars = x :: xs} fc rec stk (_ :: env) (Later p)
|
||||
= evalLocal fc rec stk env p
|
||||
|
||||
tryApply : {auto c : Ref Ctxt Defs} ->
|
||||
List Name -> Stack free -> EEnv free vars -> CDef ->
|
||||
List Name -> Stack free -> EEnv free vars -> CDef ->
|
||||
Core (Maybe (CExp free))
|
||||
tryApply {free} {vars} rec stk env (MkFun args exp)
|
||||
= do let Just (env', stk') = takeFromStack env stk args
|
||||
| Nothing => pure Nothing
|
||||
res <- eval rec env' stk'
|
||||
(rewrite sym (appendAssociative args vars free) in
|
||||
res <- eval rec env' stk'
|
||||
(rewrite sym (appendAssociative args vars free) in
|
||||
embed {vars = vars ++ free} exp)
|
||||
pure (Just res)
|
||||
tryApply rec stk env _ = pure Nothing
|
||||
|
||||
eval : {auto c : Ref Ctxt Defs} ->
|
||||
List Name -> EEnv free vars -> Stack free -> CExp (vars ++ free) ->
|
||||
List Name -> EEnv free vars -> Stack free -> CExp (vars ++ free) ->
|
||||
Core (CExp free)
|
||||
eval rec env stk (CLocal fc p) = evalLocal fc rec stk env p
|
||||
eval rec env stk (CRef fc n)
|
||||
eval rec env stk (CRef fc n)
|
||||
= do defs <- get Ctxt
|
||||
Just gdef <- lookupCtxtExact n (gamma defs)
|
||||
| Nothing => pure (unload stk (CRef fc n))
|
||||
let Just def = compexpr gdef
|
||||
| Nothing => pure (unload stk (CRef fc n))
|
||||
let arity = getArity def
|
||||
let arity = getArity def
|
||||
if (Inline `elem` flags gdef) && (not (n `elem` rec))
|
||||
then do ap <- tryApply (n :: rec) stk env def
|
||||
pure $ maybe (unloadApp arity stk (CRef fc n)) id ap
|
||||
else pure $ unloadApp arity stk (CRef fc n)
|
||||
eval {vars} {free} rec env [] (CLam fc x sc)
|
||||
= do let thinsc = thin x {outer = x :: vars} {inner = free} sc
|
||||
eval {vars} {free} rec env [] (CLam fc x sc)
|
||||
= do let thinsc = thin x {outer = x :: vars} {inner = free} sc
|
||||
sc' <- eval rec (CLocal fc First :: weakenEnv env) [] thinsc
|
||||
pure $ CLam fc x sc'
|
||||
eval rec env (e :: stk) (CLam fc x sc) = eval rec (e :: env) stk sc
|
||||
eval {vars} {free} rec env stk (CLet fc x val sc)
|
||||
= do let thinsc = thin x {outer = x :: vars} {inner = free} sc
|
||||
eval {vars} {free} rec env stk (CLet fc x val sc)
|
||||
= do let thinsc = thin x {outer = x :: vars} {inner = free} sc
|
||||
sc' <- eval rec (CLocal fc First :: weakenEnv env) [] thinsc
|
||||
pure $ CLet fc x !(eval rec env [] val) sc'
|
||||
eval rec env stk (CApp fc f args)
|
||||
eval rec env stk (CApp fc f args)
|
||||
= eval rec env (!(traverse (eval rec env []) args) ++ stk) f
|
||||
eval rec env stk (CCon fc n t args)
|
||||
eval rec env stk (CCon fc n t args)
|
||||
= pure $ unload stk $ CCon fc n t !(traverse (eval rec env []) args)
|
||||
eval rec env stk (COp fc p args)
|
||||
eval rec env stk (COp fc p args)
|
||||
= pure $ unload stk $ COp fc p !(mapV (eval rec env []) args)
|
||||
where
|
||||
mapV : (a -> Core b) -> Vect n a -> Core (Vect n b)
|
||||
mapV f [] = pure []
|
||||
mapV f (x :: xs) = pure $ !(f x) :: !(mapV f xs)
|
||||
eval rec env stk (CExtPrim fc p args)
|
||||
eval rec env stk (CExtPrim fc p args)
|
||||
= pure $ unload stk $ CExtPrim fc p !(traverse (eval rec env []) args)
|
||||
eval rec env stk (CForce fc e)
|
||||
= case !(eval rec env [] e) of
|
||||
CDelay _ e' => eval rec [] stk e'
|
||||
res => pure $ unload stk (CForce fc res)
|
||||
eval rec env stk (CDelay fc e)
|
||||
eval rec env stk (CDelay fc e)
|
||||
= pure $ unload stk (CDelay fc !(eval rec env [] e))
|
||||
eval rec env stk (CConCase fc sc alts def)
|
||||
eval rec env stk (CConCase fc sc alts def)
|
||||
= do sc' <- eval rec env [] sc
|
||||
case !(pickAlt rec env stk sc' alts def) of
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do def' <- case def of
|
||||
Nothing => pure Nothing
|
||||
Just d => pure (Just !(eval rec env stk d))
|
||||
pure $
|
||||
CConCase fc sc'
|
||||
CConCase fc sc'
|
||||
!(traverse (evalAlt fc rec env stk) alts)
|
||||
def'
|
||||
Just val => pure val
|
||||
eval rec env stk (CConstCase fc sc alts def)
|
||||
eval rec env stk (CConstCase fc sc alts def)
|
||||
= do sc' <- eval rec env [] sc
|
||||
case !(pickConstAlt rec env stk sc' alts def) of
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do def' <- case def of
|
||||
Nothing => pure Nothing
|
||||
Just d => pure (Just !(eval rec env stk d))
|
||||
pure $
|
||||
CConstCase fc sc'
|
||||
CConstCase fc sc'
|
||||
!(traverse (evalConstAlt rec env stk) alts)
|
||||
def'
|
||||
Just val => pure val
|
||||
@ -172,8 +172,8 @@ mutual
|
||||
evalAlt : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> List Name -> EEnv free vars -> Stack free -> CConAlt (vars ++ free) ->
|
||||
Core (CConAlt free)
|
||||
evalAlt {free} {vars} fc rec env stk (MkConAlt n t args sc)
|
||||
= do let sc' = thinAll {outer=args ++ vars} {inner=free} args
|
||||
evalAlt {free} {vars} fc rec env stk (MkConAlt n t args sc)
|
||||
= do let sc' = thinAll {outer=args ++ vars} {inner=free} args
|
||||
(rewrite sym (appendAssociative args vars free) in sc)
|
||||
scEval <- eval rec (extendLoc fc env args) (map (weakenNs args) stk) sc'
|
||||
pure $ MkConAlt n t args scEval
|
||||
@ -186,37 +186,37 @@ mutual
|
||||
|
||||
pickAlt : {auto c : Ref Ctxt Defs} ->
|
||||
List Name -> EEnv free vars -> Stack free ->
|
||||
CExp free -> List (CConAlt (vars ++ free)) ->
|
||||
Maybe (CExp (vars ++ free)) ->
|
||||
CExp free -> List (CConAlt (vars ++ free)) ->
|
||||
Maybe (CExp (vars ++ free)) ->
|
||||
Core (Maybe (CExp free))
|
||||
pickAlt rec env stk (CCon fc n t args) [] def
|
||||
pickAlt rec env stk (CCon fc n t args) [] def
|
||||
= case def of
|
||||
Nothing => pure Nothing
|
||||
Just d => pure $ Just !(eval rec env stk d)
|
||||
pickAlt {vars} {free} rec env stk (CCon fc n t args) (MkConAlt n' t' args' sc :: alts) def
|
||||
= if t == t'
|
||||
pickAlt {vars} {free} rec env stk (CCon fc n t args) (MkConAlt n' t' args' sc :: alts) def
|
||||
= if t == t'
|
||||
then case checkLengthMatch args args' of
|
||||
Nothing => pure Nothing
|
||||
Just m =>
|
||||
do let env' : EEnv free (args' ++ vars)
|
||||
do let env' : EEnv free (args' ++ vars)
|
||||
= extend env args args' m
|
||||
pure $ Just !(eval rec env' stk
|
||||
(rewrite sym (appendAssociative args' vars free) in
|
||||
pure $ Just !(eval rec env' stk
|
||||
(rewrite sym (appendAssociative args' vars free) in
|
||||
sc))
|
||||
else pickAlt rec env stk (CCon fc n t args) alts def
|
||||
pickAlt rec env stk _ _ _ = pure Nothing
|
||||
|
||||
pickConstAlt : {auto c : Ref Ctxt Defs} ->
|
||||
List Name -> EEnv free vars -> Stack free ->
|
||||
CExp free -> List (CConstAlt (vars ++ free)) ->
|
||||
Maybe (CExp (vars ++ free)) ->
|
||||
CExp free -> List (CConstAlt (vars ++ free)) ->
|
||||
Maybe (CExp (vars ++ free)) ->
|
||||
Core (Maybe (CExp free))
|
||||
pickConstAlt rec env stk (CPrimVal fc c) [] def
|
||||
pickConstAlt rec env stk (CPrimVal fc c) [] def
|
||||
= case def of
|
||||
Nothing => pure Nothing
|
||||
Just d => pure $ Just !(eval rec env stk d)
|
||||
pickConstAlt {vars} {free} rec env stk (CPrimVal fc c) (MkConstAlt c' sc :: alts) def
|
||||
= if c == c'
|
||||
pickConstAlt {vars} {free} rec env stk (CPrimVal fc c) (MkConstAlt c' sc :: alts) def
|
||||
= if c == c'
|
||||
then pure $ Just !(eval rec env stk sc)
|
||||
else pickConstAlt rec env stk (CPrimVal fc c) alts def
|
||||
pickConstAlt rec env stk _ _ _ = pure Nothing
|
||||
@ -241,4 +241,4 @@ inlineDef n
|
||||
inlined <- inline cexpr
|
||||
-- coreLift $ putStrLn $ show (fullname def) ++ " after: " ++ show inlined
|
||||
setCompiled n inlined
|
||||
|
||||
|
||||
|
@ -222,7 +222,7 @@ cCall fc cfn clib args ret
|
||||
argTypes <- traverse (cftySpec fc) (filter notWorld args)
|
||||
retType <- cftySpec fc retty
|
||||
pure $
|
||||
"(let ([c-code (foreign-callable #f " ++
|
||||
"(let ([c-code (foreign-callable #f " ++
|
||||
mkFun args retty n ++
|
||||
" (" ++ showSep " " argTypes ++ ") " ++ retType ++ ")])" ++
|
||||
" (lock-object c-code) (foreign-callable-entry-point c-code))"
|
||||
|
@ -59,7 +59,7 @@ mutual
|
||||
chickenPrim : Int -> SVars vars -> ExtPrim -> List (CExp vars) -> Core String
|
||||
chickenPrim i vs CCall [ret, fn, args, world]
|
||||
= throw (InternalError ("Can't compile C FFI calls to Chicken Scheme yet"))
|
||||
chickenPrim i vs prim args
|
||||
chickenPrim i vs prim args
|
||||
= schExtCommon chickenPrim chickenString i vs prim args
|
||||
|
||||
compileToSCM : Ref Ctxt Defs ->
|
||||
|
@ -49,7 +49,7 @@ data SVars : List Name -> Type where
|
||||
|
||||
extendSVars : (xs : List Name) -> SVars ns -> SVars (xs ++ ns)
|
||||
extendSVars {ns} xs vs = extSVars' (cast (length ns)) xs vs
|
||||
where
|
||||
where
|
||||
extSVars' : Int -> (xs : List Name) -> SVars ns -> SVars (xs ++ ns)
|
||||
extSVars' i [] vs = vs
|
||||
extSVars' i (x :: xs) vs = schName (MN "v" i) :: extSVars' (i + 1) xs vs
|
||||
@ -74,7 +74,7 @@ op o args = "(" ++ o ++ " " ++ showSep " " args ++ ")"
|
||||
boolop : String -> List String -> String
|
||||
boolop o args = "(or (and " ++ op o args ++ " 1) 0)"
|
||||
|
||||
||| Generate scheme for a primitive function.
|
||||
||| Generate scheme for a primitive function.
|
||||
schOp : PrimFn arity -> Vect arity String -> String
|
||||
schOp (Add IntType) [x, y] = op "b+" [x, y, "63"]
|
||||
schOp (Sub IntType) [x, y] = op "b-" [x, y, "63"]
|
||||
@ -152,7 +152,7 @@ schOp BelieveMe [_,_,x] = x
|
||||
|
||||
||| Extended primitives for the scheme backend, outside the standard set of primFn
|
||||
public export
|
||||
data ExtPrim = CCall | SchemeCall | PutStr | GetStr
|
||||
data ExtPrim = CCall | SchemeCall | PutStr | GetStr
|
||||
| FileOpen | FileClose | FileReadLine | FileWriteLine | FileEOF
|
||||
| NewIORef | ReadIORef | WriteIORef
|
||||
| Stdin | Stdout | Stderr
|
||||
@ -180,7 +180,7 @@ Show ExtPrim where
|
||||
|
||||
||| Match on a user given name to get the scheme primitive
|
||||
toPrim : Name -> ExtPrim
|
||||
toPrim pn@(NS _ n)
|
||||
toPrim pn@(NS _ n)
|
||||
= cond [(n == UN "prim__schemeCall", SchemeCall),
|
||||
(n == UN "prim__cCall", CCall),
|
||||
(n == UN "prim__putStr", PutStr),
|
||||
@ -222,7 +222,7 @@ schConstant _ WorldType = "#t"
|
||||
schCaseDef : Maybe String -> String
|
||||
schCaseDef Nothing = ""
|
||||
schCaseDef (Just tm) = "(else " ++ tm ++ ")"
|
||||
|
||||
|
||||
export
|
||||
schArglist : SVars ns -> String
|
||||
schArglist [] = ""
|
||||
@ -240,14 +240,14 @@ parameters (schExtPrim : {vars : _} -> Int -> SVars vars -> ExtPrim -> List (CEx
|
||||
where
|
||||
bindArgs : Int -> (ns : List Name) -> SVars (ns ++ vars) -> String -> String
|
||||
bindArgs i [] vs body = body
|
||||
bindArgs i (n :: ns) (v :: vs) body
|
||||
bindArgs i (n :: ns) (v :: vs) body
|
||||
= "(let ((" ++ v ++ " " ++ "(vector-ref " ++ target ++ " " ++ show i ++ "))) "
|
||||
++ bindArgs (i + 1) ns vs body ++ ")"
|
||||
|
||||
schConstAlt : Int -> SVars vars -> String -> CConstAlt vars -> Core String
|
||||
schConstAlt i vs target (MkConstAlt c exp)
|
||||
= pure $ "((equal? " ++ target ++ " " ++ schConstant schString c ++ ") " ++ !(schExp i vs exp) ++ ")"
|
||||
|
||||
|
||||
-- oops, no traverse for Vect in Core
|
||||
schArgs : Int -> SVars vars -> Vect n (CExp vars) -> Core (Vect n String)
|
||||
schArgs i vs [] = pure []
|
||||
@ -257,35 +257,35 @@ parameters (schExtPrim : {vars : _} -> Int -> SVars vars -> ExtPrim -> List (CEx
|
||||
schExp : Int -> SVars vars -> CExp vars -> Core String
|
||||
schExp i vs (CLocal fc el) = pure $ lookupSVar el vs
|
||||
schExp i vs (CRef fc n) = pure $ schName n
|
||||
schExp i vs (CLam fc x sc)
|
||||
= do let vs' = extendSVars [x] vs
|
||||
schExp i vs (CLam fc x sc)
|
||||
= do let vs' = extendSVars [x] vs
|
||||
sc' <- schExp i vs' sc
|
||||
pure $ "(lambda (" ++ lookupSVar First vs' ++ ") " ++ sc' ++ ")"
|
||||
schExp i vs (CLet fc x val sc)
|
||||
schExp i vs (CLet fc x val sc)
|
||||
= do let vs' = extendSVars [x] vs
|
||||
val' <- schExp i vs val
|
||||
sc' <- schExp i vs' sc
|
||||
pure $ "(let ((" ++ lookupSVar First vs' ++ " " ++ val' ++ ")) " ++ sc' ++ ")"
|
||||
schExp i vs (CApp fc x [])
|
||||
schExp i vs (CApp fc x [])
|
||||
= pure $ "(" ++ !(schExp i vs x) ++ ")"
|
||||
schExp i vs (CApp fc x args)
|
||||
schExp i vs (CApp fc x args)
|
||||
= pure $ "(" ++ !(schExp i vs x) ++ " " ++ showSep " " !(traverse (schExp i vs) args) ++ ")"
|
||||
schExp i vs (CCon fc x tag args)
|
||||
schExp i vs (CCon fc x tag args)
|
||||
= pure $ schConstructor tag !(traverse (schExp i vs) args)
|
||||
schExp i vs (COp fc op args)
|
||||
schExp i vs (COp fc op args)
|
||||
= pure $ schOp op !(schArgs i vs args)
|
||||
schExp i vs (CExtPrim fc p args)
|
||||
schExp i vs (CExtPrim fc p args)
|
||||
= schExtPrim i vs (toPrim p) args
|
||||
schExp i vs (CForce fc t) = pure $ "(force " ++ !(schExp i vs t) ++ ")"
|
||||
schExp i vs (CDelay fc t) = pure $ "(delay " ++ !(schExp i vs t) ++ ")"
|
||||
schExp i vs (CConCase fc sc alts def)
|
||||
schExp i vs (CConCase fc sc alts def)
|
||||
= do tcode <- schExp (i+1) vs sc
|
||||
defc <- maybe (pure Nothing) (\v => pure (Just !(schExp i vs v))) def
|
||||
let n = "sc" ++ show i
|
||||
pure $ "(let ((" ++ n ++ " " ++ tcode ++ ")) (case (get-tag " ++ n ++ ") "
|
||||
++ showSep " " !(traverse (schConAlt (i+1) vs n) alts)
|
||||
++ schCaseDef defc ++ "))"
|
||||
schExp i vs (CConstCase fc sc alts def)
|
||||
schExp i vs (CConstCase fc sc alts def)
|
||||
= do defc <- maybe (pure Nothing) (\v => pure (Just !(schExp i vs v))) def
|
||||
tcode <- schExp (i+1) vs sc
|
||||
let n = "sc" ++ show i
|
||||
@ -314,12 +314,12 @@ parameters (schExtPrim : {vars : _} -> Int -> SVars vars -> ExtPrim -> List (CEx
|
||||
schExtCommon i vs SchemeCall [ret, fn, args, world]
|
||||
= pure $ mkWorld ("(apply (eval (string->symbol " ++ !(schExp i vs fn) ++")) "
|
||||
++ !(readArgs i vs args) ++ ")")
|
||||
schExtCommon i vs PutStr [arg, world]
|
||||
schExtCommon i vs PutStr [arg, world]
|
||||
= pure $ "(display " ++ !(schExp i vs arg) ++ ") " ++ mkWorld (schConstructor 0 []) -- code for MkUnit
|
||||
schExtCommon i vs GetStr [world]
|
||||
schExtCommon i vs GetStr [world]
|
||||
= pure $ mkWorld "(blodwen-get-line (current-input-port))"
|
||||
schExtCommon i vs FileOpen [file, mode, bin, world]
|
||||
= pure $ mkWorld $ fileOp $ "(blodwen-open "
|
||||
= pure $ mkWorld $ fileOp $ "(blodwen-open "
|
||||
++ !(schExp i vs file) ++ " "
|
||||
++ !(schExp i vs mode) ++ " "
|
||||
++ !(schExp i vs bin) ++ ")"
|
||||
@ -328,7 +328,7 @@ parameters (schExtPrim : {vars : _} -> Int -> SVars vars -> ExtPrim -> List (CEx
|
||||
schExtCommon i vs FileReadLine [file, world]
|
||||
= pure $ mkWorld $ fileOp $ "(blodwen-get-line " ++ !(schExp i vs file) ++ ")"
|
||||
schExtCommon i vs FileWriteLine [file, str, world]
|
||||
= pure $ mkWorld $ fileOp $ "(blodwen-putstring "
|
||||
= pure $ mkWorld $ fileOp $ "(blodwen-putstring "
|
||||
++ !(schExp i vs file) ++ " "
|
||||
++ !(schExp i vs str) ++ ")"
|
||||
schExtCommon i vs FileEOF [file, world]
|
||||
@ -338,17 +338,17 @@ parameters (schExtPrim : {vars : _} -> Int -> SVars vars -> ExtPrim -> List (CEx
|
||||
schExtCommon i vs ReadIORef [_, ref, world]
|
||||
= pure $ mkWorld $ "(unbox " ++ !(schExp i vs ref) ++ ")"
|
||||
schExtCommon i vs WriteIORef [_, ref, val, world]
|
||||
= pure $ mkWorld $ "(set-box! "
|
||||
++ !(schExp i vs ref) ++ " "
|
||||
= pure $ mkWorld $ "(set-box! "
|
||||
++ !(schExp i vs ref) ++ " "
|
||||
++ !(schExp i vs val) ++ ")"
|
||||
schExtCommon i vs VoidElim [_, _]
|
||||
= pure "(display \"Error: Executed 'void'\")"
|
||||
schExtCommon i vs (Unknown n) args
|
||||
schExtCommon i vs (Unknown n) args
|
||||
= throw (InternalError ("Can't compile unknown external primitive " ++ show n))
|
||||
schExtCommon i vs Stdin [] = pure "(current-input-port)"
|
||||
schExtCommon i vs Stdout [] = pure "(current-output-port)"
|
||||
schExtCommon i vs Stderr [] = pure "(current-error-port)"
|
||||
schExtCommon i vs prim args
|
||||
schExtCommon i vs prim args
|
||||
= throw (InternalError ("Badly formed external primitive " ++ show prim
|
||||
++ " " ++ show args))
|
||||
|
||||
@ -362,7 +362,7 @@ parameters (schExtPrim : {vars : _} -> Int -> SVars vars -> ExtPrim -> List (CEx
|
||||
= pure $ "(define (" ++ schName !(getFullName n) ++ " . any-args) " ++ !(schExp 0 [] exp) ++ ")\n"
|
||||
schDef n (MkForeign _ _ _) = pure "" -- compiled by specific back end
|
||||
schDef n (MkCon t a) = pure "" -- Nothing to compile here
|
||||
|
||||
|
||||
-- Convert the name to scheme code
|
||||
-- (There may be no code generated, for example if it's a constructor)
|
||||
export
|
||||
|
@ -63,7 +63,7 @@ mutual
|
||||
racketPrim : Int -> SVars vars -> ExtPrim -> List (CExp vars) -> Core String
|
||||
racketPrim i vs CCall [ret, fn, args, world]
|
||||
= throw (InternalError ("Can't compile C FFI calls to Racket yet"))
|
||||
racketPrim i vs prim args
|
||||
racketPrim i vs prim args
|
||||
= schExtCommon racketPrim racketString i vs prim args
|
||||
|
||||
-- Reference label for keeping track of loaded external libraries
|
||||
@ -91,7 +91,7 @@ cftySpec fc t = throw (GenericMsg fc ("Can't pass argument of type " ++ show t +
|
||||
|
||||
loadlib : String -> String -> String
|
||||
loadlib libn ver
|
||||
= "(define-ffi-definer define-" ++ libn ++
|
||||
= "(define-ffi-definer define-" ++ libn ++
|
||||
" (ffi-lib \"" ++ libn ++ "\" " ++ ver ++ "))\n"
|
||||
|
||||
getLibVers : String -> (String, String)
|
||||
@ -131,8 +131,8 @@ cCall fc cfn libspec args ret
|
||||
argTypes <- traverse (\a => do s <- cftySpec fc (snd a)
|
||||
pure (a, s)) args
|
||||
retType <- cftySpec fc ret
|
||||
let cbind = "(define-" ++ libn ++ " " ++ cfn ++
|
||||
" (_fun " ++ showSep " " (map snd argTypes) ++ " -> " ++
|
||||
let cbind = "(define-" ++ libn ++ " " ++ cfn ++
|
||||
" (_fun " ++ showSep " " (map snd argTypes) ++ " -> " ++
|
||||
retType ++ "))\n"
|
||||
let call = "(" ++ cfn ++ " " ++
|
||||
showSep " " !(traverse useArg argTypes) ++ ")"
|
||||
@ -149,7 +149,7 @@ cCall fc cfn libspec args ret
|
||||
applyLams : String -> List (Maybe (String, CFType)) -> String
|
||||
applyLams n [] = n
|
||||
applyLams n (Nothing :: as) = applyLams ("(" ++ n ++ " #f)") as
|
||||
applyLams n (Just (a, ty) :: as)
|
||||
applyLams n (Just (a, ty) :: as)
|
||||
= applyLams ("(" ++ n ++ " " ++ cToRkt ty a ++ ")") as
|
||||
|
||||
getVal : CFType -> String -> String
|
||||
@ -199,7 +199,7 @@ useCC fc [] args ret
|
||||
useCC fc (cc :: ccs) args ret
|
||||
= case parseCC cc of
|
||||
Nothing => useCC fc ccs args ret
|
||||
Just ("scheme", [sfn]) =>
|
||||
Just ("scheme", [sfn]) =>
|
||||
do body <- schemeCall fc sfn (map fst args) ret
|
||||
pure ("", body)
|
||||
Just ("C", [cfn, clib]) => cCall fc cfn clib args ret
|
||||
@ -216,7 +216,7 @@ mkArgs i (c :: cs) = (MN "farg" i, True) :: mkArgs (i + 1) cs
|
||||
schFgnDef : {auto c : Ref Ctxt Defs} ->
|
||||
{auto l : Ref Loaded (List String)} ->
|
||||
FC -> Name -> CDef -> Core (String, String)
|
||||
schFgnDef fc n (MkForeign cs args ret)
|
||||
schFgnDef fc n (MkForeign cs args ret)
|
||||
= do let argns = mkArgs 0 args
|
||||
let allargns = map fst argns
|
||||
let useargns = map fst (filter snd argns)
|
||||
@ -251,9 +251,9 @@ compileToRKT c tm outfile
|
||||
let code = concat (map snd fgndefs) ++ concat compdefs
|
||||
main <- schExp racketPrim racketString 0 [] !(compileExp tags tm)
|
||||
support <- readDataFile "racket/support.rkt"
|
||||
let scm = schHeader (concat (map fst fgndefs)) ++
|
||||
support ++ code ++
|
||||
"(void " ++ main ++ ")\n" ++
|
||||
let scm = schHeader (concat (map fst fgndefs)) ++
|
||||
support ++ code ++
|
||||
"(void " ++ main ++ ")\n" ++
|
||||
schFooter
|
||||
Right () <- coreLift $ writeFile outfile scm
|
||||
| Left err => throw (FileErr outfile err)
|
||||
|
BIN
src/Core/.Context.idr.swp
Normal file
BIN
src/Core/.Context.idr.swp
Normal file
Binary file not shown.
BIN
src/Core/.Normalise.idr.swp
Normal file
BIN
src/Core/.Normalise.idr.swp
Normal file
Binary file not shown.
@ -18,7 +18,7 @@ searchType : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> RigCount ->
|
||||
(defaults : Bool) -> (trying : List (Term vars)) ->
|
||||
(depth : Nat) ->
|
||||
(defining : Name) ->
|
||||
(defining : Name) ->
|
||||
(checkDets : Bool) -> (topTy : ClosedTerm) ->
|
||||
Env Term vars -> (target : Term vars) -> Core (Term vars)
|
||||
|
||||
@ -35,7 +35,7 @@ export
|
||||
mkArgs : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount ->
|
||||
Env Term vars -> NF vars ->
|
||||
Env Term vars -> NF vars ->
|
||||
Core (List (ArgInfo vars), NF vars)
|
||||
mkArgs fc rigc env (NBind nfc x (Pi c p ty) sc)
|
||||
= do defs <- get Ctxt
|
||||
@ -43,25 +43,25 @@ mkArgs fc rigc env (NBind nfc x (Pi c p ty) sc)
|
||||
nm <- genName "sa"
|
||||
argTy <- quote empty env ty
|
||||
let argRig = rigMult rigc c
|
||||
(idx, arg) <- newMeta fc argRig env nm argTy
|
||||
(idx, arg) <- newMeta fc argRig env nm argTy
|
||||
(Hole (length env) False) False
|
||||
setInvertible fc (Resolved idx)
|
||||
(rest, restTy) <- mkArgs fc rigc env
|
||||
(rest, restTy) <- mkArgs fc rigc env
|
||||
!(sc defs (toClosure defaultOpts env arg))
|
||||
pure (MkArgInfo idx argRig p arg argTy :: rest, restTy)
|
||||
mkArgs fc rigc env ty = pure ([], ty)
|
||||
|
||||
searchIfHole : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC ->
|
||||
FC ->
|
||||
(defaults : Bool) -> List (Term vars) ->
|
||||
(ispair : Bool) -> (depth : Nat) ->
|
||||
(defining : Name) -> (topTy : ClosedTerm) -> Env Term vars ->
|
||||
(arg : ArgInfo vars) ->
|
||||
Core ()
|
||||
searchIfHole fc defaults trying ispair Z def top env arg
|
||||
(defining : Name) -> (topTy : ClosedTerm) -> Env Term vars ->
|
||||
(arg : ArgInfo vars) ->
|
||||
Core ()
|
||||
searchIfHole fc defaults trying ispair Z def top env arg
|
||||
= throw (CantSolveGoal fc [] top) -- possibly should say depth limit hit?
|
||||
searchIfHole fc defaults trying ispair (S depth) def top env arg
|
||||
searchIfHole fc defaults trying ispair (S depth) def top env arg
|
||||
= do let hole = holeID arg
|
||||
let rig = argRig arg
|
||||
|
||||
@ -70,11 +70,11 @@ searchIfHole fc defaults trying ispair (S depth) def top env arg
|
||||
| Nothing => throw (CantSolveGoal fc [] top)
|
||||
let Hole _ _ = definition gdef
|
||||
| _ => pure () -- already solved
|
||||
top' <- if ispair
|
||||
top' <- if ispair
|
||||
then normaliseScope defs [] (type gdef)
|
||||
else pure top
|
||||
|
||||
argdef <- searchType fc rig defaults trying depth def False top' env
|
||||
argdef <- searchType fc rig defaults trying depth def False top' env
|
||||
!(normaliseScope defs env (argType arg))
|
||||
vs <- unify InTerm fc env (metaApp arg) argdef
|
||||
let [] = constraints vs
|
||||
@ -90,7 +90,7 @@ successful [] = pure []
|
||||
successful (elab :: elabs)
|
||||
= do ust <- get UST
|
||||
defs <- branch
|
||||
catch (do -- Run the elaborator
|
||||
catch (do -- Run the elaborator
|
||||
res <- elab
|
||||
-- Record post-elaborator state
|
||||
ust' <- get UST
|
||||
@ -122,7 +122,7 @@ exactlyOne : {vars : _} ->
|
||||
FC -> Env Term vars -> (topTy : ClosedTerm) ->
|
||||
List (Core (Term vars)) ->
|
||||
Core (Term vars)
|
||||
exactlyOne fc env top [elab]
|
||||
exactlyOne fc env top [elab]
|
||||
= catch elab
|
||||
(\err => case err of
|
||||
CantSolveGoal _ _ _ => throw err
|
||||
@ -130,13 +130,13 @@ exactlyOne fc env top [elab]
|
||||
exactlyOne {vars} fc env top all
|
||||
= do elabs <- successful all
|
||||
case rights elabs of
|
||||
[(res, defs, ust)] =>
|
||||
[(res, defs, ust)] =>
|
||||
do put UST ust
|
||||
put Ctxt defs
|
||||
commit
|
||||
pure res
|
||||
[] => throw (CantSolveGoal fc [] top)
|
||||
rs => throw (AmbiguousSearch fc env
|
||||
rs => throw (AmbiguousSearch fc env
|
||||
!(traverse normRes rs))
|
||||
where
|
||||
normRes : (Term vars, Defs, UState) -> Core (Term vars)
|
||||
@ -147,22 +147,22 @@ exactlyOne {vars} fc env top all
|
||||
-- because something is apparently available now, it will be available by the
|
||||
-- time we get to linearity checking.
|
||||
-- It's also fine to use anything if we're working at multiplicity 0
|
||||
getAllEnv : FC -> RigCount -> (done : List Name) ->
|
||||
getAllEnv : FC -> RigCount -> (done : List Name) ->
|
||||
Env Term vars -> List (Term (done ++ vars), Term (done ++ vars))
|
||||
getAllEnv fc rigc done [] = []
|
||||
getAllEnv {vars = v :: vs} fc rigc done (b :: env)
|
||||
= let rest = getAllEnv fc rigc (done ++ [v]) env in
|
||||
getAllEnv {vars = v :: vs} fc rigc done (b :: env)
|
||||
= let rest = getAllEnv fc rigc (done ++ [v]) env in
|
||||
if multiplicity b == RigW || rigc == Rig0
|
||||
then let MkVar p = weakenVar {name=v} {inner=v :: vs} done First in
|
||||
(Local fc Nothing _ p,
|
||||
rewrite appendAssociative done [v] vs in
|
||||
weakenNs (done ++ [v]) (binderType b)) ::
|
||||
(Local fc Nothing _ p,
|
||||
rewrite appendAssociative done [v] vs in
|
||||
weakenNs (done ++ [v]) (binderType b)) ::
|
||||
rewrite appendAssociative done [v] vs in rest
|
||||
else rewrite appendAssociative done [v] vs in rest
|
||||
|
||||
-- A local is usable if it contains no holes in a determining argument position
|
||||
usableLocal : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> (defaults : Bool) ->
|
||||
FC -> (defaults : Bool) ->
|
||||
Env Term vars -> (locTy : NF vars) -> Core Bool
|
||||
-- pattern variables count as concrete things!
|
||||
usableLocal loc defaults env (NApp fc (NMeta (PV _ _) _ _) args)
|
||||
@ -180,24 +180,24 @@ usableLocal {vars} loc defaults env (NTCon _ n _ _ args)
|
||||
usableLocalArg i dets (c :: cs)
|
||||
= if i `elem` dets
|
||||
then do defs <- get Ctxt
|
||||
u <- usableLocal loc defaults env !(evalClosure defs c)
|
||||
u <- usableLocal loc defaults env !(evalClosure defs c)
|
||||
if u
|
||||
then usableLocalArg (1 + i) dets cs
|
||||
else pure False
|
||||
else usableLocalArg (1 + i) dets cs
|
||||
usableLocal loc defaults env (NDCon _ n _ _ args)
|
||||
= do defs <- get Ctxt
|
||||
us <- traverse (usableLocal loc defaults env)
|
||||
us <- traverse (usableLocal loc defaults env)
|
||||
!(traverse (evalClosure defs) args)
|
||||
pure (and (map Delay us))
|
||||
usableLocal loc defaults env (NApp _ (NLocal _ _ _) args)
|
||||
= do defs <- get Ctxt
|
||||
us <- traverse (usableLocal loc defaults env)
|
||||
us <- traverse (usableLocal loc defaults env)
|
||||
!(traverse (evalClosure defs) args)
|
||||
pure (and (map Delay us))
|
||||
usableLocal loc defaults env (NBind fc x (Pi _ _ _) sc)
|
||||
= do defs <- get Ctxt
|
||||
usableLocal loc defaults env
|
||||
usableLocal loc defaults env
|
||||
!(sc defs (toClosure defaultOpts env (Erased fc)))
|
||||
usableLocal loc defaults env (NErased _) = pure False
|
||||
usableLocal loc _ _ _ = pure True
|
||||
@ -208,20 +208,20 @@ searchLocalWith : {auto c : Ref Ctxt Defs} ->
|
||||
(defaults : Bool) -> List (Term vars) ->
|
||||
(depth : Nat) ->
|
||||
(defining : Name) -> (topTy : ClosedTerm) ->
|
||||
Env Term vars -> List (Term vars, Term vars) ->
|
||||
Env Term vars -> List (Term vars, Term vars) ->
|
||||
(target : NF vars) -> Core (Term vars)
|
||||
searchLocalWith fc rigc defaults trying depth def top env [] target
|
||||
= throw (CantSolveGoal fc [] top)
|
||||
searchLocalWith {vars} fc rigc defaults trying depth def top env ((prf, ty) :: rest) target
|
||||
= tryUnify
|
||||
= tryUnify
|
||||
(do defs <- get Ctxt
|
||||
nty <- nf defs env ty
|
||||
findPos defs prf id nty target)
|
||||
(searchLocalWith fc rigc defaults trying depth def top env rest target)
|
||||
where
|
||||
clearEnvType : {idx : Nat} -> .(IsVar name idx vs) ->
|
||||
clearEnvType : {idx : Nat} -> .(IsVar name idx vs) ->
|
||||
FC -> Env Term vs -> Env Term vs
|
||||
clearEnvType First fc (b :: env)
|
||||
clearEnvType First fc (b :: env)
|
||||
= Lam (multiplicity b) Explicit (Erased fc) :: env
|
||||
clearEnvType (Later p) fc (b :: env) = b :: clearEnvType p fc env
|
||||
|
||||
@ -230,10 +230,10 @@ searchLocalWith {vars} fc rigc defaults trying depth def top env ((prf, ty) :: r
|
||||
= clearEnvType p fc env
|
||||
clearEnv _ env = env
|
||||
|
||||
findDirect : Defs -> Term vars ->
|
||||
findDirect : Defs -> Term vars ->
|
||||
(Term vars -> Term vars) ->
|
||||
NF vars -> -- local's type
|
||||
(target : NF vars) ->
|
||||
(target : NF vars) ->
|
||||
Core (Term vars)
|
||||
findDirect defs prf f ty target
|
||||
= do (args, appTy) <- mkArgs fc rigc env ty
|
||||
@ -252,16 +252,16 @@ searchLocalWith {vars} fc rigc defaults trying depth def top env ((prf, ty) :: r
|
||||
let env' = clearEnv prf env
|
||||
-- Work right to left, because later arguments may solve
|
||||
-- earlier ones by unification
|
||||
traverse (searchIfHole fc defaults trying False depth def top env')
|
||||
traverse (searchIfHole fc defaults trying False depth def top env')
|
||||
(reverse args)
|
||||
pure candidate
|
||||
else do logNF 10 "Can't use " env ty
|
||||
throw (CantSolveGoal fc [] top)
|
||||
|
||||
findPos : Defs -> Term vars ->
|
||||
findPos : Defs -> Term vars ->
|
||||
(Term vars -> Term vars) ->
|
||||
NF vars -> -- local's type
|
||||
(target : NF vars) ->
|
||||
(target : NF vars) ->
|
||||
Core (Term vars)
|
||||
findPos defs prf f nty@(NTCon pfc pn _ _ [xty, yty]) target
|
||||
= tryUnify (findDirect defs prf f nty target)
|
||||
@ -277,14 +277,14 @@ searchLocalWith {vars} fc rigc defaults trying depth def top env ((prf, ty) :: r
|
||||
ytytm <- quote empty env yty
|
||||
tryUnify
|
||||
(do xtynf <- evalClosure defs xty
|
||||
findPos defs prf
|
||||
findPos defs prf
|
||||
(\arg => apply fc (Ref fc Func fname)
|
||||
[xtytm,
|
||||
ytytm,
|
||||
f arg])
|
||||
xtynf target)
|
||||
(do ytynf <- evalClosure defs yty
|
||||
findPos defs prf
|
||||
findPos defs prf
|
||||
(\arg => apply fc (Ref fc Func sname)
|
||||
[xtytm,
|
||||
ytytm,
|
||||
@ -300,10 +300,10 @@ searchLocal : {auto c : Ref Ctxt Defs} ->
|
||||
(defaults : Bool) -> List (Term vars) ->
|
||||
(depth : Nat) ->
|
||||
(defining : Name) -> (topTy : ClosedTerm) ->
|
||||
Env Term vars ->
|
||||
Env Term vars ->
|
||||
(target : NF vars) -> Core (Term vars)
|
||||
searchLocal fc rig defaults trying depth def top env target
|
||||
= searchLocalWith fc rig defaults trying depth def top env
|
||||
= searchLocalWith fc rig defaults trying depth def top env
|
||||
(getAllEnv fc rig [] env) target
|
||||
|
||||
isPairNF : {auto c : Ref Ctxt Defs} ->
|
||||
@ -320,12 +320,12 @@ searchName : {auto c : Ref Ctxt Defs} ->
|
||||
(defaults : Bool) -> List (Term vars) ->
|
||||
(depth : Nat) ->
|
||||
(defining : Name) -> (topTy : ClosedTerm) ->
|
||||
Env Term vars -> (target : NF vars) ->
|
||||
(Name, GlobalDef) ->
|
||||
Env Term vars -> (target : NF vars) ->
|
||||
(Name, GlobalDef) ->
|
||||
Core (Term vars)
|
||||
searchName fc rigc defaults trying depth def top env target (n, ndef)
|
||||
= do defs <- get Ctxt
|
||||
when (not (visibleInAny (!getNS :: !getNestedNS)
|
||||
when (not (visibleInAny (!getNS :: !getNestedNS)
|
||||
(fullname ndef) (visibility ndef))) $
|
||||
throw (CantSolveGoal fc [] top)
|
||||
when (BlockedHint `elem` flags ndef) $
|
||||
@ -348,7 +348,7 @@ searchName fc rigc defaults trying depth def top env target (n, ndef)
|
||||
logTermNF 10 "Candidate " env candidate
|
||||
-- Work right to left, because later arguments may solve earlier
|
||||
-- dependencies by unification
|
||||
traverse (searchIfHole fc defaults trying ispair depth def top env)
|
||||
traverse (searchIfHole fc defaults trying ispair depth def top env)
|
||||
(reverse args)
|
||||
pure candidate
|
||||
|
||||
@ -358,7 +358,7 @@ searchNames : {auto c : Ref Ctxt Defs} ->
|
||||
(defaults : Bool) -> List (Term vars) ->
|
||||
(depth : Nat) ->
|
||||
(defining : Name) -> (topTy : ClosedTerm) ->
|
||||
Env Term vars -> Bool -> List Name ->
|
||||
Env Term vars -> Bool -> List Name ->
|
||||
(target : NF vars) -> Core (Term vars)
|
||||
searchNames fc rigc defaults trying depth defining topty env ambig [] target
|
||||
= throw (CantSolveGoal fc [] topty)
|
||||
@ -371,7 +371,7 @@ searchNames fc rigc defaults trying depth defining topty env ambig (n :: ns) tar
|
||||
then anyOne fc env topty elabs
|
||||
else exactlyOne fc env topty elabs
|
||||
where
|
||||
visible : Context ->
|
||||
visible : Context ->
|
||||
List (List String) -> Name -> Core (Maybe (Name, GlobalDef))
|
||||
visible gam nspace n
|
||||
= do Just def <- lookupCtxtExact n gam
|
||||
@ -383,8 +383,8 @@ searchNames fc rigc defaults trying depth defining topty env ambig (n :: ns) tar
|
||||
concreteDets : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Bool ->
|
||||
Env Term vars -> (top : ClosedTerm) ->
|
||||
(pos : Nat) -> (dets : List Nat) ->
|
||||
Env Term vars -> (top : ClosedTerm) ->
|
||||
(pos : Nat) -> (dets : List Nat) ->
|
||||
(args : List (Closure vars)) ->
|
||||
Core ()
|
||||
concreteDets fc defaults env top pos dets [] = pure ()
|
||||
@ -421,28 +421,28 @@ concreteDets {vars} fc defaults env top pos dets (arg :: args)
|
||||
checkConcreteDets : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Bool ->
|
||||
Env Term vars -> (top : ClosedTerm) ->
|
||||
Env Term vars -> (top : ClosedTerm) ->
|
||||
NF vars ->
|
||||
Core ()
|
||||
checkConcreteDets fc defaults env top (NTCon tfc tyn t a args)
|
||||
checkConcreteDets fc defaults env top (NTCon tfc tyn t a args)
|
||||
= do defs <- get Ctxt
|
||||
if !(isPairType tyn)
|
||||
then case args of
|
||||
[aty, bty] =>
|
||||
[aty, bty] =>
|
||||
do anf <- evalClosure defs aty
|
||||
bnf <- evalClosure defs bty
|
||||
checkConcreteDets fc defaults env top anf
|
||||
checkConcreteDets fc defaults env top bnf
|
||||
_ => do sd <- getSearchData fc defaults tyn
|
||||
concreteDets fc defaults env top 0 (detArgs sd) args
|
||||
else
|
||||
else
|
||||
do sd <- getSearchData fc defaults tyn
|
||||
concreteDets fc defaults env top 0 (detArgs sd) args
|
||||
checkConcreteDets fc defaults env top _
|
||||
= pure ()
|
||||
|
||||
abandonIfCycle : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars -> Term vars -> List (Term vars) ->
|
||||
Env Term vars -> Term vars -> List (Term vars) ->
|
||||
Core ()
|
||||
abandonIfCycle env tm [] = pure ()
|
||||
abandonIfCycle env tm (ty :: tys)
|
||||
@ -470,7 +470,7 @@ searchType {vars} fc rigc defaults trying depth def checkdets top env target
|
||||
if a == length args
|
||||
then do logNF 10 "Next target" env nty
|
||||
sd <- getSearchData fc defaults tyn
|
||||
-- Check determining arguments are okay for 'args'
|
||||
-- Check determining arguments are okay for 'args'
|
||||
when checkdets $
|
||||
checkConcreteDets fc defaults env top
|
||||
(NTCon tfc tyn t a args)
|
||||
@ -484,7 +484,7 @@ searchType {vars} fc rigc defaults trying depth def checkdets top env target
|
||||
ambig : Error -> Bool
|
||||
ambig (AmbiguousSearch _ _ _) = True
|
||||
ambig _ = False
|
||||
|
||||
|
||||
-- Take the earliest error message (that's when we look inside pairs,
|
||||
-- typically, and it's best to be more precise)
|
||||
tryGroups : Maybe Error ->
|
||||
@ -506,14 +506,14 @@ searchType {vars} fc rigc defaults trying depth def checkdets top env target
|
||||
-- {auto u : Ref UST UState} ->
|
||||
-- FC -> RigCount ->
|
||||
-- (defaults : Bool) -> (depth : Nat) ->
|
||||
-- (defining : Name) -> (topTy : Term vars) -> Env Term vars ->
|
||||
-- (defining : Name) -> (topTy : Term vars) -> Env Term vars ->
|
||||
-- Core (Term vars)
|
||||
Core.Unify.search fc rigc defaults depth def top env
|
||||
= do defs <- get Ctxt
|
||||
logTermNF 2 "Initial target: " env top
|
||||
log 2 $ "Running search with defaults " ++ show defaults
|
||||
tm <- searchType fc rigc defaults [] depth def
|
||||
True (abstractEnvType fc env top) env
|
||||
tm <- searchType fc rigc defaults [] depth def
|
||||
True (abstractEnvType fc env top) env
|
||||
top
|
||||
logTermNF 2 "Result" env tm
|
||||
defs <- get Ctxt
|
||||
|
@ -27,7 +27,7 @@ import Data.Buffer
|
||||
-- TTC files can only be compatible if the version number is the same
|
||||
export
|
||||
ttcVersion : Int
|
||||
ttcVersion = 9
|
||||
ttcVersion = 10
|
||||
|
||||
export
|
||||
checkTTCVersion : Int -> Int -> Core ()
|
||||
@ -109,7 +109,7 @@ HasNames e => HasNames (TTCFile e) where
|
||||
!(fullRW gam rewritenames)
|
||||
!(fullPrim gam primnames)
|
||||
!(full gam namedirectives)
|
||||
cgdirectives
|
||||
cgdirectives
|
||||
!(full gam extra)
|
||||
where
|
||||
fullPair : Context -> Maybe PairNames -> Core (Maybe PairNames)
|
||||
@ -169,15 +169,15 @@ HasNames e => HasNames (TTCFile e) where
|
||||
|
||||
|
||||
asName : List String -> Maybe (List String) -> Name -> Name
|
||||
asName mod (Just ns) (NS oldns n)
|
||||
= if mod == oldns
|
||||
asName mod (Just ns) (NS oldns n)
|
||||
= if mod == oldns
|
||||
then NS ns n -- TODO: What about if there are nested namespaces in a module?
|
||||
else NS oldns n
|
||||
asName _ _ n = n
|
||||
|
||||
-- NOTE: TTC files are only compatible if the version number is the same,
|
||||
-- *and* the 'annot/extra' type are the same, or there are no holes/constraints
|
||||
writeTTCFile : (HasNames extra, TTC extra) =>
|
||||
writeTTCFile : (HasNames extra, TTC extra) =>
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
Ref Bin Binary -> TTCFile extra -> Core ()
|
||||
writeTTCFile b file_in
|
||||
@ -204,7 +204,7 @@ writeTTCFile b file_in
|
||||
toBuf b (cgdirectives file)
|
||||
toBuf b (extraData file)
|
||||
|
||||
readTTCFile : TTC extra =>
|
||||
readTTCFile : TTC extra =>
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
List String -> Maybe (List String) ->
|
||||
Ref Bin Binary -> Core (TTCFile extra)
|
||||
@ -238,12 +238,12 @@ readTTCFile modns as b
|
||||
cgds <- fromBuf b
|
||||
ex <- fromBuf b
|
||||
pure (MkTTCFile ver ifaceHash importHashes
|
||||
[] [] [] defs uholes -- holes guesses constraints defs
|
||||
[] [] [] defs uholes -- holes guesses constraints defs
|
||||
autohs typehs imp nextv cns nns
|
||||
pns rws prims nds cgds ex)
|
||||
|
||||
-- Pull out the list of GlobalDefs that we want to save
|
||||
getSaveDefs : List Name -> List (Name, Binary) -> Defs ->
|
||||
getSaveDefs : List Name -> List (Name, Binary) -> Defs ->
|
||||
Core (List (Name, Binary))
|
||||
getSaveDefs [] acc _ = pure acc
|
||||
getSaveDefs (n :: ns) acc defs
|
||||
@ -269,10 +269,10 @@ writeToTTC extradata fname
|
||||
ust <- get UST
|
||||
gdefs <- getSaveDefs (keys (toSave defs)) [] defs
|
||||
log 5 $ "Writing " ++ fname ++ " with hash " ++ show (ifaceHash defs)
|
||||
writeTTCFile buf
|
||||
writeTTCFile buf
|
||||
(MkTTCFile ttcVersion (ifaceHash defs) (importHashes defs)
|
||||
(toList (holes ust))
|
||||
(toList (guesses ust))
|
||||
(toList (holes ust))
|
||||
(toList (guesses ust))
|
||||
(toList (constraints ust))
|
||||
gdefs
|
||||
(keys (userHoles defs))
|
||||
@ -282,10 +282,10 @@ writeToTTC extradata fname
|
||||
(nextName ust)
|
||||
(currentNS defs)
|
||||
(nestedNS defs)
|
||||
(pairnames (options defs))
|
||||
(rewritenames (options defs))
|
||||
(primnames (options defs))
|
||||
(namedirectives defs)
|
||||
(pairnames (options defs))
|
||||
(rewritenames (options defs))
|
||||
(primnames (options defs))
|
||||
(namedirectives defs)
|
||||
(cgdirectives defs)
|
||||
extradata)
|
||||
Right ok <- coreLift $ writeToFile fname !(get Bin)
|
||||
@ -316,8 +316,8 @@ addGlobalDef modns as (n, def)
|
||||
|
||||
addTypeHint : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> (Name, Name, Bool) -> Core ()
|
||||
addTypeHint fc (tyn, hintn, d)
|
||||
= do logC 10 (pure (show !(getFullName hintn) ++ " for " ++
|
||||
addTypeHint fc (tyn, hintn, d)
|
||||
= do logC 10 (pure (show !(getFullName hintn) ++ " for " ++
|
||||
show !(getFullName tyn)))
|
||||
addHintFor fc tyn hintn d True
|
||||
|
||||
@ -326,14 +326,14 @@ addAutoHint : {auto c : Ref Ctxt Defs} ->
|
||||
addAutoHint (hintn, d) = addGlobalHint hintn d
|
||||
|
||||
export
|
||||
updatePair : {auto c : Ref Ctxt Defs} ->
|
||||
updatePair : {auto c : Ref Ctxt Defs} ->
|
||||
Maybe PairNames -> Core ()
|
||||
updatePair p
|
||||
= do defs <- get Ctxt
|
||||
put Ctxt (record { options->pairnames $= (p <+>) } defs)
|
||||
|
||||
export
|
||||
updateRewrite : {auto c : Ref Ctxt Defs} ->
|
||||
updateRewrite : {auto c : Ref Ctxt Defs} ->
|
||||
Maybe RewriteNames -> Core ()
|
||||
updateRewrite r
|
||||
= do defs <- get Ctxt
|
||||
@ -344,17 +344,17 @@ updatePrimNames : PrimNames -> PrimNames -> PrimNames
|
||||
updatePrimNames p
|
||||
= record { fromIntegerName $= ((fromIntegerName p) <+>),
|
||||
fromStringName $= ((fromStringName p) <+>),
|
||||
fromCharName $= ((fromCharName p) <+>) }
|
||||
fromCharName $= ((fromCharName p) <+>) }
|
||||
|
||||
export
|
||||
updatePrims : {auto c : Ref Ctxt Defs} ->
|
||||
updatePrims : {auto c : Ref Ctxt Defs} ->
|
||||
PrimNames -> Core ()
|
||||
updatePrims p
|
||||
= do defs <- get Ctxt
|
||||
put Ctxt (record { options->primnames $= updatePrimNames p } defs)
|
||||
|
||||
export
|
||||
updateNameDirectives : {auto c : Ref Ctxt Defs} ->
|
||||
updateNameDirectives : {auto c : Ref Ctxt Defs} ->
|
||||
List (Name, List String) -> Core ()
|
||||
updateNameDirectives [] = pure ()
|
||||
updateNameDirectives ((t, ns) :: nds)
|
||||
@ -363,7 +363,7 @@ updateNameDirectives ((t, ns) :: nds)
|
||||
updateNameDirectives nds
|
||||
|
||||
export
|
||||
updateCGDirectives : {auto c : Ref Ctxt Defs} ->
|
||||
updateCGDirectives : {auto c : Ref Ctxt Defs} ->
|
||||
List (CG, String) -> Core ()
|
||||
updateCGDirectives cgs
|
||||
= do defs <- get Ctxt
|
||||
@ -382,7 +382,7 @@ readFromTTC : TTC extra =>
|
||||
(fname : String) -> -- file containing the module
|
||||
(modNS : List String) -> -- module namespace
|
||||
(importAs : List String) -> -- namespace to import as
|
||||
Core (Maybe (extra, Int,
|
||||
Core (Maybe (extra, Int,
|
||||
List (List String, Bool, List String)))
|
||||
readFromTTC loc reexp fname modNS importAs
|
||||
= do defs <- get Ctxt
|
||||
@ -394,8 +394,8 @@ readFromTTC loc reexp fname modNS importAs
|
||||
Right buf <- coreLift $ readFromFile fname
|
||||
| Left err => throw (InternalError (fname ++ ": " ++ show err))
|
||||
bin <- newRef Bin buf -- for reading the file into
|
||||
let as = if importAs == modNS
|
||||
then Nothing
|
||||
let as = if importAs == modNS
|
||||
then Nothing
|
||||
else Just importAs
|
||||
ttc <- readTTCFile modNS as bin
|
||||
logTime "Adding defs" $ traverse (addGlobalDef modNS as) (context ttc)
|
||||
|
@ -15,7 +15,7 @@ data Phase = CompileTime | RunTime
|
||||
|
||||
data ArgType : List Name -> Type where
|
||||
Known : RigCount -> (ty : Term vars) -> ArgType vars -- arg has type 'ty'
|
||||
Stuck : (fty : Term vars) -> ArgType vars
|
||||
Stuck : (fty : Term vars) -> ArgType vars
|
||||
-- ^ arg will have argument type of 'fty' when we know enough to
|
||||
-- calculate it
|
||||
Unknown : ArgType vars
|
||||
@ -30,7 +30,7 @@ record PatInfo (pvar : Name) (vars : List Name) where
|
||||
constructor MkInfo
|
||||
pat : Pat
|
||||
loc : IsVar name idx vars
|
||||
argType : ArgType vars -- Type of the argument being inspected (i.e.
|
||||
argType : ArgType vars -- Type of the argument being inspected (i.e.
|
||||
-- *not* refined by this particular pattern)
|
||||
|
||||
Show (PatInfo n vars) where
|
||||
@ -52,7 +52,7 @@ data NamedPats : List Name -> -- pattern variables still to process
|
||||
-- in order
|
||||
Type where
|
||||
Nil : NamedPats vars []
|
||||
(::) : PatInfo pvar vars ->
|
||||
(::) : PatInfo pvar vars ->
|
||||
-- ^ a pattern, where its variable appears in the vars list,
|
||||
-- and its type. The type has no variable names; any names it
|
||||
-- refers to are explicit
|
||||
@ -63,7 +63,7 @@ getPatInfo [] = []
|
||||
getPatInfo (x :: xs) = pat x :: getPatInfo xs
|
||||
|
||||
updatePats : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars ->
|
||||
Env Term vars ->
|
||||
NF vars -> NamedPats vars todo -> Core (NamedPats vars todo)
|
||||
updatePats env nf [] = pure []
|
||||
updatePats {todo = pvar :: ns} env (NBind fc _ (Pi c _ farg) fsc) (p :: ps)
|
||||
@ -76,7 +76,7 @@ updatePats {todo = pvar :: ns} env (NBind fc _ (Pi c _ farg) fsc) (p :: ps)
|
||||
_ => pure (p :: ps)
|
||||
updatePats env nf (p :: ps)
|
||||
= case argType p of
|
||||
Unknown =>
|
||||
Unknown =>
|
||||
do defs <- get Ctxt
|
||||
empty <- clearDefs defs
|
||||
pure (record { argType = Stuck !(quote empty env nf) } p :: ps)
|
||||
@ -87,20 +87,20 @@ mkEnv fc [] = []
|
||||
mkEnv fc (n :: ns) = PVar RigW Explicit (Erased fc) :: mkEnv fc ns
|
||||
|
||||
substInPatInfo : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Term vars -> PatInfo pvar vars ->
|
||||
NamedPats vars todo ->
|
||||
FC -> Name -> Term vars -> PatInfo pvar vars ->
|
||||
NamedPats vars todo ->
|
||||
Core (PatInfo pvar vars, NamedPats vars todo)
|
||||
substInPatInfo {pvar} {vars} fc n tm p ps
|
||||
substInPatInfo {pvar} {vars} fc n tm p ps
|
||||
= case argType p of
|
||||
Known c ty => pure (record { argType = Known c (substName n tm ty) } p, ps)
|
||||
Stuck fty =>
|
||||
Stuck fty =>
|
||||
do defs <- get Ctxt
|
||||
empty <- clearDefs defs
|
||||
let env = mkEnv fc vars
|
||||
case !(nf defs env (substName n tm fty)) of
|
||||
NBind pfc _ (Pi c _ farg) fsc =>
|
||||
pure (record { argType = Known c !(quote empty env farg) } p,
|
||||
!(updatePats env
|
||||
!(updatePats env
|
||||
!(fsc defs (toClosure defaultOpts env
|
||||
(Ref pfc Bound pvar))) ps))
|
||||
_ => pure (p, ps)
|
||||
@ -109,10 +109,10 @@ substInPatInfo {pvar} {vars} fc n tm p ps
|
||||
-- Substitute the name with a term in the pattern types, and reduce further
|
||||
-- (this aims to resolve any 'Stuck' pattern types)
|
||||
substInPats : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Term vars -> NamedPats vars todo ->
|
||||
FC -> Name -> Term vars -> NamedPats vars todo ->
|
||||
Core (NamedPats vars todo)
|
||||
substInPats fc n tm [] = pure []
|
||||
substInPats fc n tm (p :: ps)
|
||||
substInPats fc n tm (p :: ps)
|
||||
= do (p', ps') <- substInPatInfo fc n tm p ps
|
||||
pure (p' :: !(substInPats fc n tm ps'))
|
||||
|
||||
@ -122,7 +122,7 @@ getPat First (x :: xs) = x
|
||||
getPat (Later p) (x :: xs) = getPat p xs
|
||||
|
||||
dropPat : {idx : Nat} ->
|
||||
.(el : IsVar name idx ps) ->
|
||||
.(el : IsVar name idx ps) ->
|
||||
NamedPats ns ps -> NamedPats ns (dropVar ps el)
|
||||
dropPat First (x :: xs) = xs
|
||||
dropPat (Later p) (x :: xs) = x :: dropPat p xs
|
||||
@ -133,8 +133,8 @@ Show (NamedPats vars todo) where
|
||||
showAll : NamedPats vs ts -> String
|
||||
showAll [] = ""
|
||||
showAll {ts = t :: _ } [x]
|
||||
= show t ++ " " ++ show (pat x) ++ " [" ++ show (argType x) ++ "]"
|
||||
showAll {ts = t :: _ } (x :: xs)
|
||||
= show t ++ " " ++ show (pat x) ++ " [" ++ show (argType x) ++ "]"
|
||||
showAll {ts = t :: _ } (x :: xs)
|
||||
= show t ++ " " ++ show (pat x) ++ " [" ++ show (argType x) ++ "]"
|
||||
++ ", " ++ showAll xs
|
||||
|
||||
@ -151,10 +151,10 @@ weaken : NamedPats vars todo -> NamedPats (x :: vars) todo
|
||||
weaken [] = []
|
||||
weaken (p :: ps) = weaken p :: weaken ps
|
||||
|
||||
weakenNs : (ns : List Name) ->
|
||||
weakenNs : (ns : List Name) ->
|
||||
NamedPats vars todo -> NamedPats (ns ++ vars) todo
|
||||
weakenNs ns [] = []
|
||||
weakenNs ns (p :: ps)
|
||||
weakenNs ns (p :: ps)
|
||||
= weakenNs ns p :: weakenNs ns ps
|
||||
|
||||
(++) : NamedPats vars ms -> NamedPats vars ns -> NamedPats vars (ms ++ ns)
|
||||
@ -170,18 +170,18 @@ take (x :: xs) (p :: ps) = p :: take xs ps
|
||||
|
||||
data PatClause : (vars : List Name) -> (todo : List Name) -> Type where
|
||||
MkPatClause : List Name -> -- names matched so far (from original lhs)
|
||||
NamedPats vars todo ->
|
||||
NamedPats vars todo ->
|
||||
(rhs : Term vars) -> PatClause vars todo
|
||||
|
||||
getNPs : PatClause vars todo -> NamedPats vars todo
|
||||
getNPs (MkPatClause _ lhs rhs) = lhs
|
||||
|
||||
Show (PatClause vars todo) where
|
||||
show (MkPatClause _ ps rhs)
|
||||
show (MkPatClause _ ps rhs)
|
||||
= show ps ++ " => " ++ show rhs
|
||||
|
||||
substInClause : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> PatClause vars (a :: todo) ->
|
||||
substInClause : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> PatClause vars (a :: todo) ->
|
||||
Core (PatClause vars (a :: todo))
|
||||
substInClause {vars} {a} fc (MkPatClause pvars (MkInfo pat pprf fty :: pats) rhs)
|
||||
= do pats' <- substInPats fc a (mkTerm vars pat) pats
|
||||
@ -195,7 +195,7 @@ checkLengthMatch : (xs : List a) -> (ys : List b) -> Maybe (LengthMatch xs ys)
|
||||
checkLengthMatch [] [] = Just NilMatch
|
||||
checkLengthMatch [] (x :: xs) = Nothing
|
||||
checkLengthMatch (x :: xs) [] = Nothing
|
||||
checkLengthMatch (x :: xs) (y :: ys)
|
||||
checkLengthMatch (x :: xs) (y :: ys)
|
||||
= Just (ConsMatch !(checkLengthMatch xs ys))
|
||||
|
||||
data Partitions : List (PatClause vars todo) -> Type where
|
||||
@ -255,11 +255,11 @@ clauseType phase (MkPatClause pvars (MkInfo arg _ ty :: rest) rhs)
|
||||
partition : Phase -> (ps : List (PatClause vars (a :: as))) -> Partitions ps
|
||||
partition phase [] = NoClauses
|
||||
partition phase (x :: xs) with (partition phase xs)
|
||||
partition phase (x :: (cs ++ ps)) | (ConClauses cs rest)
|
||||
partition phase (x :: (cs ++ ps)) | (ConClauses cs rest)
|
||||
= case clauseType phase x of
|
||||
ConClause => ConClauses (x :: cs) rest
|
||||
VarClause => VarClauses [x] (ConClauses cs rest)
|
||||
partition phase (x :: (vs ++ ps)) | (VarClauses vs rest)
|
||||
partition phase (x :: (vs ++ ps)) | (VarClauses vs rest)
|
||||
= case clauseType phase x of
|
||||
ConClause => ConClauses [x] (VarClauses vs rest)
|
||||
VarClause => VarClauses (x :: vs) rest
|
||||
@ -274,13 +274,13 @@ data ConType : Type where
|
||||
CConst : Constant -> ConType
|
||||
|
||||
conTypeEq : (x, y : ConType) -> Maybe (x = y)
|
||||
conTypeEq (CName x tag) (CName x' tag')
|
||||
conTypeEq (CName x tag) (CName x' tag')
|
||||
= do Refl <- nameEq x x'
|
||||
case decEq tag tag' of
|
||||
Yes Refl => Just Refl
|
||||
No contra => Nothing
|
||||
conTypeEq CDelay CDelay = Just Refl
|
||||
conTypeEq (CConst x) (CConst y)
|
||||
conTypeEq (CConst x) (CConst y)
|
||||
= case constantEq x y of
|
||||
Nothing => Nothing
|
||||
Just Refl => Just Refl
|
||||
@ -289,11 +289,11 @@ conTypeEq _ _ = Nothing
|
||||
data Group : List Name -> -- variables in scope
|
||||
List Name -> -- pattern variables still to process
|
||||
Type where
|
||||
ConGroup : Name -> (tag : Int) ->
|
||||
ConGroup : Name -> (tag : Int) ->
|
||||
List (PatClause (newargs ++ vars) (newargs ++ todo)) ->
|
||||
Group vars todo
|
||||
DelayGroup : List (PatClause (tyarg :: valarg :: vars)
|
||||
(tyarg :: valarg :: todo)) ->
|
||||
DelayGroup : List (PatClause (tyarg :: valarg :: vars)
|
||||
(tyarg :: valarg :: todo)) ->
|
||||
Group vars todo
|
||||
ConstGroup : Constant -> List (PatClause vars todo) ->
|
||||
Group vars todo
|
||||
@ -305,9 +305,9 @@ Show (Group vars todo) where
|
||||
|
||||
data GroupMatch : ConType -> List Pat -> Group vars todo -> Type where
|
||||
ConMatch : LengthMatch ps newargs ->
|
||||
GroupMatch (CName n tag) ps
|
||||
GroupMatch (CName n tag) ps
|
||||
(ConGroup {newargs} n tag (MkPatClause pvs pats rhs :: rest))
|
||||
DelayMatch : GroupMatch CDelay []
|
||||
DelayMatch : GroupMatch CDelay []
|
||||
(DelayGroup {tyarg} {valarg} (MkPatClause pvs pats rhs :: rest))
|
||||
ConstMatch : GroupMatch (CConst c) []
|
||||
(ConstGroup c (MkPatClause pvs pats rhs :: rest))
|
||||
@ -315,7 +315,7 @@ data GroupMatch : ConType -> List Pat -> Group vars todo -> Type where
|
||||
|
||||
checkGroupMatch : (c : ConType) -> (ps : List Pat) -> (g : Group vars todo) ->
|
||||
GroupMatch c ps g
|
||||
checkGroupMatch (CName x tag) ps (ConGroup {newargs} x' tag' (MkPatClause pvs pats rhs :: rest))
|
||||
checkGroupMatch (CName x tag) ps (ConGroup {newargs} x' tag' (MkPatClause pvs pats rhs :: rest))
|
||||
= case checkLengthMatch ps newargs of
|
||||
Nothing => NoMatch
|
||||
Just prf => case (nameEq x x', decEq tag tag') of
|
||||
@ -324,7 +324,7 @@ checkGroupMatch (CName x tag) ps (ConGroup {newargs} x' tag' (MkPatClause pvs pa
|
||||
checkGroupMatch (CName x tag) ps _ = NoMatch
|
||||
checkGroupMatch CDelay [] (DelayGroup (MkPatClause pvs pats rhs :: rest))
|
||||
= DelayMatch
|
||||
checkGroupMatch (CConst c) [] (ConstGroup c' (MkPatClause pvs pats rhs :: rest))
|
||||
checkGroupMatch (CConst c) [] (ConstGroup c' (MkPatClause pvs pats rhs :: rest))
|
||||
= case constantEq c c' of
|
||||
Nothing => NoMatch
|
||||
Just Refl => ConstMatch
|
||||
@ -370,9 +370,9 @@ nextNames {vars} fc root (p :: pats) fty
|
||||
-- replace the prefix of patterns with 'pargs'
|
||||
newPats : (pargs : List Pat) -> LengthMatch pargs ns ->
|
||||
NamedPats vars (ns ++ todo) ->
|
||||
NamedPats vars ns
|
||||
NamedPats vars ns
|
||||
newPats [] NilMatch rest = []
|
||||
newPats (newpat :: xs) (ConsMatch w) (pi :: rest)
|
||||
newPats (newpat :: xs) (ConsMatch w) (pi :: rest)
|
||||
= record { pat = newpat} pi :: newPats xs w rest
|
||||
|
||||
updateNames : List (Name, Pat) -> List (Name, Name)
|
||||
@ -388,7 +388,7 @@ updatePatNames ns (pi :: ps)
|
||||
= record { pat $= update } pi :: updatePatNames ns ps
|
||||
where
|
||||
update : Pat -> Pat
|
||||
update (PAs fc n p)
|
||||
update (PAs fc n p)
|
||||
= case lookup n ns of
|
||||
Nothing => PAs fc n (update p)
|
||||
Just n' => PAs fc n' (update p)
|
||||
@ -396,7 +396,7 @@ updatePatNames ns (pi :: ps)
|
||||
update (PTyCon fc n a ps) = PTyCon fc n a (map update ps)
|
||||
update (PArrow fc x s t) = PArrow fc x (update s) (update t)
|
||||
update (PDelay fc r t p) = PDelay fc r (update t) (update p)
|
||||
update (PLoc fc n)
|
||||
update (PLoc fc n)
|
||||
= case lookup n ns of
|
||||
Nothing => PLoc fc n
|
||||
Just n' => PLoc fc n'
|
||||
@ -406,12 +406,12 @@ groupCons : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name ->
|
||||
List Name ->
|
||||
List (PatClause vars (a :: todo)) ->
|
||||
List (PatClause vars (a :: todo)) ->
|
||||
Core (List (Group vars todo))
|
||||
groupCons fc fn pvars cs
|
||||
groupCons fc fn pvars cs
|
||||
= gc [] cs
|
||||
where
|
||||
addConG : Name -> (tag : Int) ->
|
||||
addConG : Name -> (tag : Int) ->
|
||||
List Pat -> NamedPats vars todo ->
|
||||
(rhs : Term vars) ->
|
||||
(acc : List (Group vars todo)) ->
|
||||
@ -420,7 +420,7 @@ groupCons fc fn pvars cs
|
||||
-- add new pattern arguments for each of that constructor's arguments.
|
||||
-- The type of 'ConGroup' ensures that we refer to the arguments by
|
||||
-- the same name in each of the clauses
|
||||
addConG {todo} n tag pargs pats rhs []
|
||||
addConG {todo} n tag pargs pats rhs []
|
||||
= do cty <- the (Core (NF vars)) $ if n == UN "->"
|
||||
then pure $ NBind fc (MN "_" 0) (Pi RigW Explicit (NType fc)) $
|
||||
(\d, a => pure $ NBind fc (MN "_" 1) (Pi RigW Explicit (NErased fc))
|
||||
@ -434,15 +434,15 @@ groupCons fc fn pvars cs
|
||||
-- explicit dependencies in types accurate)
|
||||
let pats' = updatePatNames (updateNames (zip patnames pargs))
|
||||
(weakenNs patnames pats)
|
||||
let clause = MkPatClause {todo = patnames ++ todo}
|
||||
pvars
|
||||
(newargs ++ pats')
|
||||
let clause = MkPatClause {todo = patnames ++ todo}
|
||||
pvars
|
||||
(newargs ++ pats')
|
||||
(weakenNs patnames rhs)
|
||||
pure [ConGroup n tag [clause]]
|
||||
addConG {todo} n tag pargs pats rhs (g :: gs) with (checkGroupMatch (CName n tag) pargs g)
|
||||
addConG {todo} n tag pargs pats rhs
|
||||
((ConGroup {newargs} n tag ((MkPatClause pvars ps tm) :: rest)) :: gs)
|
||||
| (ConMatch {newargs} lprf)
|
||||
| (ConMatch {newargs} lprf)
|
||||
= do let newps = newPats pargs lprf ps
|
||||
let pats' = updatePatNames (updateNames (zip newargs pargs))
|
||||
(weakenNs newargs pats)
|
||||
@ -454,7 +454,7 @@ groupCons fc fn pvars cs
|
||||
-- match the clauses top to bottom.
|
||||
pure ((ConGroup n tag (MkPatClause pvars ps tm :: rest ++ [newclause]))
|
||||
:: gs)
|
||||
addConG n tag pargs pats rhs (g :: gs) | NoMatch
|
||||
addConG n tag pargs pats rhs (g :: gs) | NoMatch
|
||||
= do gs' <- addConG n tag pargs pats rhs gs
|
||||
pure (g :: gs')
|
||||
|
||||
@ -468,11 +468,11 @@ groupCons fc fn pvars cs
|
||||
Core (List (Group vars todo))
|
||||
addDelayG {todo} pty parg pats rhs []
|
||||
= do let dty = NBind fc (MN "a" 0) (Pi Rig0 Explicit (NType fc)) $
|
||||
(\d, a =>
|
||||
(\d, a =>
|
||||
do a' <- evalClosure d a
|
||||
pure (NBind fc (MN "x" 0) (Pi RigW Explicit a')
|
||||
(\dv, av => pure (NDelayed fc LUnknown a'))))
|
||||
([tyname, argname] ** newargs) <- nextNames {vars} fc "e" [pty, parg]
|
||||
([tyname, argname] ** newargs) <- nextNames {vars} fc "e" [pty, parg]
|
||||
(Just dty)
|
||||
| _ => throw (InternalError "Error compiling Delay pattern match")
|
||||
let pats' = updatePatNames (updateNames [(tyname, pty),
|
||||
@ -490,9 +490,9 @@ groupCons fc fn pvars cs
|
||||
let pats' = updatePatNames (updateNames [(tyarg, pty),
|
||||
(valarg, parg)])
|
||||
(weakenNs [tyarg, valarg] pats)
|
||||
let newclause : PatClause (tyarg :: valarg :: vars)
|
||||
let newclause : PatClause (tyarg :: valarg :: vars)
|
||||
(tyarg :: valarg :: todo)
|
||||
= MkPatClause pvars (newps ++ pats')
|
||||
= MkPatClause pvars (newps ++ pats')
|
||||
(weakenNs [tyarg, valarg] rhs)
|
||||
pure ((DelayGroup (MkPatClause pvars ps tm :: rest ++ [newclause]))
|
||||
:: gs)
|
||||
@ -504,30 +504,30 @@ groupCons fc fn pvars cs
|
||||
(rhs : Term vars) ->
|
||||
(acc : List (Group vars todo)) ->
|
||||
Core (List (Group vars todo))
|
||||
addConstG c pats rhs []
|
||||
addConstG c pats rhs []
|
||||
= pure [ConstGroup c [MkPatClause pvars pats rhs]]
|
||||
addConstG {todo} c pats rhs (g :: gs) with (checkGroupMatch (CConst c) [] g)
|
||||
addConstG {todo} c pats rhs
|
||||
((ConstGroup c ((MkPatClause pvars ps tm) :: rest)) :: gs) | ConstMatch
|
||||
((ConstGroup c ((MkPatClause pvars ps tm) :: rest)) :: gs) | ConstMatch
|
||||
= let newclause : PatClause vars todo
|
||||
= MkPatClause pvars pats rhs in
|
||||
pure ((ConstGroup c
|
||||
pure ((ConstGroup c
|
||||
(MkPatClause pvars ps tm :: rest ++ [newclause])) :: gs)
|
||||
addConstG c pats rhs (g :: gs) | NoMatch
|
||||
addConstG c pats rhs (g :: gs) | NoMatch
|
||||
= do gs' <- addConstG c pats rhs gs
|
||||
pure (g :: gs')
|
||||
|
||||
|
||||
addGroup : {idx : Nat} -> Pat -> .(IsVar name idx vars) ->
|
||||
NamedPats vars todo -> Term vars ->
|
||||
List (Group vars todo) ->
|
||||
NamedPats vars todo -> Term vars ->
|
||||
List (Group vars todo) ->
|
||||
Core (List (Group vars todo))
|
||||
-- In 'As' replace the name on the RHS with a reference to the
|
||||
-- variable we're doing the case split on
|
||||
addGroup (PAs fc n p) pprf pats rhs acc
|
||||
addGroup (PAs fc n p) pprf pats rhs acc
|
||||
= addGroup p pprf pats (substName n (Local fc (Just True) _ pprf) rhs) acc
|
||||
addGroup (PCon _ n t a pargs) pprf pats rhs acc
|
||||
addGroup (PCon _ n t a pargs) pprf pats rhs acc
|
||||
= addConG n t pargs pats rhs acc
|
||||
addGroup (PTyCon _ n a pargs) pprf pats rhs acc
|
||||
addGroup (PTyCon _ n a pargs) pprf pats rhs acc
|
||||
= addConG n 0 pargs pats rhs acc
|
||||
addGroup (PArrow _ _ s t) pprf pats rhs acc
|
||||
= addConG (UN "->") 0 [s, t] pats rhs acc
|
||||
@ -535,16 +535,16 @@ groupCons fc fn pvars cs
|
||||
-- scrutinee (need to check in 'caseGroups below)
|
||||
addGroup (PDelay _ _ pty parg) pprf pats rhs acc
|
||||
= addDelayG pty parg pats rhs acc
|
||||
addGroup (PConst _ c) pprf pats rhs acc
|
||||
addGroup (PConst _ c) pprf pats rhs acc
|
||||
= addConstG c pats rhs acc
|
||||
addGroup _ pprf pats rhs acc = pure acc -- Can't happen, not a constructor
|
||||
-- -- FIXME: Is this possible to rule out with a type? Probably.
|
||||
|
||||
gc : List (Group vars todo) ->
|
||||
List (PatClause vars (a :: todo)) ->
|
||||
gc : List (Group vars todo) ->
|
||||
List (PatClause vars (a :: todo)) ->
|
||||
Core (List (Group vars todo))
|
||||
gc acc [] = pure acc
|
||||
gc {a} acc ((MkPatClause pvars (MkInfo pat pprf fty :: pats) rhs) :: cs)
|
||||
gc {a} acc ((MkPatClause pvars (MkInfo pat pprf fty :: pats) rhs) :: cs)
|
||||
= do acc' <- addGroup pat pprf pats rhs acc
|
||||
gc acc' cs
|
||||
|
||||
@ -555,12 +555,12 @@ getFirstArgType : NamedPats ns (p :: ps) -> ArgType ns
|
||||
getFirstArgType (p :: _) = argType p
|
||||
|
||||
-- Check whether all the initial patterns have the same concrete, known
|
||||
-- and matchable type, which is multiplicity > 0.
|
||||
-- and matchable type, which is multiplicity > 0.
|
||||
-- If so, it's okay to match on it
|
||||
sameType : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name ->
|
||||
Env Term ns -> List (NamedPats ns (p :: ps)) ->
|
||||
Env Term ns -> List (NamedPats ns (p :: ps)) ->
|
||||
Core ()
|
||||
sameType fc fn env [] = pure ()
|
||||
sameType {ns} fc fn env (p :: xs)
|
||||
@ -580,10 +580,10 @@ sameType {ns} fc fn env (p :: xs)
|
||||
|
||||
sameTypeAs : NF ns -> List (ArgType ns) -> Core ()
|
||||
sameTypeAs ty [] = pure ()
|
||||
sameTypeAs ty (Known Rig0 t :: xs)
|
||||
sameTypeAs ty (Known Rig0 t :: xs)
|
||||
= throw (CaseCompile fc fn (MatchErased (_ ** (env, mkTerm _ (firstPat p)))))
|
||||
-- Can't match on erased thing
|
||||
sameTypeAs ty (Known c t :: xs)
|
||||
sameTypeAs ty (Known c t :: xs)
|
||||
= do defs <- get Ctxt
|
||||
if headEq ty !(nf defs env t)
|
||||
then sameTypeAs ty xs
|
||||
@ -594,8 +594,8 @@ sameType {ns} fc fn env (p :: xs)
|
||||
-- If so, we'll match it to refine later types and move on
|
||||
samePat : List (NamedPats ns (p :: ps)) -> Bool
|
||||
samePat [] = True
|
||||
samePat (pi :: xs)
|
||||
= samePatAs (dropAs (getFirstPat pi))
|
||||
samePat (pi :: xs)
|
||||
= samePatAs (dropAs (getFirstPat pi))
|
||||
(map (dropAs . getFirstPat) xs)
|
||||
where
|
||||
dropAs : Pat -> Pat
|
||||
@ -613,7 +613,7 @@ samePat (pi :: xs)
|
||||
then samePatAs (PCon fc n t a args) ps
|
||||
else False
|
||||
samePatAs (PConst fc c) (PConst _ c' :: ps)
|
||||
= if c == c'
|
||||
= if c == c'
|
||||
then samePatAs (PConst fc c) ps
|
||||
else False
|
||||
samePatAs (PArrow fc x s t) (PArrow _ _ s' t' :: ps)
|
||||
@ -652,17 +652,17 @@ countDiff xs = length (distinct [] (map getFirstCon xs))
|
||||
|
||||
distinct : List Pat -> List Pat -> List Pat
|
||||
distinct acc [] = acc
|
||||
distinct acc (p :: ps)
|
||||
= if elemBy sameCase p acc
|
||||
distinct acc (p :: ps)
|
||||
= if elemBy sameCase p acc
|
||||
then distinct acc ps
|
||||
else distinct (p :: acc) ps
|
||||
|
||||
getScore : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name ->
|
||||
List (NamedPats ns (p :: ps)) ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name ->
|
||||
List (NamedPats ns (p :: ps)) ->
|
||||
Core (Either CaseError ())
|
||||
getScore fc name npss
|
||||
getScore fc name npss
|
||||
= do catch (do sameType fc name (mkEnv fc ns) npss
|
||||
pure (Right ()))
|
||||
(\err => case err of
|
||||
@ -673,10 +673,10 @@ getScore fc name npss
|
||||
-- same family, or all variables, or all the same type constructor.
|
||||
pickNext : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> List (NamedPats ns (p :: ps)) ->
|
||||
FC -> Name -> List (NamedPats ns (p :: ps)) ->
|
||||
Core (Var (p :: ps))
|
||||
-- last possible variable
|
||||
pickNext {ps = []} fc fn npss
|
||||
pickNext {ps = []} fc fn npss
|
||||
= if samePat npss
|
||||
then pure (MkVar First)
|
||||
else do Right () <- getScore fc fn npss
|
||||
@ -699,7 +699,7 @@ shuffleVars : {idx : Nat} -> .(el : IsVar name idx todo) -> PatClause vars todo
|
||||
shuffleVars el (MkPatClause pvars lhs rhs) = MkPatClause pvars (moveFirst el lhs) rhs
|
||||
|
||||
mutual
|
||||
{- 'PatClause' contains a list of patterns still to process (that's the
|
||||
{- 'PatClause' contains a list of patterns still to process (that's the
|
||||
"todo") and a right hand side with the variables we know about "vars".
|
||||
So "match" builds the remainder of the case tree for
|
||||
the unprocessed patterns. "err" is the tree for when the patterns don't
|
||||
@ -708,7 +708,7 @@ mutual
|
||||
match : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Phase ->
|
||||
List (PatClause vars todo) -> (err : Maybe (CaseTree vars)) ->
|
||||
List (PatClause vars todo) -> (err : Maybe (CaseTree vars)) ->
|
||||
Core (CaseTree vars)
|
||||
-- Before 'partition', reorder the arguments so that the one we
|
||||
-- inspect next has a concrete type that is the same in all cases, and
|
||||
@ -720,14 +720,14 @@ mutual
|
||||
maybe (pure (Unmatched "No clauses"))
|
||||
pure
|
||||
!(mixture fc fn phase ps err)
|
||||
match {todo = []} fc fn phase [] err
|
||||
match {todo = []} fc fn phase [] err
|
||||
= maybe (pure (Unmatched "No patterns"))
|
||||
pure err
|
||||
match {todo = []} fc fn phase ((MkPatClause pvars [] rhs) :: _) err
|
||||
match {todo = []} fc fn phase ((MkPatClause pvars [] rhs) :: _) err
|
||||
= pure $ STerm rhs
|
||||
|
||||
caseGroups : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Phase ->
|
||||
{idx : Nat} -> .(IsVar pvar idx vars) -> Term vars ->
|
||||
List (Group vars todo) -> Maybe (CaseTree vars) ->
|
||||
@ -737,10 +737,10 @@ mutual
|
||||
pure (Case _ el (resolveNames vars ty) g)
|
||||
where
|
||||
altGroups : List (Group vars todo) -> Core (List (CaseAlt vars))
|
||||
altGroups [] = maybe (pure [])
|
||||
(\e => pure [DefaultCase e])
|
||||
altGroups [] = maybe (pure [])
|
||||
(\e => pure [DefaultCase e])
|
||||
errorCase
|
||||
altGroups (ConGroup {newargs} cn tag rest :: cs)
|
||||
altGroups (ConGroup {newargs} cn tag rest :: cs)
|
||||
= do crest <- match fc fn phase rest (map (weakenNs newargs) errorCase)
|
||||
cs' <- altGroups cs
|
||||
pure (ConCase cn tag newargs crest :: cs')
|
||||
@ -754,17 +754,17 @@ mutual
|
||||
pure (ConstCase c crest :: cs')
|
||||
|
||||
conRule : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Phase ->
|
||||
List (PatClause vars (a :: todo)) ->
|
||||
Maybe (CaseTree vars) ->
|
||||
Maybe (CaseTree vars) ->
|
||||
Core (CaseTree vars)
|
||||
conRule fc fn phase [] err = maybe (pure (Unmatched "No constructor clauses")) pure err
|
||||
conRule fc fn phase [] err = maybe (pure (Unmatched "No constructor clauses")) pure err
|
||||
-- ASSUMPTION, not expressed in the type, that the patterns all have
|
||||
-- the same variable (pprf) for the first argument. If not, the result
|
||||
-- will be a broken case tree... so we should find a way to express this
|
||||
-- in the type if we can.
|
||||
conRule {a} fc fn phase cs@(MkPatClause pvars (MkInfo pat pprf fty :: pats) rhs :: rest) err
|
||||
conRule {a} fc fn phase cs@(MkPatClause pvars (MkInfo pat pprf fty :: pats) rhs :: rest) err
|
||||
= do refinedcs <- traverse (substInClause fc) cs
|
||||
groups <- groupCons fc fn pvars refinedcs
|
||||
ty <- case fty of
|
||||
@ -773,19 +773,19 @@ mutual
|
||||
caseGroups fc fn phase pprf ty groups err
|
||||
|
||||
varRule : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Phase ->
|
||||
List (PatClause vars (a :: todo)) ->
|
||||
Maybe (CaseTree vars) ->
|
||||
Maybe (CaseTree vars) ->
|
||||
Core (CaseTree vars)
|
||||
varRule {vars} {a} fc fn phase cs err
|
||||
varRule {vars} {a} fc fn phase cs err
|
||||
= do alts' <- traverse updateVar cs
|
||||
match fc fn phase alts' err
|
||||
where
|
||||
updateVar : PatClause vars (a :: todo) -> Core (PatClause vars todo)
|
||||
-- replace the name with the relevant variable on the rhs
|
||||
updateVar (MkPatClause pvars (MkInfo (PLoc pfc n) prf fty :: pats) rhs)
|
||||
= pure $ MkPatClause (n :: pvars)
|
||||
= pure $ MkPatClause (n :: pvars)
|
||||
!(substInPats fc a (Local pfc (Just False) _ prf) pats)
|
||||
(substName n (Local pfc (Just False) _ prf) rhs)
|
||||
-- If it's an as pattern, replace the name with the relevant variable on
|
||||
@ -797,41 +797,41 @@ mutual
|
||||
-- match anything, name won't appear in rhs but need to update
|
||||
-- LHS pattern types based on what we've learned
|
||||
updateVar (MkPatClause pvars (MkInfo pat prf fty :: pats) rhs)
|
||||
= pure $ MkPatClause pvars
|
||||
= pure $ MkPatClause pvars
|
||||
!(substInPats fc a (mkTerm vars pat) pats) rhs
|
||||
|
||||
mixture : {auto i : Ref PName Int} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{ps : List (PatClause vars (a :: todo))} ->
|
||||
FC -> Name -> Phase ->
|
||||
Partitions ps ->
|
||||
Maybe (CaseTree vars) ->
|
||||
Partitions ps ->
|
||||
Maybe (CaseTree vars) ->
|
||||
Core (Maybe (CaseTree vars))
|
||||
mixture fc fn phase (ConClauses cs rest) err
|
||||
mixture fc fn phase (ConClauses cs rest) err
|
||||
= do fallthrough <- mixture fc fn phase rest err
|
||||
pure (Just !(conRule fc fn phase cs fallthrough))
|
||||
mixture fc fn phase (VarClauses vs rest) err
|
||||
mixture fc fn phase (VarClauses vs rest) err
|
||||
= do fallthrough <- mixture fc fn phase rest err
|
||||
pure (Just !(varRule fc fn phase vs fallthrough))
|
||||
mixture fc fn {a} {todo} phase NoClauses err
|
||||
mixture fc fn {a} {todo} phase NoClauses err
|
||||
= pure err
|
||||
|
||||
mkPatClause : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name ->
|
||||
(args : List Name) -> ClosedTerm -> (List Pat, ClosedTerm) ->
|
||||
Core (PatClause args args)
|
||||
mkPatClause fc fn args ty (ps, rhs)
|
||||
mkPatClause fc fn args ty (ps, rhs)
|
||||
= maybe (throw (CaseCompile fc fn DifferingArgNumbers))
|
||||
(\eq =>
|
||||
(\eq =>
|
||||
do defs <- get Ctxt
|
||||
nty <- nf defs [] ty
|
||||
ns <- mkNames args ps eq (Just nty)
|
||||
pure (MkPatClause [] ns
|
||||
(rewrite sym (appendNilRightNeutral args) in
|
||||
(rewrite sym (appendNilRightNeutral args) in
|
||||
(weakenNs args rhs))))
|
||||
(checkLengthMatch args ps)
|
||||
where
|
||||
mkNames : (vars : List Name) -> (ps : List Pat) ->
|
||||
mkNames : (vars : List Name) -> (ps : List Pat) ->
|
||||
LengthMatch vars ps -> Maybe (NF []) ->
|
||||
Core (NamedPats vars vars)
|
||||
mkNames [] [] NilMatch fty = pure []
|
||||
@ -841,33 +841,33 @@ mkPatClause fc fn args ty (ps, rhs)
|
||||
fa_tys <-
|
||||
case fty of
|
||||
Nothing => pure (Nothing, Unknown)
|
||||
Just (NBind pfc _ (Pi c _ farg) fsc) =>
|
||||
Just (NBind pfc _ (Pi c _ farg) fsc) =>
|
||||
pure (Just !(fsc defs (toClosure defaultOpts [] (Ref pfc Bound arg))),
|
||||
Known c (embed {more = arg :: args}
|
||||
Known c (embed {more = arg :: args}
|
||||
!(quote empty [] farg)))
|
||||
Just t =>
|
||||
pure (Nothing,
|
||||
Stuck (embed {more = arg :: args}
|
||||
Just t =>
|
||||
pure (Nothing,
|
||||
Stuck (embed {more = arg :: args}
|
||||
!(quote empty [] t)))
|
||||
pure (MkInfo p First (snd fa_tys)
|
||||
:: weaken !(mkNames args ps eq (fst fa_tys)))
|
||||
|
||||
export
|
||||
patCompile : {auto c : Ref Ctxt Defs} ->
|
||||
patCompile : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Phase ->
|
||||
ClosedTerm -> List (List Pat, ClosedTerm) ->
|
||||
ClosedTerm -> List (List Pat, ClosedTerm) ->
|
||||
Maybe (CaseTree []) ->
|
||||
Core (args ** CaseTree args)
|
||||
patCompile fc fn phase ty [] def
|
||||
patCompile fc fn phase ty [] def
|
||||
= maybe (pure ([] ** Unmatched "No definition"))
|
||||
(\e => pure ([] ** e))
|
||||
def
|
||||
patCompile fc fn phase ty (p :: ps) def
|
||||
patCompile fc fn phase ty (p :: ps) def
|
||||
= do let ns = getNames 0 (fst p)
|
||||
pats <- traverse (mkPatClause fc fn ns ty) (p :: ps)
|
||||
log 5 $ "Pattern clauses " ++ show pats
|
||||
i <- newRef PName (the Int 0)
|
||||
cases <- match fc fn phase pats
|
||||
cases <- match fc fn phase pats
|
||||
(rewrite sym (appendNilRightNeutral ns) in
|
||||
map (weakenNs ns) def)
|
||||
pure (_ ** cases)
|
||||
@ -898,7 +898,7 @@ simpleCase : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Phase -> Name -> ClosedTerm -> (def : Maybe (CaseTree [])) ->
|
||||
(clauses : List (ClosedTerm, ClosedTerm)) ->
|
||||
Core (args ** CaseTree args)
|
||||
simpleCase fc phase fn ty def clauses
|
||||
simpleCase fc phase fn ty def clauses
|
||||
= do log 5 $ "Compiling clauses " ++ show clauses
|
||||
ps <- traverse (toPatClause fc fn) clauses
|
||||
defs <- get Ctxt
|
||||
@ -906,7 +906,7 @@ simpleCase fc phase fn ty def clauses
|
||||
|
||||
export
|
||||
getPMDef : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Phase -> Name -> ClosedTerm -> List Clause ->
|
||||
FC -> Phase -> Name -> ClosedTerm -> List Clause ->
|
||||
Core (args ** CaseTree args)
|
||||
-- If there's no clauses, make a definition with the right number of arguments
|
||||
-- for the type, which we can use in coverage checking to ensure that one of
|
||||
@ -928,15 +928,15 @@ getPMDef fc phase fn ty clauses
|
||||
where
|
||||
mkSubstEnv : Int -> Env Term vars -> SubstEnv vars []
|
||||
mkSubstEnv i [] = Nil
|
||||
mkSubstEnv i (v :: vs)
|
||||
mkSubstEnv i (v :: vs)
|
||||
= Ref fc Bound (MN "pat" i) :: mkSubstEnv (i + 1) vs
|
||||
|
||||
close : Env Term vars -> Term vars -> ClosedTerm
|
||||
close {vars} env tm
|
||||
= substs (mkSubstEnv 0 env)
|
||||
close {vars} env tm
|
||||
= substs (mkSubstEnv 0 env)
|
||||
(rewrite appendNilRightNeutral vars in tm)
|
||||
|
||||
toClosed : Defs -> Clause -> (ClosedTerm, ClosedTerm)
|
||||
toClosed defs (MkClause env lhs rhs)
|
||||
toClosed defs (MkClause env lhs rhs)
|
||||
= (close env lhs, close env rhs)
|
||||
|
||||
|
@ -9,7 +9,7 @@ mutual
|
||||
public export
|
||||
data CaseTree : List Name -> Type where
|
||||
Case : {name : _} ->
|
||||
(idx : Nat) ->
|
||||
(idx : Nat) ->
|
||||
IsVar name idx vars ->
|
||||
(scTy : Term vars) -> List (CaseAlt vars) ->
|
||||
CaseTree vars
|
||||
@ -21,7 +21,7 @@ mutual
|
||||
data CaseAlt : List Name -> Type where
|
||||
ConCase : Name -> (tag : Int) -> (args : List Name) ->
|
||||
CaseTree (args ++ vars) -> CaseAlt vars
|
||||
DelayCase : (ty : Name) -> (arg : Name) ->
|
||||
DelayCase : (ty : Name) -> (arg : Name) ->
|
||||
CaseTree (ty :: arg :: vars) -> CaseAlt vars
|
||||
ConstCase : Constant -> CaseTree vars -> CaseAlt vars
|
||||
DefaultCase : CaseTree vars -> CaseAlt vars
|
||||
@ -75,7 +75,7 @@ Show Pat where
|
||||
mutual
|
||||
insertCaseNames : (ns : List Name) -> CaseTree (outer ++ inner) ->
|
||||
CaseTree (outer ++ (ns ++ inner))
|
||||
insertCaseNames {inner} {outer} ns (Case idx prf scTy alts)
|
||||
insertCaseNames {inner} {outer} ns (Case idx prf scTy alts)
|
||||
= let MkVar prf' = insertVarNames {outer} {inner} {ns} _ prf in
|
||||
Case _ prf' (insertNames {outer} ns scTy)
|
||||
(map (insertCaseAltNames {outer} {inner} ns) alts)
|
||||
@ -83,26 +83,26 @@ mutual
|
||||
insertCaseNames ns (Unmatched msg) = Unmatched msg
|
||||
insertCaseNames ns Impossible = Impossible
|
||||
|
||||
insertCaseAltNames : (ns : List Name) ->
|
||||
CaseAlt (outer ++ inner) ->
|
||||
insertCaseAltNames : (ns : List Name) ->
|
||||
CaseAlt (outer ++ inner) ->
|
||||
CaseAlt (outer ++ (ns ++ inner))
|
||||
insertCaseAltNames {outer} {inner} ns (ConCase x tag args ct)
|
||||
= ConCase x tag args
|
||||
insertCaseAltNames {outer} {inner} ns (ConCase x tag args ct)
|
||||
= ConCase x tag args
|
||||
(rewrite appendAssociative args outer (ns ++ inner) in
|
||||
insertCaseNames {outer = args ++ outer} {inner} ns
|
||||
(rewrite sym (appendAssociative args outer inner) in
|
||||
ct))
|
||||
insertCaseAltNames {outer} {inner} ns (DelayCase tyn valn ct)
|
||||
= DelayCase tyn valn
|
||||
= DelayCase tyn valn
|
||||
(insertCaseNames {outer = tyn :: valn :: outer} {inner} ns ct)
|
||||
insertCaseAltNames ns (ConstCase x ct)
|
||||
insertCaseAltNames ns (ConstCase x ct)
|
||||
= ConstCase x (insertCaseNames ns ct)
|
||||
insertCaseAltNames ns (DefaultCase ct)
|
||||
insertCaseAltNames ns (DefaultCase ct)
|
||||
= DefaultCase (insertCaseNames ns ct)
|
||||
|
||||
|
||||
export
|
||||
thinTree : (n : Name) -> CaseTree (outer ++ inner) -> CaseTree (outer ++ n :: inner)
|
||||
thinTree n (Case idx prf scTy alts)
|
||||
thinTree n (Case idx prf scTy alts)
|
||||
= let MkVar prf' = insertVar {n} _ prf in
|
||||
Case _ prf' (thin n scTy) (map (insertCaseAltNames [n]) alts)
|
||||
thinTree n (STerm tm) = STerm (thin n tm)
|
||||
@ -111,7 +111,7 @@ thinTree n Impossible = Impossible
|
||||
|
||||
export
|
||||
Weaken CaseTree where
|
||||
weakenNs ns t = insertCaseNames {outer = []} ns t
|
||||
weakenNs ns t = insertCaseNames {outer = []} ns t
|
||||
|
||||
getNames : ({vs : _} -> NameMap Bool -> Term vs -> NameMap Bool) ->
|
||||
NameMap Bool -> CaseTree vars -> NameMap Bool
|
||||
@ -126,7 +126,7 @@ getNames add ns sc = getSet ns sc
|
||||
|
||||
getAltSets : NameMap Bool -> List (CaseAlt vs) -> NameMap Bool
|
||||
getAltSets ns [] = ns
|
||||
getAltSets ns (a :: as)
|
||||
getAltSets ns (a :: as)
|
||||
= assert_total $ getAltSets (getAltSet ns a) as
|
||||
|
||||
getSet : NameMap Bool -> CaseTree vs -> NameMap Bool
|
||||
@ -155,16 +155,16 @@ mkPat' args orig (Ref fc (TyCon t a) n) = PTyCon fc n a args
|
||||
mkPat' args orig (Bind fc x (Pi _ _ s) t)
|
||||
= let t' = subst (Erased fc) t in
|
||||
PArrow fc x (mkPat' [] s s) (mkPat' [] t' t')
|
||||
mkPat' args orig (App fc fn arg)
|
||||
mkPat' args orig (App fc fn arg)
|
||||
= let parg = mkPat' [] arg arg in
|
||||
mkPat' (parg :: args) orig fn
|
||||
mkPat' args orig (As fc (Ref _ Bound n) ptm)
|
||||
mkPat' args orig (As fc (Ref _ Bound n) ptm)
|
||||
= PAs fc n (mkPat' [] ptm ptm)
|
||||
mkPat' args orig (As fc _ ptm)
|
||||
mkPat' args orig (As fc _ ptm)
|
||||
= mkPat' [] orig ptm
|
||||
mkPat' args orig (TDelay fc r ty p)
|
||||
mkPat' args orig (TDelay fc r ty p)
|
||||
= PDelay fc r (mkPat' [] orig ty) (mkPat' [] orig p)
|
||||
mkPat' args orig (PrimVal fc c)
|
||||
mkPat' args orig (PrimVal fc c)
|
||||
= if constTag c == 0
|
||||
then PConst fc c
|
||||
else PTyCon fc (UN (show c)) 0 []
|
||||
@ -173,24 +173,24 @@ mkPat' args orig tm = PUnmatchable (getLoc orig) orig
|
||||
|
||||
export
|
||||
argToPat : ClosedTerm -> Pat
|
||||
argToPat tm
|
||||
argToPat tm
|
||||
= mkPat' [] tm tm
|
||||
|
||||
export
|
||||
mkTerm : (vars : List Name) -> Pat -> Term vars
|
||||
mkTerm vars (PAs fc x y) = mkTerm vars y
|
||||
mkTerm vars (PCon fc x tag arity xs)
|
||||
mkTerm vars (PCon fc x tag arity xs)
|
||||
= apply fc (Ref fc (DataCon tag arity) x)
|
||||
(map (mkTerm vars) xs)
|
||||
mkTerm vars (PTyCon fc x arity xs)
|
||||
mkTerm vars (PTyCon fc x arity xs)
|
||||
= apply fc (Ref fc (TyCon 0 arity) x)
|
||||
(map (mkTerm vars) xs)
|
||||
mkTerm vars (PConst fc c) = PrimVal fc c
|
||||
mkTerm vars (PArrow fc x s t)
|
||||
mkTerm vars (PArrow fc x s t)
|
||||
= Bind fc x (Pi RigW Explicit (mkTerm vars s)) (mkTerm (x :: vars) t)
|
||||
mkTerm vars (PDelay fc r ty p)
|
||||
= TDelay fc r (mkTerm vars ty) (mkTerm vars p)
|
||||
mkTerm vars (PLoc fc n)
|
||||
mkTerm vars (PLoc fc n)
|
||||
= case isVar n vars of
|
||||
Just (MkVar prf) => Local fc Nothing _ prf
|
||||
_ => Ref fc Bound n
|
||||
|
@ -74,8 +74,8 @@ data CDef : Type where
|
||||
-- Constructor
|
||||
MkCon : (tag : Int) -> (arity : Nat) -> CDef
|
||||
-- Foreign definition
|
||||
MkForeign : (ccs : List String) ->
|
||||
(fargs : List CFType) ->
|
||||
MkForeign : (ccs : List String) ->
|
||||
(fargs : List CFType) ->
|
||||
CFType ->
|
||||
CDef
|
||||
-- A function which will fail at runtime (usually due to being a hole) so needs
|
||||
@ -135,15 +135,15 @@ export
|
||||
Show CDef where
|
||||
show (MkFun args exp) = show args ++ ": " ++ show exp
|
||||
show (MkCon tag arity) = "Constructor tag " ++ show tag ++ " arity " ++ show arity
|
||||
show (MkForeign ccs args ret)
|
||||
= "Foreign call " ++ show ccs ++ " " ++
|
||||
show (MkForeign ccs args ret)
|
||||
= "Foreign call " ++ show ccs ++ " " ++
|
||||
show args ++ " -> " ++ show ret
|
||||
show (MkError exp) = "Error: " ++ show exp
|
||||
|
||||
mutual
|
||||
export
|
||||
thin : (n : Name) -> CExp (outer ++ inner) -> CExp (outer ++ n :: inner)
|
||||
thin n (CLocal fc prf)
|
||||
thin n (CLocal fc prf)
|
||||
= let MkVar var' = insertVar {n} _ prf in
|
||||
CLocal fc var'
|
||||
thin _ (CRef fc x) = CRef fc x
|
||||
@ -219,15 +219,15 @@ mutual
|
||||
-- in the remaining set with Erased
|
||||
export
|
||||
shrinkCExp : SubVars newvars vars -> CExp vars -> CExp newvars
|
||||
shrinkCExp sub (CLocal fc prf)
|
||||
shrinkCExp sub (CLocal fc prf)
|
||||
= case subElem prf sub of
|
||||
Nothing => CErased fc
|
||||
Just (MkVar prf') => CLocal fc prf'
|
||||
shrinkCExp _ (CRef fc x) = CRef fc x
|
||||
shrinkCExp sub (CLam fc x sc)
|
||||
shrinkCExp sub (CLam fc x sc)
|
||||
= let sc' = shrinkCExp (KeepCons sub) sc in
|
||||
CLam fc x sc'
|
||||
shrinkCExp sub (CLet fc x val sc)
|
||||
shrinkCExp sub (CLet fc x val sc)
|
||||
= let sc' = shrinkCExp (KeepCons sub) sc in
|
||||
CLet fc x (shrinkCExp sub val) sc'
|
||||
shrinkCExp sub (CApp fc x xs)
|
||||
@ -241,11 +241,11 @@ mutual
|
||||
shrinkCExp sub (CForce fc x) = CForce fc (shrinkCExp sub x)
|
||||
shrinkCExp sub (CDelay fc x) = CDelay fc (shrinkCExp sub x)
|
||||
shrinkCExp sub (CConCase fc sc xs def)
|
||||
= CConCase fc (shrinkCExp sub sc)
|
||||
= CConCase fc (shrinkCExp sub sc)
|
||||
(assert_total (map (shrinkConAlt sub) xs))
|
||||
(assert_total (map (shrinkCExp sub) def))
|
||||
shrinkCExp sub (CConstCase fc sc xs def)
|
||||
= CConstCase fc (shrinkCExp sub sc)
|
||||
= CConstCase fc (shrinkCExp sub sc)
|
||||
(assert_total (map (shrinkConstAlt sub) xs))
|
||||
(assert_total (map (shrinkCExp sub) def))
|
||||
shrinkCExp _ (CPrimVal fc x) = CPrimVal fc x
|
||||
|
@ -429,6 +429,23 @@ newDef fc n rig vars ty vis def
|
||||
= MkGlobalDef fc n ty [] rig vars vis unchecked [] empty False False False def
|
||||
Nothing []
|
||||
|
||||
-- Rewrite rules, applied after type checking, for runtime code only
|
||||
-- LHS and RHS must have the same type, but we don't (currently) require that
|
||||
-- the result still type checks (which might happen e.g. if transforming to a
|
||||
-- faster implementation with different behaviour)
|
||||
-- (Q: Do we need the 'Env' here? Usually we end up needing an 'Env' with a
|
||||
-- 'NF but we're working with terms rather than values...)
|
||||
public export
|
||||
data Transform : Type where
|
||||
MkTransform : Env Term vars -> Term vars -> Term vars -> Transform
|
||||
|
||||
export
|
||||
getFnName : Transform -> Maybe Name
|
||||
getFnName (MkTransform _ app _)
|
||||
= case getFn app of
|
||||
Ref _ _ fn => Just fn
|
||||
_ => Nothing
|
||||
|
||||
public export
|
||||
interface HasNames a where
|
||||
full : Context -> a -> Core a
|
||||
@ -469,8 +486,8 @@ HasNames (Term vars) where
|
||||
= pure (TDelayed fc x !(full gam y))
|
||||
full gam (TDelay fc x t y)
|
||||
= pure (TDelay fc x !(full gam t) !(full gam y))
|
||||
full gam (TForce fc y)
|
||||
= pure (TForce fc !(full gam y))
|
||||
full gam (TForce fc r y)
|
||||
= pure (TForce fc r !(full gam y))
|
||||
full gam tm = pure tm
|
||||
|
||||
resolved gam (Ref fc x n)
|
||||
@ -492,8 +509,8 @@ HasNames (Term vars) where
|
||||
= pure (TDelayed fc x !(resolved gam y))
|
||||
resolved gam (TDelay fc x t y)
|
||||
= pure (TDelay fc x !(resolved gam t) !(resolved gam y))
|
||||
resolved gam (TForce fc y)
|
||||
= pure (TForce fc !(resolved gam y))
|
||||
resolved gam (TForce fc r y)
|
||||
= pure (TForce fc r !(resolved gam y))
|
||||
resolved gam tm = pure tm
|
||||
|
||||
mutual
|
||||
@ -673,6 +690,15 @@ HasNames GlobalDef where
|
||||
sizeChange = !(traverse (resolved gam) (sizeChange def))
|
||||
} def
|
||||
|
||||
export
|
||||
HasNames Transform where
|
||||
full gam (MkTransform env lhs rhs)
|
||||
= pure $ MkTransform !(full gam env) !(full gam lhs) !(full gam rhs)
|
||||
|
||||
resolved gam (MkTransform env lhs rhs)
|
||||
= pure $ MkTransform !(resolved gam env)
|
||||
!(resolved gam lhs) !(resolved gam rhs)
|
||||
|
||||
public export
|
||||
record Defs where
|
||||
constructor MkDefs
|
||||
@ -701,6 +727,10 @@ record Defs where
|
||||
-- We don't look up anything in here, it's merely for saving out to TTC.
|
||||
-- We save the hints in the 'GlobalDef' itself for faster lookup.
|
||||
saveAutoHints : List (Name, Bool)
|
||||
transforms : NameMap Transform
|
||||
-- ^ A mapping from names to transformation rules which update applications
|
||||
-- of that name
|
||||
saveTransforms : List (Name, Transform)
|
||||
namedirectives : List (Name, List String)
|
||||
ifaceHash : Int
|
||||
importHashes : List (List String, Int)
|
||||
@ -737,9 +767,9 @@ initDefs : Core Defs
|
||||
initDefs
|
||||
= do gam <- initCtxt
|
||||
pure (MkDefs gam [] ["Main"] [] defaults empty 100
|
||||
empty empty empty [] [] [] 5381 [] [] [] [] [] empty
|
||||
empty)
|
||||
|
||||
empty empty empty [] [] empty []
|
||||
[] 5381 [] [] [] [] [] empty empty)
|
||||
|
||||
-- Reset the context, except for the options
|
||||
export
|
||||
clearCtxt : {auto c : Ref Ctxt Defs} ->
|
||||
@ -1298,6 +1328,17 @@ setOpenHints hs
|
||||
= do d <- get Ctxt
|
||||
put Ctxt (record { openHints = hs } d)
|
||||
|
||||
export
|
||||
addTransform : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Transform -> Core ()
|
||||
addTransform fc t
|
||||
= do defs <- get Ctxt
|
||||
let Just fn = getFnName t
|
||||
| Nothing =>
|
||||
throw (GenericMsg fc "LHS of a transformation must be a function application")
|
||||
put Ctxt (record { transforms $= insert fn t,
|
||||
saveTransforms $= ((fn, t) ::) } defs)
|
||||
|
||||
export
|
||||
clearSavedHints : {auto c : Ref Ctxt Defs} -> Core ()
|
||||
clearSavedHints
|
||||
@ -1898,7 +1939,7 @@ logTimeOver nsecs str act
|
||||
assert_total $ -- We're not dividing by 0
|
||||
do str' <- str
|
||||
coreLift $ putStrLn $ "TIMING " ++ str' ++ ": " ++
|
||||
show (time `div` nano) ++ "." ++
|
||||
show (time `div` nano) ++ "." ++
|
||||
addZeros (unpack (show ((time `mod` nano) `div` 1000000))) ++
|
||||
"s"
|
||||
pure res
|
||||
@ -1973,7 +2014,7 @@ showTimeRecord
|
||||
= do coreLift $ putStr (key ++ ": ")
|
||||
let nano = 1000000000
|
||||
assert_total $ -- We're not dividing by 0
|
||||
coreLift $ putStrLn $ show (time `div` nano) ++ "." ++
|
||||
coreLift $ putStrLn $ show (time `div` nano) ++ "." ++
|
||||
addZeros (unpack (show ((time `mod` nano) `div` 1000000))) ++
|
||||
"s"
|
||||
|
||||
|
@ -66,7 +66,7 @@ freeEnv fc (n :: ns) = PVar RigW Explicit (Erased fc) :: freeEnv fc ns
|
||||
getCons : Defs -> NF vars -> Core (List (NF [], Name, Int, Nat))
|
||||
getCons defs (NTCon _ tn _ _ _)
|
||||
= case !(lookupDefExact tn (gamma defs)) of
|
||||
Just (TCon _ _ _ _ _ cons) =>
|
||||
Just (TCon _ _ _ _ _ cons) =>
|
||||
do cs' <- traverse addTy cons
|
||||
pure (mapMaybe id cs')
|
||||
_ => pure []
|
||||
@ -76,7 +76,7 @@ getCons defs (NTCon _ tn _ _ _)
|
||||
= do Just gdef <- lookupCtxtExact cn (gamma defs)
|
||||
| _ => pure Nothing
|
||||
case (definition gdef, type gdef) of
|
||||
(DCon t arity, ty) =>
|
||||
(DCon t arity, ty) =>
|
||||
pure (Just (!(nf defs [] ty), cn, t, arity))
|
||||
_ => pure Nothing
|
||||
getCons defs _ = pure []
|
||||
@ -94,7 +94,7 @@ emptyRHS fc sc = sc
|
||||
|
||||
mkAlt : FC -> CaseTree vars -> (Name, Int, Nat) -> CaseAlt vars
|
||||
mkAlt fc sc (cn, t, ar)
|
||||
= ConCase cn t (map (MN "m") (take ar [0..]))
|
||||
= ConCase cn t (map (MN "m") (take ar [0..]))
|
||||
(weakenNs _ (emptyRHS fc sc))
|
||||
|
||||
altMatch : CaseAlt vars -> CaseAlt vars -> Bool
|
||||
@ -106,9 +106,9 @@ altMatch _ _ = False
|
||||
|
||||
-- Given a type and a list of case alternatives, return the
|
||||
-- well-typed alternatives which were *not* in the list
|
||||
getMissingAlts : FC -> Defs -> NF vars -> List (CaseAlt vars) ->
|
||||
getMissingAlts : FC -> Defs -> NF vars -> List (CaseAlt vars) ->
|
||||
Core (List (CaseAlt vars))
|
||||
-- If it's a primitive, there's too many to reasonably check, so require a
|
||||
-- If it's a primitive, there's too many to reasonably check, so require a
|
||||
-- catch all
|
||||
getMissingAlts fc defs (NPrimVal _ c) alts
|
||||
= if any isDefault alts
|
||||
@ -128,13 +128,13 @@ getMissingAlts fc defs (NType _) alts
|
||||
isDefault (DefaultCase _) = True
|
||||
isDefault _ = False
|
||||
getMissingAlts fc defs nfty alts
|
||||
= do allCons <- getCons defs nfty
|
||||
pure (filter (noneOf alts)
|
||||
= do allCons <- getCons defs nfty
|
||||
pure (filter (noneOf alts)
|
||||
(map (mkAlt fc (Unmatched "Coverage check") . snd) allCons))
|
||||
where
|
||||
-- Return whether the alternative c matches none of the given cases in alts
|
||||
noneOf : List (CaseAlt vars) -> CaseAlt vars -> Bool
|
||||
noneOf alts c = not $ any (altMatch c) alts
|
||||
noneOf alts c = not $ any (altMatch c) alts
|
||||
|
||||
-- Mapping of variable to constructor tag already matched for it
|
||||
KnownVars : List Name -> Type -> Type
|
||||
@ -152,8 +152,8 @@ showK {a} xs = show (map aString xs)
|
||||
|
||||
weakenNs : (args : List Name) -> KnownVars vars a -> KnownVars (args ++ vars) a
|
||||
weakenNs args [] = []
|
||||
weakenNs {vars} args ((MkVar p, t) :: xs)
|
||||
= (insertVarNames _ {outer = []} {ns=args} {inner=vars} p, t)
|
||||
weakenNs {vars} args ((MkVar p, t) :: xs)
|
||||
= (insertVarNames _ {outer = []} {ns=args} {inner=vars} p, t)
|
||||
:: weakenNs args xs
|
||||
|
||||
findTag : IsVar n idx vars -> KnownVars vars a -> Maybe a
|
||||
@ -186,13 +186,13 @@ tagIsNot ts (DefaultCase _) = False
|
||||
-- Replace a default case with explicit branches for the constructors.
|
||||
-- This is easier than checking whether a default is needed when traversing
|
||||
-- the tree (just one constructor lookup up front).
|
||||
replaceDefaults : FC -> Defs -> NF vars -> List (CaseAlt vars) ->
|
||||
replaceDefaults : FC -> Defs -> NF vars -> List (CaseAlt vars) ->
|
||||
Core (List (CaseAlt vars))
|
||||
-- Leave it alone if it's a primitive type though, since we need the catch
|
||||
-- all case there
|
||||
replaceDefaults fc defs (NPrimVal _ _) cs = pure cs
|
||||
replaceDefaults fc defs (NType _) cs = pure cs
|
||||
replaceDefaults fc defs nfty cs
|
||||
replaceDefaults fc defs nfty cs
|
||||
= do cs' <- traverse rep cs
|
||||
pure (dropRep (concat cs'))
|
||||
where
|
||||
@ -226,50 +226,50 @@ buildArgs fc defs known not ps cs@(Case {name = var} idx el ty altsIn)
|
||||
-- the ones it can't possibly be (the 'not') because a previous case
|
||||
-- has matched.
|
||||
= do let fenv = freeEnv fc _
|
||||
nfty <- nf defs fenv ty
|
||||
nfty <- nf defs fenv ty
|
||||
alts <- replaceDefaults fc defs nfty altsIn
|
||||
let alts' = alts ++ !(getMissingAlts fc defs nfty alts)
|
||||
let altsK = maybe alts' (\t => filter (tagIs t) alts')
|
||||
(findTag el known)
|
||||
(findTag el known)
|
||||
let altsN = maybe altsK (\ts => filter (tagIsNot ts) altsK)
|
||||
(findTag el not)
|
||||
buildArgsAlt not altsN
|
||||
where
|
||||
buildArgAlt : KnownVars vars (List Int) ->
|
||||
CaseAlt vars -> Core (List (List ClosedTerm))
|
||||
buildArgAlt not' (ConCase n t args sc)
|
||||
buildArgAlt not' (ConCase n t args sc)
|
||||
= do let con = Ref fc (DataCon t (length args)) n
|
||||
let ps' = map (substName var
|
||||
let ps' = map (substName var
|
||||
(apply fc
|
||||
con (map (Ref fc Bound) args))) ps
|
||||
buildArgs fc defs (weakenNs args ((MkVar el, t) :: known))
|
||||
buildArgs fc defs (weakenNs args ((MkVar el, t) :: known))
|
||||
(weakenNs args not') ps' sc
|
||||
buildArgAlt not' (DelayCase t a sc)
|
||||
= let ps' = map (substName var (TDelay fc LUnknown
|
||||
= let ps' = map (substName var (TDelay fc LUnknown
|
||||
(Ref fc Bound t)
|
||||
(Ref fc Bound a))) ps in
|
||||
buildArgs fc defs (weakenNs [t,a] known) (weakenNs [t,a] not')
|
||||
ps' sc
|
||||
buildArgAlt not' (ConstCase c sc)
|
||||
buildArgAlt not' (ConstCase c sc)
|
||||
= do let ps' = map (substName var (PrimVal fc c)) ps
|
||||
buildArgs fc defs known not' ps' sc
|
||||
buildArgAlt not' (DefaultCase sc)
|
||||
buildArgAlt not' (DefaultCase sc)
|
||||
= buildArgs fc defs known not' ps sc
|
||||
|
||||
buildArgsAlt : KnownVars vars (List Int) -> List (CaseAlt vars) ->
|
||||
Core (List (List ClosedTerm))
|
||||
buildArgsAlt not' [] = pure []
|
||||
buildArgsAlt not' (c@(ConCase _ t _ _) :: cs)
|
||||
= pure $ !(buildArgAlt not' c) ++
|
||||
= pure $ !(buildArgAlt not' c) ++
|
||||
!(buildArgsAlt (addNot el t not') cs)
|
||||
buildArgsAlt not' (c :: cs)
|
||||
= pure $ !(buildArgAlt not' c) ++ !(buildArgsAlt not' cs)
|
||||
|
||||
buildArgs fc defs known not ps (STerm vs)
|
||||
buildArgs fc defs known not ps (STerm vs)
|
||||
= pure [] -- matched, so return nothing
|
||||
buildArgs fc defs known not ps (Unmatched msg)
|
||||
buildArgs fc defs known not ps (Unmatched msg)
|
||||
= pure [ps] -- unmatched, so return it
|
||||
buildArgs fc defs known not ps Impossible
|
||||
buildArgs fc defs known not ps Impossible
|
||||
= pure [] -- not a possible match, so return nothing
|
||||
|
||||
-- Traverse a case tree and return pattern clauses which are not
|
||||
@ -278,7 +278,7 @@ buildArgs fc defs known not ps Impossible
|
||||
-- checked
|
||||
export
|
||||
getMissing : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> CaseTree vars ->
|
||||
FC -> Name -> CaseTree vars ->
|
||||
Core (List ClosedTerm)
|
||||
getMissing fc n ctree
|
||||
= do defs <- get Ctxt
|
||||
|
@ -10,7 +10,7 @@ data Env : (tm : List Name -> Type) -> List Name -> Type where
|
||||
|
||||
export
|
||||
extend : (x : Name) -> Binder (tm vars) -> Env tm vars -> Env tm (x :: vars)
|
||||
extend x = (::) {x}
|
||||
extend x = (::) {x}
|
||||
|
||||
export
|
||||
length : Env tm xs -> Nat
|
||||
@ -24,7 +24,7 @@ data IsDefined : Name -> List Name -> Type where
|
||||
|
||||
export
|
||||
defined : {vars : _} ->
|
||||
(n : Name) -> Env Term vars ->
|
||||
(n : Name) -> Env Term vars ->
|
||||
Maybe (IsDefined n vars)
|
||||
defined n [] = Nothing
|
||||
defined {vars = x :: xs} n (b :: env)
|
||||
@ -36,19 +36,19 @@ defined {vars = x :: xs} n (b :: env)
|
||||
export
|
||||
bindEnv : FC -> Env Term vars -> (tm : Term vars) -> ClosedTerm
|
||||
bindEnv loc [] tm = tm
|
||||
bindEnv loc (b :: env) tm
|
||||
bindEnv loc (b :: env) tm
|
||||
= bindEnv loc env (Bind loc _ b tm)
|
||||
|
||||
revOnto : (xs, vs : _) -> reverseOnto xs vs = reverse vs ++ xs
|
||||
revOnto xs [] = Refl
|
||||
revOnto xs (v :: vs)
|
||||
= rewrite revOnto (v :: xs) vs in
|
||||
revOnto xs (v :: vs)
|
||||
= rewrite revOnto (v :: xs) vs in
|
||||
rewrite appendAssociative (reverse vs) [v] xs in
|
||||
rewrite revOnto [v] vs in Refl
|
||||
|
||||
revNs : (vs, ns : List a) -> reverse ns ++ reverse vs = reverse (vs ++ ns)
|
||||
revNs [] ns = rewrite appendNilRightNeutral (reverse ns) in Refl
|
||||
revNs (v :: vs) ns
|
||||
revNs (v :: vs) ns
|
||||
= rewrite revOnto [v] vs in
|
||||
rewrite revOnto [v] (vs ++ ns) in
|
||||
rewrite sym (revNs vs ns) in
|
||||
@ -59,22 +59,22 @@ revNs (v :: vs) ns
|
||||
-- in big environments
|
||||
-- Also reversing the names at the end saves significant time over concatenating
|
||||
-- when environments get fairly big.
|
||||
getBinderUnder : Weaken tm =>
|
||||
getBinderUnder : Weaken tm =>
|
||||
{idx : Nat} ->
|
||||
(ns : List Name) ->
|
||||
.(IsVar x idx vars) -> Env tm vars ->
|
||||
(ns : List Name) ->
|
||||
.(IsVar x idx vars) -> Env tm vars ->
|
||||
Binder (tm (reverse ns ++ vars))
|
||||
getBinderUnder {idx = Z} {vars = v :: vs} ns First (b :: env)
|
||||
getBinderUnder {idx = Z} {vars = v :: vs} ns First (b :: env)
|
||||
= rewrite appendAssociative (reverse ns) [v] vs in
|
||||
rewrite revNs [v] ns in
|
||||
map (weakenNs (reverse (v :: ns))) b
|
||||
getBinderUnder {idx = S k} {vars = v :: vs} ns (Later lp) (b :: env)
|
||||
getBinderUnder {idx = S k} {vars = v :: vs} ns (Later lp) (b :: env)
|
||||
= rewrite appendAssociative (reverse ns) [v] vs in
|
||||
rewrite revNs [v] ns in
|
||||
getBinderUnder (v :: ns) lp env
|
||||
|
||||
export
|
||||
getBinder : Weaken tm =>
|
||||
getBinder : Weaken tm =>
|
||||
{idx : Nat} ->
|
||||
.(IsVar x idx vars) -> Env tm vars -> Binder (tm vars)
|
||||
getBinder el env = getBinderUnder [] el env
|
||||
@ -87,10 +87,10 @@ abstractEnvType : FC -> Env Term vars -> (tm : Term vars) -> ClosedTerm
|
||||
abstractEnvType fc [] tm = tm
|
||||
abstractEnvType fc (Let c val ty :: env) tm
|
||||
= abstractEnvType fc env (Bind fc _ (Let c val ty) tm)
|
||||
abstractEnvType fc (Pi c e ty :: env) tm
|
||||
abstractEnvType fc (Pi c e ty :: env) tm
|
||||
= abstractEnvType fc env (Bind fc _ (Pi c e ty) tm)
|
||||
abstractEnvType fc (b :: env) tm
|
||||
= abstractEnvType fc env (Bind fc _
|
||||
abstractEnvType fc (b :: env) tm
|
||||
= abstractEnvType fc env (Bind fc _
|
||||
(Pi (multiplicity b) Explicit (binderType b)) tm)
|
||||
|
||||
-- As above, for the corresponding term
|
||||
@ -99,16 +99,16 @@ abstractEnv : FC -> Env Term vars -> (tm : Term vars) -> ClosedTerm
|
||||
abstractEnv fc [] tm = tm
|
||||
abstractEnv fc (Let c val ty :: env) tm
|
||||
= abstractEnv fc env (Bind fc _ (Let c val ty) tm)
|
||||
abstractEnv fc (b :: env) tm
|
||||
= abstractEnv fc env (Bind fc _
|
||||
abstractEnv fc (b :: env) tm
|
||||
= abstractEnv fc env (Bind fc _
|
||||
(Lam (multiplicity b) Explicit (binderType b)) tm)
|
||||
|
||||
-- As above, but abstract over all binders including lets
|
||||
export
|
||||
abstractFullEnvType : FC -> Env Term vars -> (tm : Term vars) -> ClosedTerm
|
||||
abstractFullEnvType fc [] tm = tm
|
||||
abstractFullEnvType fc (b :: env) tm
|
||||
= abstractFullEnvType fc env (Bind fc _
|
||||
abstractFullEnvType fc (b :: env) tm
|
||||
= abstractFullEnvType fc env (Bind fc _
|
||||
(Pi (multiplicity b) Explicit (binderType b)) tm)
|
||||
|
||||
export
|
||||
@ -120,8 +120,8 @@ letToLam (b :: env) = b :: letToLam env
|
||||
mutual
|
||||
-- Quicker, if less safe, to store variables as a Nat, for quick comparison
|
||||
findUsed : Env Term vars -> List Nat -> Term vars -> List Nat
|
||||
findUsed env used (Local fc r idx p)
|
||||
= if elemBy eqNat idx used
|
||||
findUsed env used (Local fc r idx p)
|
||||
= if elemBy eqNat idx used
|
||||
then used
|
||||
else assert_total (findUsedInBinder env (idx :: used)
|
||||
(getBinder p env))
|
||||
@ -135,31 +135,31 @@ mutual
|
||||
findUsedArgs env u [] = u
|
||||
findUsedArgs env u (a :: as)
|
||||
= findUsedArgs env (findUsed env u a) as
|
||||
findUsed env used (Bind fc x b tm)
|
||||
findUsed env used (Bind fc x b tm)
|
||||
= assert_total $
|
||||
dropS (findUsed (b :: env)
|
||||
(map S (findUsedInBinder env used b))
|
||||
tm)
|
||||
where
|
||||
where
|
||||
dropS : List Nat -> List Nat
|
||||
dropS [] = []
|
||||
dropS (Z :: xs) = dropS xs
|
||||
dropS (S p :: xs) = p :: dropS xs
|
||||
findUsed env used (App fc fn arg)
|
||||
findUsed env used (App fc fn arg)
|
||||
= findUsed env (findUsed env used fn) arg
|
||||
findUsed env used (As fc a p)
|
||||
findUsed env used (As fc a p)
|
||||
= findUsed env (findUsed env used a) p
|
||||
findUsed env used (TDelayed fc r tm)
|
||||
= findUsed env used tm
|
||||
findUsed env used (TDelay fc r ty tm)
|
||||
= findUsed env (findUsed env used ty) tm
|
||||
findUsed env used (TForce fc tm)
|
||||
findUsed env used (TForce fc r tm)
|
||||
= findUsed env used tm
|
||||
findUsed env used _ = used
|
||||
|
||||
|
||||
findUsedInBinder : Env Term vars -> List Nat ->
|
||||
Binder (Term vars) -> List Nat
|
||||
findUsedInBinder env used (Let _ val ty)
|
||||
findUsedInBinder env used (Let _ val ty)
|
||||
= findUsed env (findUsed env used val) ty
|
||||
findUsedInBinder env used (PLet _ val ty)
|
||||
= findUsed env (findUsed env used val) ty
|
||||
@ -168,32 +168,32 @@ mutual
|
||||
toVar : (vars : List Name) -> Nat -> Maybe (Var vars)
|
||||
toVar (v :: vs) Z = Just (MkVar First)
|
||||
toVar (v :: vs) (S k)
|
||||
= do MkVar prf <- toVar vs k
|
||||
= do MkVar prf <- toVar vs k
|
||||
Just (MkVar (Later prf))
|
||||
toVar _ _ = Nothing
|
||||
|
||||
export
|
||||
findUsedLocs : Env Term vars -> Term vars -> List (Var vars)
|
||||
findUsedLocs env tm
|
||||
findUsedLocs env tm
|
||||
= mapMaybe (toVar _) (findUsed env [] tm)
|
||||
|
||||
isUsed : Nat -> List (Var vars) -> Bool
|
||||
isUsed n [] = False
|
||||
isUsed n (v :: vs) = n == varIdx v || isUsed n vs
|
||||
|
||||
mkShrinkSub : (vars : _) -> List (Var (n :: vars)) ->
|
||||
mkShrinkSub : (vars : _) -> List (Var (n :: vars)) ->
|
||||
(newvars ** SubVars newvars (n :: vars))
|
||||
mkShrinkSub [] els
|
||||
mkShrinkSub [] els
|
||||
= if isUsed 0 els
|
||||
then (_ ** KeepCons SubRefl)
|
||||
else (_ ** DropCons SubRefl)
|
||||
mkShrinkSub (x :: xs) els
|
||||
mkShrinkSub (x :: xs) els
|
||||
= let (_ ** subRest) = mkShrinkSub xs (dropFirst els) in
|
||||
if isUsed 0 els
|
||||
then (_ ** KeepCons subRest)
|
||||
else (_ ** DropCons subRest)
|
||||
|
||||
mkShrink : List (Var vars) ->
|
||||
mkShrink : List (Var vars) ->
|
||||
(newvars ** SubVars newvars vars)
|
||||
mkShrink {vars = []} xs = (_ ** SubRefl)
|
||||
mkShrink {vars = v :: vs} xs = mkShrinkSub _ xs
|
||||
@ -201,7 +201,7 @@ mkShrink {vars = v :: vs} xs = mkShrinkSub _ xs
|
||||
-- Find the smallest subset of the environment which is needed to type check
|
||||
-- the given term
|
||||
export
|
||||
findSubEnv : Env Term vars -> Term vars ->
|
||||
findSubEnv : Env Term vars -> Term vars ->
|
||||
(vars' : List Name ** SubVars vars' vars)
|
||||
findSubEnv env tm = mkShrink (findUsedLocs env tm)
|
||||
|
||||
@ -209,7 +209,7 @@ export
|
||||
shrinkEnv : Env Term vars -> SubVars newvars vars -> Maybe (Env Term newvars)
|
||||
shrinkEnv env SubRefl = Just env
|
||||
shrinkEnv (b :: env) (DropCons p) = shrinkEnv env p
|
||||
shrinkEnv (b :: env) (KeepCons p)
|
||||
shrinkEnv (b :: env) (KeepCons p)
|
||||
= do env' <- shrinkEnv env p
|
||||
b' <- shrinkBinder b p
|
||||
pure (b' :: env')
|
||||
|
@ -11,7 +11,7 @@ import Core.Value
|
||||
|
||||
-- Get the type of an already typechecked thing.
|
||||
-- We need this (occasionally) because we don't store types in subterms (e.g. on
|
||||
-- applications) and we don't keep the type of suterms up to date throughout
|
||||
-- applications) and we don't keep the type of suterms up to date throughout
|
||||
-- unification. Perhaps we should? There's a trade off here, and recalculating on
|
||||
-- the rare occasions it's necessary doesn't seem to cost too much, but keep an
|
||||
-- eye on it...
|
||||
@ -19,26 +19,26 @@ import Core.Value
|
||||
mutual
|
||||
chk : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars -> Term vars -> Core (Glued vars)
|
||||
chk env (Local fc r idx p)
|
||||
chk env (Local fc r idx p)
|
||||
= pure $ gnf env (binderType (getBinder p env))
|
||||
chk env (Ref fc nt n)
|
||||
= do defs <- get Ctxt
|
||||
Just ty <- lookupTyExact n (gamma defs)
|
||||
| Nothing => throw (UndefinedName fc n)
|
||||
pure $ gnf env (embed ty)
|
||||
chk env (Meta fc n i args)
|
||||
chk env (Meta fc n i args)
|
||||
= do defs <- get Ctxt
|
||||
Just mty <- lookupTyExact (Resolved i) (gamma defs)
|
||||
| Nothing => throw (UndefinedName fc n)
|
||||
chkMeta fc env !(nf defs env (embed mty)) args
|
||||
chk env (Bind fc nm b sc)
|
||||
chk env (Bind fc nm b sc)
|
||||
= do bt <- chkBinder env b
|
||||
sct <- chk {vars = nm :: _} (b :: env) sc
|
||||
pure $ gnf env (discharge fc nm b !(getTerm bt) !(getTerm sct))
|
||||
chk env (App fc f a)
|
||||
chk env (App fc f a)
|
||||
= do fty <- chk env f
|
||||
case !(getNF fty) of
|
||||
NBind _ _ (Pi _ _ ty) scdone =>
|
||||
NBind _ _ (Pi _ _ ty) scdone =>
|
||||
do defs <- get Ctxt
|
||||
aty <- chk env a
|
||||
sc' <- scdone defs (toClosure defaultOpts env a)
|
||||
@ -47,15 +47,15 @@ mutual
|
||||
throw (NotFunctionType fc env fty')
|
||||
chk env (As fc n p) = chk env p
|
||||
chk env (TDelayed fc r tm) = pure (gType fc)
|
||||
chk env (TDelay fc r dty tm)
|
||||
chk env (TDelay fc r dty tm)
|
||||
= do gtm <- chk env tm
|
||||
tm' <- getNF gtm
|
||||
defs <- get Ctxt
|
||||
pure $ glueBack defs env (NDelayed fc r tm')
|
||||
chk env (TForce fc tm)
|
||||
chk env (TForce fc r tm)
|
||||
= do tm' <- chk env tm
|
||||
case !(getNF tm') of
|
||||
NDelayed fc r fty =>
|
||||
NDelayed fc _ fty =>
|
||||
do defs <- get Ctxt
|
||||
pure $ glueBack defs env fty
|
||||
chk env (PrimVal fc x) = pure $ gnf env (chkConstant fc x)
|
||||
@ -65,7 +65,7 @@ mutual
|
||||
chkMeta : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Env Term vars -> NF vars -> List (Term vars) ->
|
||||
Core (Glued vars)
|
||||
chkMeta fc env ty []
|
||||
chkMeta fc env ty []
|
||||
= do defs <- get Ctxt
|
||||
pure $ glueBack defs env ty
|
||||
chkMeta fc env (NBind _ _ (Pi _ _ ty) scdone) (a :: args)
|
||||
@ -88,9 +88,9 @@ mutual
|
||||
= Bind fc n (Pi c x ty) scopety
|
||||
discharge fc n (Let c val ty) bindty scopety
|
||||
= Bind fc n (Let c val ty) scopety
|
||||
discharge fc n (Pi c x ty) bindty scopety
|
||||
discharge fc n (Pi c x ty) bindty scopety
|
||||
= bindty
|
||||
discharge fc n (PVar c p ty) bindty scopety
|
||||
discharge fc n (PVar c p ty) bindty scopety
|
||||
= Bind fc n (PVTy c ty) scopety
|
||||
discharge fc n (PLet c val ty) bindty scopety
|
||||
= Bind fc n (PLet c val ty) scopety
|
||||
|
@ -15,7 +15,7 @@ public export
|
||||
interface Hashable a where
|
||||
hash : a -> Int
|
||||
hashWithSalt : Int -> a -> Int
|
||||
|
||||
|
||||
hash = hashWithSalt 5381
|
||||
hashWithSalt h i = h * 33 + hash i
|
||||
|
||||
@ -54,9 +54,9 @@ Hashable String where
|
||||
hashChars : Int -> Int -> Int -> String -> Int
|
||||
hashChars h p len str
|
||||
= assert_total $
|
||||
if p == len
|
||||
if p == len
|
||||
then h
|
||||
else hashChars (h * 33 + cast (strIndex str p))
|
||||
else hashChars (h * 33 + cast (strIndex str p))
|
||||
(p + 1) len str
|
||||
|
||||
export
|
||||
@ -79,7 +79,7 @@ Hashable PiInfo where
|
||||
|
||||
export
|
||||
Hashable ty => Hashable (Binder ty) where
|
||||
hashWithSalt h (Lam c p ty)
|
||||
hashWithSalt h (Lam c p ty)
|
||||
= h `hashWithSalt` 0 `hashWithSalt` c `hashWithSalt` p `hashWithSalt` ty
|
||||
hashWithSalt h (Let c val ty)
|
||||
= h `hashWithSalt` 1 `hashWithSalt` c `hashWithSalt` val `hashWithSalt` ty
|
||||
@ -98,72 +98,72 @@ Hashable (Var vars) where
|
||||
mutual
|
||||
export
|
||||
Hashable (Term vars) where
|
||||
hashWithSalt h (Local fc x idx y)
|
||||
hashWithSalt h (Local fc x idx y)
|
||||
= h `hashWithSalt` 0 `hashWithSalt` idx
|
||||
hashWithSalt h (Ref fc x name)
|
||||
hashWithSalt h (Ref fc x name)
|
||||
= h `hashWithSalt` 1 `hashWithSalt` name
|
||||
hashWithSalt h (Meta fc x y xs)
|
||||
hashWithSalt h (Meta fc x y xs)
|
||||
= h `hashWithSalt` 2 `hashWithSalt` y `hashWithSalt` xs
|
||||
hashWithSalt h (Bind fc x b scope)
|
||||
hashWithSalt h (Bind fc x b scope)
|
||||
= h `hashWithSalt` 3 `hashWithSalt` b `hashWithSalt` scope
|
||||
hashWithSalt h (App fc fn arg)
|
||||
hashWithSalt h (App fc fn arg)
|
||||
= h `hashWithSalt` 4 `hashWithSalt` fn `hashWithSalt` arg
|
||||
hashWithSalt h (As fc nm pat)
|
||||
= h `hashWithSalt` 5 `hashWithSalt` nm `hashWithSalt` pat
|
||||
hashWithSalt h (TDelayed fc x y)
|
||||
hashWithSalt h (TDelayed fc x y)
|
||||
= h `hashWithSalt` 6 `hashWithSalt` y
|
||||
hashWithSalt h (TDelay fc x t y)
|
||||
= h `hashWithSalt` 7 `hashWithSalt` t `hashWithSalt` y
|
||||
hashWithSalt h (TForce fc x)
|
||||
hashWithSalt h (TForce fc r x)
|
||||
= h `hashWithSalt` 8 `hashWithSalt` x
|
||||
hashWithSalt h (PrimVal fc c)
|
||||
hashWithSalt h (PrimVal fc c)
|
||||
= h `hashWithSalt` 9 `hashWithSalt` (show c)
|
||||
hashWithSalt h (Erased fc)
|
||||
hashWithSalt h (Erased fc)
|
||||
= hashWithSalt h 10
|
||||
hashWithSalt h (TType fc)
|
||||
= hashWithSalt h 11
|
||||
|
||||
export
|
||||
Hashable Pat where
|
||||
hashWithSalt h (PAs fc nm pat)
|
||||
hashWithSalt h (PAs fc nm pat)
|
||||
= h `hashWithSalt` 0 `hashWithSalt` nm `hashWithSalt` pat
|
||||
hashWithSalt h (PCon fc x tag arity xs)
|
||||
hashWithSalt h (PCon fc x tag arity xs)
|
||||
= h `hashWithSalt` 1 `hashWithSalt` x `hashWithSalt` xs
|
||||
hashWithSalt h (PTyCon fc x arity xs)
|
||||
hashWithSalt h (PTyCon fc x arity xs)
|
||||
= h `hashWithSalt` 2 `hashWithSalt` x `hashWithSalt` xs
|
||||
hashWithSalt h (PConst fc c)
|
||||
hashWithSalt h (PConst fc c)
|
||||
= h `hashWithSalt` 3 `hashWithSalt` (show c)
|
||||
hashWithSalt h (PArrow fc x s t)
|
||||
hashWithSalt h (PArrow fc x s t)
|
||||
= h `hashWithSalt` 4 `hashWithSalt` s `hashWithSalt` t
|
||||
hashWithSalt h (PDelay fc r t p)
|
||||
= h `hashWithSalt` 5 `hashWithSalt` t `hashWithSalt` p
|
||||
hashWithSalt h (PLoc fc x)
|
||||
hashWithSalt h (PLoc fc x)
|
||||
= h `hashWithSalt` 6 `hashWithSalt` x
|
||||
hashWithSalt h (PUnmatchable fc x)
|
||||
hashWithSalt h (PUnmatchable fc x)
|
||||
= h `hashWithSalt` 7 `hashWithSalt` x
|
||||
|
||||
export
|
||||
Hashable (CaseTree vars) where
|
||||
hashWithSalt h (Case idx x scTy xs)
|
||||
hashWithSalt h (Case idx x scTy xs)
|
||||
= h `hashWithSalt` 0 `hashWithSalt` idx `hashWithSalt` xs
|
||||
hashWithSalt h (STerm x)
|
||||
hashWithSalt h (STerm x)
|
||||
= h `hashWithSalt` 1 `hashWithSalt` x
|
||||
hashWithSalt h (Unmatched msg)
|
||||
hashWithSalt h (Unmatched msg)
|
||||
= h `hashWithSalt` 2
|
||||
hashWithSalt h Impossible
|
||||
= h `hashWithSalt` 3
|
||||
|
||||
export
|
||||
Hashable (CaseAlt vars) where
|
||||
hashWithSalt h (ConCase x tag args y)
|
||||
hashWithSalt h (ConCase x tag args y)
|
||||
= h `hashWithSalt` 0 `hashWithSalt` x `hashWithSalt` args
|
||||
`hashWithSalt` y
|
||||
hashWithSalt h (DelayCase t x y)
|
||||
= h `hashWithSalt` 2 `hashWithSalt` (show t)
|
||||
hashWithSalt h (DelayCase t x y)
|
||||
= h `hashWithSalt` 2 `hashWithSalt` (show t)
|
||||
`hashWithSalt` (show x) `hashWithSalt` y
|
||||
hashWithSalt h (ConstCase x y)
|
||||
hashWithSalt h (ConstCase x y)
|
||||
= h `hashWithSalt` 3 `hashWithSalt` (show x) `hashWithSalt` y
|
||||
hashWithSalt h (DefaultCase x)
|
||||
hashWithSalt h (DefaultCase x)
|
||||
= h `hashWithSalt` 4 `hashWithSalt` x
|
||||
|
||||
|
||||
|
@ -10,11 +10,11 @@ import Core.TT
|
||||
|
||||
addPrim : {auto c : Ref Ctxt Defs} ->
|
||||
Prim -> Core ()
|
||||
addPrim p
|
||||
addPrim p
|
||||
= do addBuiltin (opName (fn p)) (type p) (totality p) (fn p)
|
||||
-- compileDef empty (opName (fn p))
|
||||
|
||||
export
|
||||
addPrimitives : {auto c : Ref Ctxt Defs} -> Core ()
|
||||
addPrimitives
|
||||
addPrimitives
|
||||
= traverse_ addPrim allPrimitives
|
||||
|
@ -37,7 +37,7 @@ doneScope (MkVar (Later p) :: xs) = MkVar p :: doneScope xs
|
||||
|
||||
count : Nat -> Usage ns -> Nat
|
||||
count p [] = 0
|
||||
count p (v :: xs)
|
||||
count p (v :: xs)
|
||||
= if p == varIdx v then 1 + count p xs else count p xs
|
||||
|
||||
localPrf : {later : _} -> Var (later ++ n :: vars)
|
||||
@ -50,7 +50,7 @@ mutual
|
||||
updateHoleUsageArgs : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
(useInHole : Bool) ->
|
||||
Var vars -> List (Term vars) -> Core Bool
|
||||
Var vars -> List (Term vars) -> Core Bool
|
||||
updateHoleUsageArgs useInHole var [] = pure False
|
||||
updateHoleUsageArgs useInHole var (a :: as)
|
||||
= do h <- updateHoleUsage useInHole var a
|
||||
@ -79,10 +79,10 @@ mutual
|
||||
= do updateHoleUsage False var a
|
||||
scty <- updateHoleType useInHole var sc as
|
||||
pure (Bind bfc nm (Pi c e ty) scty)
|
||||
updateHoleType useInHole var ty as
|
||||
updateHoleType useInHole var ty as
|
||||
= do updateHoleUsageArgs False var as
|
||||
pure ty
|
||||
|
||||
|
||||
updateHoleUsagePats : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
(useInHole : Bool) ->
|
||||
@ -111,11 +111,11 @@ mutual
|
||||
findLocal (Local _ _ _ p :: _) Z = Just (MkVar p)
|
||||
findLocal (_ :: els) (S k) = findLocal els k
|
||||
findLocal _ _ = Nothing
|
||||
|
||||
|
||||
updateHoleUsage : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
(useInHole : Bool) ->
|
||||
Var vars -> Term vars -> Core Bool
|
||||
Var vars -> Term vars -> Core Bool
|
||||
updateHoleUsage useInHole (MkVar var) (Bind _ n (Let c val ty) sc)
|
||||
= do h <- updateHoleUsage useInHole (MkVar var) val
|
||||
h' <- updateHoleUsage useInHole (MkVar (Later var)) sc
|
||||
@ -133,23 +133,23 @@ mutual
|
||||
ty' <- updateHoleType useInHole var ty args
|
||||
updateTy i ty'
|
||||
pure True
|
||||
_ => updateHoleUsageArgs useInHole var args
|
||||
_ => updateHoleUsageArgs useInHole var args
|
||||
updateHoleUsage useInHole var (As _ a p)
|
||||
= do h <- updateHoleUsage useInHole var a
|
||||
h' <- updateHoleUsage useInHole var a
|
||||
pure (h || h')
|
||||
updateHoleUsage useInHole var (TDelayed _ _ t)
|
||||
updateHoleUsage useInHole var (TDelayed _ _ t)
|
||||
= updateHoleUsage useInHole var t
|
||||
updateHoleUsage useInHole var (TDelay _ _ _ t)
|
||||
updateHoleUsage useInHole var (TDelay _ _ _ t)
|
||||
= updateHoleUsage useInHole var t
|
||||
updateHoleUsage useInHole var (TForce _ t)
|
||||
updateHoleUsage useInHole var (TForce _ _ t)
|
||||
= updateHoleUsage useInHole var t
|
||||
updateHoleUsage useInHole var tm
|
||||
updateHoleUsage useInHole var tm
|
||||
= case getFnArgs tm of
|
||||
(Ref _ _ fn, args) =>
|
||||
(Ref _ _ fn, args) =>
|
||||
do aup <- updateHoleUsageArgs useInHole var args
|
||||
defs <- get Ctxt
|
||||
Just (NS _ (CaseBlock _ _), PMDef _ _ _ _ pats) <-
|
||||
Just (NS _ (CaseBlock _ _), PMDef _ _ _ _ pats) <-
|
||||
lookupExactBy (\d => (fullname d, definition d))
|
||||
fn (gamma defs)
|
||||
| _ => pure aup
|
||||
@ -166,10 +166,10 @@ mutual
|
||||
lcheck : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState } ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars -> Term vars ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars -> Term vars ->
|
||||
Core (Term vars, Glued vars, Usage vars)
|
||||
lcheck {vars} rig erase env (Local {name} fc x idx prf)
|
||||
= let b = getBinder prf env
|
||||
lcheck {vars} rig erase env (Local {name} fc x idx prf)
|
||||
= let b = getBinder prf env
|
||||
rigb = multiplicity b
|
||||
ty = binderType b in
|
||||
do when (not erase) $ rigSafe rigb rig
|
||||
@ -237,8 +237,8 @@ mutual
|
||||
-- if there's a hole, assume it will contain the missing usage
|
||||
-- if there is none already
|
||||
let used = case rigMult (multiplicity b) rig of
|
||||
Rig1 => if holeFound && used_in == 0
|
||||
then 1
|
||||
Rig1 => if holeFound && used_in == 0
|
||||
then 1
|
||||
else used_in
|
||||
_ => used_in
|
||||
|
||||
@ -256,11 +256,11 @@ mutual
|
||||
checkUsageOK used Rig0 = pure ()
|
||||
checkUsageOK used RigW = pure ()
|
||||
checkUsageOK used Rig1
|
||||
= if used == 1
|
||||
= if used == 1
|
||||
then pure ()
|
||||
else throw (LinearUsed fc used nm)
|
||||
|
||||
lcheck rig erase env (App fc f a)
|
||||
lcheck rig erase env (App fc f a)
|
||||
= do (f', gfty, fused) <- lcheck rig erase env f
|
||||
defs <- get Ctxt
|
||||
fty <- getNF gfty
|
||||
@ -285,45 +285,45 @@ mutual
|
||||
when (not !(convert defs env aty ty)) $
|
||||
do ty' <- quote defs env ty
|
||||
aty' <- quote defs env aty
|
||||
throw (CantConvert fc env ty' aty')
|
||||
pure (App fc f' aerased,
|
||||
glueBack defs env sc',
|
||||
throw (CantConvert fc env ty' aty')
|
||||
pure (App fc f' aerased,
|
||||
glueBack defs env sc',
|
||||
fused ++ aused)
|
||||
_ => do tfty <- getTerm gfty
|
||||
throw (GenericMsg fc ("Linearity checking failed on " ++ show f' ++
|
||||
throw (GenericMsg fc ("Linearity checking failed on " ++ show f' ++
|
||||
" (" ++ show tfty ++ " not a function type)"))
|
||||
|
||||
lcheck rig erase env (As fc as pat)
|
||||
lcheck rig erase env (As fc as pat)
|
||||
= do (as', _, _) <- lcheck rig erase env as
|
||||
(pat', pty, u) <- lcheck rig erase env pat
|
||||
pure (As fc as' pat', pty, u)
|
||||
lcheck rig erase env (TDelayed fc r ty)
|
||||
lcheck rig erase env (TDelayed fc r ty)
|
||||
= do (ty', _, u) <- lcheck rig erase env ty
|
||||
pure (TDelayed fc r ty', gType fc, u)
|
||||
lcheck rig erase env (TDelay fc r ty val)
|
||||
lcheck rig erase env (TDelay fc r ty val)
|
||||
= do (ty', _, _) <- lcheck Rig0 erase env ty
|
||||
(val', gty, u) <- lcheck rig erase env val
|
||||
ty <- getTerm gty
|
||||
pure (TDelay fc r ty' val', gnf env (TDelayed fc r ty), u)
|
||||
lcheck rig erase env (TForce fc val)
|
||||
lcheck rig erase env (TForce fc r val)
|
||||
= do (val', gty, u) <- lcheck rig erase env val
|
||||
tynf <- getNF gty
|
||||
case tynf of
|
||||
NDelayed _ r narg
|
||||
=> do defs <- get Ctxt
|
||||
pure (TForce fc val', glueBack defs env narg, u)
|
||||
pure (TForce fc r val', glueBack defs env narg, u)
|
||||
_ => throw (GenericMsg fc "Not a delayed tyoe")
|
||||
lcheck rig erase env (PrimVal fc c)
|
||||
lcheck rig erase env (PrimVal fc c)
|
||||
= pure (PrimVal fc c, gErased fc, [])
|
||||
lcheck rig erase env (Erased fc)
|
||||
lcheck rig erase env (Erased fc)
|
||||
= pure (Erased fc, gErased fc, [])
|
||||
lcheck rig erase env (TType fc)
|
||||
lcheck rig erase env (TType fc)
|
||||
= pure (TType fc, gType fc, [])
|
||||
|
||||
lcheckBinder : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars ->
|
||||
Binder (Term vars) ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars ->
|
||||
Binder (Term vars) ->
|
||||
Core (Binder (Term vars), Glued vars, Usage vars)
|
||||
lcheckBinder rig erase env (Lam c x ty)
|
||||
= do (tyv, tyt, _) <- lcheck Rig0 erase env ty
|
||||
@ -352,7 +352,7 @@ mutual
|
||||
Core (Term vars, Glued vars, Usage vars)
|
||||
discharge defs env fc nm (Lam c x ty) gbindty scope gscopety used
|
||||
= do scty <- getTerm gscopety
|
||||
pure (Bind fc nm (Lam c x ty) scope,
|
||||
pure (Bind fc nm (Lam c x ty) scope,
|
||||
gnf env (Bind fc nm (Pi c x ty) scty), used)
|
||||
discharge defs env fc nm (Let c val ty) gbindty scope gscopety used
|
||||
= do scty <- getTerm gscopety
|
||||
@ -371,7 +371,7 @@ mutual
|
||||
discharge defs env fc nm (PVTy c ty) gbindty scope gscopety used
|
||||
= pure (Bind fc nm (PVTy c ty) scope, gbindty, used)
|
||||
|
||||
data ArgUsage
|
||||
data ArgUsage
|
||||
= UseAny -- RigW so we don't care
|
||||
| Use0 -- argument position not used
|
||||
| Use1 -- argument position used exactly once
|
||||
@ -397,7 +397,7 @@ mutual
|
||||
= do us <- traverse (getPUsage ty) pats
|
||||
pure (map snd !(combine us))
|
||||
where
|
||||
getCaseUsage : Term ns -> Env Term vs -> List (Term vs) ->
|
||||
getCaseUsage : Term ns -> Env Term vs -> List (Term vs) ->
|
||||
Usage vs -> Term vs ->
|
||||
Core (List (Name, ArgUsage))
|
||||
getCaseUsage ty env (As _ _ p :: args) used rhs
|
||||
@ -420,7 +420,7 @@ mutual
|
||||
Rig1 => pure ((n, UseKeep) :: rest)
|
||||
_ => pure ((n, UseKeep) :: rest)
|
||||
getCaseUsage tm env args used rhs = pure []
|
||||
|
||||
|
||||
checkUsageOK : FC -> Nat -> Name -> Bool -> RigCount -> Core ()
|
||||
checkUsageOK fc used nm isloc Rig0 = pure ()
|
||||
checkUsageOK fc used nm isloc RigW = pure ()
|
||||
@ -429,7 +429,7 @@ mutual
|
||||
then throw (LinearUsed fc used nm)
|
||||
else pure ()
|
||||
checkUsageOK fc used nm isloc Rig1
|
||||
= if used == 1
|
||||
= if used == 1
|
||||
then pure ()
|
||||
else throw (LinearUsed fc used nm)
|
||||
|
||||
@ -441,15 +441,15 @@ mutual
|
||||
= if idx == varIdx p
|
||||
then True
|
||||
else isLocArg p args
|
||||
isLocArg p (As _ tm pat :: args)
|
||||
isLocArg p (As _ tm pat :: args)
|
||||
= isLocArg p (tm :: pat :: args)
|
||||
isLocArg p (_ :: args) = isLocArg p args
|
||||
|
||||
-- As checkEnvUsage in general, but it's okay for local variables to
|
||||
-- remain unused (since in that case, they must be used outside the
|
||||
-- case block)
|
||||
checkEnvUsage : RigCount ->
|
||||
Env Term vars -> Usage (done ++ vars) ->
|
||||
checkEnvUsage : RigCount ->
|
||||
Env Term vars -> Usage (done ++ vars) ->
|
||||
List (Term (done ++ vars)) ->
|
||||
Term (done ++ vars) -> Core ()
|
||||
checkEnvUsage rig [] usage args tm = pure ()
|
||||
@ -461,14 +461,14 @@ mutual
|
||||
then updateHoleUsage (used_in == 0) pos tm
|
||||
else pure False
|
||||
let used = case rigMult (multiplicity b) rig of
|
||||
Rig1 => if holeFound && used_in == 0
|
||||
then 1
|
||||
Rig1 => if holeFound && used_in == 0
|
||||
then 1
|
||||
else used_in
|
||||
_ => used_in
|
||||
checkUsageOK (getLoc (binderType b))
|
||||
used nm (isLocArg pos args)
|
||||
used nm (isLocArg pos args)
|
||||
(rigMult (multiplicity b) rig)
|
||||
checkEnvUsage {done = done ++ [nm]} rig env
|
||||
checkEnvUsage {done = done ++ [nm]} rig env
|
||||
(rewrite sym (appendAssociative done [nm] xs) in usage)
|
||||
(rewrite sym (appendAssociative done [nm] xs) in args)
|
||||
(rewrite sym (appendAssociative done [nm] xs) in tm)
|
||||
@ -487,7 +487,7 @@ mutual
|
||||
log 10 $ "Arg usage: " ++ show ause
|
||||
pure ause
|
||||
|
||||
combineUsage : (Name, ArgUsage) -> (Name, ArgUsage) ->
|
||||
combineUsage : (Name, ArgUsage) -> (Name, ArgUsage) ->
|
||||
Core (Name, ArgUsage)
|
||||
combineUsage (n, Use0) (_, Use1)
|
||||
= throw (GenericMsg topfc ("Inconsistent usage of " ++ show n ++ " in case branches"))
|
||||
@ -510,7 +510,7 @@ mutual
|
||||
pure (u' :: us')
|
||||
combineUsages _ _ = throw (InternalError "Argument usage lists inconsistent")
|
||||
|
||||
combine : List (List (Name, ArgUsage)) ->
|
||||
combine : List (List (Name, ArgUsage)) ->
|
||||
Core (List (Name, ArgUsage))
|
||||
combine [] = pure []
|
||||
combine [x] = pure x
|
||||
@ -520,7 +520,7 @@ mutual
|
||||
|
||||
lcheckDef : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars -> Name ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars -> Name ->
|
||||
Core ClosedTerm
|
||||
lcheckDef rig True env n
|
||||
= do defs <- get Ctxt
|
||||
@ -533,17 +533,17 @@ mutual
|
||||
| Nothing => throw (InternalError ("Linearity checking failed on " ++ show n))
|
||||
Just def <- lookupCtxtExact (Resolved idx) (gamma defs)
|
||||
| Nothing => throw (InternalError ("Linearity checking failed on " ++ show n))
|
||||
if linearChecked def
|
||||
if linearChecked def
|
||||
then pure (type def)
|
||||
else do case definition def of
|
||||
PMDef _ _ _ _ pats =>
|
||||
PMDef _ _ _ _ pats =>
|
||||
do u <- getArgUsage (getLoc (type def))
|
||||
rig (type def) pats
|
||||
log 10 $ "Overall arg usage " ++ show u
|
||||
let ty' = updateUsage u (type def)
|
||||
updateTy idx ty'
|
||||
setLinearCheck idx True
|
||||
logTerm 5 ("New type of " ++
|
||||
logTerm 5 ("New type of " ++
|
||||
show (fullname def)) ty'
|
||||
pure ty'
|
||||
_ => pure (type def)
|
||||
@ -559,10 +559,10 @@ mutual
|
||||
UseAny => c in -- no constraint, so leave alone
|
||||
Bind bfc n (Pi c' e ty) sc'
|
||||
updateUsage _ ty = ty
|
||||
|
||||
|
||||
expandMeta : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars ->
|
||||
RigCount -> (erase : Bool) -> Env Term vars ->
|
||||
Name -> Int -> Def -> List (Term vars) ->
|
||||
Core (Term vars, Glued vars, Usage vars)
|
||||
expandMeta rig erase env n idx (PMDef _ [] (STerm fn) _ _) args
|
||||
@ -581,28 +581,28 @@ mutual
|
||||
lcheckMeta : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
RigCount -> Bool -> Env Term vars ->
|
||||
FC -> Name -> Int ->
|
||||
FC -> Name -> Int ->
|
||||
(args : List (Term vars)) ->
|
||||
(checked : List (Term vars)) ->
|
||||
NF vars -> Core (Term vars, Glued vars, Usage vars)
|
||||
lcheckMeta rig erase env fc n idx
|
||||
lcheckMeta rig erase env fc n idx
|
||||
(arg :: args) chk (NBind _ _ (Pi rigf _ ty) sc)
|
||||
= do let checkRig = rigMult rigf rig
|
||||
(arg', gargTy, aused) <- lcheck checkRig erase env arg
|
||||
defs <- get Ctxt
|
||||
sc' <- sc defs (toClosure defaultOpts env arg')
|
||||
let aerased = if erase && rigf == Rig0 then Erased fc else arg'
|
||||
(tm, gty, u) <- lcheckMeta rig erase env fc n idx args
|
||||
(tm, gty, u) <- lcheckMeta rig erase env fc n idx args
|
||||
(aerased :: chk) sc'
|
||||
pure (tm, gty, aused ++ u)
|
||||
lcheckMeta rig erase env fc n idx (arg :: args) chk nty
|
||||
= do defs <- get Ctxt
|
||||
empty <- clearDefs defs
|
||||
ty <- quote empty env nty
|
||||
throw (GenericMsg fc ("Linearity checking failed on metavar
|
||||
" ++ show n ++ " (" ++ show ty ++
|
||||
throw (GenericMsg fc ("Linearity checking failed on metavar
|
||||
" ++ show n ++ " (" ++ show ty ++
|
||||
" not a function type)"))
|
||||
lcheckMeta rig erase env fc n idx [] chk nty
|
||||
lcheckMeta rig erase env fc n idx [] chk nty
|
||||
= do defs <- get Ctxt
|
||||
pure (Meta fc n idx (reverse chk), glueBack defs env nty, [])
|
||||
|
||||
@ -610,9 +610,9 @@ mutual
|
||||
checkEnvUsage : {done : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount ->
|
||||
Env Term vars -> Usage (done ++ vars) ->
|
||||
Term (done ++ vars) ->
|
||||
FC -> RigCount ->
|
||||
Env Term vars -> Usage (done ++ vars) ->
|
||||
Term (done ++ vars) ->
|
||||
Core ()
|
||||
checkEnvUsage fc rig [] usage tm = pure ()
|
||||
checkEnvUsage fc rig {done} {vars = nm :: xs} (b :: env) usage tm
|
||||
@ -623,12 +623,12 @@ checkEnvUsage fc rig {done} {vars = nm :: xs} (b :: env) usage tm
|
||||
then updateHoleUsage (used_in == 0) pos tm
|
||||
else pure False
|
||||
let used = case rigMult (multiplicity b) rig of
|
||||
Rig1 => if holeFound && used_in == 0
|
||||
then 1
|
||||
Rig1 => if holeFound && used_in == 0
|
||||
then 1
|
||||
else used_in
|
||||
_ => used_in
|
||||
checkUsageOK used (rigMult (multiplicity b) rig)
|
||||
checkEnvUsage {done = done ++ [nm]} fc rig env
|
||||
checkEnvUsage {done = done ++ [nm]} fc rig env
|
||||
(rewrite sym (appendAssociative done [nm] xs) in usage)
|
||||
(rewrite sym (appendAssociative done [nm] xs) in tm)
|
||||
where
|
||||
@ -636,7 +636,7 @@ checkEnvUsage fc rig {done} {vars = nm :: xs} (b :: env) usage tm
|
||||
checkUsageOK used Rig0 = pure ()
|
||||
checkUsageOK used RigW = pure ()
|
||||
checkUsageOK used Rig1
|
||||
= if used == 1
|
||||
= if used == 1
|
||||
then pure ()
|
||||
else throw (LinearUsed fc used nm)
|
||||
|
||||
@ -648,7 +648,7 @@ export
|
||||
linearCheck : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount -> (erase : Bool) ->
|
||||
Env Term vars -> Term vars ->
|
||||
Env Term vars -> Term vars ->
|
||||
Core (Term vars)
|
||||
linearCheck fc rig erase env tm
|
||||
= do logTerm 5 "Linearity check on " tm
|
||||
|
@ -64,8 +64,8 @@ addLHS : {auto c : Ref Ctxt Defs} ->
|
||||
addLHS loc outerenvlen env tm
|
||||
= do meta <- get MD
|
||||
tm' <- toFullNames (bindEnv loc (toPat env) tm)
|
||||
put MD (record {
|
||||
lhsApps $= ((loc, outerenvlen, tm') ::)
|
||||
put MD (record {
|
||||
lhsApps $= ((loc, outerenvlen, tm') ::)
|
||||
} meta)
|
||||
where
|
||||
toPat : Env Term vs -> Env Term vs
|
||||
@ -84,7 +84,7 @@ addLHS loc outerenvlen env tm
|
||||
substEnv : {vars : _} ->
|
||||
FC -> Env Term vars -> (tm : Term vars) -> ClosedTerm
|
||||
substEnv loc [] tm = tm
|
||||
substEnv {vars = x :: _} loc (b :: env) tm
|
||||
substEnv {vars = x :: _} loc (b :: env) tm
|
||||
= substEnv loc env (subst (Ref loc Bound x) tm)
|
||||
|
||||
export
|
||||
@ -94,8 +94,8 @@ addNameType : {auto c : Ref Ctxt Defs} ->
|
||||
addNameType loc n env tm
|
||||
= do meta <- get MD
|
||||
n' <- getFullName n
|
||||
put MD (record {
|
||||
names $= ((loc, (n', 0, substEnv loc env tm)) ::)
|
||||
put MD (record {
|
||||
names $= ((loc, (n', 0, substEnv loc env tm)) ::)
|
||||
} meta)
|
||||
|
||||
export
|
||||
@ -105,8 +105,8 @@ addTyDecl : {auto c : Ref Ctxt Defs} ->
|
||||
addTyDecl loc n env tm
|
||||
= do meta <- get MD
|
||||
n' <- getFullName n
|
||||
put MD (record {
|
||||
tydecls $= ((loc, (n', length env, bindEnv loc env tm)) ::)
|
||||
put MD (record {
|
||||
tydecls $= ((loc, (n', length env, bindEnv loc env tm)) ::)
|
||||
} meta)
|
||||
|
||||
export
|
||||
@ -143,15 +143,15 @@ findEntryWith p ((l, x) :: xs)
|
||||
|
||||
export
|
||||
findLHSAt : {auto m : Ref MD Metadata} ->
|
||||
(FC -> ClosedTerm -> Bool) ->
|
||||
(FC -> ClosedTerm -> Bool) ->
|
||||
Core (Maybe (FC, Nat, ClosedTerm))
|
||||
findLHSAt p
|
||||
findLHSAt p
|
||||
= do meta <- get MD
|
||||
pure (findEntryWith (\ loc, tm => p loc (snd tm)) (lhsApps meta))
|
||||
|
||||
export
|
||||
findTypeAt : {auto m : Ref MD Metadata} ->
|
||||
(FC -> (Name, Nat, ClosedTerm) -> Bool) ->
|
||||
(FC -> (Name, Nat, ClosedTerm) -> Bool) ->
|
||||
Core (Maybe (Name, Nat, ClosedTerm))
|
||||
findTypeAt p
|
||||
= do meta <- get MD
|
||||
@ -159,7 +159,7 @@ findTypeAt p
|
||||
|
||||
export
|
||||
findTyDeclAt : {auto m : Ref MD Metadata} ->
|
||||
(FC -> (Name, Nat, ClosedTerm) -> Bool) ->
|
||||
(FC -> (Name, Nat, ClosedTerm) -> Bool) ->
|
||||
Core (Maybe (FC, Name, Nat, ClosedTerm))
|
||||
findTyDeclAt p
|
||||
= do meta <- get MD
|
||||
@ -183,9 +183,9 @@ normaliseTypes
|
||||
ns' <- traverse (nfType defs) (names meta)
|
||||
put MD (record { names = ns' } meta)
|
||||
where
|
||||
nfType : Defs -> (FC, (Name, Nat, ClosedTerm)) ->
|
||||
nfType : Defs -> (FC, (Name, Nat, ClosedTerm)) ->
|
||||
Core (FC, (Name, Nat, ClosedTerm))
|
||||
nfType defs (loc, (n, len, ty))
|
||||
nfType defs (loc, (n, len, ty))
|
||||
= pure (loc, (n, len, !(normaliseArgHoles defs [] ty)))
|
||||
|
||||
record TTMFile where
|
||||
|
@ -95,7 +95,7 @@ nameTag (Resolved _) = 8
|
||||
|
||||
export
|
||||
Ord Name where
|
||||
compare (NS x y) (NS x' y')
|
||||
compare (NS x y) (NS x' y')
|
||||
= case compare y y' of -- Compare base name first (more likely to differ)
|
||||
EQ => compare x x'
|
||||
-- Because of the terrible way Idris 1 compiles 'case', this
|
||||
@ -103,7 +103,7 @@ Ord Name where
|
||||
GT => GT
|
||||
LT => LT
|
||||
compare (UN x) (UN y) = compare x y
|
||||
compare (MN x y) (MN x' y')
|
||||
compare (MN x y) (MN x' y')
|
||||
= case compare y y' of
|
||||
EQ => compare x x'
|
||||
GT => GT
|
||||
@ -176,5 +176,5 @@ nameEq (WithBlock x y) (WithBlock x' y') with (decEq x x')
|
||||
nameEq (Resolved x) (Resolved y) with (decEq x y)
|
||||
nameEq (Resolved y) (Resolved y) | (Yes Refl) = Just Refl
|
||||
nameEq (Resolved x) (Resolved y) | (No contra) = Nothing
|
||||
nameEq _ _ = Nothing
|
||||
nameEq _ _ = Nothing
|
||||
|
||||
|
@ -40,7 +40,7 @@ Stack vars = List (Closure vars)
|
||||
|
||||
evalWithOpts : {vars : _} ->
|
||||
Defs -> EvalOpts ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
Term (vars ++ free) -> Stack free -> Core (NF free)
|
||||
|
||||
export
|
||||
@ -67,11 +67,11 @@ useMeta fc n defs opts
|
||||
parameters (defs : Defs, topopts : EvalOpts)
|
||||
mutual
|
||||
eval : {vars : _} ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
Term (vars ++ free) -> Stack free -> Core (NF free)
|
||||
eval env locs (Local fc mrig idx prf) stk
|
||||
= evalLocal env fc mrig idx prf stk locs
|
||||
eval env locs (Ref fc nt fn) stk
|
||||
eval env locs (Local fc mrig idx prf) stk
|
||||
= evalLocal env fc mrig idx prf stk locs
|
||||
eval env locs (Ref fc nt fn) stk
|
||||
= evalRef env locs False fc nt fn stk (NApp fc (NRef nt fn) stk)
|
||||
eval {vars} {free} env locs (Meta fc name idx args) stk
|
||||
= evalMeta env locs fc name idx (closeArgs args) stk
|
||||
@ -88,46 +88,46 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
= if holesOnly topopts || argHolesOnly topopts
|
||||
then do b' <- traverse (\tm => eval env locs tm []) b
|
||||
pure $ NBind fc x b'
|
||||
(\defs', arg => evalWithOpts defs' topopts
|
||||
(\defs', arg => evalWithOpts defs' topopts
|
||||
env (arg :: locs) scope stk)
|
||||
else eval env (MkClosure topopts locs env val :: locs) scope stk
|
||||
eval env locs (Bind fc x b scope) stk
|
||||
eval env locs (Bind fc x b scope) stk
|
||||
= do b' <- traverse (\tm => eval env locs tm []) b
|
||||
pure $ NBind fc x b'
|
||||
(\defs', arg => evalWithOpts defs' topopts
|
||||
(\defs', arg => evalWithOpts defs' topopts
|
||||
env (arg :: locs) scope stk)
|
||||
eval env locs (App fc fn arg) stk
|
||||
eval env locs (App fc fn arg) stk
|
||||
= eval env locs fn (MkClosure topopts locs env arg :: stk)
|
||||
eval env locs (As fc n tm) stk
|
||||
eval env locs (As fc n tm) stk
|
||||
= if removeAs topopts
|
||||
then eval env locs tm stk
|
||||
else do n' <- eval env locs n stk
|
||||
tm' <- eval env locs tm stk
|
||||
tm' <- eval env locs tm stk
|
||||
pure (NAs fc n' tm')
|
||||
eval env locs (TDelayed fc r ty) stk
|
||||
eval env locs (TDelayed fc r ty) stk
|
||||
= do ty' <- eval env locs ty stk
|
||||
pure (NDelayed fc r ty')
|
||||
eval env locs (TDelay fc r ty tm) stk
|
||||
eval env locs (TDelay fc r ty tm) stk
|
||||
= pure (NDelay fc r (MkClosure topopts locs env ty)
|
||||
(MkClosure topopts locs env tm))
|
||||
eval env locs (TForce fc tm) stk
|
||||
eval env locs (TForce fc r tm) stk
|
||||
= do tm' <- eval env locs tm []
|
||||
case tm' of
|
||||
NDelay fc r _ arg =>
|
||||
NDelay fc r _ arg =>
|
||||
eval env (arg :: locs) (Local {name = UN "fvar"} fc Nothing _ First) stk
|
||||
_ => pure (NForce fc tm' stk)
|
||||
_ => pure (NForce fc r tm' stk)
|
||||
eval env locs (PrimVal fc c) stk = pure $ NPrimVal fc c
|
||||
eval env locs (Erased fc) stk = pure $ NErased fc
|
||||
eval env locs (TType fc) stk = pure $ NType fc
|
||||
|
||||
|
||||
evalLocal : {vars : _} ->
|
||||
Env Term free ->
|
||||
FC -> Maybe Bool ->
|
||||
Env Term free ->
|
||||
FC -> Maybe Bool ->
|
||||
(idx : Nat) -> .(IsVar name idx (vars ++ free)) ->
|
||||
Stack free ->
|
||||
LocalEnv free vars ->
|
||||
LocalEnv free vars ->
|
||||
Core (NF free)
|
||||
evalLocal {vars = []} env fc mrig idx prf stk locs
|
||||
evalLocal {vars = []} env fc mrig idx prf stk locs
|
||||
= if not (holesOnly topopts || argHolesOnly topopts) && isLet mrig idx env
|
||||
then
|
||||
case getBinder prf env of
|
||||
@ -146,7 +146,7 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
isLet _ n env = isLet' n env
|
||||
evalLocal env fc mrig Z First stk (MkClosure opts locs' env' tm' :: locs)
|
||||
= evalWithOpts defs opts env' locs' tm' stk
|
||||
evalLocal {free} {vars = x :: xs}
|
||||
evalLocal {free} {vars = x :: xs}
|
||||
env fc mrig Z First stk (MkNFClosure nf :: locs)
|
||||
= applyToStack nf stk
|
||||
where
|
||||
@ -159,19 +159,19 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
(NApp fc (NRef nt fn) args)
|
||||
applyToStack (NApp fc (NLocal mrig idx p) args) stk
|
||||
= let MkVar p' = insertVarNames {outer=[]} {ns = xs} idx p in
|
||||
evalLocal env fc mrig _ p' (args ++ stk) locs
|
||||
applyToStack (NDCon fc n t a args) stk
|
||||
evalLocal env fc mrig _ p' (args ++ stk) locs
|
||||
applyToStack (NDCon fc n t a args) stk
|
||||
= pure $ NDCon fc n t a (args ++ stk)
|
||||
applyToStack (NTCon fc n t a args) stk
|
||||
applyToStack (NTCon fc n t a args) stk
|
||||
= pure $ NTCon fc n t a (args ++ stk)
|
||||
applyToStack nf _ = pure nf
|
||||
|
||||
evalLocal {vars = x :: xs} {free}
|
||||
env fc mrig (S idx) (Later p) stk (_ :: locs)
|
||||
= evalLocal {vars = xs} env fc mrig idx p stk locs
|
||||
|
||||
= evalLocal {vars = xs} env fc mrig idx p stk locs
|
||||
|
||||
evalMeta : {vars : _} ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
FC -> Name -> Int -> List (Closure free) ->
|
||||
Stack free -> Core (NF free)
|
||||
evalMeta {vars} env locs fc nm i args stk
|
||||
@ -179,7 +179,7 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
(NApp fc (NMeta nm i args) stk)
|
||||
|
||||
evalRef : {vars : _} ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
Env Term free -> LocalEnv free vars ->
|
||||
(isMeta : Bool) ->
|
||||
FC -> NameType -> Name -> Stack free -> (def : Lazy (NF free)) ->
|
||||
Core (NF free)
|
||||
@ -189,19 +189,19 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
= pure $ NTCon fc fn tag arity stk
|
||||
evalRef env locs meta fc Bound fn stk def
|
||||
= pure def
|
||||
evalRef env locs meta fc nt n stk def
|
||||
evalRef env locs meta fc nt n stk def
|
||||
= do Just res <- lookupCtxtExact n (gamma defs)
|
||||
| Nothing => pure def
|
||||
| Nothing => pure def
|
||||
let redok = evalAll topopts ||
|
||||
reducibleInAny (currentNS defs :: nestedNS defs)
|
||||
(fullname res)
|
||||
reducibleInAny (currentNS defs :: nestedNS defs)
|
||||
(fullname res)
|
||||
(visibility res)
|
||||
if redok
|
||||
then do
|
||||
opts' <- if noCycles res
|
||||
then useMeta fc n defs topopts
|
||||
else pure topopts
|
||||
evalDef env locs opts' meta fc
|
||||
evalDef env locs opts' meta fc
|
||||
(multiplicity res) (definition res) (flags res) stk def
|
||||
else pure def
|
||||
|
||||
@ -212,28 +212,28 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
getCaseBound [] [] loc = Just loc
|
||||
getCaseBound [] (x :: xs) loc = Nothing -- mismatched arg length
|
||||
getCaseBound (arg :: args) [] loc = Nothing -- mismatched arg length
|
||||
getCaseBound (arg :: args) (n :: ns) loc
|
||||
getCaseBound (arg :: args) (n :: ns) loc
|
||||
= do loc' <- getCaseBound args ns loc
|
||||
pure (arg :: loc')
|
||||
|
||||
evalConAlt : Env Term free ->
|
||||
evalConAlt : Env Term free ->
|
||||
LocalEnv free (more ++ vars) -> EvalOpts -> FC ->
|
||||
Stack free ->
|
||||
(args : List Name) ->
|
||||
List (Closure free) ->
|
||||
Stack free ->
|
||||
(args : List Name) ->
|
||||
List (Closure free) ->
|
||||
CaseTree (args ++ more) ->
|
||||
(default : Core (NF free)) ->
|
||||
(default : Core (NF free)) ->
|
||||
Core (NF free)
|
||||
evalConAlt {more} {vars} env loc opts fc stk args args' sc def
|
||||
= maybe def (\bound =>
|
||||
let loc' : LocalEnv _ ((args ++ more) ++ vars)
|
||||
= maybe def (\bound =>
|
||||
let loc' : LocalEnv _ ((args ++ more) ++ vars)
|
||||
= rewrite sym (appendAssociative args more vars) in
|
||||
bound in
|
||||
evalTree env loc' opts fc stk sc def)
|
||||
(getCaseBound args' args loc)
|
||||
|
||||
tryAlt : Env Term free ->
|
||||
LocalEnv free (more ++ vars) -> EvalOpts -> FC ->
|
||||
LocalEnv free (more ++ vars) -> EvalOpts -> FC ->
|
||||
Stack free -> NF free -> CaseAlt more ->
|
||||
(default : Core (NF free)) -> Core (NF free)
|
||||
-- Ordinary constructor matching
|
||||
@ -255,10 +255,10 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
tryAlt env loc opts fc stk (NType _) (ConCase (UN "Type") tag [] sc) def
|
||||
= evalTree env loc opts fc stk sc def
|
||||
-- Arrow matching, in typecase
|
||||
tryAlt {more} {vars}
|
||||
tryAlt {more} {vars}
|
||||
env loc opts fc stk (NBind pfc x (Pi r e aty) scty) (ConCase (UN "->") tag [s,t] sc) def
|
||||
= evalConAlt {more} {vars} env loc opts fc stk [s,t]
|
||||
[MkNFClosure aty,
|
||||
[MkNFClosure aty,
|
||||
MkNFClosure (NBind pfc x (Lam r e aty) scty)]
|
||||
sc def
|
||||
-- Delay matching
|
||||
@ -270,7 +270,7 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
else def
|
||||
-- Default case matches against any *concrete* value
|
||||
tryAlt env loc opts fc stk val (DefaultCase sc) def
|
||||
= if concrete val
|
||||
= if concrete val
|
||||
then evalTree env loc opts fc stk sc def
|
||||
else def
|
||||
where
|
||||
@ -292,24 +292,24 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
= tryAlt env loc opts fc stk val x (findAlt env loc opts fc stk val xs def)
|
||||
|
||||
evalTree : {vars : _} ->
|
||||
Env Term free -> LocalEnv free (args ++ vars) ->
|
||||
Env Term free -> LocalEnv free (args ++ vars) ->
|
||||
EvalOpts -> FC ->
|
||||
Stack free -> CaseTree args ->
|
||||
(default : Core (NF free)) -> Core (NF free)
|
||||
evalTree {args} {vars} {free} env loc opts fc stk (Case idx x _ alts) def
|
||||
= do let x' : IsVar _ _ ((args ++ vars) ++ free)
|
||||
= do let x' : IsVar _ _ ((args ++ vars) ++ free)
|
||||
= rewrite sym (appendAssociative args vars free) in
|
||||
varExtend x
|
||||
xval <- evalLocal env fc Nothing idx x' [] loc
|
||||
xval <- evalLocal env fc Nothing idx x' [] loc
|
||||
findAlt env loc opts fc stk xval alts def
|
||||
evalTree {args} {vars} {free} env loc opts fc stk (STerm tm) def
|
||||
= do let tm' : Term ((args ++ vars) ++ free)
|
||||
= do let tm' : Term ((args ++ vars) ++ free)
|
||||
= rewrite sym (appendAssociative args vars free) in
|
||||
embed tm
|
||||
embed tm
|
||||
case fuel opts of
|
||||
Nothing => evalWithOpts defs opts env loc tm' stk
|
||||
Just Z => def
|
||||
Just (S k) =>
|
||||
Just (S k) =>
|
||||
do let opts' = record { fuel = Just k } opts
|
||||
evalWithOpts defs opts' env loc tm' stk
|
||||
evalTree env loc opts fc stk _ def = def
|
||||
@ -320,22 +320,22 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
Maybe (Vect arity (Closure free), Stack free)
|
||||
takeFromStack arity stk = takeStk arity stk []
|
||||
where
|
||||
takeStk : (remain : Nat) -> Stack free ->
|
||||
Vect got (Closure free) ->
|
||||
takeStk : (remain : Nat) -> Stack free ->
|
||||
Vect got (Closure free) ->
|
||||
Maybe (Vect (got + remain) (Closure free), Stack free)
|
||||
takeStk {got} Z stk acc = Just (rewrite plusZeroRightNeutral got in
|
||||
reverse acc, stk)
|
||||
takeStk (S k) [] acc = Nothing
|
||||
takeStk {got} (S k) (arg :: stk) acc
|
||||
takeStk {got} (S k) (arg :: stk) acc
|
||||
= rewrite sym (plusSuccRightSucc got k) in
|
||||
takeStk k stk (arg :: acc)
|
||||
|
||||
extendFromStack : (args : List Name) ->
|
||||
extendFromStack : (args : List Name) ->
|
||||
LocalEnv free vars -> Stack free ->
|
||||
Maybe (LocalEnv free (args ++ vars), Stack free)
|
||||
extendFromStack [] loc stk = Just (loc, stk)
|
||||
extendFromStack (n :: ns) loc [] = Nothing
|
||||
extendFromStack (n :: ns) loc (arg :: args)
|
||||
extendFromStack (n :: ns) loc (arg :: args)
|
||||
= do (loc', stk') <- extendFromStack ns loc args
|
||||
pure (arg :: loc', stk')
|
||||
|
||||
@ -345,7 +345,7 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
evalOp {arity} fn stk def
|
||||
= case takeFromStack arity stk of
|
||||
-- Stack must be exactly the right height
|
||||
Just (args, []) =>
|
||||
Just (args, []) =>
|
||||
do argsnf <- evalAll args
|
||||
case fn argsnf of
|
||||
Nothing => pure def
|
||||
@ -356,11 +356,11 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
evalAll : Vect n (Closure free) -> Core (Vect n (NF free))
|
||||
evalAll [] = pure []
|
||||
evalAll (c :: cs) = pure $ !(evalClosure defs c) :: !(evalAll cs)
|
||||
|
||||
|
||||
evalDef : {vars : _} ->
|
||||
Env Term free -> LocalEnv free vars -> EvalOpts ->
|
||||
(isMeta : Bool) -> FC ->
|
||||
RigCount -> Def -> List DefFlag ->
|
||||
RigCount -> Def -> List DefFlag ->
|
||||
Stack free -> (def : Lazy (NF free)) ->
|
||||
Core (NF free)
|
||||
evalDef {vars} env locs opts meta fc rigd (PMDef r args tree _ _) flags stk def
|
||||
@ -373,14 +373,14 @@ parameters (defs : Defs, topopts : EvalOpts)
|
||||
-- + It's a metavariable and not in Rig0
|
||||
-- + It's a metavariable and we're not in 'argHolesOnly'
|
||||
-- + It's inlinable and and we're in 'tcInline'
|
||||
= if r
|
||||
= if r
|
||||
|| (not (holesOnly opts) && not (argHolesOnly opts) && not (tcInline opts))
|
||||
|| (meta && rigd /= Rig0)
|
||||
|| (meta && holesOnly opts)
|
||||
|| (tcInline opts && elem TCInline flags)
|
||||
then case extendFromStack args locs stk of
|
||||
Nothing => pure def
|
||||
Just (locs', stk') =>
|
||||
Just (locs', stk') =>
|
||||
evalTree env locs' opts fc stk' tree (pure def)
|
||||
else pure def
|
||||
evalDef {vars} env locs opts meta fc rigd (Builtin op) flags stk def
|
||||
@ -407,17 +407,17 @@ nfOpts opts defs env tm = eval defs opts env [] tm []
|
||||
|
||||
export
|
||||
gnf : Env Term vars -> Term vars -> Glued vars
|
||||
gnf env tm
|
||||
gnf env tm
|
||||
= MkGlue True
|
||||
(pure tm)
|
||||
(pure tm)
|
||||
(\c => do defs <- get Ctxt
|
||||
nf defs env tm)
|
||||
|
||||
export
|
||||
gnfOpts : EvalOpts -> Env Term vars -> Term vars -> Glued vars
|
||||
gnfOpts opts env tm
|
||||
gnfOpts opts env tm
|
||||
= MkGlue True
|
||||
(pure tm)
|
||||
(pure tm)
|
||||
(\c => do defs <- get Ctxt
|
||||
nfOpts opts defs env tm)
|
||||
|
||||
@ -450,7 +450,7 @@ genName n
|
||||
mutual
|
||||
quoteArgs : {bound : _} ->
|
||||
Ref QVar Int -> Defs -> Bounds bound ->
|
||||
Env Term free -> List (Closure free) ->
|
||||
Env Term free -> List (Closure free) ->
|
||||
Core (List (Term (bound ++ free)))
|
||||
quoteArgs q defs bounds env [] = pure []
|
||||
quoteArgs q defs bounds env (a :: args)
|
||||
@ -458,27 +458,27 @@ mutual
|
||||
!(quoteArgs q defs bounds env args))
|
||||
|
||||
quoteHead : {bound : _} ->
|
||||
Ref QVar Int -> Defs ->
|
||||
FC -> Bounds bound -> Env Term free -> NHead free ->
|
||||
Ref QVar Int -> Defs ->
|
||||
FC -> Bounds bound -> Env Term free -> NHead free ->
|
||||
Core (Term (bound ++ free))
|
||||
quoteHead {bound} q defs fc bounds env (NLocal mrig _ prf)
|
||||
quoteHead {bound} q defs fc bounds env (NLocal mrig _ prf)
|
||||
= let MkVar prf' = addLater bound prf in
|
||||
pure $ Local fc mrig _ prf'
|
||||
where
|
||||
addLater : (ys : List Name) -> IsVar n idx xs ->
|
||||
addLater : (ys : List Name) -> IsVar n idx xs ->
|
||||
Var (ys ++ xs)
|
||||
addLater [] isv = MkVar isv
|
||||
addLater (x :: xs) isv
|
||||
addLater (x :: xs) isv
|
||||
= let MkVar isv' = addLater xs isv in
|
||||
MkVar (Later isv')
|
||||
quoteHead q defs fc bounds env (NRef Bound (MN n i))
|
||||
quoteHead q defs fc bounds env (NRef Bound (MN n i))
|
||||
= case findName bounds of
|
||||
Just (MkVar p) => pure $ Local fc Nothing _ (varExtend p)
|
||||
Nothing => pure $ Ref fc Bound (MN n i)
|
||||
where
|
||||
findName : Bounds bound' -> Maybe (Var bound')
|
||||
findName None = Nothing
|
||||
findName (Add x (MN n' i') ns)
|
||||
findName (Add x (MN n' i') ns)
|
||||
= if i == i' -- this uniquely identifies it, given how we
|
||||
-- generated the names, and is a faster test!
|
||||
then Just (MkVar First)
|
||||
@ -494,9 +494,9 @@ mutual
|
||||
|
||||
quoteBinder : {bound : _} ->
|
||||
Ref QVar Int -> Defs -> Bounds bound ->
|
||||
Env Term free -> Binder (NF free) ->
|
||||
Env Term free -> Binder (NF free) ->
|
||||
Core (Binder (Term (bound ++ free)))
|
||||
quoteBinder q defs bounds env (Lam r p ty)
|
||||
quoteBinder q defs bounds env (Lam r p ty)
|
||||
= do ty' <- quoteGenNF q defs bounds env ty
|
||||
pure (Lam r p ty')
|
||||
quoteBinder q defs bounds env (Let r val ty)
|
||||
@ -519,11 +519,11 @@ mutual
|
||||
|
||||
quoteGenNF : {bound : _} ->
|
||||
Ref QVar Int ->
|
||||
Defs -> Bounds bound ->
|
||||
Defs -> Bounds bound ->
|
||||
Env Term vars -> NF vars -> Core (Term (bound ++ vars))
|
||||
quoteGenNF q defs bound env (NBind fc n b sc)
|
||||
= do var <- genName "qv"
|
||||
sc' <- quoteGenNF q defs (Add n var bound) env
|
||||
sc' <- quoteGenNF q defs (Add n var bound) env
|
||||
!(sc defs (toClosure defaultOpts env (Ref fc Bound var)))
|
||||
b' <- quoteBinder q defs bound env b
|
||||
pure (Bind fc n b' sc')
|
||||
@ -531,10 +531,10 @@ mutual
|
||||
= do f' <- quoteHead q defs fc bound env f
|
||||
args' <- quoteArgs q defs bound env args
|
||||
pure $ apply fc f' args'
|
||||
quoteGenNF q defs bound env (NDCon fc n t ar args)
|
||||
quoteGenNF q defs bound env (NDCon fc n t ar args)
|
||||
= do args' <- quoteArgs q defs bound env args
|
||||
pure $ apply fc (Ref fc (DataCon t ar) n) args'
|
||||
quoteGenNF q defs bound env (NTCon fc n t ar args)
|
||||
quoteGenNF q defs bound env (NTCon fc n t ar args)
|
||||
= do args' <- quoteArgs q defs bound env args
|
||||
pure $ apply fc (Ref fc (TyCon t ar) n) args'
|
||||
quoteGenNF q defs bound env (NAs fc n pat)
|
||||
@ -552,17 +552,17 @@ mutual
|
||||
pure (TDelay fc r tyQ argQ)
|
||||
where
|
||||
toHolesOnly : Closure vs -> Closure vs
|
||||
toHolesOnly (MkClosure _ locs env tm)
|
||||
toHolesOnly (MkClosure _ locs env tm)
|
||||
= MkClosure withHoles locs env tm
|
||||
toHolesOnly c = c
|
||||
quoteGenNF q defs bound env (NForce fc arg args)
|
||||
= do args' <- quoteArgs q defs bound env args
|
||||
quoteGenNF q defs bound env (NForce fc r arg args)
|
||||
= do args' <- quoteArgs q defs bound env args
|
||||
case arg of
|
||||
NDelay fc _ _ arg =>
|
||||
do argNF <- evalClosure defs arg
|
||||
pure $ apply fc !(quoteGenNF q defs bound env argNF) args'
|
||||
t => do arg' <- quoteGenNF q defs bound env arg
|
||||
pure $ apply fc (TForce fc arg') args'
|
||||
pure $ apply fc (TForce fc r arg') args'
|
||||
quoteGenNF q defs bound env (NPrimVal fc c) = pure $ PrimVal fc c
|
||||
quoteGenNF q defs bound env (NErased fc) = pure $ Erased fc
|
||||
quoteGenNF q defs bound env (NType fc) = pure $ TType fc
|
||||
@ -581,10 +581,10 @@ Quote Closure where
|
||||
|
||||
export
|
||||
glueBack : Defs -> Env Term vars -> NF vars -> Glued vars
|
||||
glueBack defs env nf
|
||||
glueBack defs env nf
|
||||
= MkGlue False
|
||||
(do empty <- clearDefs defs
|
||||
quote empty env nf)
|
||||
quote empty env nf)
|
||||
(const (pure nf))
|
||||
|
||||
export
|
||||
@ -593,27 +593,27 @@ normalise defs env tm = quote defs env !(nf defs env tm)
|
||||
|
||||
export
|
||||
normaliseOpts : EvalOpts -> Defs -> Env Term free -> Term free -> Core (Term free)
|
||||
normaliseOpts opts defs env tm
|
||||
normaliseOpts opts defs env tm
|
||||
= quote defs env !(nfOpts opts defs env tm)
|
||||
|
||||
export
|
||||
normaliseHoles : Defs -> Env Term free -> Term free -> Core (Term free)
|
||||
normaliseHoles defs env tm
|
||||
normaliseHoles defs env tm
|
||||
= quote defs env !(nfOpts withHoles defs env tm)
|
||||
|
||||
export
|
||||
normaliseLHS : Defs -> Env Term free -> Term free -> Core (Term free)
|
||||
normaliseLHS defs env tm
|
||||
normaliseLHS defs env tm
|
||||
= quote defs env !(nfOpts onLHS defs env tm)
|
||||
|
||||
export
|
||||
normaliseArgHoles : Defs -> Env Term free -> Term free -> Core (Term free)
|
||||
normaliseArgHoles defs env tm
|
||||
normaliseArgHoles defs env tm
|
||||
= quote defs env !(nfOpts withArgHoles defs env tm)
|
||||
|
||||
export
|
||||
normaliseAll : Defs -> Env Term free -> Term free -> Core (Term free)
|
||||
normaliseAll defs env tm
|
||||
normaliseAll defs env tm
|
||||
= quote defs env !(nfOpts withAll defs env tm)
|
||||
|
||||
-- Normalise, but without normalising the types of binders. Dealing with
|
||||
@ -621,19 +621,19 @@ normaliseAll defs env tm
|
||||
-- a big win
|
||||
export
|
||||
normaliseScope : Defs -> Env Term vars -> Term vars -> Core (Term vars)
|
||||
normaliseScope defs env (Bind fc n b sc)
|
||||
normaliseScope defs env (Bind fc n b sc)
|
||||
= pure $ Bind fc n b !(normaliseScope defs (b :: env) sc)
|
||||
normaliseScope defs env tm = normalise defs env tm
|
||||
|
||||
public export
|
||||
interface Convert (tm : List Name -> Type) where
|
||||
convert : Defs -> Env Term vars ->
|
||||
convert : Defs -> Env Term vars ->
|
||||
tm vars -> tm vars -> Core Bool
|
||||
convGen : Ref QVar Int ->
|
||||
Defs -> Env Term vars ->
|
||||
Defs -> Env Term vars ->
|
||||
tm vars -> tm vars -> Core Bool
|
||||
|
||||
convert defs env tm tm'
|
||||
convert defs env tm tm'
|
||||
= do q <- newRef QVar 0
|
||||
convGen q defs env tm tm'
|
||||
|
||||
@ -646,10 +646,10 @@ mutual
|
||||
allConv q defs env _ _ = pure False
|
||||
|
||||
chkConvHead : Ref QVar Int -> Defs -> Env Term vars ->
|
||||
NHead vars -> NHead vars -> Core Bool
|
||||
NHead vars -> NHead vars -> Core Bool
|
||||
chkConvHead q defs env (NLocal _ idx _) (NLocal _ idx' _) = pure $ idx == idx'
|
||||
chkConvHead q defs env (NRef _ n) (NRef _ n') = pure $ n == n'
|
||||
chkConvHead q defs env (NMeta n i args) (NMeta n' i' args')
|
||||
chkConvHead q defs env (NMeta n i args) (NMeta n' i' args')
|
||||
= if i == i'
|
||||
then allConv q defs env args args'
|
||||
else pure False
|
||||
@ -678,7 +678,7 @@ mutual
|
||||
|
||||
export
|
||||
Convert NF where
|
||||
convGen q defs env (NBind fc x b sc) (NBind _ x' b' sc')
|
||||
convGen q defs env (NBind fc x b sc) (NBind _ x' b' sc')
|
||||
= do var <- genName "conv"
|
||||
let c = MkClosure defaultOpts [] env (Ref fc Bound var)
|
||||
bok <- convBinders q defs env b b'
|
||||
@ -688,16 +688,16 @@ mutual
|
||||
convGen q defs env bsc bsc'
|
||||
else pure False
|
||||
|
||||
convGen q defs env tmx@(NBind fc x (Lam c ix tx) scx) tmy
|
||||
convGen q defs env tmx@(NBind fc x (Lam c ix tx) scx) tmy
|
||||
= do empty <- clearDefs defs
|
||||
etay <- nf defs env
|
||||
etay <- nf defs env
|
||||
(Bind fc x (Lam c ix !(quote empty env tx))
|
||||
(App fc (weaken !(quote empty env tmy))
|
||||
(Local fc Nothing _ First)))
|
||||
convGen q defs env tmx etay
|
||||
convGen q defs env tmx tmy@(NBind fc y (Lam c iy ty) scy)
|
||||
= do empty <- clearDefs defs
|
||||
etax <- nf defs env
|
||||
etax <- nf defs env
|
||||
(Bind fc y (Lam c iy !(quote empty env ty))
|
||||
(App fc (weaken !(quote empty env tmx))
|
||||
(Local fc Nothing _ First)))
|
||||
@ -727,9 +727,11 @@ mutual
|
||||
= if compatible r r'
|
||||
then convGen q defs env arg arg'
|
||||
else pure False
|
||||
convGen q defs env (NForce _ arg args) (NForce _ arg' args')
|
||||
= if !(convGen q defs env arg arg')
|
||||
then allConv q defs env args args'
|
||||
convGen q defs env (NForce _ r arg args) (NForce _ r' arg' args')
|
||||
= if compatible r r'
|
||||
then if !(convGen q defs env arg arg')
|
||||
then allConv q defs env args args'
|
||||
else pure False
|
||||
else pure False
|
||||
|
||||
convGen q defs env (NPrimVal _ c) (NPrimVal _ c') = pure (c == c')
|
||||
@ -750,7 +752,7 @@ mutual
|
||||
|
||||
export
|
||||
getValArity : Defs -> Env Term vars -> NF vars -> Core Nat
|
||||
getValArity defs env (NBind fc x (Pi _ _ _) sc)
|
||||
getValArity defs env (NBind fc x (Pi _ _ _) sc)
|
||||
= pure (S !(getValArity defs env !(sc defs (toClosure defaultOpts env (Erased fc)))))
|
||||
getValArity defs env val = pure 0
|
||||
|
||||
@ -768,7 +770,7 @@ logNF lvl msg env tmnf
|
||||
then do defs <- get Ctxt
|
||||
tm <- quote defs env tmnf
|
||||
tm' <- toFullNames tm
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
++ ": " ++ show tm'
|
||||
else pure ()
|
||||
|
||||
@ -783,7 +785,7 @@ logTermNF lvl msg env tm
|
||||
then do defs <- get Ctxt
|
||||
tmnf <- normaliseHoles defs env tm
|
||||
tm' <- toFullNames tmnf
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
++ ": " ++ show tm'
|
||||
else pure ()
|
||||
|
||||
@ -796,7 +798,7 @@ logGlue lvl msg env gtm
|
||||
then do defs <- get Ctxt
|
||||
tm <- getTerm gtm
|
||||
tm' <- toFullNames tm
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
++ ": " ++ show tm'
|
||||
else pure ()
|
||||
|
||||
@ -810,7 +812,7 @@ logGlueNF lvl msg env gtm
|
||||
tm <- getTerm gtm
|
||||
tmnf <- normaliseHoles defs env tm
|
||||
tm' <- toFullNames tmnf
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
coreLift $ putStrLn $ "LOG " ++ show lvl ++ ": " ++ msg
|
||||
++ ": " ++ show tm'
|
||||
else pure ()
|
||||
|
||||
@ -828,11 +830,11 @@ logEnv lvl msg env
|
||||
dumpEnv [] = pure ()
|
||||
dumpEnv {vs = x :: _} (Let c val ty :: bs)
|
||||
= do logTermNF lvl (msg ++ ": let " ++ show x) bs val
|
||||
logTermNF lvl (msg ++ ":" ++ show c ++ " " ++
|
||||
logTermNF lvl (msg ++ ":" ++ show c ++ " " ++
|
||||
show x) bs ty
|
||||
dumpEnv bs
|
||||
dumpEnv {vs = x :: _} (b :: bs)
|
||||
= do logTermNF lvl (msg ++ ":" ++ show (multiplicity b) ++ " " ++
|
||||
= do logTermNF lvl (msg ++ ":" ++ show (multiplicity b) ++ " " ++
|
||||
show x) bs (binderType b)
|
||||
dumpEnv bs
|
||||
|
||||
@ -851,29 +853,29 @@ replace' {vars} tmpi defs env lhs parg tm
|
||||
|
||||
repSub : NF vars -> Core (Term vars)
|
||||
repSub (NBind fc x b scfn)
|
||||
= do b' <- traverse repSub b
|
||||
= do b' <- traverse repSub b
|
||||
let x' = MN "tmp" tmpi
|
||||
sc' <- replace' (tmpi + 1) defs env lhs parg
|
||||
sc' <- replace' (tmpi + 1) defs env lhs parg
|
||||
!(scfn defs (toClosure defaultOpts env (Ref fc Bound x')))
|
||||
pure (Bind fc x b' (refsToLocals (Add x x' None) sc'))
|
||||
repSub (NApp fc hd [])
|
||||
repSub (NApp fc hd [])
|
||||
= do empty <- clearDefs defs
|
||||
quote empty env (NApp fc hd [])
|
||||
repSub (NApp fc hd args)
|
||||
repSub (NApp fc hd args)
|
||||
= do args' <- traverse repArg args
|
||||
pure $ apply fc
|
||||
pure $ apply fc
|
||||
!(replace' tmpi defs env lhs parg (NApp fc hd []))
|
||||
args'
|
||||
repSub (NDCon fc n t a args)
|
||||
repSub (NDCon fc n t a args)
|
||||
= do args' <- traverse repArg args
|
||||
empty <- clearDefs defs
|
||||
pure $ apply fc
|
||||
pure $ apply fc
|
||||
!(quote empty env (NDCon fc n t a []))
|
||||
args'
|
||||
repSub (NTCon fc n t a args)
|
||||
repSub (NTCon fc n t a args)
|
||||
= do args' <- traverse repArg args
|
||||
empty <- clearDefs defs
|
||||
pure $ apply fc
|
||||
pure $ apply fc
|
||||
!(quote empty env (NTCon fc n t a []))
|
||||
args'
|
||||
repSub (NAs fc a p)
|
||||
@ -887,10 +889,10 @@ replace' {vars} tmpi defs env lhs parg tm
|
||||
= do ty' <- replace' tmpi defs env lhs parg !(evalClosure defs ty)
|
||||
tm' <- replace' tmpi defs env lhs parg !(evalClosure defs tm)
|
||||
pure (TDelay fc r ty' tm')
|
||||
repSub (NForce fc tm args)
|
||||
repSub (NForce fc r tm args)
|
||||
= do args' <- traverse repArg args
|
||||
tm' <- repSub tm
|
||||
pure $ apply fc (TForce fc tm') args'
|
||||
pure $ apply fc (TForce fc r tm') args'
|
||||
repSub tm = do empty <- clearDefs defs
|
||||
quote empty env tm
|
||||
|
||||
|
@ -65,40 +65,40 @@ strLength _ = Nothing
|
||||
|
||||
strHead : Vect 1 (NF vars) -> Maybe (NF vars)
|
||||
strHead [NPrimVal fc (Str "")] = Nothing
|
||||
strHead [NPrimVal fc (Str str)]
|
||||
strHead [NPrimVal fc (Str str)]
|
||||
= Just (NPrimVal fc (Ch (assert_total (strHead str))))
|
||||
strHead _ = Nothing
|
||||
|
||||
strTail : Vect 1 (NF vars) -> Maybe (NF vars)
|
||||
strTail [NPrimVal fc (Str "")] = Nothing
|
||||
strTail [NPrimVal fc (Str str)]
|
||||
strTail [NPrimVal fc (Str str)]
|
||||
= Just (NPrimVal fc (Str (assert_total (strTail str))))
|
||||
strTail _ = Nothing
|
||||
|
||||
strIndex : Vect 2 (NF vars) -> Maybe (NF vars)
|
||||
strIndex [NPrimVal fc (Str str), NPrimVal _ (I i)]
|
||||
strIndex [NPrimVal fc (Str str), NPrimVal _ (I i)]
|
||||
= if i >= 0 && cast i < length str
|
||||
then Just (NPrimVal fc (Ch (assert_total (prim__strIndex str i))))
|
||||
else Nothing
|
||||
strIndex _ = Nothing
|
||||
|
||||
strCons : Vect 2 (NF vars) -> Maybe (NF vars)
|
||||
strCons [NPrimVal fc (Ch x), NPrimVal _ (Str y)]
|
||||
strCons [NPrimVal fc (Ch x), NPrimVal _ (Str y)]
|
||||
= Just (NPrimVal fc (Str (strCons x y)))
|
||||
strCons _ = Nothing
|
||||
|
||||
strAppend : Vect 2 (NF vars) -> Maybe (NF vars)
|
||||
strAppend [NPrimVal fc (Str x), NPrimVal _ (Str y)]
|
||||
strAppend [NPrimVal fc (Str x), NPrimVal _ (Str y)]
|
||||
= Just (NPrimVal fc (Str (x ++ y)))
|
||||
strAppend _ = Nothing
|
||||
|
||||
strReverse : Vect 1 (NF vars) -> Maybe (NF vars)
|
||||
strReverse [NPrimVal fc (Str x)]
|
||||
strReverse [NPrimVal fc (Str x)]
|
||||
= Just (NPrimVal fc (Str (reverse x)))
|
||||
strReverse _ = Nothing
|
||||
|
||||
strSubstr : Vect 3 (NF vars) -> Maybe (NF vars)
|
||||
strSubstr [NPrimVal fc (I start), NPrimVal _ (I len), NPrimVal _ (Str str)]
|
||||
strSubstr [NPrimVal fc (I start), NPrimVal _ (I len), NPrimVal _ (Str str)]
|
||||
= Just (NPrimVal fc (Str (prim__strSubstr start len str)))
|
||||
strSubstr _ = Nothing
|
||||
|
||||
@ -241,14 +241,14 @@ believeMe [_, _, NType fc] = Just (NType fc)
|
||||
believeMe [_, _, val] = Nothing
|
||||
|
||||
constTy : Constant -> Constant -> Constant -> ClosedTerm
|
||||
constTy a b c
|
||||
= PrimVal emptyFC a `linFnType`
|
||||
constTy a b c
|
||||
= PrimVal emptyFC a `linFnType`
|
||||
(PrimVal emptyFC b `linFnType` PrimVal emptyFC c)
|
||||
|
||||
constTy3 : Constant -> Constant -> Constant -> Constant -> ClosedTerm
|
||||
constTy3 a b c d
|
||||
= PrimVal emptyFC a `linFnType`
|
||||
(PrimVal emptyFC b `linFnType`
|
||||
constTy3 a b c d
|
||||
= PrimVal emptyFC a `linFnType`
|
||||
(PrimVal emptyFC b `linFnType`
|
||||
(PrimVal emptyFC c `linFnType` PrimVal emptyFC d))
|
||||
|
||||
predTy : Constant -> Constant -> ClosedTerm
|
||||
@ -264,7 +264,7 @@ doubleTy : ClosedTerm
|
||||
doubleTy = predTy DoubleType DoubleType
|
||||
|
||||
believeMeTy : ClosedTerm
|
||||
believeMeTy
|
||||
believeMeTy
|
||||
= Bind emptyFC (UN "a") (Pi Rig0 Explicit (TType emptyFC)) $
|
||||
Bind emptyFC (UN "b") (Pi Rig0 Explicit (TType emptyFC)) $
|
||||
Bind emptyFC (UN "x") (Pi RigW Explicit (Local emptyFC Nothing _ (Later First))) $
|
||||
@ -279,7 +279,7 @@ castTo DoubleType = castDouble
|
||||
castTo _ = const Nothing
|
||||
|
||||
export
|
||||
getOp : PrimFn arity ->
|
||||
getOp : PrimFn arity ->
|
||||
{vars : List Name} -> Vect arity (NF vars) -> Maybe (NF vars)
|
||||
getOp (Add ty) = binOp add
|
||||
getOp (Sub ty) = binOp sub
|
||||
@ -373,7 +373,7 @@ allPrimitives =
|
||||
map (\t => MkPrim (Neg t) (predTy t t) isTotal) [IntType, IntegerType, DoubleType] ++
|
||||
map (\t => MkPrim (ShiftL t) (arithTy t) notCovering) [IntType] ++
|
||||
map (\t => MkPrim (ShiftR t) (arithTy t) notCovering) [IntType] ++
|
||||
|
||||
|
||||
map (\t => MkPrim (LT t) (cmpTy t) isTotal) [IntType, IntegerType, CharType, DoubleType, StringType] ++
|
||||
map (\t => MkPrim (LTE t) (cmpTy t) isTotal) [IntType, IntegerType, CharType, DoubleType, StringType] ++
|
||||
map (\t => MkPrim (EQ t) (cmpTy t) isTotal) [IntType, IntegerType, CharType, DoubleType, StringType] ++
|
||||
|
@ -3,8 +3,8 @@ module Core.Reflect
|
||||
import Core.Context
|
||||
import Core.Env
|
||||
import Core.Normalise
|
||||
import Core.Value
|
||||
import Core.TT
|
||||
import Core.Value
|
||||
|
||||
%default covering
|
||||
|
||||
@ -49,12 +49,12 @@ reflection n = NS ["Reflection", "Language"] (UN n)
|
||||
|
||||
export
|
||||
cantReify : NF vars -> String -> Core a
|
||||
cantReify val ty
|
||||
cantReify val ty
|
||||
= throw (GenericMsg (getLoc val) ("Can't reify as " ++ ty))
|
||||
|
||||
export
|
||||
cantReflect : FC -> String -> Core a
|
||||
cantReflect fc ty
|
||||
cantReflect fc ty
|
||||
= throw (GenericMsg fc ("Can't reflect as " ++ ty))
|
||||
|
||||
export
|
||||
@ -155,6 +155,23 @@ Reflect a => Reflect (List a) where
|
||||
xs' <- reflect fc defs env xs
|
||||
appCon fc defs (prelude "::") [Erased fc, x', xs']
|
||||
|
||||
export
|
||||
Reify a => Reify (Maybe a) where
|
||||
reify defs (NDCon _ (NS _ (UN "Nothing")) _ _ _)
|
||||
= pure Nothing
|
||||
reify defs (NDCon _ (NS _ (UN "Just")) _ _ [_, x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (Just x')
|
||||
reify defs val = cantReify val "Maybe"
|
||||
|
||||
export
|
||||
Reflect a => Reflect (Maybe a) where
|
||||
reflect fc defs env Nothing = appCon fc defs (prelude "Nothing") [Erased fc]
|
||||
reflect fc defs env (Just x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (prelude "Just") [Erased fc, x']
|
||||
|
||||
|
||||
export
|
||||
(Reify a, Reify b) => Reify (a, b) where
|
||||
reify defs (NDCon _ (NS _ (UN "MkPair")) _ _ [_, _, x, y])
|
||||
@ -187,14 +204,14 @@ Reify Name where
|
||||
|
||||
export
|
||||
Reflect Name where
|
||||
reflect fc defs env (UN x)
|
||||
reflect fc defs env (UN x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "UN") [x']
|
||||
reflect fc defs env (MN x i)
|
||||
reflect fc defs env (MN x i)
|
||||
= do x' <- reflect fc defs env x
|
||||
i' <- reflect fc defs env i
|
||||
appCon fc defs (reflection "MN") [x', i']
|
||||
reflect fc defs env (NS ns n)
|
||||
reflect fc defs env (NS ns n)
|
||||
= do ns' <- reflect fc defs env ns
|
||||
n' <- reflect fc defs env n
|
||||
appCon fc defs (reflection "NS") [ns', n']
|
||||
@ -385,9 +402,9 @@ export
|
||||
Reflect (IsVar name idx vs) where
|
||||
reflect fc defs env First
|
||||
= appCon fc defs (reflection "First") [Erased fc, Erased fc]
|
||||
reflect fc defs env (Later p)
|
||||
reflect fc defs env (Later p)
|
||||
= do p' <- reflect fc defs env p
|
||||
appCon fc defs (reflection "Later")
|
||||
appCon fc defs (reflection "Later")
|
||||
[Erased fc, Erased fc, Erased fc, Erased fc, p']
|
||||
|
||||
-- Assume terms are normalised so there's not Let bindings in particular
|
||||
@ -443,11 +460,12 @@ Reflect (Term vs) where
|
||||
tm' <- reflect fc defs env tm
|
||||
appCon fc defs (reflection "TDelay")
|
||||
[Erased fc, dfc', r', ty', tm']
|
||||
reflect fc defs env (TForce dfc tm)
|
||||
reflect fc defs env (TForce dfc r tm)
|
||||
= do dfc' <- reflect fc defs env dfc
|
||||
r' <- reflect fc defs env r
|
||||
tm' <- reflect fc defs env tm
|
||||
appCon fc defs (reflection "TForce")
|
||||
[Erased fc, dfc', tm']
|
||||
[Erased fc, r', dfc', tm']
|
||||
reflect fc defs env (PrimVal pfc c)
|
||||
= do pfc' <- reflect fc defs env pfc
|
||||
c' <- reflect fc defs env c
|
||||
|
288
src/Core/TT.idr
288
src/Core/TT.idr
@ -19,7 +19,7 @@ data NameType : Type where
|
||||
TyCon : (tag : Int) -> (arity : Nat) -> NameType
|
||||
|
||||
public export
|
||||
data Constant
|
||||
data Constant
|
||||
= I Int
|
||||
| BI Integer
|
||||
| Str String
|
||||
@ -260,11 +260,11 @@ data Binder : Type -> Type where
|
||||
-- pattern bound variables. The PiInfo gives the implicitness at the
|
||||
-- point it was bound (Explicit if it was explicitly named in the
|
||||
-- program)
|
||||
PVar : RigCount -> PiInfo -> (ty : type) -> Binder type
|
||||
PVar : RigCount -> PiInfo -> (ty : type) -> Binder type
|
||||
-- variable bound for an as pattern (Like a let, but no computational
|
||||
-- force, and only used on the lhs. Converted to a let on the rhs because
|
||||
-- we want the computational behaviour.)
|
||||
PLet : RigCount -> (val : type) -> (ty : type) -> Binder type
|
||||
PLet : RigCount -> (val : type) -> (ty : type) -> Binder type
|
||||
-- the type of pattern bound variables
|
||||
PVTy : RigCount -> (ty : type) -> Binder type
|
||||
|
||||
@ -285,7 +285,7 @@ multiplicity (Pi c x ty) = c
|
||||
multiplicity (PVar c p ty) = c
|
||||
multiplicity (PLet c val ty) = c
|
||||
multiplicity (PVTy c ty) = c
|
||||
|
||||
|
||||
export
|
||||
setMultiplicity : Binder tm -> RigCount -> Binder tm
|
||||
setMultiplicity (Lam c x ty) c' = Lam c' x ty
|
||||
@ -299,7 +299,7 @@ showCount : RigCount -> String
|
||||
showCount Rig0 = "0 "
|
||||
showCount Rig1 = "1 "
|
||||
showCount RigW = ""
|
||||
|
||||
|
||||
Show ty => Show (Binder ty) where
|
||||
show (Lam c _ t) = "\\" ++ showCount c ++ show t
|
||||
show (Pi c _ t) = "Pi " ++ showCount c ++ show t
|
||||
@ -374,13 +374,13 @@ data LazyReason = LInf | LLazy | LUnknown
|
||||
public export
|
||||
data Term : List Name -> Type where
|
||||
Local : {name : _} ->
|
||||
FC -> Maybe Bool ->
|
||||
FC -> Maybe Bool ->
|
||||
(idx : Nat) -> .(IsVar name idx vars) -> Term vars
|
||||
Ref : FC -> NameType -> (name : Name) -> Term vars
|
||||
-- Metavariables and the scope they are applied to
|
||||
Meta : FC -> Name -> Int -> List (Term vars) -> Term vars
|
||||
Bind : FC -> (x : Name) ->
|
||||
(b : Binder (Term vars)) ->
|
||||
Bind : FC -> (x : Name) ->
|
||||
(b : Binder (Term vars)) ->
|
||||
(scope : Term (x :: vars)) -> Term vars
|
||||
App : FC -> (fn : Term vars) -> (arg : Term vars) -> Term vars
|
||||
-- as patterns; since we check LHS patterns as terms before turning
|
||||
@ -394,7 +394,7 @@ data Term : List Name -> Type where
|
||||
-- Typed laziness annotations
|
||||
TDelayed : FC -> LazyReason -> Term vars -> Term vars
|
||||
TDelay : FC -> LazyReason -> (ty : Term vars) -> (arg : Term vars) -> Term vars
|
||||
TForce : FC -> Term vars -> Term vars
|
||||
TForce : FC -> LazyReason -> Term vars -> Term vars
|
||||
PrimVal : FC -> (c : Constant) -> Term vars
|
||||
Erased : FC -> Term vars
|
||||
TType : FC -> Term vars
|
||||
@ -409,7 +409,7 @@ getLoc (App fc _ _) = fc
|
||||
getLoc (As fc _ _) = fc
|
||||
getLoc (TDelayed fc _ _) = fc
|
||||
getLoc (TDelay fc _ _ _) = fc
|
||||
getLoc (TForce fc _) = fc
|
||||
getLoc (TForce fc _ _) = fc
|
||||
getLoc (PrimVal fc _) = fc
|
||||
getLoc (Erased fc) = fc
|
||||
getLoc (TType fc) = fc
|
||||
@ -441,15 +441,15 @@ export
|
||||
Eq (Term vars) where
|
||||
(==) (Local _ _ idx _) (Local _ _ idx' _) = idx == idx'
|
||||
(==) (Ref _ _ n) (Ref _ _ n') = n == n'
|
||||
(==) (Meta _ _ i args) (Meta _ _ i' args')
|
||||
(==) (Meta _ _ i args) (Meta _ _ i' args')
|
||||
= assert_total (i == i' && args == args')
|
||||
(==) (Bind _ _ b sc) (Bind _ _ b' sc')
|
||||
(==) (Bind _ _ b sc) (Bind _ _ b' sc')
|
||||
= assert_total (b == b' && sc == believe_me sc')
|
||||
(==) (App _ f a) (App _ f' a') = f == f' && a == a'
|
||||
(==) (As _ a p) (As _ a' p') = a == a' && p == p'
|
||||
(==) (TDelayed _ _ t) (TDelayed _ _ t') = t == t'
|
||||
(==) (TDelay _ _ t x) (TDelay _ _ t' x') = t == t' && x == x'
|
||||
(==) (TForce _ t) (TForce _ t') = t == t'
|
||||
(==) (TForce _ _ t) (TForce _ _ t') = t == t'
|
||||
(==) (PrimVal _ c) (PrimVal _ c') = c == c'
|
||||
(==) (Erased _) (Erased _) = True
|
||||
(==) (TType _) (TType _) = True
|
||||
@ -496,20 +496,20 @@ Ord Visibility where
|
||||
compare Public Export = GT
|
||||
|
||||
public export
|
||||
data PartialReason
|
||||
= NotStrictlyPositive
|
||||
data PartialReason
|
||||
= NotStrictlyPositive
|
||||
| BadCall (List Name)
|
||||
| RecPath (List Name)
|
||||
|
||||
export
|
||||
Show PartialReason where
|
||||
show NotStrictlyPositive = "not strictly positive"
|
||||
show (BadCall [n])
|
||||
show (BadCall [n])
|
||||
= "not terminating due to call to " ++ show n
|
||||
show (BadCall ns)
|
||||
= "not terminating due to calls to " ++ showSep ", " (map show ns)
|
||||
show (RecPath ns)
|
||||
= "not terminating due to recursive path " ++ showSep " -> " (map show ns)
|
||||
show (BadCall ns)
|
||||
= "not terminating due to calls to " ++ showSep ", " (map show ns)
|
||||
show (RecPath ns)
|
||||
= "not terminating due to recursive path " ++ showSep " -> " (map show ns)
|
||||
|
||||
public export
|
||||
data Terminating
|
||||
@ -524,7 +524,7 @@ Show Terminating where
|
||||
show (NotTerminating p) = show p
|
||||
|
||||
public export
|
||||
data Covering
|
||||
data Covering
|
||||
= IsCovering
|
||||
| MissingCases (List (Term []))
|
||||
| NonCoveringCall (List Name)
|
||||
@ -533,9 +533,9 @@ export
|
||||
Show Covering where
|
||||
show IsCovering = "covering"
|
||||
show (MissingCases c) = "not covering all cases"
|
||||
show (NonCoveringCall [f])
|
||||
show (NonCoveringCall [f])
|
||||
= "not covering due to call to function " ++ show f
|
||||
show (NonCoveringCall cs)
|
||||
show (NonCoveringCall cs)
|
||||
= "not covering due to calls to functions " ++ showSep ", " (map show cs)
|
||||
|
||||
-- Totality status of a definition. We separate termination checking from
|
||||
@ -573,12 +573,12 @@ notCovering = MkTotality Unchecked (MissingCases [])
|
||||
|
||||
export
|
||||
insertVar : {outer : _} ->
|
||||
(idx : Nat) ->
|
||||
(idx : Nat) ->
|
||||
.(IsVar name idx (outer ++ inner)) ->
|
||||
Var (outer ++ n :: inner)
|
||||
insertVar {outer = []} idx x = MkVar (Later x)
|
||||
insertVar {outer = (name :: xs)} Z First = MkVar First
|
||||
insertVar {n} {outer = (x :: xs)} (S i) (Later y)
|
||||
insertVar {n} {outer = (x :: xs)} (S i) (Later y)
|
||||
= let MkVar prf = insertVar {n} i y in
|
||||
MkVar (Later prf)
|
||||
|
||||
@ -586,30 +586,30 @@ export
|
||||
weakenVar : (ns : List Name) -> {idx : Nat} -> .(IsVar name idx inner) ->
|
||||
Var (ns ++ inner)
|
||||
weakenVar [] x = MkVar x
|
||||
weakenVar (y :: xs) x
|
||||
weakenVar (y :: xs) x
|
||||
= let MkVar x' = weakenVar xs x in
|
||||
MkVar (Later x')
|
||||
|
||||
export
|
||||
insertVarNames : {outer, ns : _} ->
|
||||
(idx : Nat) ->
|
||||
(idx : Nat) ->
|
||||
.(IsVar name idx (outer ++ inner)) ->
|
||||
Var (outer ++ (ns ++ inner))
|
||||
insertVarNames {ns} {outer = []} idx prf = weakenVar ns prf
|
||||
insertVarNames {outer = (y :: xs)} Z First = MkVar First
|
||||
insertVarNames {ns} {outer = (y :: xs)} (S i) (Later x)
|
||||
insertVarNames {ns} {outer = (y :: xs)} (S i) (Later x)
|
||||
= let MkVar prf = insertVarNames {ns} i x in
|
||||
MkVar (Later prf)
|
||||
|
||||
export
|
||||
thin : {outer, inner : _} ->
|
||||
(n : Name) -> Term (outer ++ inner) -> Term (outer ++ n :: inner)
|
||||
thin n (Local fc r idx prf)
|
||||
thin n (Local fc r idx prf)
|
||||
= let MkVar var' = insertVar {n} idx prf in
|
||||
Local fc r _ var'
|
||||
thin n (Ref fc nt name) = Ref fc nt name
|
||||
thin n (Meta fc name idx args) = Meta fc name idx (map (thin n) args)
|
||||
thin {outer} {inner} n (Bind fc x b scope)
|
||||
thin {outer} {inner} n (Bind fc x b scope)
|
||||
= let sc' = thin {outer = x :: outer} {inner} n scope in
|
||||
Bind fc x (thinBinder n b) sc'
|
||||
where
|
||||
@ -625,7 +625,7 @@ thin n (App fc fn arg) = App fc (thin n fn) (thin n arg)
|
||||
thin n (As fc nm tm) = As fc (thin n nm) (thin n tm)
|
||||
thin n (TDelayed fc r ty) = TDelayed fc r (thin n ty)
|
||||
thin n (TDelay fc r ty tm) = TDelay fc r (thin n ty) (thin n tm)
|
||||
thin n (TForce fc tm) = TForce fc (thin n tm)
|
||||
thin n (TForce fc r tm) = TForce fc r (thin n tm)
|
||||
thin n (PrimVal fc c) = PrimVal fc c
|
||||
thin n (Erased fc) = Erased fc
|
||||
thin n (TType fc) = TType fc
|
||||
@ -634,28 +634,28 @@ export
|
||||
insertNames : {outer, inner : _} ->
|
||||
(ns : List Name) -> Term (outer ++ inner) ->
|
||||
Term (outer ++ (ns ++ inner))
|
||||
insertNames ns (Local fc r idx prf)
|
||||
insertNames ns (Local fc r idx prf)
|
||||
= let MkVar prf' = insertVarNames {ns} idx prf in
|
||||
Local fc r _ prf'
|
||||
insertNames ns (Ref fc nt name) = Ref fc nt name
|
||||
insertNames ns (Meta fc name idx args)
|
||||
= Meta fc name idx (map (insertNames ns) args)
|
||||
insertNames {outer} {inner} ns (Bind fc x b scope)
|
||||
= Bind fc x (assert_total (map (insertNames ns) b))
|
||||
insertNames {outer} {inner} ns (Bind fc x b scope)
|
||||
= Bind fc x (assert_total (map (insertNames ns) b))
|
||||
(insertNames {outer = x :: outer} {inner} ns scope)
|
||||
insertNames ns (App fc fn arg)
|
||||
insertNames ns (App fc fn arg)
|
||||
= App fc (insertNames ns fn) (insertNames ns arg)
|
||||
insertNames ns (As fc as tm)
|
||||
insertNames ns (As fc as tm)
|
||||
= As fc (insertNames ns as) (insertNames ns tm)
|
||||
insertNames ns (TDelayed fc r ty) = TDelayed fc r (insertNames ns ty)
|
||||
insertNames ns (TDelay fc r ty tm)
|
||||
insertNames ns (TDelay fc r ty tm)
|
||||
= TDelay fc r (insertNames ns ty) (insertNames ns tm)
|
||||
insertNames ns (TForce fc tm) = TForce fc (insertNames ns tm)
|
||||
insertNames ns (TForce fc r tm) = TForce fc r (insertNames ns tm)
|
||||
insertNames ns (PrimVal fc c) = PrimVal fc c
|
||||
insertNames ns (Erased fc) = Erased fc
|
||||
insertNames ns (TType fc) = TType fc
|
||||
|
||||
export
|
||||
export
|
||||
Weaken Term where
|
||||
weaken tm = thin {outer = []} _ tm
|
||||
weakenNs ns tm = insertNames {outer = []} ns tm
|
||||
@ -698,7 +698,7 @@ export
|
||||
getFnArgs : Term vars -> (Term vars, List (Term vars))
|
||||
getFnArgs tm = getFA [] tm
|
||||
where
|
||||
getFA : List (Term vars) -> Term vars ->
|
||||
getFA : List (Term vars) -> Term vars ->
|
||||
(Term vars, List (Term vars))
|
||||
getFA args (App _ f a) = getFA (a :: args) f
|
||||
getFA args tm = (tm, args)
|
||||
@ -718,7 +718,7 @@ data CompatibleVars : List Name -> List Name -> Type where
|
||||
CompatExt : CompatibleVars xs ys -> CompatibleVars (n :: xs) (m :: ys)
|
||||
|
||||
export
|
||||
areVarsCompatible : (xs : List Name) -> (ys : List Name) ->
|
||||
areVarsCompatible : (xs : List Name) -> (ys : List Name) ->
|
||||
Maybe (CompatibleVars xs ys)
|
||||
areVarsCompatible [] [] = pure CompatPre
|
||||
areVarsCompatible (x :: xs) (y :: ys)
|
||||
@ -732,15 +732,15 @@ extendCompats : (args : List Name) ->
|
||||
extendCompats [] prf = prf
|
||||
extendCompats (x :: xs) prf = CompatExt (extendCompats xs prf)
|
||||
|
||||
renameLocalRef : CompatibleVars xs ys ->
|
||||
{idx : Nat} ->
|
||||
.(IsVar name idx xs) ->
|
||||
renameLocalRef : CompatibleVars xs ys ->
|
||||
{idx : Nat} ->
|
||||
.(IsVar name idx xs) ->
|
||||
Var ys
|
||||
renameLocalRef prf p = believe_me (MkVar p)
|
||||
-- renameLocalRef CompatPre First = (MkVar First)
|
||||
-- renameLocalRef (CompatExt x) First = (MkVar First)
|
||||
-- renameLocalRef CompatPre (Later p) = (MkVar (Later p))
|
||||
-- renameLocalRef (CompatExt y) (Later p)
|
||||
-- renameLocalRef (CompatExt y) (Later p)
|
||||
-- = let (MkVar p') = renameLocalRef y p in MkVar (Later p')
|
||||
|
||||
renameVarList : CompatibleVars xs ys -> Var xs -> Var ys
|
||||
@ -748,24 +748,24 @@ renameVarList prf (MkVar p) = renameLocalRef prf p
|
||||
|
||||
-- TODO: Surely identity at run time, can we replace with 'believe_me'?
|
||||
export
|
||||
renameVars : CompatibleVars xs ys -> Term xs -> Term ys
|
||||
renameVars : CompatibleVars xs ys -> Term xs -> Term ys
|
||||
renameVars CompatPre tm = tm
|
||||
renameVars prf (Local fc r idx vprf)
|
||||
renameVars prf (Local fc r idx vprf)
|
||||
= let MkVar vprf' = renameLocalRef prf vprf in
|
||||
Local fc r _ vprf'
|
||||
renameVars prf (Ref fc x name) = Ref fc x name
|
||||
renameVars prf (Meta fc n i args)
|
||||
renameVars prf (Meta fc n i args)
|
||||
= Meta fc n i (map (renameVars prf) args)
|
||||
renameVars prf (Bind fc x b scope)
|
||||
renameVars prf (Bind fc x b scope)
|
||||
= Bind fc x (map (renameVars prf) b) (renameVars (CompatExt prf) scope)
|
||||
renameVars prf (App fc fn arg)
|
||||
renameVars prf (App fc fn arg)
|
||||
= App fc (renameVars prf fn) (renameVars prf arg)
|
||||
renameVars prf (As fc as tm)
|
||||
= As fc (renameVars prf as) (renameVars prf tm)
|
||||
renameVars prf (TDelayed fc r ty) = TDelayed fc r (renameVars prf ty)
|
||||
renameVars prf (TDelay fc r ty tm)
|
||||
renameVars prf (TDelay fc r ty tm)
|
||||
= TDelay fc r (renameVars prf ty) (renameVars prf tm)
|
||||
renameVars prf (TForce fc x) = TForce fc (renameVars prf x)
|
||||
renameVars prf (TForce fc r x) = TForce fc r (renameVars prf x)
|
||||
renameVars prf (PrimVal fc c) = PrimVal fc c
|
||||
renameVars prf (Erased fc) = Erased fc
|
||||
renameVars prf (TType fc) = TType fc
|
||||
@ -781,15 +781,15 @@ data SubVars : List Name -> List Name -> Type where
|
||||
KeepCons : SubVars xs ys -> SubVars (x :: xs) (x :: ys)
|
||||
|
||||
export
|
||||
subElem : {idx : Nat} -> .(IsVar name idx xs) ->
|
||||
subElem : {idx : Nat} -> .(IsVar name idx xs) ->
|
||||
SubVars ys xs -> Maybe (Var ys)
|
||||
subElem prf SubRefl = Just (MkVar prf)
|
||||
subElem First (DropCons p) = Nothing
|
||||
subElem (Later x) (DropCons p)
|
||||
subElem (Later x) (DropCons p)
|
||||
= do MkVar prf' <- subElem x p
|
||||
Just (MkVar prf')
|
||||
subElem First (KeepCons p) = Just (MkVar First)
|
||||
subElem (Later x) (KeepCons p)
|
||||
subElem (Later x) (KeepCons p)
|
||||
= do MkVar prf' <- subElem x p
|
||||
Just (MkVar (Later prf'))
|
||||
|
||||
@ -806,9 +806,9 @@ subInclude ns (KeepCons p) = KeepCons (subInclude ns p)
|
||||
|
||||
mutual
|
||||
export
|
||||
shrinkBinder : Binder (Term vars) -> SubVars newvars vars ->
|
||||
shrinkBinder : Binder (Term vars) -> SubVars newvars vars ->
|
||||
Maybe (Binder (Term newvars))
|
||||
shrinkBinder (Lam c p ty) prf
|
||||
shrinkBinder (Lam c p ty) prf
|
||||
= Just (Lam c p !(shrinkTerm ty prf))
|
||||
shrinkBinder (Let c val ty) prf
|
||||
= Just (Let c !(shrinkTerm val prf) !(shrinkTerm ty prf))
|
||||
@ -827,26 +827,26 @@ mutual
|
||||
|
||||
export
|
||||
shrinkTerm : Term vars -> SubVars newvars vars -> Maybe (Term newvars)
|
||||
shrinkTerm (Local fc r idx loc) prf
|
||||
shrinkTerm (Local fc r idx loc) prf
|
||||
= case subElem loc prf of
|
||||
Nothing => Nothing
|
||||
Just (MkVar loc') => Just (Local fc r _ loc')
|
||||
shrinkTerm (Ref fc x name) prf = Just (Ref fc x name)
|
||||
shrinkTerm (Meta fc x y xs) prf
|
||||
shrinkTerm (Meta fc x y xs) prf
|
||||
= do xs' <- traverse (\x => shrinkTerm x prf) xs
|
||||
Just (Meta fc x y xs')
|
||||
shrinkTerm (Bind fc x b scope) prf
|
||||
shrinkTerm (Bind fc x b scope) prf
|
||||
= Just (Bind fc x !(shrinkBinder b prf) !(shrinkTerm scope (KeepCons prf)))
|
||||
shrinkTerm (App fc fn arg) prf
|
||||
shrinkTerm (App fc fn arg) prf
|
||||
= Just (App fc !(shrinkTerm fn prf) !(shrinkTerm arg prf))
|
||||
shrinkTerm (As fc as tm) prf
|
||||
shrinkTerm (As fc as tm) prf
|
||||
= Just (As fc !(shrinkTerm as prf) !(shrinkTerm tm prf))
|
||||
shrinkTerm (TDelayed fc x y) prf
|
||||
shrinkTerm (TDelayed fc x y) prf
|
||||
= Just (TDelayed fc x !(shrinkTerm y prf))
|
||||
shrinkTerm (TDelay fc x t y) prf
|
||||
= Just (TDelay fc x !(shrinkTerm t prf) !(shrinkTerm y prf))
|
||||
shrinkTerm (TForce fc x) prf
|
||||
= Just (TForce fc !(shrinkTerm x prf))
|
||||
shrinkTerm (TForce fc r x) prf
|
||||
= Just (TForce fc r !(shrinkTerm x prf))
|
||||
shrinkTerm (PrimVal fc c) prf = Just (PrimVal fc c)
|
||||
shrinkTerm (Erased fc) prf = Just (Erased fc)
|
||||
shrinkTerm (TType fc) prf = Just (TType fc)
|
||||
@ -863,14 +863,14 @@ addVars : {later, bound : _} ->
|
||||
Var (later ++ (bound ++ vars))
|
||||
addVars {later = []} {bound} bs p = weakenVar bound p
|
||||
addVars {later = (x :: xs)} bs First = MkVar First
|
||||
addVars {later = (x :: xs)} bs (Later p)
|
||||
addVars {later = (x :: xs)} bs (Later p)
|
||||
= let MkVar p' = addVars {later = xs} bs p in
|
||||
MkVar (Later p')
|
||||
|
||||
resolveRef : (done : List Name) -> Bounds bound -> FC -> Name ->
|
||||
resolveRef : (done : List Name) -> Bounds bound -> FC -> Name ->
|
||||
Maybe (Term (later ++ (done ++ bound ++ vars)))
|
||||
resolveRef done None fc n = Nothing
|
||||
resolveRef {later} {vars} done (Add {xs} new old bs) fc n
|
||||
resolveRef {later} {vars} done (Add {xs} new old bs) fc n
|
||||
= if n == old
|
||||
then rewrite appendAssociative later done (new :: xs ++ vars) in
|
||||
let MkVar p = weakenVar {inner = new :: xs ++ vars}
|
||||
@ -880,30 +880,30 @@ resolveRef {later} {vars} done (Add {xs} new old bs) fc n
|
||||
in resolveRef (done ++ [new]) bs fc n
|
||||
|
||||
mkLocals : {later, bound : _} ->
|
||||
Bounds bound ->
|
||||
Bounds bound ->
|
||||
Term (later ++ vars) -> Term (later ++ (bound ++ vars))
|
||||
mkLocals bs (Local fc r idx p)
|
||||
mkLocals bs (Local fc r idx p)
|
||||
= let MkVar p' = addVars bs p in Local fc r _ p'
|
||||
mkLocals bs (Ref fc Bound name)
|
||||
mkLocals bs (Ref fc Bound name)
|
||||
= maybe (Ref fc Bound name) id (resolveRef [] bs fc name)
|
||||
mkLocals bs (Ref fc nt name)
|
||||
mkLocals bs (Ref fc nt name)
|
||||
= Ref fc nt name
|
||||
mkLocals bs (Meta fc name y xs)
|
||||
mkLocals bs (Meta fc name y xs)
|
||||
= maybe (Meta fc name y (map (mkLocals bs) xs))
|
||||
id (resolveRef [] bs fc name)
|
||||
mkLocals {later} bs (Bind fc x b scope)
|
||||
= Bind fc x (map (mkLocals bs) b)
|
||||
mkLocals {later} bs (Bind fc x b scope)
|
||||
= Bind fc x (map (mkLocals bs) b)
|
||||
(mkLocals {later = x :: later} bs scope)
|
||||
mkLocals bs (App fc fn arg)
|
||||
mkLocals bs (App fc fn arg)
|
||||
= App fc (mkLocals bs fn) (mkLocals bs arg)
|
||||
mkLocals bs (As fc as tm)
|
||||
mkLocals bs (As fc as tm)
|
||||
= As fc (mkLocals bs as) (mkLocals bs tm)
|
||||
mkLocals bs (TDelayed fc x y)
|
||||
mkLocals bs (TDelayed fc x y)
|
||||
= TDelayed fc x (mkLocals bs y)
|
||||
mkLocals bs (TDelay fc x t y)
|
||||
= TDelay fc x (mkLocals bs t) (mkLocals bs y)
|
||||
mkLocals bs (TForce fc x)
|
||||
= TForce fc (mkLocals bs x)
|
||||
mkLocals bs (TForce fc r x)
|
||||
= TForce fc r (mkLocals bs x)
|
||||
mkLocals bs (PrimVal fc c) = PrimVal fc c
|
||||
mkLocals bs (Erased fc) = Erased fc
|
||||
mkLocals bs (TType fc) = TType fc
|
||||
@ -921,7 +921,7 @@ refToLocal x new tm = refsToLocals (Add new x None) tm
|
||||
export
|
||||
isVar : (n : Name) -> (ns : List Name) -> Maybe (Var ns)
|
||||
isVar n [] = Nothing
|
||||
isVar n (m :: ms)
|
||||
isVar n (m :: ms)
|
||||
= case nameEq n m of
|
||||
Nothing => do MkVar p <- isVar n ms
|
||||
pure (MkVar (Later p))
|
||||
@ -934,20 +934,20 @@ resolveNames vars (Ref fc Bound name)
|
||||
= case isVar name vars of
|
||||
Just (MkVar prf) => Local fc (Just False) _ prf
|
||||
_ => Ref fc Bound name
|
||||
resolveNames vars (Meta fc n i xs)
|
||||
resolveNames vars (Meta fc n i xs)
|
||||
= Meta fc n i (map (resolveNames vars) xs)
|
||||
resolveNames vars (Bind fc x b scope)
|
||||
resolveNames vars (Bind fc x b scope)
|
||||
= Bind fc x (map (resolveNames vars) b) (resolveNames (x :: vars) scope)
|
||||
resolveNames vars (App fc fn arg)
|
||||
resolveNames vars (App fc fn arg)
|
||||
= App fc (resolveNames vars fn) (resolveNames vars arg)
|
||||
resolveNames vars (As fc as pat)
|
||||
resolveNames vars (As fc as pat)
|
||||
= As fc (resolveNames vars as) (resolveNames vars pat)
|
||||
resolveNames vars (TDelayed fc x y)
|
||||
resolveNames vars (TDelayed fc x y)
|
||||
= TDelayed fc x (resolveNames vars y)
|
||||
resolveNames vars (TDelay fc x t y)
|
||||
= TDelay fc x (resolveNames vars t) (resolveNames vars y)
|
||||
resolveNames vars (TForce fc x)
|
||||
= TForce fc (resolveNames vars x)
|
||||
resolveNames vars (TForce fc r x)
|
||||
= TForce fc r (resolveNames vars x)
|
||||
resolveNames vars tm = tm
|
||||
|
||||
|
||||
@ -959,15 +959,15 @@ namespace SubstEnv
|
||||
public export
|
||||
data SubstEnv : List Name -> List Name -> Type where
|
||||
Nil : SubstEnv [] vars
|
||||
(::) : Term vars ->
|
||||
(::) : Term vars ->
|
||||
SubstEnv ds vars -> SubstEnv (d :: ds) vars
|
||||
|
||||
findDrop : {drop : _} -> {idx : Nat} ->
|
||||
FC -> Maybe Bool -> .(IsVar name idx (drop ++ vars)) ->
|
||||
FC -> Maybe Bool -> .(IsVar name idx (drop ++ vars)) ->
|
||||
SubstEnv drop vars -> Term vars
|
||||
findDrop {drop = []} fc r var env = Local fc r _ var
|
||||
findDrop {drop = x :: xs} fc r First (tm :: env) = tm
|
||||
findDrop {drop = x :: xs} fc r (Later p) (tm :: env)
|
||||
findDrop {drop = x :: xs} fc r (Later p) (tm :: env)
|
||||
= findDrop fc r p env
|
||||
|
||||
find : {outer : _} -> {idx : Nat} ->
|
||||
@ -979,24 +979,24 @@ namespace SubstEnv
|
||||
find {outer = x :: xs} fc r (Later p) env = weaken (find fc r p env)
|
||||
|
||||
substEnv : {outer : _} ->
|
||||
SubstEnv drop vars -> Term (outer ++ (drop ++ vars)) ->
|
||||
SubstEnv drop vars -> Term (outer ++ (drop ++ vars)) ->
|
||||
Term (outer ++ vars)
|
||||
substEnv env (Local fc r _ prf)
|
||||
substEnv env (Local fc r _ prf)
|
||||
= find fc r prf env
|
||||
substEnv env (Ref fc x name) = Ref fc x name
|
||||
substEnv env (Meta fc n i xs)
|
||||
substEnv env (Meta fc n i xs)
|
||||
= Meta fc n i (map (substEnv env) xs)
|
||||
substEnv {outer} env (Bind fc x b scope)
|
||||
= Bind fc x (map (substEnv env) b)
|
||||
substEnv {outer} env (Bind fc x b scope)
|
||||
= Bind fc x (map (substEnv env) b)
|
||||
(substEnv {outer = x :: outer} env scope)
|
||||
substEnv env (App fc fn arg)
|
||||
substEnv env (App fc fn arg)
|
||||
= App fc (substEnv env fn) (substEnv env arg)
|
||||
substEnv env (As fc as pat)
|
||||
substEnv env (As fc as pat)
|
||||
= As fc (substEnv env as) (substEnv env pat)
|
||||
substEnv env (TDelayed fc x y) = TDelayed fc x (substEnv env y)
|
||||
substEnv env (TDelay fc x t y)
|
||||
substEnv env (TDelay fc x t y)
|
||||
= TDelay fc x (substEnv env t) (substEnv env y)
|
||||
substEnv env (TForce fc x) = TForce fc (substEnv env x)
|
||||
substEnv env (TForce fc r x) = TForce fc r (substEnv env x)
|
||||
substEnv env (PrimVal fc c) = PrimVal fc c
|
||||
substEnv env (Erased fc) = Erased fc
|
||||
substEnv env (TType fc) = TType fc
|
||||
@ -1016,24 +1016,24 @@ substName x new (Ref fc nt name)
|
||||
= case nameEq x name of
|
||||
Nothing => Ref fc nt name
|
||||
Just Refl => new
|
||||
substName x new (Meta fc n i xs)
|
||||
substName x new (Meta fc n i xs)
|
||||
= Meta fc n i (map (substName x new) xs)
|
||||
-- ASSUMPTION: When we substitute under binders, the name has always been
|
||||
-- resolved to a Local, so no need to check that x isn't shadowing
|
||||
substName x new (Bind fc y b scope)
|
||||
substName x new (Bind fc y b scope)
|
||||
= Bind fc y (map (substName x new) b) (substName x (weaken new) scope)
|
||||
substName x new (App fc fn arg)
|
||||
substName x new (App fc fn arg)
|
||||
= App fc (substName x new fn) (substName x new arg)
|
||||
substName x new (As fc as pat)
|
||||
substName x new (As fc as pat)
|
||||
= As fc as (substName x new pat)
|
||||
substName x new (TDelayed fc y z)
|
||||
substName x new (TDelayed fc y z)
|
||||
= TDelayed fc y (substName x new z)
|
||||
substName x new (TDelay fc y t z)
|
||||
= TDelay fc y (substName x new t) (substName x new z)
|
||||
substName x new (TForce fc y)
|
||||
= TForce fc (substName x new y)
|
||||
substName x new (TForce fc r y)
|
||||
= TForce fc r (substName x new y)
|
||||
substName x new tm = tm
|
||||
|
||||
|
||||
export
|
||||
addMetas : NameMap Bool -> Term vars -> NameMap Bool
|
||||
addMetas ns (Local fc x idx y) = ns
|
||||
@ -1043,17 +1043,17 @@ addMetas ns (Meta fc n i xs) = addMetaArgs (insert n False ns) xs
|
||||
addMetaArgs : NameMap Bool -> List (Term vars) -> NameMap Bool
|
||||
addMetaArgs ns [] = ns
|
||||
addMetaArgs ns (t :: ts) = addMetaArgs (addMetas ns t) ts
|
||||
addMetas ns (Bind fc x (Let c val ty) scope)
|
||||
addMetas ns (Bind fc x (Let c val ty) scope)
|
||||
= addMetas (addMetas (addMetas ns val) ty) scope
|
||||
addMetas ns (Bind fc x b scope)
|
||||
addMetas ns (Bind fc x b scope)
|
||||
= addMetas (addMetas ns (binderType b)) scope
|
||||
addMetas ns (App fc fn arg)
|
||||
addMetas ns (App fc fn arg)
|
||||
= addMetas (addMetas ns fn) arg
|
||||
addMetas ns (As fc as tm) = addMetas ns tm
|
||||
addMetas ns (TDelayed fc x y) = addMetas ns y
|
||||
addMetas ns (TDelay fc x t y)
|
||||
addMetas ns (TDelay fc x t y)
|
||||
= addMetas (addMetas ns t) y
|
||||
addMetas ns (TForce fc x) = addMetas ns x
|
||||
addMetas ns (TForce fc r x) = addMetas ns x
|
||||
addMetas ns (PrimVal fc c) = ns
|
||||
addMetas ns (Erased fc) = ns
|
||||
addMetas ns (TType fc) = ns
|
||||
@ -1062,33 +1062,33 @@ addMetas ns (TType fc) = ns
|
||||
export
|
||||
getMetas : Term vars -> NameMap Bool
|
||||
getMetas tm = addMetas empty tm
|
||||
|
||||
|
||||
export
|
||||
addRefs : (underAssert : Bool) -> (aTotal : Name) ->
|
||||
addRefs : (underAssert : Bool) -> (aTotal : Name) ->
|
||||
NameMap Bool -> Term vars -> NameMap Bool
|
||||
addRefs ua at ns (Local fc x idx y) = ns
|
||||
addRefs ua at ns (Ref fc x name) = insert name ua ns
|
||||
addRefs ua at ns (Meta fc n i xs)
|
||||
addRefs ua at ns (Meta fc n i xs)
|
||||
= addRefsArgs ns xs
|
||||
where
|
||||
addRefsArgs : NameMap Bool -> List (Term vars) -> NameMap Bool
|
||||
addRefsArgs ns [] = ns
|
||||
addRefsArgs ns (t :: ts) = addRefsArgs (addRefs ua at ns t) ts
|
||||
addRefs ua at ns (Bind fc x (Let c val ty) scope)
|
||||
addRefs ua at ns (Bind fc x (Let c val ty) scope)
|
||||
= addRefs ua at (addRefs ua at (addRefs ua at ns val) ty) scope
|
||||
addRefs ua at ns (Bind fc x b scope)
|
||||
addRefs ua at ns (Bind fc x b scope)
|
||||
= addRefs ua at (addRefs ua at ns (binderType b)) scope
|
||||
addRefs ua at ns (App _ (App _ (Ref fc _ name) x) y)
|
||||
= if name == at
|
||||
then addRefs True at (insert name True ns) y
|
||||
else addRefs ua at (addRefs ua at (insert name ua ns) x) y
|
||||
addRefs ua at ns (App fc fn arg)
|
||||
addRefs ua at ns (App fc fn arg)
|
||||
= addRefs ua at (addRefs ua at ns fn) arg
|
||||
addRefs ua at ns (As fc as tm) = addRefs ua at ns tm
|
||||
addRefs ua at ns (TDelayed fc x y) = addRefs ua at ns y
|
||||
addRefs ua at ns (TDelay fc x t y)
|
||||
addRefs ua at ns (TDelay fc x t y)
|
||||
= addRefs ua at (addRefs ua at ns t) y
|
||||
addRefs ua at ns (TForce fc x) = addRefs ua at ns x
|
||||
addRefs ua at ns (TForce fc r x) = addRefs ua at ns x
|
||||
addRefs ua at ns (PrimVal fc c) = ns
|
||||
addRefs ua at ns (Erased fc) = ns
|
||||
addRefs ua at ns (TType fc) = ns
|
||||
@ -1104,46 +1104,46 @@ export Show (Term vars) where
|
||||
show tm = let (fn, args) = getFnArgs tm in showApp fn args
|
||||
where
|
||||
showApp : Term vars -> List (Term vars) -> String
|
||||
showApp (Local {name} _ c idx _) []
|
||||
showApp (Local {name} _ c idx _) []
|
||||
= show name ++ "[" ++ show idx ++ "]"
|
||||
showApp (Ref _ _ n) [] = show n
|
||||
showApp (Meta _ n i args) []
|
||||
showApp (Meta _ n i args) []
|
||||
= "?" ++ show n ++ "_" ++ show args
|
||||
showApp (Bind _ x (Lam c p ty) sc) []
|
||||
= "\\" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (Lam c p ty) sc) []
|
||||
= "\\" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
" => " ++ show sc
|
||||
showApp (Bind _ x (Let c val ty) sc) []
|
||||
= "let " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (Let c val ty) sc) []
|
||||
= "let " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
" = " ++ show val ++ " in " ++ show sc
|
||||
showApp (Bind _ x (Pi c Explicit ty) sc) []
|
||||
= "((" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (Pi c Explicit ty) sc) []
|
||||
= "((" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
") -> " ++ show sc ++ ")"
|
||||
showApp (Bind _ x (Pi c Implicit ty) sc) []
|
||||
= "{" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (Pi c Implicit ty) sc) []
|
||||
= "{" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
"} -> " ++ show sc
|
||||
showApp (Bind _ x (Pi c AutoImplicit ty) sc) []
|
||||
= "{auto" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (Pi c AutoImplicit ty) sc) []
|
||||
= "{auto" ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
"} -> " ++ show sc
|
||||
showApp (Bind _ x (PVar c Explicit ty) sc) []
|
||||
= "pat " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (PVar c Explicit ty) sc) []
|
||||
= "pat " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
" => " ++ show sc
|
||||
showApp (Bind _ x (PVar c Implicit ty) sc) []
|
||||
= "{pat " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (PVar c Implicit ty) sc) []
|
||||
= "{pat " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
"} => " ++ show sc
|
||||
showApp (Bind _ x (PVar c AutoImplicit ty) sc) []
|
||||
= "{auto pat " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (PVar c AutoImplicit ty) sc) []
|
||||
= "{auto pat " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
"} => " ++ show sc
|
||||
showApp (Bind _ x (PLet c val ty) sc) []
|
||||
= "plet " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (PLet c val ty) sc) []
|
||||
= "plet " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
" = " ++ show val ++ " in " ++ show sc
|
||||
showApp (Bind _ x (PVTy c ty) sc) []
|
||||
= "pty " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
showApp (Bind _ x (PVTy c ty) sc) []
|
||||
= "pty " ++ showCount c ++ show x ++ " : " ++ show ty ++
|
||||
" => " ++ show sc
|
||||
showApp (App _ _ _) [] = "[can't happen]"
|
||||
showApp (As _ n tm) [] = show n ++ "@" ++ show tm
|
||||
showApp (TDelayed _ _ tm) [] = "%Delayed " ++ show tm
|
||||
showApp (TDelay _ _ _ tm) [] = "%Delay " ++ show tm
|
||||
showApp (TForce _ tm) [] = "%Force " ++ show tm
|
||||
showApp (TForce _ _ tm) [] = "%Force " ++ show tm
|
||||
showApp (PrimVal _ c) [] = show c
|
||||
showApp (Erased _) [] = "[__]"
|
||||
showApp (TType _) [] = "Type"
|
||||
|
136
src/Core/TTC.idr
136
src/Core/TTC.idr
@ -19,13 +19,13 @@ import Utils.Binary
|
||||
|
||||
export
|
||||
TTC FC where
|
||||
toBuf b (MkFC file startPos endPos)
|
||||
toBuf b (MkFC file startPos endPos)
|
||||
= do tag 0; toBuf b file; toBuf b startPos; toBuf b endPos
|
||||
toBuf b EmptyFC = tag 1
|
||||
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => do f <- fromBuf b;
|
||||
0 => do f <- fromBuf b;
|
||||
s <- fromBuf b; e <- fromBuf b
|
||||
pure (MkFC f s e)
|
||||
1 => pure EmptyFC
|
||||
@ -41,7 +41,7 @@ TTC Name where
|
||||
toBuf b (Nested x y) = do tag 5; toBuf b x; toBuf b y
|
||||
toBuf b (CaseBlock x y) = do tag 6; toBuf b x; toBuf b y
|
||||
toBuf b (WithBlock x y) = do tag 7; toBuf b x; toBuf b y
|
||||
toBuf b (Resolved x)
|
||||
toBuf b (Resolved x)
|
||||
= throw (InternalError ("Can't write resolved name " ++ show x))
|
||||
|
||||
fromBuf b
|
||||
@ -70,8 +70,8 @@ TTC Name where
|
||||
y <- fromBuf b
|
||||
pure (WithBlock x y)
|
||||
_ => corrupt "Name"
|
||||
|
||||
export
|
||||
|
||||
export
|
||||
TTC RigCount where
|
||||
toBuf b Rig0 = tag 0
|
||||
toBuf b Rig1 = tag 1
|
||||
@ -104,10 +104,10 @@ TTC Constant where
|
||||
toBuf b (Str x) = do tag 2; toBuf b x
|
||||
toBuf b (Ch x) = do tag 3; toBuf b x
|
||||
toBuf b (Db x) = do tag 4; toBuf b x
|
||||
|
||||
|
||||
toBuf b WorldVal = tag 5
|
||||
toBuf b IntType = tag 6
|
||||
toBuf b IntegerType = tag 7
|
||||
toBuf b IntegerType = tag 7
|
||||
toBuf b StringType = tag 8
|
||||
toBuf b CharType = tag 9
|
||||
toBuf b DoubleType = tag 10
|
||||
@ -198,7 +198,7 @@ mutual
|
||||
|
||||
export
|
||||
TTC (Term vars) where
|
||||
toBuf b (Local {name} fc c idx y)
|
||||
toBuf b (Local {name} fc c idx y)
|
||||
= if idx < 244
|
||||
then do toBuf b (prim__truncBigInt_B8 (12 + cast idx))
|
||||
toBuf b c
|
||||
@ -207,43 +207,43 @@ mutual
|
||||
toBuf b c
|
||||
toBuf b name
|
||||
toBuf b idx
|
||||
toBuf b (Ref fc nt name)
|
||||
toBuf b (Ref fc nt name)
|
||||
= do tag 1;
|
||||
toBuf b nt; toBuf b name
|
||||
toBuf b (Meta fc n i xs)
|
||||
toBuf b (Meta fc n i xs)
|
||||
= do tag 2;
|
||||
toBuf b n; toBuf b xs
|
||||
toBuf b (Bind fc x bnd scope)
|
||||
toBuf b (Bind fc x bnd scope)
|
||||
= do tag 3;
|
||||
toBuf b x;
|
||||
toBuf b x;
|
||||
toBuf b bnd; toBuf b scope
|
||||
toBuf b (App fc fn arg)
|
||||
toBuf b (App fc fn arg)
|
||||
= do tag 4;
|
||||
toBuf b fn; toBuf b arg
|
||||
-- let (fn, args) = getFnArgs (App fc fn arg)
|
||||
-- toBuf b fn; -- toBuf b p;
|
||||
-- toBuf b fn; -- toBuf b p;
|
||||
-- toBuf b args
|
||||
toBuf b (As fc as tm)
|
||||
= do tag 5;
|
||||
toBuf b as; toBuf b tm
|
||||
toBuf b (TDelayed fc r tm)
|
||||
toBuf b (TDelayed fc r tm)
|
||||
= do tag 6;
|
||||
toBuf b r; toBuf b tm
|
||||
toBuf b (TDelay fc r ty tm)
|
||||
= do tag 7;
|
||||
toBuf b r; toBuf b ty; toBuf b tm
|
||||
toBuf b (TForce fc tm)
|
||||
toBuf b (TForce fc r tm)
|
||||
= do tag 8;
|
||||
toBuf b tm
|
||||
toBuf b (PrimVal fc c)
|
||||
toBuf b r; toBuf b tm
|
||||
toBuf b (PrimVal fc c)
|
||||
= do tag 9;
|
||||
toBuf b c
|
||||
toBuf b (Erased fc)
|
||||
toBuf b (Erased fc)
|
||||
= tag 10
|
||||
toBuf b (TType fc)
|
||||
= tag 11
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => do c <- fromBuf b
|
||||
name <- fromBuf b
|
||||
@ -264,11 +264,11 @@ mutual
|
||||
pure (As emptyFC as tm)
|
||||
6 => do lr <- fromBuf b; tm <- fromBuf b
|
||||
pure (TDelayed emptyFC lr tm)
|
||||
7 => do lr <- fromBuf b;
|
||||
7 => do lr <- fromBuf b;
|
||||
ty <- fromBuf b; tm <- fromBuf b
|
||||
pure (TDelay emptyFC lr ty tm)
|
||||
8 => do tm <- fromBuf b
|
||||
pure (TForce emptyFC tm)
|
||||
8 => do lr <- fromBuf b; tm <- fromBuf b
|
||||
pure (TForce emptyFC lr tm)
|
||||
9 => do c <- fromBuf b
|
||||
pure (PrimVal emptyFC c)
|
||||
10 => pure (Erased emptyFC)
|
||||
@ -280,24 +280,24 @@ mutual
|
||||
|
||||
export
|
||||
TTC Pat where
|
||||
toBuf b (PAs fc x y)
|
||||
toBuf b (PAs fc x y)
|
||||
= do tag 0; toBuf b fc; toBuf b x; toBuf b y
|
||||
toBuf b (PCon fc x t arity xs)
|
||||
toBuf b (PCon fc x t arity xs)
|
||||
= do tag 1; toBuf b fc; toBuf b x; toBuf b t; toBuf b arity; toBuf b xs
|
||||
toBuf b (PTyCon fc x arity xs)
|
||||
toBuf b (PTyCon fc x arity xs)
|
||||
= do tag 2; toBuf b fc; toBuf b x; toBuf b arity; toBuf b xs
|
||||
toBuf b (PConst fc c)
|
||||
= do tag 3; toBuf b fc; toBuf b c
|
||||
toBuf b (PArrow fc x s t)
|
||||
= do tag 4; toBuf b fc; toBuf b x; toBuf b s; toBuf b t
|
||||
toBuf b (PDelay fc x t y)
|
||||
toBuf b (PDelay fc x t y)
|
||||
= do tag 5; toBuf b fc; toBuf b x; toBuf b t; toBuf b y
|
||||
toBuf b (PLoc fc x)
|
||||
toBuf b (PLoc fc x)
|
||||
= do tag 6; toBuf b fc; toBuf b x
|
||||
toBuf b (PUnmatchable fc x)
|
||||
toBuf b (PUnmatchable fc x)
|
||||
= do tag 7; toBuf b fc; toBuf b x
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => do fc <- fromBuf b; x <- fromBuf b;
|
||||
y <- fromBuf b
|
||||
@ -327,15 +327,15 @@ TTC Pat where
|
||||
mutual
|
||||
export
|
||||
TTC (CaseTree vars) where
|
||||
toBuf b (Case {name} idx x scTy xs)
|
||||
toBuf b (Case {name} idx x scTy xs)
|
||||
= do tag 0; toBuf b name; toBuf b idx; toBuf b xs
|
||||
toBuf b (STerm x)
|
||||
toBuf b (STerm x)
|
||||
= do tag 1; toBuf b x
|
||||
toBuf b (Unmatched msg)
|
||||
toBuf b (Unmatched msg)
|
||||
= do tag 2; toBuf b msg
|
||||
toBuf b Impossible = tag 3
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => do name <- fromBuf b; idx <- fromBuf b
|
||||
xs <- fromBuf b
|
||||
@ -349,16 +349,16 @@ mutual
|
||||
|
||||
export
|
||||
TTC (CaseAlt vars) where
|
||||
toBuf b (ConCase x t args y)
|
||||
toBuf b (ConCase x t args y)
|
||||
= do tag 0; toBuf b x; toBuf b t; toBuf b args; toBuf b y
|
||||
toBuf b (DelayCase ty arg y)
|
||||
toBuf b (DelayCase ty arg y)
|
||||
= do tag 1; toBuf b ty; toBuf b arg; toBuf b y
|
||||
toBuf b (ConstCase x y)
|
||||
= do tag 2; toBuf b x; toBuf b y
|
||||
toBuf b (DefaultCase x)
|
||||
= do tag 3; toBuf b x
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => do x <- fromBuf b; t <- fromBuf b
|
||||
args <- fromBuf b; y <- fromBuf b
|
||||
@ -374,7 +374,7 @@ mutual
|
||||
export
|
||||
TTC (Env Term vars) where
|
||||
toBuf b [] = pure ()
|
||||
toBuf b ((::) bnd env)
|
||||
toBuf b ((::) bnd env)
|
||||
= do toBuf b bnd; toBuf b env
|
||||
|
||||
-- Length has to correspond to length of 'vars'
|
||||
@ -390,7 +390,7 @@ TTC Visibility where
|
||||
toBuf b Export = tag 1
|
||||
toBuf b Public = tag 2
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => pure Private
|
||||
1 => pure Export
|
||||
@ -403,7 +403,7 @@ TTC PartialReason where
|
||||
toBuf b (BadCall xs) = do tag 1; toBuf b xs
|
||||
toBuf b (RecPath xs) = do tag 2; toBuf b xs
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => pure NotStrictlyPositive
|
||||
1 => do xs <- fromBuf b
|
||||
@ -429,14 +429,14 @@ TTC Terminating where
|
||||
export
|
||||
TTC Covering where
|
||||
toBuf b IsCovering = tag 0
|
||||
toBuf b (MissingCases ms)
|
||||
toBuf b (MissingCases ms)
|
||||
= do tag 1
|
||||
toBuf b ms
|
||||
toBuf b (NonCoveringCall ns)
|
||||
toBuf b (NonCoveringCall ns)
|
||||
= do tag 2
|
||||
toBuf b ns
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => pure IsCovering
|
||||
1 => do ms <- fromBuf b
|
||||
@ -544,7 +544,7 @@ TTC (PrimFn n) where
|
||||
35 => do ty <- fromBuf b; pure (ShiftL ty)
|
||||
36 => do ty <- fromBuf b; pure (ShiftR ty)
|
||||
_ => corrupt "PrimFn 2"
|
||||
|
||||
|
||||
fromBuf3 : Ref Bin Binary ->
|
||||
Core (PrimFn 3)
|
||||
fromBuf3 b
|
||||
@ -552,7 +552,7 @@ TTC (PrimFn n) where
|
||||
18 => pure StrSubstr
|
||||
100 => pure BelieveMe
|
||||
_ => corrupt "PrimFn 3"
|
||||
|
||||
|
||||
mutual
|
||||
export
|
||||
TTC (CExp vars) where
|
||||
@ -671,7 +671,7 @@ TTC CDef where
|
||||
toBuf b (MkForeign cs args ret) = do tag 2; toBuf b cs; toBuf b args; toBuf b ret
|
||||
toBuf b (MkError cexpr) = do tag 3; toBuf b cexpr
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => do args <- fromBuf b; cexpr <- fromBuf b
|
||||
pure (MkFun args cexpr)
|
||||
@ -734,7 +734,7 @@ TTC PrimNames where
|
||||
export
|
||||
TTC Def where
|
||||
toBuf b None = tag 0
|
||||
toBuf b (PMDef r args ct rt pats)
|
||||
toBuf b (PMDef r args ct rt pats)
|
||||
= do tag 1; toBuf b args; toBuf b ct; toBuf b rt; toBuf b pats
|
||||
toBuf b (ExternDef a)
|
||||
= do tag 2; toBuf b a
|
||||
@ -743,22 +743,22 @@ TTC Def where
|
||||
toBuf b (Builtin a)
|
||||
= throw (InternalError "Trying to serialise a Builtin")
|
||||
toBuf b (DCon t arity) = do tag 4; toBuf b t; toBuf b arity
|
||||
toBuf b (TCon t arity parampos detpos ms datacons)
|
||||
toBuf b (TCon t arity parampos detpos ms datacons)
|
||||
= do tag 5; toBuf b t; toBuf b arity; toBuf b parampos
|
||||
toBuf b detpos; toBuf b ms; toBuf b datacons
|
||||
toBuf b (Hole locs p)
|
||||
toBuf b (Hole locs p)
|
||||
= do tag 6; toBuf b locs; toBuf b p
|
||||
toBuf b (BySearch c depth def)
|
||||
toBuf b (BySearch c depth def)
|
||||
= do tag 7; toBuf b c; toBuf b depth; toBuf b def
|
||||
toBuf b (Guess guess envb constraints)
|
||||
toBuf b (Guess guess envb constraints)
|
||||
= do tag 8; toBuf b guess; toBuf b envb; toBuf b constraints
|
||||
toBuf b ImpBind = tag 9
|
||||
toBuf b Delayed = tag 10
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => pure None
|
||||
1 => do args <- fromBuf b
|
||||
1 => do args <- fromBuf b
|
||||
ct <- fromBuf b
|
||||
rt <- fromBuf b
|
||||
pats <- fromBuf b
|
||||
@ -771,7 +771,7 @@ TTC Def where
|
||||
4 => do t <- fromBuf b; a <- fromBuf b
|
||||
pure (DCon t a)
|
||||
5 => do t <- fromBuf b; a <- fromBuf b
|
||||
ps <- fromBuf b; dets <- fromBuf b;
|
||||
ps <- fromBuf b; dets <- fromBuf b;
|
||||
ms <- fromBuf b; cs <- fromBuf b
|
||||
pure (TCon t a ps dets ms cs)
|
||||
6 => do l <- fromBuf b
|
||||
@ -839,7 +839,7 @@ TTC SCCall where
|
||||
|
||||
export
|
||||
TTC GlobalDef where
|
||||
toBuf b gdef
|
||||
toBuf b gdef
|
||||
= -- Only write full details for user specified names. The others will
|
||||
-- be holes where all we will ever need after loading is the definition
|
||||
do toBuf b (fullname gdef)
|
||||
@ -859,30 +859,44 @@ TTC GlobalDef where
|
||||
toBuf b (noCycles gdef)
|
||||
toBuf b (sizeChange gdef)
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= do name <- fromBuf b
|
||||
def <- fromBuf b
|
||||
cdef <- fromBuf b
|
||||
refsList <- fromBuf b;
|
||||
refsList <- fromBuf b;
|
||||
let refs = map fromList refsList
|
||||
if isUserName name
|
||||
then do loc <- fromBuf b;
|
||||
ty <- fromBuf b; eargs <- fromBuf b;
|
||||
then do loc <- fromBuf b;
|
||||
ty <- fromBuf b; eargs <- fromBuf b;
|
||||
mul <- fromBuf b; vars <- fromBuf b
|
||||
vis <- fromBuf b; tot <- fromBuf b
|
||||
fl <- fromBuf b
|
||||
inv <- fromBuf b
|
||||
c <- fromBuf b
|
||||
sc <- fromBuf b
|
||||
pure (MkGlobalDef loc name ty eargs mul vars vis
|
||||
pure (MkGlobalDef loc name ty eargs mul vars vis
|
||||
tot fl refs inv c True def cdef sc)
|
||||
else do let fc = emptyFC
|
||||
pure (MkGlobalDef fc name (Erased fc) []
|
||||
RigW [] Public unchecked [] refs
|
||||
False False True def cdef [])
|
||||
|
||||
TTC Transform where
|
||||
toBuf b (MkTransform {vars} env lhs rhs)
|
||||
= do toBuf b vars
|
||||
toBuf b env
|
||||
toBuf b lhs
|
||||
toBuf b rhs
|
||||
|
||||
fromBuf b
|
||||
= do vars <- fromBuf b
|
||||
env <- fromBuf b
|
||||
lhs <- fromBuf b
|
||||
rhs <- fromBuf b
|
||||
pure (MkTransform {vars} env lhs rhs)
|
||||
|
||||
-- decode : Context -> Int -> ContextEntry -> Core GlobalDef
|
||||
Core.Context.decode gam idx (Coded bin)
|
||||
Core.Context.decode gam idx (Coded bin)
|
||||
= do b <- newRef Bin bin
|
||||
def <- fromBuf b
|
||||
let a = getContent gam
|
||||
|
@ -41,15 +41,15 @@ totRefsIn defs ty = totRefs defs (keys (getRefs (Resolved (-1)) ty))
|
||||
scEq : Term vars -> Term vars -> Bool
|
||||
scEq (Local _ _ idx _) (Local _ _ idx' _) = idx == idx'
|
||||
scEq (Ref _ _ n) (Ref _ _ n') = n == n'
|
||||
scEq (Meta _ _ i args) _ = True
|
||||
scEq _ (Meta _ _ i args) = True
|
||||
scEq (Meta _ _ i args) _ = True
|
||||
scEq _ (Meta _ _ i args) = True
|
||||
scEq (Bind _ _ b sc) (Bind _ _ b' sc') = False -- not checkable
|
||||
scEq (App _ f a) (App _ f' a') = scEq f f' && scEq a a'
|
||||
scEq (As _ a p) p' = scEq p p'
|
||||
scEq p (As _ a p') = scEq p p'
|
||||
scEq (TDelayed _ _ t) (TDelayed _ _ t') = scEq t t'
|
||||
scEq (TDelay _ _ t x) (TDelay _ _ t' x') = scEq t t' && scEq x x'
|
||||
scEq (TForce _ t) (TForce _ t') = scEq t t'
|
||||
scEq (TForce _ _ t) (TForce _ _ t') = scEq t t'
|
||||
scEq (PrimVal _ c) (PrimVal _ c') = c == c'
|
||||
scEq (Erased _) (Erased _) = True
|
||||
scEq (TType _) (TType _) = True
|
||||
@ -69,8 +69,8 @@ mutual
|
||||
List (Nat, Term vars) -> -- LHS args and their position
|
||||
Term vars -> -- Right hand side
|
||||
Core (List SCCall)
|
||||
findSC {vars} defs env g pats (Bind fc n b sc)
|
||||
= pure $
|
||||
findSC {vars} defs env g pats (Bind fc n b sc)
|
||||
= pure $
|
||||
!(findSCbinder b) ++
|
||||
!(findSC defs (b :: env) g (map (\ (p, tm) => (p, weaken tm)) pats) sc)
|
||||
where
|
||||
@ -97,7 +97,7 @@ mutual
|
||||
pure (concat scs)
|
||||
(_, Ref fc Func fn, args) =>
|
||||
do Just ty <- lookupTyExact fn (gamma defs)
|
||||
| Nothing =>
|
||||
| Nothing =>
|
||||
findSCcall defs env Unguarded pats fc fn 0 args
|
||||
arity <- getArity defs [] ty
|
||||
findSCcall defs env Unguarded pats fc fn arity args
|
||||
@ -109,7 +109,7 @@ mutual
|
||||
-- arity (i.e. the arity of the function we're calling) to ensure that
|
||||
-- it's noted that we don't know the size change relationship with the
|
||||
-- extra arguments.
|
||||
expandToArity : Nat -> List (Maybe (Nat, SizeChange)) ->
|
||||
expandToArity : Nat -> List (Maybe (Nat, SizeChange)) ->
|
||||
List (Maybe (Nat, SizeChange))
|
||||
expandToArity Z xs = xs
|
||||
expandToArity (S k) (x :: xs) = x :: expandToArity k xs
|
||||
@ -144,29 +144,29 @@ mutual
|
||||
= if assertedSmaller big tm
|
||||
then True
|
||||
else case getFnArgs tm of
|
||||
(Ref _ (DataCon t a) cn, args)
|
||||
(Ref _ (DataCon t a) cn, args)
|
||||
=> any (smaller True defs big s) args
|
||||
_ => case s of
|
||||
App _ f _ => smaller inc defs big f tm
|
||||
App _ f _ => smaller inc defs big f tm
|
||||
-- Higher order recursive argument
|
||||
_ => False
|
||||
|
||||
-- if the argument is an 'assert_smaller', return the thing it's smaller than
|
||||
asserted : Name -> Term vars -> Maybe (Term vars)
|
||||
asserted aSmaller tm
|
||||
asserted aSmaller tm
|
||||
= case getFnArgs tm of
|
||||
(Ref _ nt fn, [_, _, b, _])
|
||||
(Ref _ nt fn, [_, _, b, _])
|
||||
=> if fn == aSmaller
|
||||
then Just b
|
||||
else Nothing
|
||||
_ => Nothing
|
||||
|
||||
-- Calculate the size change for the given argument.
|
||||
-- i.e., return the size relationship of the given argument with an entry
|
||||
-- i.e., return the size relationship of the given argument with an entry
|
||||
-- in 'pats'; the position in 'pats' and the size change.
|
||||
-- Nothing if there is no relation with any of them.
|
||||
mkChange : Defs -> Name ->
|
||||
(pats : List (Nat, Term vars)) ->
|
||||
(pats : List (Nat, Term vars)) ->
|
||||
(arg : Term vars) ->
|
||||
Maybe (Nat, SizeChange)
|
||||
mkChange defs aSmaller [] arg = Nothing
|
||||
@ -184,7 +184,7 @@ mutual
|
||||
-- rather than treating the definitions separately.
|
||||
getCasePats : Defs -> Name -> List (Nat, Term vars) ->
|
||||
List (Term vars) ->
|
||||
Core (Maybe (List (vs ** (Env Term vs,
|
||||
Core (Maybe (List (vs ** (Env Term vs,
|
||||
List (Nat, Term vs), Term vs))))
|
||||
getCasePats {vars} defs n pats args
|
||||
= case !(lookupDefExact n (gamma defs)) of
|
||||
@ -216,11 +216,11 @@ mutual
|
||||
urhs (App fc f a) = App fc (updateRHS ms f) (updateRHS ms a)
|
||||
urhs (As fc a p) = As fc (updateRHS ms a) (updateRHS ms p)
|
||||
urhs (TDelayed fc r ty) = TDelayed fc r (updateRHS ms ty)
|
||||
urhs (TDelay fc r ty tm)
|
||||
urhs (TDelay fc r ty tm)
|
||||
= TDelay fc r (updateRHS ms ty) (updateRHS ms tm)
|
||||
urhs (TForce fc tm) = TForce fc (updateRHS ms tm)
|
||||
urhs (TForce fc r tm) = TForce fc r (updateRHS ms tm)
|
||||
urhs (Bind fc x b sc)
|
||||
= Bind fc x (map (updateRHS ms) b)
|
||||
= Bind fc x (map (updateRHS ms) b)
|
||||
(updateRHS (map (\vt => (weaken (fst vt), weaken (snd vt))) ms) sc)
|
||||
urhs (PrimVal fc c) = PrimVal fc c
|
||||
urhs (Erased fc) = Erased fc
|
||||
@ -229,7 +229,7 @@ mutual
|
||||
updatePat : List (Term vs, Term vs') -> (Nat, Term vs) -> (Nat, Term vs')
|
||||
updatePat ms (n, tm) = (n, updateRHS ms tm)
|
||||
|
||||
matchArgs : (vs ** (Env Term vs, Term vs, Term vs)) ->
|
||||
matchArgs : (vs ** (Env Term vs, Term vs, Term vs)) ->
|
||||
(vs ** (Env Term vs, List (Nat, Term vs), Term vs))
|
||||
matchArgs (_ ** (env', lhs, rhs))
|
||||
= let patMatch = reverse (zip args (getArgs lhs)) in
|
||||
@ -246,7 +246,7 @@ mutual
|
||||
List (Nat, Term vars) ->
|
||||
FC -> Name -> Nat -> List (Term vars) ->
|
||||
Core (List SCCall)
|
||||
findSCcall defs env g pats fc fn_in arity args
|
||||
findSCcall defs env g pats fc fn_in arity args
|
||||
-- Under 'assert_total' we assume that all calls are fine, so leave
|
||||
-- the size change list empty
|
||||
= do Just gdef <- lookupCtxtExact fn_in (gamma defs)
|
||||
@ -262,45 +262,49 @@ mutual
|
||||
Just ps => do scs <- traverse (findInCase defs g) ps
|
||||
pure (concat scs))]
|
||||
(do scs <- traverse (findSC defs env g pats) args
|
||||
pure ([MkSCCall fn
|
||||
(expandToArity arity
|
||||
(map (mkChange defs aSmaller pats) args))]
|
||||
pure ([MkSCCall fn
|
||||
(expandToArity arity
|
||||
(map (mkChange defs aSmaller pats) args))]
|
||||
++ concat scs))
|
||||
|
||||
findInCase : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> Guardedness ->
|
||||
Defs -> Guardedness ->
|
||||
(vs ** (Env Term vs, List (Nat, Term vs), Term vs)) ->
|
||||
Core (List SCCall)
|
||||
findInCase defs g (_ ** (env, pats, tm))
|
||||
findInCase defs g (_ ** (env, pats, tm))
|
||||
= do logC 10 (do ps <- traverse toFullNames (map snd pats)
|
||||
pure ("Looking in case args " ++ show ps))
|
||||
logTermNF 10 " =" env tm
|
||||
rhs <- normaliseOpts tcOnly defs env tm
|
||||
findSC defs env g pats rhs
|
||||
|
||||
-- Remove all laziness annotations which are nothing to do with coinduction,
|
||||
-- meaning that all only Force/Delay left is to guard coinductive calls.
|
||||
-- Remove all force and delay annotations which are nothing to do with
|
||||
-- coinduction meaning that all Delays left guard coinductive calls.
|
||||
delazy : Defs -> Term vars -> Term vars
|
||||
delazy defs (TDelayed fc r tm)
|
||||
delazy defs (TDelayed fc r tm)
|
||||
= let tm' = delazy defs tm in
|
||||
case r of
|
||||
LInf => TDelayed fc r tm'
|
||||
_ => tm'
|
||||
delazy defs (TDelay fc r ty tm)
|
||||
delazy defs (TDelay fc r ty tm)
|
||||
= let ty' = delazy defs ty
|
||||
tm' = delazy defs tm in
|
||||
case r of
|
||||
LInf => TDelay fc r ty' tm'
|
||||
_ => tm'
|
||||
delazy defs (TForce fc r t)
|
||||
= case r of
|
||||
LInf => TForce fc r (delazy defs t)
|
||||
_ => delazy defs t
|
||||
delazy defs (Meta fc n i args) = Meta fc n i (map (delazy defs) args)
|
||||
delazy defs (Bind fc x b sc)
|
||||
delazy defs (Bind fc x b sc)
|
||||
= Bind fc x (map (delazy defs) b) (delazy defs sc)
|
||||
delazy defs (App fc f a) = App fc (delazy defs f) (delazy defs a)
|
||||
delazy defs (As fc a p) = As fc (delazy defs a) (delazy defs p)
|
||||
delazy defs tm = tm
|
||||
|
||||
findCalls : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> (vars ** (Env Term vars, Term vars, Term vars)) ->
|
||||
Defs -> (vars ** (Env Term vars, Term vars, Term vars)) ->
|
||||
Core (List SCCall)
|
||||
findCalls defs (_ ** (env, lhs, rhs_in))
|
||||
= do let pargs = getArgs (delazy defs lhs)
|
||||
@ -310,7 +314,7 @@ findCalls defs (_ ** (env, lhs, rhs_in))
|
||||
|
||||
getSC : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> Def -> Core (List SCCall)
|
||||
getSC defs (PMDef _ args _ _ pats)
|
||||
getSC defs (PMDef _ args _ _ pats)
|
||||
= do sc <- traverse (findCalls defs) pats
|
||||
pure (concat sc)
|
||||
getSC defs _ = pure []
|
||||
@ -338,7 +342,7 @@ nextArg x = x + 1
|
||||
initArgs : {auto a : Ref APos Arg} ->
|
||||
Nat -> Core (List (Maybe (Arg, SizeChange)))
|
||||
initArgs Z = pure []
|
||||
initArgs (S k)
|
||||
initArgs (S k)
|
||||
= do arg <- get APos
|
||||
put APos (nextArg arg)
|
||||
args' <- initArgs k
|
||||
@ -351,7 +355,7 @@ initArgs (S k)
|
||||
-- use that rather than continuing to traverse the graph!
|
||||
checkSC : {auto a : Ref APos Arg} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
Defs ->
|
||||
Defs ->
|
||||
Name -> -- function we're checking
|
||||
List (Maybe (Arg, SizeChange)) -> -- functions arguments and change
|
||||
List (Name, List (Maybe Arg)) -> -- calls we've seen so far
|
||||
@ -398,7 +402,7 @@ checkSC defs f args path
|
||||
then case term of
|
||||
NotTerminating (RecPath _) =>
|
||||
-- might have lost information while assuming this
|
||||
-- was mutually recursive, so start again with new
|
||||
-- was mutually recursive, so start again with new
|
||||
-- arguments (that is, where we'd start if the
|
||||
-- function was the top level thing we were checking)
|
||||
do args' <- initArgs (length (fnArgs sc))
|
||||
@ -419,11 +423,11 @@ checkSC defs f args path
|
||||
|
||||
calcTerminating : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Core Terminating
|
||||
calcTerminating loc n
|
||||
calcTerminating loc n
|
||||
= do defs <- get Ctxt
|
||||
case !(lookupCtxtExact n (gamma defs)) of
|
||||
Nothing => throw (UndefinedName loc n)
|
||||
Just def =>
|
||||
Just def =>
|
||||
case !(totRefs defs (nub !(addCases defs (keys (refersTo def))))) of
|
||||
IsTerminating =>
|
||||
do let ty = type def
|
||||
@ -455,7 +459,7 @@ checkTerminating : {auto c : Ref Ctxt Defs} ->
|
||||
checkTerminating loc n
|
||||
= do tot <- getTotality loc n
|
||||
case isTerminating tot of
|
||||
Unchecked =>
|
||||
Unchecked =>
|
||||
do tot' <- calcTerminating loc n
|
||||
setTerminating loc n tot'
|
||||
pure tot'
|
||||
@ -468,7 +472,7 @@ nameIn defs tyns (NBind fc x b sc)
|
||||
else do sc' <- sc defs (toClosure defaultOpts [] (Erased fc))
|
||||
nameIn defs tyns sc'
|
||||
nameIn defs tyns (NApp _ _ args)
|
||||
= anyM (nameIn defs tyns)
|
||||
= anyM (nameIn defs tyns)
|
||||
!(traverse (evalClosure defs) args)
|
||||
nameIn defs tyns (NTCon _ n _ _ args)
|
||||
= if n `elem` tyns
|
||||
@ -476,7 +480,7 @@ nameIn defs tyns (NTCon _ n _ _ args)
|
||||
else do args' <- traverse (evalClosure defs) args
|
||||
anyM (nameIn defs tyns) args'
|
||||
nameIn defs tyns (NDCon _ n _ _ args)
|
||||
= anyM (nameIn defs tyns)
|
||||
= anyM (nameIn defs tyns)
|
||||
!(traverse (evalClosure defs) args)
|
||||
nameIn defs tyns _ = pure False
|
||||
|
||||
@ -485,10 +489,10 @@ nameIn defs tyns _ = pure False
|
||||
posArg : Defs -> List Name -> NF [] -> Core Terminating
|
||||
-- a tyn can only appear in the parameter positions of
|
||||
-- tc; report positivity failure if it appears anywhere else
|
||||
posArg defs tyns (NTCon _ tc _ _ args)
|
||||
posArg defs tyns (NTCon _ tc _ _ args)
|
||||
= let testargs : List (Closure [])
|
||||
= case !(lookupDefExact tc (gamma defs)) of
|
||||
Just (TCon _ _ params _ _ _) =>
|
||||
Just (TCon _ _ params _ _ _) =>
|
||||
dropParams 0 params args
|
||||
_ => args in
|
||||
if !(anyM (nameIn defs tyns)
|
||||
@ -503,7 +507,7 @@ posArg defs tyns (NTCon _ tc _ _ args)
|
||||
then dropParams (S i) ps xs
|
||||
else x :: dropParams (S i) ps xs
|
||||
-- a tyn can not appear as part of ty
|
||||
posArg defs tyns (NBind fc x (Pi c e ty) sc)
|
||||
posArg defs tyns (NBind fc x (Pi c e ty) sc)
|
||||
= if !(nameIn defs tyns ty)
|
||||
then pure (NotTerminating NotStrictlyPositive)
|
||||
else do sc' <- sc defs (toClosure defaultOpts [] (Erased fc))
|
||||
@ -511,20 +515,20 @@ posArg defs tyns (NBind fc x (Pi c e ty) sc)
|
||||
posArg defs tyn _ = pure IsTerminating
|
||||
|
||||
checkPosArgs : Defs -> List Name -> NF [] -> Core Terminating
|
||||
checkPosArgs defs tyns (NBind fc x (Pi c e ty) sc)
|
||||
checkPosArgs defs tyns (NBind fc x (Pi c e ty) sc)
|
||||
= case !(posArg defs tyns ty) of
|
||||
IsTerminating =>
|
||||
checkPosArgs defs tyns
|
||||
IsTerminating =>
|
||||
checkPosArgs defs tyns
|
||||
!(sc defs (toClosure defaultOpts [] (Erased fc)))
|
||||
bad => pure bad
|
||||
checkPosArgs defs tyns _ = pure IsTerminating
|
||||
|
||||
checkCon : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> List Name -> Name -> Core Terminating
|
||||
checkCon defs tyns cn
|
||||
checkCon defs tyns cn
|
||||
= case !(lookupTyExact cn (gamma defs)) of
|
||||
Nothing => pure Unchecked
|
||||
Just ty =>
|
||||
Just ty =>
|
||||
case !(totRefsIn defs ty) of
|
||||
IsTerminating => checkPosArgs defs tyns !(nf defs [] ty)
|
||||
bad => pure bad
|
||||
@ -541,12 +545,12 @@ checkData defs tyns (c :: cs)
|
||||
-- return whether it's terminating, along with its data constructors
|
||||
calcPositive : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Core (Terminating, List Name)
|
||||
calcPositive loc n
|
||||
calcPositive loc n
|
||||
= do defs <- get Ctxt
|
||||
case !(lookupDefTyExact n (gamma defs)) of
|
||||
Just (TCon _ _ _ _ tns dcons, ty) =>
|
||||
Just (TCon _ _ _ _ tns dcons, ty) =>
|
||||
case !(totRefsIn defs ty) of
|
||||
IsTerminating =>
|
||||
IsTerminating =>
|
||||
do t <- checkData defs (n :: tns) dcons
|
||||
pure (t , dcons)
|
||||
bad => pure (bad, dcons)
|
||||
@ -570,7 +574,7 @@ checkPositive loc n_in
|
||||
t => pure t
|
||||
|
||||
-- Check and record totality of the given name; positivity if it's a data
|
||||
-- type, termination if it's a function
|
||||
-- type, termination if it's a function
|
||||
export
|
||||
checkTotal : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Core Terminating
|
||||
|
10
src/Core/Transform.idr
Normal file
10
src/Core/Transform.idr
Normal file
@ -0,0 +1,10 @@
|
||||
module Core.Transform
|
||||
|
||||
import Core.Context
|
||||
import Core.Env
|
||||
import Core.TT
|
||||
|
||||
export
|
||||
applyTransforms : {auto c : Ref Ctxt Defs} ->
|
||||
Term vars -> Core (Term vars)
|
||||
applyTransforms t = pure t -- TODO!
|
@ -32,12 +32,12 @@ Eq UnifyMode where
|
||||
-- explicit force or delay to the first argument to unification. This says
|
||||
-- which to add, if any. Can only added at the very top level.
|
||||
public export
|
||||
data AddLazy = NoLazy | AddForce | AddDelay LazyReason
|
||||
data AddLazy = NoLazy | AddForce LazyReason | AddDelay LazyReason
|
||||
|
||||
export
|
||||
Show AddLazy where
|
||||
show NoLazy = "NoLazy"
|
||||
show AddForce = "AddForce"
|
||||
show (AddForce _) = "AddForce"
|
||||
show (AddDelay _) = "AddDelay"
|
||||
|
||||
public export
|
||||
@ -73,7 +73,7 @@ interface Unify (tm : List Name -> Type) where
|
||||
Ref UST UState ->
|
||||
UnifyMode ->
|
||||
FC -> Env Term vars ->
|
||||
tm vars -> tm vars ->
|
||||
tm vars -> tm vars ->
|
||||
Core UnifyResult
|
||||
-- As unify but at the top level can allow lazy/non-lazy to be mixed in
|
||||
-- order to infer annotations
|
||||
@ -81,7 +81,7 @@ interface Unify (tm : List Name -> Type) where
|
||||
Ref UST UState ->
|
||||
UnifyMode ->
|
||||
FC -> Env Term vars ->
|
||||
tm vars -> tm vars ->
|
||||
tm vars -> tm vars ->
|
||||
Core UnifyResult
|
||||
unifyWithLazyD = unifyD
|
||||
|
||||
@ -89,22 +89,22 @@ interface Unify (tm : List Name -> Type) where
|
||||
-- In calls to unification, the first argument is the given type, and the second
|
||||
-- argument is the expected type.
|
||||
export
|
||||
unify : Unify tm =>
|
||||
unify : Unify tm =>
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
UnifyMode ->
|
||||
FC -> Env Term vars ->
|
||||
tm vars -> tm vars ->
|
||||
tm vars -> tm vars ->
|
||||
Core UnifyResult
|
||||
unify {c} {u} = unifyD c u
|
||||
|
||||
export
|
||||
unifyWithLazy : Unify tm =>
|
||||
unifyWithLazy : Unify tm =>
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
UnifyMode ->
|
||||
FC -> Env Term vars ->
|
||||
tm vars -> tm vars ->
|
||||
tm vars -> tm vars ->
|
||||
Core UnifyResult
|
||||
unifyWithLazy {c} {u} = unifyWithLazyD c u
|
||||
|
||||
@ -114,7 +114,7 @@ search : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount ->
|
||||
(defaults : Bool) -> (depth : Nat) ->
|
||||
(defining : Name) -> (topTy : Term vars) -> Env Term vars ->
|
||||
(defining : Name) -> (topTy : Term vars) -> Env Term vars ->
|
||||
Core (Term vars)
|
||||
|
||||
ufail : FC -> String -> Core a
|
||||
@ -122,15 +122,15 @@ ufail loc msg = throw (GenericMsg loc msg)
|
||||
|
||||
convertError : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Env Term vars -> NF vars -> NF vars -> Core a
|
||||
convertError loc env x y
|
||||
convertError loc env x y
|
||||
= do defs <- get Ctxt
|
||||
empty <- clearDefs defs
|
||||
throw (CantConvert loc env !(quote empty env x)
|
||||
throw (CantConvert loc env !(quote empty env x)
|
||||
!(quote empty env y))
|
||||
|
||||
convertErrorS : {auto c : Ref Ctxt Defs} ->
|
||||
Bool -> FC -> Env Term vars -> NF vars -> NF vars -> Core a
|
||||
convertErrorS s loc env x y
|
||||
convertErrorS s loc env x y
|
||||
= if s then convertError loc env y x
|
||||
else convertError loc env x y
|
||||
|
||||
@ -144,9 +144,9 @@ postpone loc logstr env x y
|
||||
logC 10 $
|
||||
do xq <- quote defs env x
|
||||
yq <- quote defs env y
|
||||
pure (logstr ++ ": " ++ show !(toFullNames xq) ++
|
||||
pure (logstr ++ ": " ++ show !(toFullNames xq) ++
|
||||
" =?= " ++ show !(toFullNames yq))
|
||||
c <- addConstraint (MkConstraint loc env !(quote empty env x)
|
||||
c <- addConstraint (MkConstraint loc env !(quote empty env x)
|
||||
!(quote empty env y))
|
||||
pure (constrain c)
|
||||
|
||||
@ -173,20 +173,20 @@ unifyArgs mode loc env (cx :: cxs) (cy :: cys)
|
||||
pure (union res cs)
|
||||
_ => do cs <- unifyArgs mode loc env cxs cys
|
||||
-- TODO: Fix this bit! See p59 Ulf's thesis
|
||||
-- c <- addConstraint
|
||||
-- (MkSeqConstraint loc env
|
||||
-- c <- addConstraint
|
||||
-- (MkSeqConstraint loc env
|
||||
-- (map (quote gam env) (cx :: cxs))
|
||||
-- (map (quote gam env) (cy :: cys)))
|
||||
pure (union res cs) -- [c]
|
||||
unifyArgs mode loc env _ _ = ufail loc ""
|
||||
|
||||
-- Get the variables in an application argument list; fail if any arguments
|
||||
-- Get the variables in an application argument list; fail if any arguments
|
||||
-- are not variables, fail if there's any repetition of variables
|
||||
-- We use this to check that the pattern unification rule is applicable
|
||||
-- when solving a metavariable applied to arguments
|
||||
getVars : List Nat -> List (NF vars) -> Maybe (List (Var vars))
|
||||
getVars got [] = Just []
|
||||
getVars got (NApp fc (NLocal r idx v) [] :: xs)
|
||||
getVars got (NApp fc (NLocal r idx v) [] :: xs)
|
||||
= if inArgs idx got then Nothing
|
||||
else do xs' <- getVars (idx :: got) xs
|
||||
pure (MkVar v :: xs')
|
||||
@ -195,9 +195,9 @@ getVars got (NApp fc (NLocal r idx v) [] :: xs)
|
||||
-- Nat is linear time in Idris 1!
|
||||
inArgs : Nat -> List Nat -> Bool
|
||||
inArgs n [] = False
|
||||
inArgs n (n' :: ns)
|
||||
inArgs n (n' :: ns)
|
||||
= if toIntegerNat n == toIntegerNat n' then True else inArgs n ns
|
||||
getVars got (NAs _ _ p :: xs) = getVars got (p :: xs)
|
||||
getVars got (NAs _ _ p :: xs) = getVars got (p :: xs)
|
||||
getVars _ (_ :: xs) = Nothing
|
||||
|
||||
-- Make a sublist representing the variables used in the application.
|
||||
@ -206,13 +206,13 @@ getVars _ (_ :: xs) = Nothing
|
||||
toSubVars : (vars : List Name) -> List (Var vars) ->
|
||||
(newvars ** SubVars newvars vars)
|
||||
toSubVars [] xs = ([] ** SubRefl)
|
||||
toSubVars (n :: ns) xs
|
||||
toSubVars (n :: ns) xs
|
||||
-- If there's a proof 'First' in 'xs', then 'n' should be kept,
|
||||
-- otherwise dropped
|
||||
-- (Remember: 'n' might be shadowed; looking for 'First' ensures we
|
||||
-- get the *right* proof that the name is in scope!)
|
||||
= let (_ ** svs) = toSubVars ns (dropFirst xs) in
|
||||
if anyFirst xs
|
||||
if anyFirst xs
|
||||
then (_ ** KeepCons svs)
|
||||
else (_ ** DropCons svs)
|
||||
where
|
||||
@ -239,7 +239,7 @@ toSubVars (n :: ns) xs
|
||||
patternEnv : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{vars : _} ->
|
||||
Env Term vars -> List (Closure vars) ->
|
||||
Env Term vars -> List (Closure vars) ->
|
||||
Core (Maybe (newvars ** (List (Var newvars),
|
||||
SubVars newvars vars)))
|
||||
patternEnv {vars} env args
|
||||
@ -248,9 +248,9 @@ patternEnv {vars} env args
|
||||
args' <- traverse (evalArg empty) args
|
||||
case getVars [] args' of
|
||||
Nothing => pure Nothing
|
||||
Just vs =>
|
||||
Just vs =>
|
||||
let (newvars ** svs) = toSubVars _ vs in
|
||||
pure (Just (newvars **
|
||||
pure (Just (newvars **
|
||||
(updateVars vs svs, svs)))
|
||||
where
|
||||
-- Update the variable list to point into the sub environment
|
||||
@ -262,10 +262,10 @@ patternEnv {vars} env args
|
||||
= case subElem p svs of
|
||||
Nothing => updateVars ps svs
|
||||
Just p' => p' :: updateVars ps svs
|
||||
|
||||
|
||||
getVarsBelowTm : Nat -> List (Term vars) -> Maybe (List (Var vars))
|
||||
getVarsBelowTm max [] = Just []
|
||||
getVarsBelowTm max (Local fc r idx v :: xs)
|
||||
getVarsBelowTm max (Local fc r idx v :: xs)
|
||||
= if idx >= max then Nothing
|
||||
else do xs' <- getVarsBelowTm idx xs
|
||||
pure (MkVar v :: xs')
|
||||
@ -275,7 +275,7 @@ export
|
||||
patternEnvTm : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{vars : _} ->
|
||||
Env Term vars -> List (Term vars) ->
|
||||
Env Term vars -> List (Term vars) ->
|
||||
Core (Maybe (newvars ** (List (Var newvars),
|
||||
SubVars newvars vars)))
|
||||
patternEnvTm {vars} env args
|
||||
@ -283,9 +283,9 @@ patternEnvTm {vars} env args
|
||||
empty <- clearDefs defs
|
||||
case getVarsBelowTm 1000000 args of
|
||||
Nothing => pure Nothing
|
||||
Just vs =>
|
||||
Just vs =>
|
||||
let (newvars ** svs) = toSubVars _ vs in
|
||||
pure (Just (newvars **
|
||||
pure (Just (newvars **
|
||||
(updateVars vs svs, svs)))
|
||||
where
|
||||
-- Update the variable list to point into the sub environment
|
||||
@ -305,7 +305,7 @@ patternEnvTm {vars} env args
|
||||
instantiate : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{newvars : _} ->
|
||||
FC -> Env Term vars ->
|
||||
FC -> Env Term vars ->
|
||||
(metavar : Name) -> (mref : Int) -> (mdef : GlobalDef) ->
|
||||
List (Var newvars) -> -- Variable each argument maps to
|
||||
Term vars -> -- original, just for error message
|
||||
@ -324,14 +324,14 @@ instantiate {newvars} loc env mname mref mdef locs otm tm
|
||||
defs <- get Ctxt
|
||||
rhs <- mkDef [] newvars (snocList newvars) CompatPre
|
||||
(rewrite appendNilRightNeutral newvars in locs)
|
||||
(rewrite appendNilRightNeutral newvars in tm)
|
||||
(rewrite appendNilRightNeutral newvars in tm)
|
||||
ty
|
||||
logTerm 5 ("Instantiated: " ++ show mname) ty
|
||||
log 5 ("From vars: " ++ show newvars)
|
||||
logTerm 5 "Definition" rhs
|
||||
let simpleDef = isSimple rhs
|
||||
let newdef = record { definition =
|
||||
PMDef simpleDef [] (STerm rhs) (STerm rhs) []
|
||||
let newdef = record { definition =
|
||||
PMDef simpleDef [] (STerm rhs) (STerm rhs) []
|
||||
} mdef
|
||||
addDef (Resolved mref) newdef
|
||||
removeHole mref
|
||||
@ -345,7 +345,7 @@ instantiate {newvars} loc env mname mref mdef locs otm tm
|
||||
isSimple (TType _) = True
|
||||
isSimple _ = False
|
||||
|
||||
updateLoc : {v : Nat} -> List (Var vs) -> .(IsVar name v vs') ->
|
||||
updateLoc : {v : Nat} -> List (Var vs) -> .(IsVar name v vs') ->
|
||||
Maybe (Var vs)
|
||||
updateLoc [] el = Nothing
|
||||
updateLoc (p :: ps) First = Just p
|
||||
@ -360,8 +360,8 @@ instantiate {newvars} loc env mname mref mdef locs otm tm
|
||||
Just (Local fc r _ p')
|
||||
updateLocs {vs} locs (Bind fc x b sc)
|
||||
= do b' <- updateLocsB b
|
||||
sc' <- updateLocs
|
||||
(MkVar First :: map (\ (MkVar p) => (MkVar (Later p))) locs)
|
||||
sc' <- updateLocs
|
||||
(MkVar First :: map (\ (MkVar p) => (MkVar (Later p))) locs)
|
||||
sc
|
||||
Just (Bind fc x b' sc')
|
||||
where
|
||||
@ -379,15 +379,15 @@ instantiate {newvars} loc env mname mref mdef locs otm tm
|
||||
|
||||
mkDef : (got : List Name) -> (vs : List Name) -> SnocList vs ->
|
||||
CompatibleVars got rest ->
|
||||
List (Var (vs ++ got)) -> Term (vs ++ got) ->
|
||||
List (Var (vs ++ got)) -> Term (vs ++ got) ->
|
||||
Term ts -> Core (Term rest)
|
||||
mkDef {rest} got [] Empty cvs locs tm ty
|
||||
mkDef {rest} got [] Empty cvs locs tm ty
|
||||
= do let Just tm' = updateLocs (reverse locs) tm
|
||||
| Nothing => ufail loc ("Can't make solution for " ++ show mname)
|
||||
pure (renameVars cvs tm')
|
||||
mkDef got vs rec cvs locs tm (Bind _ _ (Let _ _ _) sc)
|
||||
= mkDef got vs rec cvs locs tm sc
|
||||
mkDef got (vs ++ [v]) (Snoc rec) cvs locs tm (Bind bfc x (Pi c _ ty) sc)
|
||||
mkDef got (vs ++ [v]) (Snoc rec) cvs locs tm (Bind bfc x (Pi c _ ty) sc)
|
||||
= do defs <- get Ctxt
|
||||
sc' <- mkDef (v :: got) vs rec (CompatExt cvs)
|
||||
(rewrite appendAssociative vs [v] got in locs)
|
||||
@ -395,8 +395,8 @@ instantiate {newvars} loc env mname mref mdef locs otm tm
|
||||
sc
|
||||
pure (Bind bfc x (Lam c Explicit (Erased bfc)) sc')
|
||||
mkDef got (vs ++ [v]) (Snoc rec) cvs locs tm ty
|
||||
= ufail loc $ "Can't make solution for " ++ show mname
|
||||
|
||||
= ufail loc $ "Can't make solution for " ++ show mname
|
||||
|
||||
export
|
||||
solveIfUndefined : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
@ -416,7 +416,7 @@ solveIfUndefined env (Meta fc mname idx args) soln
|
||||
| Nothing => throw (InternalError "Can't happen: no definition")
|
||||
instantiate fc env mname idx hdef locs soln stm
|
||||
pure True
|
||||
solveIfUndefined env metavar soln
|
||||
solveIfUndefined env metavar soln
|
||||
= pure False
|
||||
|
||||
isDefInvertible : {auto c : Ref Ctxt Defs} ->
|
||||
@ -432,17 +432,17 @@ mutual
|
||||
{auto u : Ref UST UState} ->
|
||||
{vars : _} ->
|
||||
(postpone : Bool) ->
|
||||
FC -> Env Term vars -> NF vars -> NF vars ->
|
||||
FC -> Env Term vars -> NF vars -> NF vars ->
|
||||
Core UnifyResult
|
||||
unifyIfEq post loc env x y
|
||||
unifyIfEq post loc env x y
|
||||
= do defs <- get Ctxt
|
||||
if !(convert defs env x y)
|
||||
then pure success
|
||||
else if post
|
||||
else if post
|
||||
then postpone loc "Postponing unifyIfEq" env x y
|
||||
else convertError loc env x y
|
||||
|
||||
getArgTypes : Defs -> (fnType : NF vars) -> List (Closure vars) ->
|
||||
getArgTypes : Defs -> (fnType : NF vars) -> List (Closure vars) ->
|
||||
Core (Maybe (List (NF vars)))
|
||||
getArgTypes defs (NBind _ n (Pi _ _ ty) sc) (a :: as)
|
||||
= do Just scTys <- getArgTypes defs !(sc defs a) as
|
||||
@ -452,18 +452,18 @@ mutual
|
||||
getArgTypes _ _ _ = pure Nothing
|
||||
|
||||
headsConvert : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars ->
|
||||
Env Term vars ->
|
||||
Maybe (List (NF vars)) -> Maybe (List (NF vars)) ->
|
||||
Core Bool
|
||||
headsConvert env (Just vs) (Just ns)
|
||||
= case (reverse vs, reverse ns) of
|
||||
(v :: _, n :: _) =>
|
||||
(v :: _, n :: _) =>
|
||||
do logNF 10 "Converting" env v
|
||||
logNF 10 "......with" env n
|
||||
defs <- get Ctxt
|
||||
convert defs env v n
|
||||
_ => pure False
|
||||
headsConvert env _ _
|
||||
headsConvert env _ _
|
||||
= do log 10 "Nothing to convert"
|
||||
pure True
|
||||
|
||||
@ -492,7 +492,7 @@ mutual
|
||||
-- If the rightmost arguments have the same type, or we don't
|
||||
-- know the types of the arguments, we'll get on with it.
|
||||
if !(headsConvert env vargTys nargTys)
|
||||
then
|
||||
then
|
||||
-- Unify the rightmost arguments, with the goal of turning the
|
||||
-- hole application into a pattern form
|
||||
case (reverse margs', reverse args') of
|
||||
@ -502,7 +502,7 @@ mutual
|
||||
do log 10 "Unifying invertible"
|
||||
ures <- unify mode fc env h f
|
||||
log 10 $ "Constraints " ++ show (constraints ures)
|
||||
uargs <- unify mode fc env
|
||||
uargs <- unify mode fc env
|
||||
(NApp fc (NMeta mname mref margs) (reverse hargs))
|
||||
(con (reverse fargs))
|
||||
pure (union ures uargs)
|
||||
@ -510,7 +510,7 @@ mutual
|
||||
do log 10 "Unifying invertible"
|
||||
ures <- unify mode fc env f h
|
||||
log 10 $ "Constraints " ++ show (constraints ures)
|
||||
uargs <- unify mode fc env
|
||||
uargs <- unify mode fc env
|
||||
(con (reverse fargs))
|
||||
(NApp fc (NMeta mname mref margs) (reverse hargs))
|
||||
pure (union ures uargs))
|
||||
@ -521,7 +521,7 @@ mutual
|
||||
(NApp fc (NMeta mname mref margs) margs')
|
||||
(con args')
|
||||
else -- TODO: Cancellable function applications
|
||||
postpone fc "Postponing hole application [3]" env
|
||||
postpone fc "Postponing hole application [3]" env
|
||||
(NApp fc (NMeta mname mref margs) margs') (con args')
|
||||
|
||||
-- Unify a hole application - we have already checked that the hole is
|
||||
@ -546,7 +546,7 @@ mutual
|
||||
mty <- lookupTyExact n (gamma defs)
|
||||
unifyInvertible swap mode loc env mname mref margs margs' mty (NTCon nfc n t a) args'
|
||||
unifyHoleApp swap mode loc env mname mref margs margs' (NApp nfc (NLocal r idx p) args')
|
||||
= unifyInvertible swap mode loc env mname mref margs margs' Nothing
|
||||
= unifyInvertible swap mode loc env mname mref margs margs' Nothing
|
||||
(NApp nfc (NLocal r idx p)) args'
|
||||
unifyHoleApp swap mode loc env mname mref margs margs' tm@(NApp nfc (NMeta n i margs2) args2')
|
||||
= do defs <- get Ctxt
|
||||
@ -556,15 +556,15 @@ mutual
|
||||
if inv
|
||||
then unifyInvertible swap mode loc env mname mref margs margs' Nothing
|
||||
(NApp nfc (NMeta n i margs2)) args2'
|
||||
else postponeS swap loc "Postponing hole application" env
|
||||
else postponeS swap loc "Postponing hole application" env
|
||||
(NApp loc (NMeta mname mref margs) margs') tm
|
||||
where
|
||||
isPatName : Name -> Bool
|
||||
isPatName (PV _ _) = True
|
||||
isPatName _ = False
|
||||
|
||||
|
||||
unifyHoleApp swap mode loc env mname mref margs margs' tm
|
||||
= postponeS swap loc "Postponing hole application" env
|
||||
= postponeS swap loc "Postponing hole application" env
|
||||
(NApp loc (NMeta mname mref margs) margs') tm
|
||||
|
||||
postponePatVar : {auto c : Ref Ctxt Defs} ->
|
||||
@ -578,7 +578,7 @@ mutual
|
||||
(soln : NF vars) ->
|
||||
Core UnifyResult
|
||||
postponePatVar swap mode loc env mname mref margs margs' tm
|
||||
= postponeS swap loc "Not in pattern fragment" env
|
||||
= postponeS swap loc "Not in pattern fragment" env
|
||||
(NApp loc (NMeta mname mref margs) margs') tm
|
||||
|
||||
solveHole : {auto c : Ref Ctxt Defs} ->
|
||||
@ -599,7 +599,7 @@ mutual
|
||||
empty <- clearDefs defs
|
||||
-- if the terms are the same, this isn't a solution
|
||||
-- but they are already unifying, so just return
|
||||
if solutionHeadSame solnf
|
||||
if solutionHeadSame solnf
|
||||
then pure success
|
||||
else -- Rather than doing the occurs check here immediately,
|
||||
-- we'll wait until all metavariables are resolved, and in
|
||||
@ -638,7 +638,7 @@ mutual
|
||||
pure $ "Unifying: " ++ show mname ++ " " ++ show qargs ++
|
||||
" with " ++ show qtm)
|
||||
case !(patternEnv env args) of
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do Just hdef <- lookupCtxtExact (Resolved mref) (gamma defs)
|
||||
| _ => postponePatVar swap mode loc env mname mref margs margs' tmnf
|
||||
let Hole _ _ = definition hdef
|
||||
@ -646,20 +646,20 @@ mutual
|
||||
if invertible hdef
|
||||
then unifyHoleApp swap mode loc env mname mref margs margs' tmnf
|
||||
else postponePatVar swap mode loc env mname mref margs margs' tmnf
|
||||
Just (newvars ** (locs, submv)) =>
|
||||
Just (newvars ** (locs, submv)) =>
|
||||
do tm <- quote empty env tmnf
|
||||
case shrinkTerm tm submv of
|
||||
Just stm => solveHole fc env mname mref
|
||||
margs margs' locs submv
|
||||
Just stm => solveHole fc env mname mref
|
||||
margs margs' locs submv
|
||||
tm stm tmnf
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do tm' <- normalise defs env tm
|
||||
case shrinkTerm tm' submv of
|
||||
Nothing => postponeS swap loc "Can't shrink" env
|
||||
(NApp loc (NMeta mname mref margs) margs')
|
||||
tmnf
|
||||
Just stm => solveHole fc env mname mref
|
||||
margs margs' locs submv
|
||||
Just stm => solveHole fc env mname mref
|
||||
margs margs' locs submv
|
||||
tm stm tmnf
|
||||
|
||||
-- Unify an application with something else
|
||||
@ -684,8 +684,8 @@ mutual
|
||||
unifyApp swap mode loc env xfc (NLocal rx x xp) [] (NApp yfc (NLocal ry y yp) [])
|
||||
= do gam <- get Ctxt
|
||||
if x == y then pure success
|
||||
else postponeS swap loc "Postponing var"
|
||||
env (NApp xfc (NLocal rx x xp) [])
|
||||
else postponeS swap loc "Postponing var"
|
||||
env (NApp xfc (NLocal rx x xp) [])
|
||||
(NApp yfc (NLocal ry y yp) [])
|
||||
-- A local against something canonical (binder or constructor) is bad
|
||||
unifyApp swap mode loc env xfc (NLocal rx x xp) args y@(NBind _ _ _ _)
|
||||
@ -700,13 +700,13 @@ mutual
|
||||
= convertErrorS swap loc env (NApp xfc (NLocal rx x xp) args) y
|
||||
-- If they're already convertible without metavariables, we're done,
|
||||
-- otherwise postpone
|
||||
unifyApp False mode loc env fc hd args tm
|
||||
unifyApp False mode loc env fc hd args tm
|
||||
= do gam <- get Ctxt
|
||||
if !(convert gam env (NApp fc hd args) tm)
|
||||
then pure success
|
||||
else postponeS False loc "Postponing constraint"
|
||||
env (NApp fc hd args) tm
|
||||
unifyApp True mode loc env fc hd args tm
|
||||
unifyApp True mode loc env fc hd args tm
|
||||
= do gam <- get Ctxt
|
||||
if !(convert gam env tm (NApp fc hd args))
|
||||
then pure success
|
||||
@ -717,7 +717,7 @@ mutual
|
||||
{auto u : Ref UST UState} ->
|
||||
{vars : _} ->
|
||||
UnifyMode -> FC -> Env Term vars ->
|
||||
FC -> NHead vars -> List (Closure vars) ->
|
||||
FC -> NHead vars -> List (Closure vars) ->
|
||||
FC -> NHead vars -> List (Closure vars) ->
|
||||
Core UnifyResult
|
||||
unifyBothApps mode loc env xfc (NLocal xr x xp) [] yfc (NLocal yr y yp) []
|
||||
@ -763,7 +763,7 @@ mutual
|
||||
case !(evalClosure defs c) of
|
||||
NApp _ (NLocal _ _ _) _ => pure $ S !(localsIn cs)
|
||||
_ => localsIn cs
|
||||
|
||||
|
||||
unifyBothApps mode loc env xfc (NMeta xn xi xargs) xargs' yfc fy yargs'
|
||||
= unifyApp False mode loc env xfc (NMeta xn xi xargs) xargs'
|
||||
(NApp yfc fy yargs')
|
||||
@ -794,15 +794,15 @@ mutual
|
||||
{auto u : Ref UST UState} ->
|
||||
{vars : _} ->
|
||||
UnifyMode -> FC -> Env Term vars ->
|
||||
FC -> Name -> Binder (NF vars) ->
|
||||
FC -> Name -> Binder (NF vars) ->
|
||||
(Defs -> Closure vars -> Core (NF vars)) ->
|
||||
FC -> Name -> Binder (NF vars) ->
|
||||
FC -> Name -> Binder (NF vars) ->
|
||||
(Defs -> Closure vars -> Core (NF vars)) ->
|
||||
Core UnifyResult
|
||||
unifyBothBinders mode loc env xfc x (Pi cx ix tx) scx yfc y (Pi cy iy ty) scy
|
||||
= do defs <- get Ctxt
|
||||
if ix /= iy || not (subRig cx cy)
|
||||
then convertError loc env
|
||||
then convertError loc env
|
||||
(NBind xfc x (Pi cx ix tx) scx)
|
||||
(NBind yfc y (Pi cy iy ty) scy)
|
||||
else
|
||||
@ -839,7 +839,7 @@ mutual
|
||||
unifyBothBinders mode loc env xfc x (Lam cx ix tx) scx yfc y (Lam cy iy ty) scy
|
||||
= do defs <- get Ctxt
|
||||
if ix /= iy || not (subRig cx cy)
|
||||
then convertError loc env
|
||||
then convertError loc env
|
||||
(NBind xfc x (Lam cx ix tx) scx)
|
||||
(NBind yfc y (Lam cy iy ty) scy)
|
||||
else
|
||||
@ -861,7 +861,7 @@ mutual
|
||||
pure (union ct cs')
|
||||
|
||||
unifyBothBinders mode loc env xfc x bx scx yfc y by scy
|
||||
= convertError loc env
|
||||
= convertError loc env
|
||||
(NBind xfc x bx scx)
|
||||
(NBind yfc y by scy)
|
||||
|
||||
@ -876,7 +876,7 @@ mutual
|
||||
= do gam <- get Ctxt
|
||||
if tagx == tagy
|
||||
then unifyArgs mode loc env xs ys
|
||||
else convertError loc env
|
||||
else convertError loc env
|
||||
(NDCon xfc x tagx ax xs)
|
||||
(NDCon yfc y tagy ay ys)
|
||||
unifyNoEta mode loc env (NTCon xfc x tagx ax xs) (NTCon yfc y tagy ay ys)
|
||||
@ -888,27 +888,27 @@ mutual
|
||||
-- what's injective...
|
||||
-- then postpone loc env (quote empty env (NTCon x tagx ax xs))
|
||||
-- (quote empty env (NTCon y tagy ay ys))
|
||||
else convertError loc env
|
||||
else convertError loc env
|
||||
(NTCon xfc x tagx ax xs)
|
||||
(NTCon yfc y tagy ay ys)
|
||||
unifyNoEta mode loc env (NDelayed xfc _ x) (NDelayed yfc _ y)
|
||||
= unify mode loc env x y
|
||||
unifyNoEta mode loc env (NDelay xfc _ xty x) (NDelay yfc _ yty y)
|
||||
= unifyArgs mode loc env [xty, x] [yty, y]
|
||||
unifyNoEta mode loc env (NForce xfc x axs) (NForce yfc y ays)
|
||||
unifyNoEta mode loc env (NForce xfc _ x axs) (NForce yfc _ y ays)
|
||||
= do cs <- unify mode loc env x y
|
||||
cs' <- unifyArgs mode loc env axs ays
|
||||
pure (union cs cs')
|
||||
unifyNoEta mode loc env (NApp xfc fx axs) (NApp yfc fy ays)
|
||||
= unifyBothApps mode loc env xfc fx axs yfc fy ays
|
||||
unifyNoEta mode loc env (NApp xfc hd args) y
|
||||
unifyNoEta mode loc env (NApp xfc hd args) y
|
||||
= unifyApp False mode loc env xfc hd args y
|
||||
unifyNoEta mode loc env y (NApp yfc hd args)
|
||||
= unifyApp True mode loc env yfc hd args y
|
||||
-- Only try stripping as patterns as a last resort
|
||||
unifyNoEta mode loc env x (NAs _ _ y) = unifyNoEta mode loc env x y
|
||||
unifyNoEta mode loc env (NAs _ _ x) y = unifyNoEta mode loc env x y
|
||||
unifyNoEta mode loc env x y
|
||||
unifyNoEta mode loc env x y
|
||||
= do defs <- get Ctxt
|
||||
empty <- clearDefs defs
|
||||
unifyIfEq False loc env x y
|
||||
@ -940,7 +940,7 @@ mutual
|
||||
|
||||
export
|
||||
Unify NF where
|
||||
unifyD _ _ mode loc env (NBind xfc x bx scx) (NBind yfc y by scy)
|
||||
unifyD _ _ mode loc env (NBind xfc x bx scx) (NBind yfc y by scy)
|
||||
= unifyBothBinders mode loc env xfc x bx scx yfc y by scy
|
||||
unifyD _ _ mode loc env tmx@(NBind xfc x (Lam cx ix tx) scx) tmy
|
||||
= do defs <- get Ctxt
|
||||
@ -952,10 +952,10 @@ mutual
|
||||
ety <- getEtaType env !(quote empty env tmx)
|
||||
case ety of
|
||||
Just argty =>
|
||||
do etay <- nf defs env
|
||||
do etay <- nf defs env
|
||||
(Bind xfc x (Lam cx ix argty)
|
||||
(App xfc
|
||||
(weaken !(quote empty env tmy))
|
||||
(App xfc
|
||||
(weaken !(quote empty env tmy))
|
||||
(Local xfc Nothing 0 First)))
|
||||
logNF 10 "Expand" env etay
|
||||
unify mode loc env tmx etay
|
||||
@ -970,10 +970,10 @@ mutual
|
||||
ety <- getEtaType env !(quote empty env tmy)
|
||||
case ety of
|
||||
Just argty =>
|
||||
do etax <- nf defs env
|
||||
do etax <- nf defs env
|
||||
(Bind yfc y (Lam cy iy argty)
|
||||
(App yfc
|
||||
(weaken !(quote empty env tmx))
|
||||
(App yfc
|
||||
(weaken !(quote empty env tmx))
|
||||
(Local yfc Nothing 0 First)))
|
||||
logNF 10 "Expand" env etax
|
||||
unify mode loc env etax tmy
|
||||
@ -984,7 +984,7 @@ mutual
|
||||
= unify mode loc env tmx tmy
|
||||
unifyWithLazyD _ _ mode loc env (NDelayed _ r tmx) tmy
|
||||
= do vs <- unify mode loc env tmx tmy
|
||||
pure (record { addLazy = AddForce } vs)
|
||||
pure (record { addLazy = AddForce r } vs)
|
||||
unifyWithLazyD _ _ mode loc env tmx (NDelayed _ r tmy)
|
||||
= do vs <- unify mode loc env tmx tmy
|
||||
pure (record { addLazy = AddDelay r } vs)
|
||||
@ -993,7 +993,7 @@ mutual
|
||||
|
||||
export
|
||||
Unify Term where
|
||||
unifyD _ _ mode loc env x y
|
||||
unifyD _ _ mode loc env x y
|
||||
= do defs <- get Ctxt
|
||||
if x == y
|
||||
then do log 10 $ "Skipped unification (equal already): "
|
||||
@ -1002,7 +1002,7 @@ mutual
|
||||
else do xnf <- nf defs env x
|
||||
ynf <- nf defs env y
|
||||
unify mode loc env xnf ynf
|
||||
unifyWithLazyD _ _ mode loc env x y
|
||||
unifyWithLazyD _ _ mode loc env x y
|
||||
= do defs <- get Ctxt
|
||||
if x == y
|
||||
then do log 10 $ "Skipped unification (equal already): "
|
||||
@ -1014,7 +1014,7 @@ mutual
|
||||
|
||||
export
|
||||
Unify Closure where
|
||||
unifyD _ _ mode loc env x y
|
||||
unifyD _ _ mode loc env x y
|
||||
= do defs <- get Ctxt
|
||||
unify mode loc env !(evalClosure defs x) !(evalClosure defs y)
|
||||
|
||||
@ -1050,20 +1050,20 @@ retry withLazy mode c
|
||||
Nothing => pure success
|
||||
Just Resolved => pure success
|
||||
Just (MkConstraint loc env x y)
|
||||
=> catch (do logTermNF 5 "Retrying" env x
|
||||
=> catch (do logTermNF 5 "Retrying" env x
|
||||
logTermNF 5 "....with" env y
|
||||
cs <- if withLazy
|
||||
then unifyWithLazy mode loc env x y
|
||||
else unify mode loc env x y
|
||||
then unifyWithLazy mode loc env x y
|
||||
else unify mode loc env x y
|
||||
case constraints cs of
|
||||
[] => do deleteConstraint c
|
||||
pure cs
|
||||
_ => pure cs)
|
||||
(\err => throw (WhenUnifying loc env x y err))
|
||||
(\err => throw (WhenUnifying loc env x y err))
|
||||
Just (MkSeqConstraint loc env xs ys)
|
||||
=> do cs <- unifyArgs mode loc env xs ys
|
||||
case constraints cs of
|
||||
[] => do deleteConstraint c
|
||||
[] => do deleteConstraint c
|
||||
pure cs
|
||||
_ => pure cs
|
||||
|
||||
@ -1072,10 +1072,10 @@ delayMeta r (S k) ty (Bind fc n b sc)
|
||||
= Bind fc n b (delayMeta r k (weaken ty) sc)
|
||||
delayMeta r envb ty tm = TDelay (getLoc tm) r ty tm
|
||||
|
||||
forceMeta : Nat -> Term vars -> Term vars
|
||||
forceMeta (S k) (Bind fc n b sc)
|
||||
= Bind fc n b (forceMeta k sc)
|
||||
forceMeta envb tm = TForce (getLoc tm) tm
|
||||
forceMeta : LazyReason -> Nat -> Term vars -> Term vars
|
||||
forceMeta r (S k) (Bind fc n b sc)
|
||||
= Bind fc n b (forceMeta r k sc)
|
||||
forceMeta r envb tm = TForce (getLoc tm) r tm
|
||||
|
||||
-- Retry the given constraint, return True if progress was made
|
||||
retryGuess : {auto c : Ref Ctxt Defs} ->
|
||||
@ -1098,15 +1098,15 @@ retryGuess mode smode (hid, (loc, hname))
|
||||
removeGuess hid
|
||||
pure True)
|
||||
(\err => case err of
|
||||
DeterminingArg _ n i _ _ =>
|
||||
DeterminingArg _ n i _ _ =>
|
||||
do logTerm 5 ("Failed (det " ++ show hname ++ " " ++ show n ++ ")")
|
||||
(type def)
|
||||
setInvertible loc (Resolved i)
|
||||
pure False -- progress made!
|
||||
_ => do logTermNF 5 ("Search failed at " ++ show rig ++ " for " ++ show hname)
|
||||
_ => do logTermNF 5 ("Search failed at " ++ show rig ++ " for " ++ show hname)
|
||||
[] (type def)
|
||||
case smode of
|
||||
LastChance =>
|
||||
LastChance =>
|
||||
throw !(normaliseErr err)
|
||||
_ => pure False) -- Postpone again
|
||||
Guess tm envb [constr] =>
|
||||
@ -1117,8 +1117,8 @@ retryGuess mode smode (hid, (loc, hname))
|
||||
case constraints cs of
|
||||
[] => do tm' <- case addLazy cs of
|
||||
NoLazy => pure tm
|
||||
AddForce => pure $ forceMeta envb tm
|
||||
AddDelay r =>
|
||||
AddForce r => pure $ forceMeta r envb tm
|
||||
AddDelay r =>
|
||||
do ty <- getType [] tm
|
||||
pure $ delayMeta r envb !(getTerm ty) tm
|
||||
let gdef = record { definition = PMDef True [] (STerm tm') (STerm tm') [] } def
|
||||
@ -1129,7 +1129,7 @@ retryGuess mode smode (hid, (loc, hname))
|
||||
newcs => do let gdef = record { definition = Guess tm envb newcs } def
|
||||
addDef (Resolved hid) gdef
|
||||
pure False
|
||||
Guess tm envb constrs =>
|
||||
Guess tm envb constrs =>
|
||||
do let umode = case smode of
|
||||
MatchArgs => InMatch
|
||||
_ => mode
|
||||
@ -1156,10 +1156,10 @@ solveConstraints : {auto c : Ref Ctxt Defs} ->
|
||||
solveConstraints umode smode
|
||||
= do ust <- get UST
|
||||
progress <- traverse (retryGuess umode smode) (toList (guesses ust))
|
||||
when (or (map Delay progress)) $
|
||||
when (or (map Delay progress)) $
|
||||
solveConstraints umode smode
|
||||
|
||||
-- Replace any 'BySearch' with 'Hole', so that we don't keep searching
|
||||
-- Replace any 'BySearch' with 'Hole', so that we don't keep searching
|
||||
-- fruitlessly while elaborating the rest of a source file
|
||||
export
|
||||
giveUpConstraints : {auto c : Ref Ctxt Defs} ->
|
||||
@ -1205,12 +1205,12 @@ checkDots
|
||||
let h = case ndef of
|
||||
Hole _ _ => True
|
||||
_ => False
|
||||
|
||||
|
||||
when (not (isNil (constraints cs)) || holesSolved cs || h) $
|
||||
throw (InternalError "Dot pattern match fail"))
|
||||
(\err => do defs <- get Ctxt
|
||||
throw (BadDotPattern fc env reason
|
||||
!(normaliseHoles defs env x)
|
||||
throw (BadDotPattern fc env reason
|
||||
!(normaliseHoles defs env x)
|
||||
!(normaliseHoles defs env y)))
|
||||
checkConstraint _ = pure ()
|
||||
|
||||
|
@ -21,9 +21,9 @@ data Constraint : Type where
|
||||
-- An unsolved constraint, noting two terms which need to be convertible
|
||||
-- in a particular environment
|
||||
MkConstraint : {vars : _} ->
|
||||
FC ->
|
||||
FC ->
|
||||
(env : Env Term vars) ->
|
||||
(x : Term vars) -> (y : Term vars) ->
|
||||
(x : Term vars) -> (y : Term vars) ->
|
||||
Constraint
|
||||
-- An unsolved sequence of constraints, arising from arguments in an
|
||||
-- application where solving later constraints relies on solving earlier
|
||||
@ -39,13 +39,13 @@ data Constraint : Type where
|
||||
|
||||
export
|
||||
TTC Constraint where
|
||||
toBuf b (MkConstraint {vars} fc env x y)
|
||||
toBuf b (MkConstraint {vars} fc env x y)
|
||||
= do tag 0; toBuf b vars; toBuf b fc; toBuf b env; toBuf b x; toBuf b y
|
||||
toBuf b (MkSeqConstraint {vars} fc env xs ys)
|
||||
= do tag 1; toBuf b vars; toBuf b fc; toBuf b env; toBuf b xs; toBuf b ys
|
||||
toBuf b Resolved = tag 2
|
||||
|
||||
fromBuf b
|
||||
fromBuf b
|
||||
= case !getTag of
|
||||
0 => do vars <- fromBuf b
|
||||
fc <- fromBuf b; env <- fromBuf b
|
||||
@ -273,18 +273,18 @@ addDot : {auto u : Ref UST UState} ->
|
||||
Core ()
|
||||
addDot fc env dotarg x reason y
|
||||
= do ust <- get UST
|
||||
put UST (record { dotConstraints $=
|
||||
((dotarg, reason, MkConstraint fc env x y) ::)
|
||||
put UST (record { dotConstraints $=
|
||||
((dotarg, reason, MkConstraint fc env x y) ::)
|
||||
} ust)
|
||||
|
||||
mkConstantAppArgs : Bool -> FC -> Env Term vars ->
|
||||
mkConstantAppArgs : Bool -> FC -> Env Term vars ->
|
||||
(wkns : List Name) ->
|
||||
List (Term (wkns ++ (vars ++ done)))
|
||||
mkConstantAppArgs lets fc [] wkns = []
|
||||
mkConstantAppArgs {done} {vars = x :: xs} lets fc (b :: env) wkns
|
||||
= let rec = mkConstantAppArgs {done} lets fc env (wkns ++ [x]) in
|
||||
if lets || not (isLet b)
|
||||
then Local fc (Just (isLet b)) (length wkns) (mkVar wkns) ::
|
||||
then Local fc (Just (isLet b)) (length wkns) (mkVar wkns) ::
|
||||
rewrite (appendAssociative wkns [x] (xs ++ done)) in rec
|
||||
else rewrite (appendAssociative wkns [x] (xs ++ done)) in rec
|
||||
where
|
||||
@ -302,19 +302,19 @@ mkConstantAppArgsOthers : Bool -> FC -> Env Term vars ->
|
||||
(wkns : List Name) ->
|
||||
List (Term (wkns ++ (vars ++ done)))
|
||||
mkConstantAppArgsOthers lets fc [] p wkns = []
|
||||
mkConstantAppArgsOthers {done} {vars = x :: xs}
|
||||
mkConstantAppArgsOthers {done} {vars = x :: xs}
|
||||
lets fc (b :: env) SubRefl wkns
|
||||
= rewrite appendAssociative wkns [x] (xs ++ done) in
|
||||
mkConstantAppArgsOthers lets fc env SubRefl (wkns ++ [x])
|
||||
mkConstantAppArgsOthers {done} {vars = x :: xs}
|
||||
mkConstantAppArgsOthers {done} {vars = x :: xs}
|
||||
lets fc (b :: env) (KeepCons p) wkns
|
||||
= rewrite appendAssociative wkns [x] (xs ++ done) in
|
||||
mkConstantAppArgsOthers lets fc env p (wkns ++ [x])
|
||||
mkConstantAppArgsOthers {done} {vars = x :: xs}
|
||||
mkConstantAppArgsOthers {done} {vars = x :: xs}
|
||||
lets fc (b :: env) (DropCons p) wkns
|
||||
= let rec = mkConstantAppArgsOthers {done} lets fc env p (wkns ++ [x]) in
|
||||
if lets || not (isLet b)
|
||||
then Local fc (Just (isLet b)) (length wkns) (mkVar wkns) ::
|
||||
then Local fc (Just (isLet b)) (length wkns) (mkVar wkns) ::
|
||||
rewrite appendAssociative wkns [x] (xs ++ done) in rec
|
||||
else rewrite appendAssociative wkns [x] (xs ++ done) in rec
|
||||
where
|
||||
@ -341,7 +341,7 @@ applyToFull {vars} fc tm env
|
||||
apply fc tm (rewrite sym (appendNilRightNeutral vars) in args)
|
||||
|
||||
export
|
||||
applyToOthers : FC -> Term vars -> Env Term vars ->
|
||||
applyToOthers : FC -> Term vars -> Env Term vars ->
|
||||
SubVars smaller vars -> Term vars
|
||||
applyToOthers {vars} fc tm env sub
|
||||
= let args = reverse (mkConstantAppArgsOthers {done = []} True fc env sub []) in
|
||||
@ -367,7 +367,7 @@ newMetaLets {vars} fc rig env n ty def nocyc lets
|
||||
log 5 $ "Adding new meta " ++ show (n, fc, rig)
|
||||
logTerm 10 ("New meta type " ++ show n) hty
|
||||
defs <- get Ctxt
|
||||
idx <- addDef n hole
|
||||
idx <- addDef n hole
|
||||
addHoleName fc n idx
|
||||
pure (idx, Meta fc n idx envArgs)
|
||||
where
|
||||
@ -379,7 +379,7 @@ export
|
||||
newMeta : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount ->
|
||||
Env Term vars -> Name -> Term vars -> Def ->
|
||||
Env Term vars -> Name -> Term vars -> Def ->
|
||||
Bool ->
|
||||
Core (Int, Term vars)
|
||||
newMeta fc r env n ty def cyc = newMetaLets fc r env n ty def cyc False
|
||||
@ -388,7 +388,7 @@ mkConstant : FC -> Env Term vars -> Term vars -> ClosedTerm
|
||||
mkConstant fc [] tm = tm
|
||||
-- mkConstant {vars = x :: _} fc (Let c val ty :: env) tm
|
||||
-- = mkConstant fc env (Bind fc x (Let c val ty) tm)
|
||||
mkConstant {vars = x :: _} fc (b :: env) tm
|
||||
mkConstant {vars = x :: _} fc (b :: env) tm
|
||||
= let ty = binderType b in
|
||||
mkConstant fc env (Bind fc x (Lam (multiplicity b) Explicit ty) tm)
|
||||
|
||||
@ -398,7 +398,7 @@ mkConstant {vars = x :: _} fc (b :: env) tm
|
||||
export
|
||||
newConstant : {auto u : Ref UST UState} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> RigCount -> Env Term vars ->
|
||||
FC -> RigCount -> Env Term vars ->
|
||||
(tm : Term vars) -> (ty : Term vars) ->
|
||||
(constrs : List Int) ->
|
||||
Core (Term vars)
|
||||
@ -406,7 +406,7 @@ newConstant {vars} fc rig env tm ty constrs
|
||||
= do let def = mkConstant fc env tm
|
||||
let defty = abstractFullEnvType fc env ty
|
||||
cn <- genName "postpone"
|
||||
let guess = newDef fc cn rig [] defty Public
|
||||
let guess = newDef fc cn rig [] defty Public
|
||||
(Guess def (length env) constrs)
|
||||
log 5 $ "Adding new constant " ++ show (cn, fc, rig)
|
||||
logTerm 10 ("New constant type " ++ show cn) defty
|
||||
@ -430,7 +430,7 @@ newSearch {vars} fc rig depth def env n ty
|
||||
= do let hty = abstractEnvType fc env ty
|
||||
let hole = newDef fc n rig [] hty Public (BySearch rig depth def)
|
||||
log 10 $ "Adding new search " ++ show n
|
||||
idx <- addDef n hole
|
||||
idx <- addDef n hole
|
||||
addGuessName fc n idx
|
||||
pure (idx, Meta fc n idx envArgs)
|
||||
where
|
||||
@ -441,8 +441,8 @@ newSearch {vars} fc rig depth def env n ty
|
||||
-- Add a hole which stands for a delayed elaborator
|
||||
export
|
||||
newDelayed : {auto u : Ref UST UState} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> RigCount ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> RigCount ->
|
||||
Env Term vars -> Name ->
|
||||
(ty : Term vars) -> Core (Int, Term vars)
|
||||
newDelayed {vars} fc rig env n ty
|
||||
@ -481,7 +481,7 @@ tryUnify elab1 elab2
|
||||
pure ok
|
||||
|
||||
export
|
||||
handleUnify : {auto c : Ref Ctxt Defs} ->
|
||||
handleUnify : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
Core a -> (Error -> Core a) -> Core a
|
||||
handleUnify elab1 elab2
|
||||
@ -505,7 +505,7 @@ checkDelayedHoles : {auto u : Ref UST UState} ->
|
||||
checkDelayedHoles
|
||||
= do ust <- get UST
|
||||
let hs = toList (delayedHoles ust)
|
||||
if (not (isNil hs))
|
||||
if (not (isNil hs))
|
||||
then do pure (Just (UnsolvedHoles (map snd hs)))
|
||||
else pure Nothing
|
||||
|
||||
@ -528,18 +528,18 @@ checkValidHole (idx, (fc, n))
|
||||
Just ty <- lookupTyExact n (gamma defs)
|
||||
| Nothing => pure ()
|
||||
throw (CantSolveGoal fc [] ty)
|
||||
Guess tm envb (con :: _) =>
|
||||
Guess tm envb (con :: _) =>
|
||||
do ust <- get UST
|
||||
let Just c = lookup con (constraints ust)
|
||||
| Nothing => pure ()
|
||||
case c of
|
||||
MkConstraint fc env x y =>
|
||||
do put UST (record { guesses = empty } ust)
|
||||
do put UST (record { guesses = empty } ust)
|
||||
xnf <- normaliseHoles defs env x
|
||||
ynf <- normaliseHoles defs env y
|
||||
throw (CantSolveEq fc env xnf ynf)
|
||||
MkSeqConstraint fc env (x :: _) (y :: _) =>
|
||||
do put UST (record { guesses = empty } ust)
|
||||
do put UST (record { guesses = empty } ust)
|
||||
xnf <- normaliseHoles defs env x
|
||||
ynf <- normaliseHoles defs env y
|
||||
throw (CantSolveEq fc env xnf ynf)
|
||||
@ -549,7 +549,7 @@ checkValidHole (idx, (fc, n))
|
||||
where
|
||||
checkRef : Name -> Core ()
|
||||
checkRef (PV n f)
|
||||
= throw (GenericMsg fc
|
||||
= throw (GenericMsg fc
|
||||
("Hole cannot depend on an unbound implicit " ++ show n))
|
||||
checkRef _ = pure ()
|
||||
|
||||
@ -569,9 +569,9 @@ checkUserHoles now
|
||||
traverse_ checkValidHole gs
|
||||
hs_map <- getCurrentHoles
|
||||
let hs = toList hs_map
|
||||
let hs' = if any isUserName (map (snd . snd) hs)
|
||||
let hs' = if any isUserName (map (snd . snd) hs)
|
||||
then [] else hs
|
||||
when (now && not (isNil hs')) $
|
||||
when (now && not (isNil hs')) $
|
||||
throw (UnsolvedHoles (map snd (nubBy nameEq hs)))
|
||||
-- Note the hole names, to ensure they are resolved
|
||||
-- by the end of elaborating the current source file
|
||||
@ -600,30 +600,30 @@ dumpHole lvl hole
|
||||
case !(lookupCtxtExact (Resolved hole) (gamma defs)) of
|
||||
Nothing => pure ()
|
||||
Just gdef => case (definition gdef, type gdef) of
|
||||
(Guess tm envb constraints, ty) =>
|
||||
do log lvl $ "!" ++ show !(getFullName (Resolved hole)) ++ " : " ++
|
||||
(Guess tm envb constraints, ty) =>
|
||||
do log lvl $ "!" ++ show !(getFullName (Resolved hole)) ++ " : " ++
|
||||
show !(toFullNames !(normaliseHoles defs [] ty))
|
||||
log lvl $ "\t = " ++ show !(normaliseHoles defs [] tm)
|
||||
++ "\n\twhen"
|
||||
traverse dumpConstraint constraints
|
||||
traverse dumpConstraint constraints
|
||||
pure ()
|
||||
(Hole _ p, ty) =>
|
||||
log lvl $ "?" ++ show (fullname gdef) ++ " : " ++
|
||||
log lvl $ "?" ++ show (fullname gdef) ++ " : " ++
|
||||
show !(normaliseHoles defs [] ty)
|
||||
++ if p then " (ImplBind)" else ""
|
||||
++ if invertible gdef then " (Invertible)" else ""
|
||||
(BySearch _ _ _, ty) =>
|
||||
log lvl $ "Search " ++ show hole ++ " : " ++
|
||||
log lvl $ "Search " ++ show hole ++ " : " ++
|
||||
show !(toFullNames !(normaliseHoles defs [] ty))
|
||||
(PMDef _ args t _ _, ty) =>
|
||||
log 4 $ "Solved: " ++ show hole ++ " : " ++
|
||||
log 4 $ "Solved: " ++ show hole ++ " : " ++
|
||||
show !(normalise defs [] ty) ++
|
||||
" = " ++ show !(normalise defs [] (Ref emptyFC Func (Resolved hole)))
|
||||
(ImpBind, ty) =>
|
||||
log 4 $ "Bound: " ++ show hole ++ " : " ++
|
||||
log 4 $ "Bound: " ++ show hole ++ " : " ++
|
||||
show !(normalise defs [] ty)
|
||||
(Delayed, ty) =>
|
||||
log 4 $ "Delayed elaborator : " ++
|
||||
log 4 $ "Delayed elaborator : " ++
|
||||
show !(normalise defs [] ty)
|
||||
_ => pure ()
|
||||
where
|
||||
@ -643,7 +643,7 @@ dumpHole lvl hole
|
||||
log lvl $ "\t\t" ++ show xs ++ " =?= " ++ show ys
|
||||
|
||||
export
|
||||
dumpConstraints : {auto u : Ref UST UState} ->
|
||||
dumpConstraints : {auto u : Ref UST UState} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
(loglevel : Nat) ->
|
||||
(all : Bool) ->
|
||||
|
@ -51,7 +51,7 @@ mutual
|
||||
public export
|
||||
data Closure : List Name -> Type where
|
||||
MkClosure : (opts : EvalOpts) ->
|
||||
LocalEnv free vars ->
|
||||
LocalEnv free vars ->
|
||||
Env Term free ->
|
||||
Term (vars ++ free) -> Closure free
|
||||
MkNFClosure : NF free -> Closure free
|
||||
@ -64,21 +64,21 @@ mutual
|
||||
NRef : NameType -> Name -> NHead vars
|
||||
NMeta : Name -> Int -> List (Closure vars) -> NHead vars
|
||||
|
||||
-- Values themselves. 'Closure' is an unevaluated thunk, which means
|
||||
-- Values themselves. 'Closure' is an unevaluated thunk, which means
|
||||
-- we can wait until necessary to reduce constructor arguments
|
||||
public export
|
||||
data NF : List Name -> Type where
|
||||
NBind : FC -> (x : Name) -> Binder (NF vars) ->
|
||||
(Defs -> Closure vars -> Core (NF vars)) -> NF vars
|
||||
NApp : FC -> NHead vars -> List (Closure vars) -> NF vars
|
||||
NDCon : FC -> Name -> (tag : Int) -> (arity : Nat) ->
|
||||
NDCon : FC -> Name -> (tag : Int) -> (arity : Nat) ->
|
||||
List (Closure vars) -> NF vars
|
||||
NTCon : FC -> Name -> (tag : Int) -> (arity : Nat) ->
|
||||
NTCon : FC -> Name -> (tag : Int) -> (arity : Nat) ->
|
||||
List (Closure vars) -> NF vars
|
||||
NAs : FC -> NF vars -> NF vars -> NF vars
|
||||
NDelayed : FC -> LazyReason -> NF vars -> NF vars
|
||||
NDelay : FC -> LazyReason -> Closure vars -> Closure vars -> NF vars
|
||||
NForce : FC -> NF vars -> List (Closure vars) -> NF vars
|
||||
NForce : FC -> LazyReason -> NF vars -> List (Closure vars) -> NF vars
|
||||
NPrimVal : FC -> Constant -> NF vars
|
||||
NErased : FC -> NF vars
|
||||
NType : FC -> NF vars
|
||||
@ -92,7 +92,7 @@ getLoc (NTCon fc _ _ _ _) = fc
|
||||
getLoc (NAs fc _ _) = fc
|
||||
getLoc (NDelayed fc _ _) = fc
|
||||
getLoc (NDelay fc _ _ _) = fc
|
||||
getLoc (NForce fc _ _) = fc
|
||||
getLoc (NForce fc _ _ _) = fc
|
||||
getLoc (NPrimVal fc _) = fc
|
||||
getLoc (NErased fc) = fc
|
||||
getLoc (NType fc) = fc
|
||||
|
@ -9,7 +9,7 @@ export
|
||||
record ANameMap a where
|
||||
constructor MkANameMap
|
||||
-- for looking up by exact (completely qualified) names
|
||||
exactNames : NameMap a
|
||||
exactNames : NameMap a
|
||||
-- for looking up by name root or partially qualified (so possibly
|
||||
-- ambiguous) names. This doesn't store machine generated names.
|
||||
hierarchy : StringMap (List (Name, a))
|
||||
@ -62,10 +62,10 @@ addToHier n val hier
|
||||
|
||||
export
|
||||
addName : Name -> a -> ANameMap a -> ANameMap a
|
||||
addName n val (MkANameMap dict hier)
|
||||
addName n val (MkANameMap dict hier)
|
||||
= let dict' = insert n val dict
|
||||
hier' = addToHier n val hier in
|
||||
MkANameMap dict' hier'
|
||||
MkANameMap dict' hier'
|
||||
|
||||
export
|
||||
toList : ANameMap a -> List (Name, a)
|
||||
|
@ -24,7 +24,7 @@ newRawArray size default
|
||||
vm size (MkRaw default)
|
||||
pure (MkIORawArray p)
|
||||
|
||||
||| Write an element at a location in an array.
|
||||
||| Write an element at a location in an array.
|
||||
||| There is *no* bounds checking, hence this is unsafe. Safe interfaces can
|
||||
||| be implemented on top of this, either with a run time or compile time
|
||||
||| check.
|
||||
@ -35,7 +35,7 @@ unsafeWriteArray (MkIORawArray p) i val
|
||||
(Raw (ArrayData elem) -> Int -> Raw elem -> IO ())
|
||||
(MkRaw p) i (MkRaw val)
|
||||
|
||||
||| Read the element at a location in an array.
|
||||
||| Read the element at a location in an array.
|
||||
||| There is *no* bounds checking, hence this is unsafe. Safe interfaces can
|
||||
||| be implemented on top of this, either with a run time or compile time
|
||||
||| check.
|
||||
@ -114,7 +114,7 @@ fromList ns
|
||||
where
|
||||
addToArray : Int -> List (Maybe elem) -> IOArray elem -> IO ()
|
||||
addToArray loc [] arr = pure ()
|
||||
addToArray loc (Nothing :: ns) arr
|
||||
addToArray loc (Nothing :: ns) arr
|
||||
= assert_total (addToArray (loc + 1) ns arr)
|
||||
addToArray loc (Just el :: ns) arr
|
||||
= do unsafeWriteArray (content arr) loc (Just el)
|
||||
|
@ -11,16 +11,16 @@ Key = Int
|
||||
|
||||
private
|
||||
data Tree : Nat -> Type -> Type where
|
||||
Leaf : Key -> v -> Tree Z v
|
||||
Branch2 : Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
Branch3 : Tree n v -> Key -> Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
Leaf : Key -> v -> Tree Z v
|
||||
Branch2 : Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
Branch3 : Tree n v -> Key -> Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
|
||||
branch4 :
|
||||
Tree n v -> Key ->
|
||||
Tree n v -> Key ->
|
||||
Tree n v -> Key ->
|
||||
Tree n v ->
|
||||
Tree (S (S n)) v
|
||||
Tree (S (S n)) v
|
||||
branch4 a b c d e f g =
|
||||
Branch2 (Branch2 a b c) d (Branch2 e f g)
|
||||
|
||||
@ -196,9 +196,9 @@ treeToList = treeToList' []
|
||||
where
|
||||
treeToList' : List (Key, v) -> Tree n v -> List (Key, v)
|
||||
treeToList' rest (Leaf k v) = (k, v) :: rest
|
||||
treeToList' rest (Branch2 t1 _ t2)
|
||||
treeToList' rest (Branch2 t1 _ t2)
|
||||
= treeToList' (treeToList' rest t2) t1
|
||||
treeToList' rest (Branch3 t1 _ t2 _ t3)
|
||||
treeToList' rest (Branch3 t1 _ t2 _ t3)
|
||||
= treeToList' (treeToList' (treeToList' rest t3) t2) t1
|
||||
|
||||
export
|
||||
@ -258,7 +258,7 @@ export
|
||||
values : IntMap v -> List v
|
||||
values = map snd . toList
|
||||
|
||||
treeMap : (a -> b) -> Tree n a -> Tree n b
|
||||
treeMap : (a -> b) -> Tree n a -> Tree n b
|
||||
treeMap f (Leaf k v) = Leaf k (f v)
|
||||
treeMap f (Branch2 t1 k t2) = Branch2 (treeMap f t1) k (treeMap f t2)
|
||||
treeMap f (Branch3 t1 k1 t2 k2 t3)
|
||||
|
@ -13,16 +13,16 @@ Key = Name
|
||||
|
||||
private
|
||||
data Tree : Nat -> Type -> Type where
|
||||
Leaf : Key -> v -> Tree Z v
|
||||
Branch2 : Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
Branch3 : Tree n v -> Key -> Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
Leaf : Key -> v -> Tree Z v
|
||||
Branch2 : Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
Branch3 : Tree n v -> Key -> Tree n v -> Key -> Tree n v -> Tree (S n) v
|
||||
|
||||
branch4 :
|
||||
Tree n v -> Key ->
|
||||
Tree n v -> Key ->
|
||||
Tree n v -> Key ->
|
||||
Tree n v ->
|
||||
Tree (S (S n)) v
|
||||
Tree (S (S n)) v
|
||||
branch4 a b c d e f g =
|
||||
Branch2 (Branch2 a b c) d (Branch2 e f g)
|
||||
|
||||
@ -198,9 +198,9 @@ treeToList = treeToList' []
|
||||
where
|
||||
treeToList' : List (Key, v) -> Tree n v -> List (Key, v)
|
||||
treeToList' rest (Leaf k v) = (k, v) :: rest
|
||||
treeToList' rest (Branch2 t1 _ t2)
|
||||
treeToList' rest (Branch2 t1 _ t2)
|
||||
= treeToList' (treeToList' rest t2) t1
|
||||
treeToList' rest (Branch3 t1 _ t2 _ t3)
|
||||
treeToList' rest (Branch3 t1 _ t2 _ t3)
|
||||
= treeToList' (treeToList' (treeToList' rest t3) t2) t1
|
||||
|
||||
export
|
||||
@ -260,7 +260,7 @@ export
|
||||
values : NameMap v -> List v
|
||||
values = map snd . toList
|
||||
|
||||
treeMap : (a -> b) -> Tree n a -> Tree n b
|
||||
treeMap : (a -> b) -> Tree n a -> Tree n b
|
||||
treeMap f (Leaf k v) = Leaf k (f v)
|
||||
treeMap f (Branch2 t1 k t2) = Branch2 (treeMap f t1) k (treeMap f t2)
|
||||
treeMap f (Branch3 t1 k1 t2 k2 t3)
|
||||
|
BIN
src/Idris/.Desugar.idr.swp
Normal file
BIN
src/Idris/.Desugar.idr.swp
Normal file
Binary file not shown.
@ -59,7 +59,7 @@ extendAs old as newsyn
|
||||
put Syn (record { infixes $= mergeLeft (infixes newsyn),
|
||||
prefixes $= mergeLeft (prefixes newsyn),
|
||||
ifaces $= mergeAs old as (ifaces newsyn),
|
||||
bracketholes $= ((bracketholes newsyn) ++) }
|
||||
bracketholes $= ((bracketholes newsyn) ++) }
|
||||
syn)
|
||||
|
||||
mkPrec : Fixity -> Nat -> OpPrec
|
||||
@ -74,7 +74,7 @@ toTokList (POp fc opn l r)
|
||||
= do syn <- get Syn
|
||||
let op = nameRoot opn
|
||||
case lookup op (infixes syn) of
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
let ops = unpack opChars in
|
||||
if any (\x => x `elem` ops) (unpack op)
|
||||
then throw (GenericMsg fc $ "Unknown operator '" ++ op ++ "'")
|
||||
@ -107,36 +107,36 @@ mutual
|
||||
{auto m : Ref MD Metadata} ->
|
||||
Side -> List Name -> PTerm -> Core RawImp
|
||||
desugar side ps (PRef fc x) = pure $ IVar fc x
|
||||
desugar side ps (PPi fc rig p mn argTy retTy)
|
||||
desugar side ps (PPi fc rig p mn argTy retTy)
|
||||
= let ps' = maybe ps (:: ps) mn in
|
||||
pure $ IPi fc rig p mn !(desugar side ps argTy)
|
||||
pure $ IPi fc rig p mn !(desugar side ps argTy)
|
||||
!(desugar side ps' retTy)
|
||||
desugar side ps (PLam fc rig p (PRef _ n@(UN _)) argTy scope)
|
||||
= pure $ ILam fc rig p (Just n) !(desugar side ps argTy)
|
||||
desugar side ps (PLam fc rig p (PRef _ n@(UN _)) argTy scope)
|
||||
= pure $ ILam fc rig p (Just n) !(desugar side ps argTy)
|
||||
!(desugar side (n :: ps) scope)
|
||||
desugar side ps (PLam fc rig p (PRef _ n@(MN _ _)) argTy scope)
|
||||
= pure $ ILam fc rig p (Just n) !(desugar side ps argTy)
|
||||
desugar side ps (PLam fc rig p (PRef _ n@(MN _ _)) argTy scope)
|
||||
= pure $ ILam fc rig p (Just n) !(desugar side ps argTy)
|
||||
!(desugar side (n :: ps) scope)
|
||||
desugar side ps (PLam fc rig p (PImplicit _) argTy scope)
|
||||
= pure $ ILam fc rig p Nothing !(desugar side ps argTy)
|
||||
desugar side ps (PLam fc rig p (PImplicit _) argTy scope)
|
||||
= pure $ ILam fc rig p Nothing !(desugar side ps argTy)
|
||||
!(desugar side ps scope)
|
||||
desugar side ps (PLam fc rig p pat argTy scope)
|
||||
= pure $ ILam fc rig p (Just (MN "lamc" 0)) !(desugar side ps argTy) $
|
||||
ICase fc (IVar fc (MN "lamc" 0)) (Implicit fc False)
|
||||
[!(desugarClause ps True (MkPatClause fc pat scope []))]
|
||||
desugar side ps (PLet fc rig (PRef _ n) nTy nVal scope [])
|
||||
= pure $ ILet fc rig n !(desugar side ps nTy) !(desugar side ps nVal)
|
||||
= pure $ ILet fc rig n !(desugar side ps nTy) !(desugar side ps nVal)
|
||||
!(desugar side (n :: ps) scope)
|
||||
desugar side ps (PLet fc rig pat nTy nVal scope alts)
|
||||
desugar side ps (PLet fc rig pat nTy nVal scope alts)
|
||||
= pure $ ICase fc !(desugar side ps nVal) !(desugar side ps nTy)
|
||||
!(traverse (desugarClause ps True)
|
||||
!(traverse (desugarClause ps True)
|
||||
(MkPatClause fc pat scope [] :: alts))
|
||||
desugar side ps (PCase fc x xs)
|
||||
= pure $ ICase fc !(desugar side ps x)
|
||||
desugar side ps (PCase fc x xs)
|
||||
= pure $ ICase fc !(desugar side ps x)
|
||||
(Implicit fc False)
|
||||
!(traverse (desugarClause ps True) xs)
|
||||
desugar side ps (PLocal fc xs scope)
|
||||
= pure $ ILocal fc (concat !(traverse (desugarDecl ps) xs))
|
||||
desugar side ps (PLocal fc xs scope)
|
||||
= pure $ ILocal fc (concat !(traverse (desugarDecl ps) xs))
|
||||
!(desugar side (definedIn xs ++ ps) scope)
|
||||
desugar side ps (PApp pfc (PUpdate fc fs) rec)
|
||||
= pure $ IUpdate pfc !(traverse (desugarUpdate side ps) fs)
|
||||
@ -144,11 +144,11 @@ mutual
|
||||
desugar side ps (PUpdate fc fs)
|
||||
= desugar side ps (PLam fc RigW Explicit (PRef fc (MN "rec" 0)) (PImplicit fc)
|
||||
(PApp fc (PUpdate fc fs) (PRef fc (MN "rec" 0))))
|
||||
desugar side ps (PApp fc x y)
|
||||
desugar side ps (PApp fc x y)
|
||||
= pure $ IApp fc !(desugar side ps x) !(desugar side ps y)
|
||||
desugar side ps (PWithApp fc x y)
|
||||
desugar side ps (PWithApp fc x y)
|
||||
= pure $ IWithApp fc !(desugar side ps x) !(desugar side ps y)
|
||||
desugar side ps (PImplicitApp fc x argn y)
|
||||
desugar side ps (PImplicitApp fc x argn y)
|
||||
= pure $ IImplicitApp fc !(desugar side ps x) argn !(desugar side ps y)
|
||||
desugar side ps (PDelayed fc r ty)
|
||||
= pure $ IDelayed fc r !(desugar side ps ty)
|
||||
@ -163,18 +163,18 @@ mutual
|
||||
[apply (IVar fc (UN "===")) [l', r'],
|
||||
apply (IVar fc (UN "~=~")) [l', r']]
|
||||
desugar side ps (PBracketed fc e) = desugar side ps e
|
||||
desugar side ps (POp fc op l r)
|
||||
desugar side ps (POp fc op l r)
|
||||
= do ts <- toTokList (POp fc op l r)
|
||||
desugarTree side ps !(parseOps ts)
|
||||
desugar side ps (PPrefixOp fc op arg)
|
||||
desugar side ps (PPrefixOp fc op arg)
|
||||
= do ts <- toTokList (PPrefixOp fc op arg)
|
||||
desugarTree side ps !(parseOps ts)
|
||||
desugar side ps (PSectionL fc op arg)
|
||||
desugar side ps (PSectionL fc op arg)
|
||||
= do syn <- get Syn
|
||||
-- It might actually be a prefix argument rather than a section
|
||||
-- so check that first, otherwise desugar as a lambda
|
||||
case lookup (nameRoot op) (prefixes syn) of
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
desugar side ps (PLam fc RigW Explicit (PRef fc (MN "arg" 0)) (PImplicit fc)
|
||||
(POp fc op (PRef fc (MN "arg" 0)) arg))
|
||||
Just prec => desugar side ps (PPrefixOp fc op arg)
|
||||
@ -188,36 +188,36 @@ mutual
|
||||
pure $ IAlternative fc (UniqueDefault (IPrimVal fc (BI x)))
|
||||
[IPrimVal fc (BI x),
|
||||
IPrimVal fc (I (fromInteger x))]
|
||||
Just fi => pure $ IApp fc (IVar fc fi)
|
||||
Just fi => pure $ IApp fc (IVar fc fi)
|
||||
(IPrimVal fc (BI x))
|
||||
desugar side ps (PPrimVal fc (Str x))
|
||||
= case !fromStringName of
|
||||
Nothing =>
|
||||
pure $ IPrimVal fc (Str x)
|
||||
Just f => pure $ IApp fc (IVar fc f)
|
||||
Just f => pure $ IApp fc (IVar fc f)
|
||||
(IPrimVal fc (Str x))
|
||||
desugar side ps (PPrimVal fc (Ch x))
|
||||
= case !fromCharName of
|
||||
Nothing =>
|
||||
pure $ IPrimVal fc (Ch x)
|
||||
Just f => pure $ IApp fc (IVar fc f)
|
||||
Just f => pure $ IApp fc (IVar fc f)
|
||||
(IPrimVal fc (Ch x))
|
||||
desugar side ps (PPrimVal fc x) = pure $ IPrimVal fc x
|
||||
desugar side ps (PQuote fc x)
|
||||
desugar side ps (PQuote fc x)
|
||||
= throw (GenericMsg fc "Reflection not implemeted yet")
|
||||
-- = pure $ IQuote fc !(desugar side ps x)
|
||||
desugar side ps (PUnquote fc x)
|
||||
desugar side ps (PUnquote fc x)
|
||||
= throw (GenericMsg fc "Reflection not implemeted yet")
|
||||
-- = pure $ IUnquote fc !(desugar side ps x)
|
||||
desugar side ps (PHole fc br holename)
|
||||
desugar side ps (PHole fc br holename)
|
||||
= do when br $
|
||||
do syn <- get Syn
|
||||
put Syn (record { bracketholes $= ((UN holename) ::) } syn)
|
||||
pure $ IHole fc holename
|
||||
desugar side ps (PType fc) = pure $ IType fc
|
||||
desugar side ps (PAs fc vname pattern)
|
||||
desugar side ps (PAs fc vname pattern)
|
||||
= pure $ IAs fc UseRight vname !(desugar side ps pattern)
|
||||
desugar side ps (PDotted fc x)
|
||||
desugar side ps (PDotted fc x)
|
||||
= pure $ IMustUnify fc "User dotted" !(desugar side ps x)
|
||||
desugar side ps (PImplicit fc) = pure $ Implicit fc True
|
||||
desugar side ps (PInfer fc) = pure $ Implicit fc False
|
||||
@ -225,35 +225,35 @@ mutual
|
||||
= expandDo side ps fc block
|
||||
desugar side ps (PList fc args)
|
||||
= expandList side ps fc args
|
||||
desugar side ps (PPair fc l r)
|
||||
desugar side ps (PPair fc l r)
|
||||
= do l' <- desugar side ps l
|
||||
r' <- desugar side ps r
|
||||
let pval = apply (IVar fc (UN "MkPair")) [l', r']
|
||||
pure $ IAlternative fc (UniqueDefault pval)
|
||||
[apply (IVar fc (UN "Pair")) [l', r'], pval]
|
||||
desugar side ps (PDPair fc (PRef nfc (UN n)) (PImplicit _) r)
|
||||
desugar side ps (PDPair fc (PRef nfc (UN n)) (PImplicit _) r)
|
||||
= do r' <- desugar side ps r
|
||||
let pval = apply (IVar fc (UN "MkDPair")) [IVar nfc (UN n), r']
|
||||
pure $ IAlternative fc (UniqueDefault pval)
|
||||
[apply (IVar fc (UN "DPair"))
|
||||
[Implicit nfc False,
|
||||
[apply (IVar fc (UN "DPair"))
|
||||
[Implicit nfc False,
|
||||
ILam nfc RigW Explicit (Just (UN n)) (Implicit nfc False) r'],
|
||||
pval]
|
||||
desugar side ps (PDPair fc (PRef nfc (UN n)) ty r)
|
||||
desugar side ps (PDPair fc (PRef nfc (UN n)) ty r)
|
||||
= do ty' <- desugar side ps ty
|
||||
r' <- desugar side ps r
|
||||
pure $ apply (IVar fc (UN "DPair"))
|
||||
[ty',
|
||||
[ty',
|
||||
ILam nfc RigW Explicit (Just (UN n)) ty' r']
|
||||
desugar side ps (PDPair fc l (PImplicit _) r)
|
||||
desugar side ps (PDPair fc l (PImplicit _) r)
|
||||
= do l' <- desugar side ps l
|
||||
r' <- desugar side ps r
|
||||
pure $ apply (IVar fc (UN "MkDPair")) [l', r']
|
||||
desugar side ps (PDPair fc l ty r)
|
||||
desugar side ps (PDPair fc l ty r)
|
||||
= throw (GenericMsg fc "Invalid dependent pair type")
|
||||
desugar side ps (PUnit fc)
|
||||
desugar side ps (PUnit fc)
|
||||
= pure $ IAlternative fc (UniqueDefault (IVar fc (UN "MkUnit")))
|
||||
[IVar fc (UN "Unit"),
|
||||
[IVar fc (UN "Unit"),
|
||||
IVar fc (UN "MkUnit")]
|
||||
desugar side ps (PIfThenElse fc x t e)
|
||||
= pure $ ICase fc !(desugar side ps x) (Implicit fc False)
|
||||
@ -273,12 +273,12 @@ mutual
|
||||
desugar side ps (PRange fc start next end)
|
||||
= case next of
|
||||
Nothing =>
|
||||
desugar side ps (PApp fc
|
||||
desugar side ps (PApp fc
|
||||
(PApp fc (PRef fc (UN "rangeFromTo"))
|
||||
start) end)
|
||||
Just n =>
|
||||
desugar side ps (PApp fc
|
||||
(PApp fc
|
||||
desugar side ps (PApp fc
|
||||
(PApp fc
|
||||
(PApp fc (PRef fc (UN "rangeFromThenTo"))
|
||||
start) n) end)
|
||||
desugar side ps (PRangeStream fc start next)
|
||||
@ -305,7 +305,7 @@ mutual
|
||||
Side -> List Name -> FC -> List PTerm -> Core RawImp
|
||||
expandList side ps fc [] = pure (IVar fc (UN "Nil"))
|
||||
expandList side ps fc (x :: xs)
|
||||
= pure $ apply (IVar fc (UN "::"))
|
||||
= pure $ apply (IVar fc (UN "::"))
|
||||
[!(desugar side ps x), !(expandList side ps fc xs)]
|
||||
|
||||
expandDo : {auto s : Ref Syn SyntaxInfo} ->
|
||||
@ -315,21 +315,21 @@ mutual
|
||||
Side -> List Name -> FC -> List PDo -> Core RawImp
|
||||
expandDo side ps fc [] = throw (GenericMsg fc "Do block cannot be empty")
|
||||
expandDo side ps _ [DoExp fc tm] = desugar side ps tm
|
||||
expandDo side ps fc [e]
|
||||
= throw (GenericMsg (getLoc e)
|
||||
"Last statement in do block must be an expression")
|
||||
expandDo side ps fc [e]
|
||||
= throw (GenericMsg (getLoc e)
|
||||
"Last statement in do block must be an expression")
|
||||
expandDo side ps topfc (DoExp fc tm :: rest)
|
||||
= do tm' <- desugar side ps tm
|
||||
rest' <- expandDo side ps topfc rest
|
||||
gam <- get Ctxt
|
||||
pure $ IApp fc (IApp fc (IVar fc (UN ">>=")) tm')
|
||||
(ILam fc RigW Explicit Nothing
|
||||
(ILam fc RigW Explicit Nothing
|
||||
(Implicit fc False) rest')
|
||||
expandDo side ps topfc (DoBind fc n tm :: rest)
|
||||
= do tm' <- desugar side ps tm
|
||||
rest' <- expandDo side ps topfc rest
|
||||
pure $ IApp fc (IApp fc (IVar fc (UN ">>=")) tm')
|
||||
(ILam fc RigW Explicit (Just n)
|
||||
(ILam fc RigW Explicit (Just n)
|
||||
(Implicit fc False) rest')
|
||||
expandDo side ps topfc (DoBindPat fc pat exp alts :: rest)
|
||||
= do pat' <- desugar LHS ps pat
|
||||
@ -339,24 +339,24 @@ mutual
|
||||
let ps' = newps ++ ps
|
||||
rest' <- expandDo side ps' topfc rest
|
||||
pure $ IApp fc (IApp fc (IVar fc (UN ">>=")) exp')
|
||||
(ILam fc RigW Explicit (Just (MN "_" 0))
|
||||
(ILam fc RigW Explicit (Just (MN "_" 0))
|
||||
(Implicit fc False)
|
||||
(ICase fc (IVar fc (MN "_" 0))
|
||||
(Implicit fc False)
|
||||
(PatClause fc bpat rest'
|
||||
(PatClause fc bpat rest'
|
||||
:: alts')))
|
||||
expandDo side ps topfc (DoLet fc n rig tm :: rest)
|
||||
expandDo side ps topfc (DoLet fc n rig tm :: rest)
|
||||
= do tm' <- desugar side ps tm
|
||||
rest' <- expandDo side ps topfc rest
|
||||
pure $ ILet fc rig n (Implicit fc False) tm' rest'
|
||||
expandDo side ps topfc (DoLetPat fc pat tm alts :: rest)
|
||||
expandDo side ps topfc (DoLetPat fc pat tm alts :: rest)
|
||||
= do pat' <- desugar LHS ps pat
|
||||
(newps, bpat) <- bindNames False pat'
|
||||
tm' <- desugar side ps tm
|
||||
alts' <- traverse (desugarClause ps True) alts
|
||||
let ps' = newps ++ ps
|
||||
rest' <- expandDo side ps' topfc rest
|
||||
pure $ ICase fc tm' (Implicit fc False)
|
||||
pure $ ICase fc tm' (Implicit fc False)
|
||||
(PatClause fc bpat rest'
|
||||
:: alts')
|
||||
expandDo side ps topfc (DoLetLocal fc decls :: rest)
|
||||
@ -402,7 +402,7 @@ mutual
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
List Name -> PTypeDecl -> Core ImpTy
|
||||
desugarType ps (MkPTy fc n ty)
|
||||
desugarType ps (MkPTy fc n ty)
|
||||
= pure $ MkImpTy fc n !(bindTypeNames ps !(desugar AnyExpr ps ty))
|
||||
|
||||
desugarClause : {auto s : Ref Syn SyntaxInfo} ->
|
||||
@ -414,7 +414,7 @@ mutual
|
||||
= do ws <- traverse (desugarDecl ps) wheres
|
||||
(bound, blhs) <- bindNames arg !(desugar LHS ps lhs)
|
||||
rhs' <- desugar AnyExpr (bound ++ ps) rhs
|
||||
pure $ PatClause fc blhs
|
||||
pure $ PatClause fc blhs
|
||||
(case ws of
|
||||
[] => rhs'
|
||||
_ => ILocal fc (concat ws) rhs')
|
||||
@ -432,24 +432,24 @@ mutual
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
List Name -> PDataDecl -> Core ImpData
|
||||
desugarData ps (MkPData fc n tycon opts datacons)
|
||||
desugarData ps (MkPData fc n tycon opts datacons)
|
||||
= pure $ MkImpData fc n !(bindTypeNames ps !(desugar AnyExpr ps tycon))
|
||||
opts
|
||||
!(traverse (desugarType ps) datacons)
|
||||
desugarData ps (MkPLater fc n tycon)
|
||||
desugarData ps (MkPLater fc n tycon)
|
||||
= pure $ MkImpLater fc n !(bindTypeNames ps !(desugar AnyExpr ps tycon))
|
||||
|
||||
desugarField : {auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
List Name -> PField ->
|
||||
List Name -> PField ->
|
||||
Core IField
|
||||
desugarField ps (MkField fc rig p n ty)
|
||||
= pure (MkIField fc rig p n !(bindTypeNames ps !(desugar AnyExpr ps ty)))
|
||||
|
||||
-- Get the declaration to process on each pass of a mutual block
|
||||
-- Essentially: types on the first pass
|
||||
-- Essentially: types on the first pass
|
||||
-- i.e. type constructors of data declarations
|
||||
-- function types
|
||||
-- interfaces (in full, since it includes function types)
|
||||
@ -479,8 +479,8 @@ mutual
|
||||
getDecl AsDef (PFixity _ _ _ _) = Nothing
|
||||
getDecl AsDef (PDirective _ _) = Nothing
|
||||
getDecl AsDef d = Just d
|
||||
|
||||
getDecl p (PParameters fc ps pds)
|
||||
|
||||
getDecl p (PParameters fc ps pds)
|
||||
= Just (PParameters fc ps (mapMaybe (getDecl p) pds))
|
||||
|
||||
getDecl Single d = Just d
|
||||
@ -493,14 +493,14 @@ mutual
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
List Name -> PDecl -> Core (List ImpDecl)
|
||||
desugarDecl ps (PClaim fc rig vis opts ty)
|
||||
desugarDecl ps (PClaim fc rig vis opts ty)
|
||||
= pure [IClaim fc rig vis opts !(desugarType ps ty)]
|
||||
desugarDecl ps (PDef fc clauses)
|
||||
desugarDecl ps (PDef fc clauses)
|
||||
-- The clauses won't necessarily all be from the same function, so split
|
||||
-- after desugaring, by function name, using collectDefs from RawImp
|
||||
= do cs <- traverse (desugarClause ps False) clauses
|
||||
defs <- traverse toIDef cs
|
||||
pure (collectDefs defs)
|
||||
pure (collectDefs defs)
|
||||
where
|
||||
getFn : RawImp -> Core Name
|
||||
getFn (IVar _ n) = pure n
|
||||
@ -509,14 +509,14 @@ mutual
|
||||
getFn tm = throw (InternalError (show tm ++ " is not a function application"))
|
||||
|
||||
toIDef : ImpClause -> Core ImpDecl
|
||||
toIDef (PatClause fc lhs rhs)
|
||||
toIDef (PatClause fc lhs rhs)
|
||||
= pure $ IDef fc !(getFn lhs) [PatClause fc lhs rhs]
|
||||
toIDef (WithClause fc lhs rhs cs)
|
||||
toIDef (WithClause fc lhs rhs cs)
|
||||
= pure $ IDef fc !(getFn lhs) [WithClause fc lhs rhs cs]
|
||||
toIDef (ImpossibleClause fc lhs)
|
||||
toIDef (ImpossibleClause fc lhs)
|
||||
= pure $ IDef fc !(getFn lhs) [ImpossibleClause fc lhs]
|
||||
|
||||
desugarDecl ps (PData fc vis ddecl)
|
||||
desugarDecl ps (PData fc vis ddecl)
|
||||
= pure [IData fc vis !(desugarData ps ddecl)]
|
||||
desugarDecl ps (PParameters fc params pds)
|
||||
= do pds' <- traverse (desugarDecl (ps ++ map fst params)) pds
|
||||
@ -540,19 +540,19 @@ mutual
|
||||
pure (fst ntm, tm')) params
|
||||
-- Look for bindable names in all the constraints and parameters
|
||||
let mnames = map dropNS (definedIn body)
|
||||
let bnames = concatMap (findBindableNames True
|
||||
(ps ++ mnames ++ map fst params) [])
|
||||
let bnames = concatMap (findBindableNames True
|
||||
(ps ++ mnames ++ map fst params) [])
|
||||
(map snd cons') ++
|
||||
concatMap (findBindableNames True
|
||||
(ps ++ mnames ++ map fst params) [])
|
||||
concatMap (findBindableNames True
|
||||
(ps ++ mnames ++ map fst params) [])
|
||||
(map snd params')
|
||||
let paramsb = map (\ (n, tm) => (n, doBind bnames tm)) params'
|
||||
let consb = map (\ (n, tm) => (n, doBind bnames tm)) cons'
|
||||
|
||||
body' <- traverse (desugarDecl (ps ++ mnames ++ map fst params)) body
|
||||
pure [IPragma (\c, nest, env =>
|
||||
pure [IPragma (\c, nest, env =>
|
||||
elabInterface fc vis env nest consb
|
||||
tn paramsb det conname
|
||||
tn paramsb det conname
|
||||
(concat body'))]
|
||||
where
|
||||
-- Turns pairs in the constraints to individual constraints. This
|
||||
@ -585,15 +585,15 @@ mutual
|
||||
pure (Just (concat b'))) body
|
||||
pure [IPragma (\c, nest, env =>
|
||||
elabImplementation fc vis pass env nest isb consb
|
||||
tn paramsb impname
|
||||
tn paramsb impname
|
||||
body')]
|
||||
desugarDecl ps (PRecord fc vis tn params conname fields)
|
||||
= do params' <- traverse (\ ntm => do tm' <- desugar AnyExpr ps (snd ntm)
|
||||
pure (fst ntm, tm')) params
|
||||
let fnames = map fname fields
|
||||
-- Look for bindable names in the parameters
|
||||
let bnames = concatMap (findBindableNames True
|
||||
(ps ++ fnames ++ map fst params) [])
|
||||
let bnames = concatMap (findBindableNames True
|
||||
(ps ++ fnames ++ map fst params) [])
|
||||
(map snd params')
|
||||
fields' <- traverse (desugarField (ps ++ fnames ++ map fst params))
|
||||
fields
|
||||
@ -607,11 +607,11 @@ mutual
|
||||
where
|
||||
fname : PField -> Name
|
||||
fname (MkField _ _ _ n _) = n
|
||||
desugarDecl ps (PFixity fc Prefix prec (UN n))
|
||||
desugarDecl ps (PFixity fc Prefix prec (UN n))
|
||||
= do syn <- get Syn
|
||||
put Syn (record { prefixes $= insert n prec } syn)
|
||||
pure []
|
||||
desugarDecl ps (PFixity fc fix prec (UN n))
|
||||
desugarDecl ps (PFixity fc fix prec (UN n))
|
||||
= do syn <- get Syn
|
||||
put Syn (record { infixes $= insert n (fix, prec) } syn)
|
||||
pure []
|
||||
@ -624,7 +624,11 @@ mutual
|
||||
desugarDecl ps (PNamespace fc ns decls)
|
||||
= do ds <- traverse (desugarDecl ps) decls
|
||||
pure [INamespace fc False ns (concat ds)]
|
||||
desugarDecl ps (PDirective fc d)
|
||||
desugarDecl ps (PTransform fc lhs rhs)
|
||||
= do (bound, blhs) <- bindNames False !(desugar LHS ps lhs)
|
||||
rhs' <- desugar AnyExpr (bound ++ ps) rhs
|
||||
pure [ITransform fc blhs rhs']
|
||||
desugarDecl ps (PDirective fc d)
|
||||
= case d of
|
||||
Hide n => pure [IPragma (\c, nest, env => hide fc n)]
|
||||
Logging i => pure [ILog i]
|
||||
|
@ -26,12 +26,12 @@ import Data.NameMap
|
||||
%default covering
|
||||
|
||||
mkImpl : FC -> Name -> List RawImp -> Name
|
||||
mkImpl fc n ps
|
||||
mkImpl fc n ps
|
||||
= DN (show n ++ " implementation at " ++ show fc)
|
||||
(UN ("__Impl_" ++ show n ++ "_" ++
|
||||
showSep "_" (map show ps)))
|
||||
|
||||
bindConstraints : FC -> PiInfo ->
|
||||
bindConstraints : FC -> PiInfo ->
|
||||
List (Maybe Name, RawImp) -> RawImp -> RawImp
|
||||
bindConstraints fc p [] ty = ty
|
||||
bindConstraints fc p ((n, ty) :: rest) sc
|
||||
@ -43,7 +43,7 @@ bindImpls fc ((n, r, ty) :: rest) sc
|
||||
= IPi fc r Implicit (Just n) ty (bindImpls fc rest sc)
|
||||
|
||||
addDefaults : FC -> Name -> List Name -> List (Name, List ImpClause) ->
|
||||
List ImpDecl ->
|
||||
List ImpDecl ->
|
||||
(List ImpDecl, List Name) -- Updated body, list of missing methods
|
||||
addDefaults fc impName allms defs body
|
||||
= let missing = dropGot allms body in
|
||||
@ -51,25 +51,25 @@ addDefaults fc impName allms defs body
|
||||
where
|
||||
-- Given the list of missing names, if any are among the default definitions,
|
||||
-- add them to the body
|
||||
extendBody : List Name -> List Name -> List ImpDecl ->
|
||||
extendBody : List Name -> List Name -> List ImpDecl ->
|
||||
(List ImpDecl, List Name)
|
||||
extendBody ms [] body = (body, ms)
|
||||
extendBody ms (n :: ns) body
|
||||
= case lookup n defs of
|
||||
Nothing => extendBody (n :: ms) ns body
|
||||
Just cs =>
|
||||
Just cs =>
|
||||
-- If any method names appear in the clauses, they should
|
||||
-- be applied to the constraint name __con because they
|
||||
-- are going to be referring to the present implementation.
|
||||
-- That is, default method implementations could depend on
|
||||
-- other methods.
|
||||
-- (See test idris2/interface014 for an example!)
|
||||
let mupdates
|
||||
let mupdates
|
||||
= map (\n => (n, IImplicitApp fc (IVar fc n)
|
||||
(Just (UN "__con"))
|
||||
(IVar fc impName))) allms
|
||||
cs' = map (substNamesClause [] mupdates) cs in
|
||||
extendBody ms ns
|
||||
extendBody ms ns
|
||||
(IDef fc n (map (substLocClause fc) cs') :: body)
|
||||
|
||||
-- Find which names are missing from the body
|
||||
@ -119,7 +119,7 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
Just conty <- lookupTyExact (iconstructor cdata) (gamma defs)
|
||||
| Nothing => throw (UndefinedName fc (iconstructor cdata))
|
||||
|
||||
let impsp = nub (concatMap findIBinds ps ++
|
||||
let impsp = nub (concatMap findIBinds ps ++
|
||||
concatMap findIBinds (map snd cons))
|
||||
|
||||
logTerm 3 ("Found interface " ++ show cn) ity
|
||||
@ -138,7 +138,7 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
-- Don't make it a hint if it's a named implementation
|
||||
let opts = maybe [Inline, Hint True] (const [Inline]) impln
|
||||
|
||||
let initTy = bindImpls fc is $ bindConstraints fc AutoImplicit cons
|
||||
let initTy = bindImpls fc is $ bindConstraints fc AutoImplicit cons
|
||||
(apply (IVar fc iname) ps)
|
||||
let paramBinds = findBindableNames True vars [] initTy
|
||||
let impTy = doBind paramBinds initTy
|
||||
@ -147,7 +147,7 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
log 5 $ "Implementation type: " ++ show impTy
|
||||
|
||||
when (typePass pass) $ processDecl [] nest env impTyDecl
|
||||
|
||||
|
||||
-- If the body is empty, we're done for now (just declaring that
|
||||
-- the implementation exists and define it later)
|
||||
when (defPass pass) $ maybe (pure ())
|
||||
@ -160,7 +160,7 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
|
||||
-- 1.5. Lookup default definitions and add them to to body
|
||||
let (body, missing)
|
||||
= addDefaults fc impName (map (dropNS . fst) (methods cdata))
|
||||
= addDefaults fc impName (map (dropNS . fst) (methods cdata))
|
||||
(defaults cdata) body_in
|
||||
|
||||
log 5 $ "Added defaults: body is " ++ show body
|
||||
@ -169,15 +169,15 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
-- 2. Elaborate top level function types for this interface
|
||||
defs <- get Ctxt
|
||||
fns <- topMethTypes [] impName methImps impsp (params cdata)
|
||||
(map fst (methods cdata))
|
||||
(map fst (methods cdata))
|
||||
(methods cdata)
|
||||
traverse (processDecl [] nest env) (map mkTopMethDecl fns)
|
||||
|
||||
-- 3. Build the record for the implementation
|
||||
let mtops = map (Basics.fst . snd) fns
|
||||
let con = iconstructor cdata
|
||||
let ilhs = impsApply (IVar fc impName)
|
||||
(map (\x => (x, IBindVar fc (show x)))
|
||||
let ilhs = impsApply (IVar fc impName)
|
||||
(map (\x => (x, IBindVar fc (show x)))
|
||||
(map fst methImps))
|
||||
-- RHS is the constructor applied to a search for the necessary
|
||||
-- parent constraints, then the method implementations
|
||||
@ -230,12 +230,12 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
|
||||
impsApply : RawImp -> List (Name, RawImp) -> RawImp
|
||||
impsApply fn [] = fn
|
||||
impsApply fn ((n, arg) :: ns)
|
||||
impsApply fn ((n, arg) :: ns)
|
||||
= impsApply (IImplicitApp fc fn (Just n) arg) ns
|
||||
|
||||
mkLam : List (Name, RigCount, PiInfo) -> RawImp -> RawImp
|
||||
mkLam [] tm = tm
|
||||
mkLam ((x, c, p) :: xs) tm
|
||||
mkLam ((x, c, p) :: xs) tm
|
||||
= ILam fc c p (Just x) (Implicit fc False) (mkLam xs tm)
|
||||
|
||||
applyTo : FC -> RawImp -> List (Name, RigCount, PiInfo) -> RawImp
|
||||
@ -248,10 +248,10 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
= applyTo fc (IImplicitApp fc tm (Just x) (IVar fc x)) xs
|
||||
|
||||
-- When applying the method in the field for the record, eta expand
|
||||
-- the expected arguments based on the field type, so that implicits get
|
||||
-- the expected arguments based on the field type, so that implicits get
|
||||
-- inserted in the right place
|
||||
mkMethField : List (Name, RigCount, RawImp) ->
|
||||
List (Name, List (Name, RigCount, PiInfo)) ->
|
||||
List (Name, List (Name, RigCount, PiInfo)) ->
|
||||
(Name, Name, List (String, String), RigCount, RawImp) -> RawImp
|
||||
mkMethField methImps fldTys (topn, n, upds, c, ty)
|
||||
= let argns = map applyUpdate (maybe [] id (lookup (dropNS topn) fldTys))
|
||||
@ -264,22 +264,22 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
(map (\n => (n, IVar fc (UN (show n)))) imps))
|
||||
where
|
||||
applyUpdate : (Name, RigCount, PiInfo) -> (Name, RigCount, PiInfo)
|
||||
applyUpdate (UN n, c, p)
|
||||
applyUpdate (UN n, c, p)
|
||||
= maybe (UN n, c, p) (\n' => (UN n', c, p)) (lookup n upds)
|
||||
applyUpdate t = t
|
||||
|
||||
methName : Name -> Name
|
||||
methName (NS _ n) = methName n
|
||||
methName n
|
||||
methName n
|
||||
= DN (show n) (UN (show n ++ "_" ++ show iname ++ "_" ++
|
||||
maybe "" show impln ++ "_" ++
|
||||
showSep "_" (map show ps)))
|
||||
|
||||
|
||||
applyCon : Name -> Name -> Core (Name, RawImp)
|
||||
applyCon impl n
|
||||
applyCon impl n
|
||||
= do mn <- inCurrentNS (methName n)
|
||||
pure (dropNS n, IVar fc mn)
|
||||
|
||||
|
||||
bindImps : List (Name, RigCount, RawImp) -> RawImp -> RawImp
|
||||
bindImps [] ty = ty
|
||||
bindImps ((n, c, t) :: ts) ty
|
||||
@ -291,9 +291,9 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
topMethType : List (Name, RawImp) ->
|
||||
Name -> List (Name, RigCount, RawImp) ->
|
||||
List String -> List Name -> List Name ->
|
||||
(Name, RigCount, (Bool, RawImp)) ->
|
||||
(Name, RigCount, (Bool, RawImp)) ->
|
||||
Core ((Name, Name, List (String, String), RigCount, RawImp),
|
||||
List (Name, RawImp))
|
||||
List (Name, RawImp))
|
||||
topMethType methupds impName methImps impsp pnames allmeths (mn, c, (d, mty_in))
|
||||
= do -- Get the specialised type by applying the method to the
|
||||
-- parameters
|
||||
@ -319,20 +319,20 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
log 3 $ "Method " ++ show mn ++ " ==> " ++
|
||||
show n ++ " : " ++ show mty
|
||||
log 5 $ "Updates " ++ show methupds
|
||||
log 5 $ "From " ++ show mbase
|
||||
log 3 $ "Name updates " ++ show upds
|
||||
log 3 $ "Param names: " ++ show pnames
|
||||
log 5 $ "From " ++ show mbase
|
||||
log 3 $ "Name updates " ++ show upds
|
||||
log 3 $ "Param names: " ++ show pnames
|
||||
log 10 $ "Used names " ++ show ibound
|
||||
let ibinds = map fst methImps
|
||||
let methupds' = if isNil ibinds then []
|
||||
else [(n, impsApply (IVar fc n)
|
||||
(map (\x => (x, IBindVar fc (show x))) ibinds))]
|
||||
pure ((mn, n, upds, c, mty), methupds')
|
||||
|
||||
|
||||
topMethTypes : List (Name, RawImp) ->
|
||||
Name -> List (Name, RigCount, RawImp) ->
|
||||
List String -> List Name -> List Name ->
|
||||
List (Name, RigCount, (Bool, RawImp)) ->
|
||||
List (Name, RigCount, (Bool, RawImp)) ->
|
||||
Core (List (Name, Name, List (String, String), RigCount, RawImp))
|
||||
topMethTypes upds impName methImps impsp pnames allmeths [] = pure []
|
||||
topMethTypes upds impName methImps impsp pnames allmeths (m :: ms)
|
||||
@ -341,7 +341,7 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
pure (m' :: ms')
|
||||
|
||||
mkTopMethDecl : (Name, Name, List (String, String), RigCount, RawImp) -> ImpDecl
|
||||
mkTopMethDecl (mn, n, upds, c, mty)
|
||||
mkTopMethDecl (mn, n, upds, c, mty)
|
||||
= IClaim fc c vis [] (MkImpTy fc n mty)
|
||||
|
||||
-- Given the method type (result of topMethType) return the mapping from
|
||||
@ -352,8 +352,8 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
findMethName : List (Name, Name) -> FC -> Name -> Core Name
|
||||
findMethName ns fc n
|
||||
= case lookup n ns of
|
||||
Nothing => throw (GenericMsg fc
|
||||
(show n ++ " is not a method of " ++
|
||||
Nothing => throw (GenericMsg fc
|
||||
(show n ++ " is not a method of " ++
|
||||
show iname))
|
||||
Just n' => pure n'
|
||||
|
||||
@ -370,9 +370,9 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
updateApp ns tm
|
||||
= throw (GenericMsg (getFC tm) "Invalid method definition")
|
||||
|
||||
updateClause : List (Name, Name) -> ImpClause ->
|
||||
updateClause : List (Name, Name) -> ImpClause ->
|
||||
Core ImpClause
|
||||
updateClause ns (PatClause fc lhs rhs)
|
||||
updateClause ns (PatClause fc lhs rhs)
|
||||
= do lhs' <- updateApp ns lhs
|
||||
pure (PatClause fc lhs' rhs)
|
||||
updateClause ns (WithClause fc lhs wval cs)
|
||||
@ -384,10 +384,10 @@ elabImplementation {vars} fc vis pass env nest is cons iname ps impln mbody
|
||||
pure (ImpossibleClause fc lhs')
|
||||
|
||||
updateBody : List (Name, Name) -> ImpDecl -> Core ImpDecl
|
||||
updateBody ns (IDef fc n cs)
|
||||
updateBody ns (IDef fc n cs)
|
||||
= do cs' <- traverse (updateClause ns) cs
|
||||
n' <- findMethName ns fc n
|
||||
pure (IDef fc n' cs')
|
||||
updateBody ns _
|
||||
= throw (GenericMsg fc
|
||||
updateBody ns _
|
||||
= throw (GenericMsg fc
|
||||
"Implementation body can only contain definitions")
|
||||
|
@ -26,25 +26,25 @@ import Data.ANameMap
|
||||
|
||||
mkDataTy : FC -> List (Name, RawImp) -> RawImp
|
||||
mkDataTy fc [] = IType fc
|
||||
mkDataTy fc ((n, ty) :: ps)
|
||||
mkDataTy fc ((n, ty) :: ps)
|
||||
= IPi fc RigW Explicit (Just n) ty (mkDataTy fc ps)
|
||||
|
||||
mkIfaceData : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Visibility -> Env Term vars ->
|
||||
List (Maybe Name, RigCount, RawImp) ->
|
||||
Name -> Name -> List (Name, RawImp) ->
|
||||
Name -> Name -> List (Name, RawImp) ->
|
||||
List Name -> List (Name, RigCount, RawImp) -> Core ImpDecl
|
||||
mkIfaceData {vars} fc vis env constraints n conName ps dets meths
|
||||
= let opts = if isNil dets
|
||||
= let opts = if isNil dets
|
||||
then [NoHints]
|
||||
else [NoHints, SearchBy dets]
|
||||
else [NoHints, SearchBy dets]
|
||||
retty = apply (IVar fc n) (map (IVar fc) (map fst ps))
|
||||
conty = mkTy Implicit (map jname ps) $
|
||||
mkTy Explicit (map bhere constraints ++ map bname meths) retty
|
||||
con = MkImpTy fc conName !(bindTypeNames (map fst ps ++ map fst meths ++ vars) conty) in
|
||||
pure $ IData fc vis (MkImpData fc n
|
||||
pure $ IData fc vis (MkImpData fc n
|
||||
!(bindTypeNames (map fst ps ++ map fst meths ++ vars)
|
||||
(mkDataTy fc ps))
|
||||
(mkDataTy fc ps))
|
||||
opts [con])
|
||||
where
|
||||
jname : (Name, RawImp) -> (Maybe Name, RigCount, RawImp)
|
||||
@ -56,7 +56,7 @@ mkIfaceData {vars} fc vis env constraints n conName ps dets meths
|
||||
bhere : (Maybe Name, RigCount, RawImp) -> (Maybe Name, RigCount, RawImp)
|
||||
bhere (n, c, t) = (n, c, IBindHere (getFC t) (PI Rig0) t)
|
||||
|
||||
mkTy : PiInfo ->
|
||||
mkTy : PiInfo ->
|
||||
List (Maybe Name, RigCount, RawImp) -> RawImp -> RawImp
|
||||
mkTy imp [] ret = ret
|
||||
mkTy imp ((n, c, argty) :: args) ret
|
||||
@ -68,7 +68,7 @@ getMethDecl : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars -> NestedNames vars ->
|
||||
(params : List (Name, RawImp)) ->
|
||||
(mnames : List Name) ->
|
||||
(FC, RigCount, List FnOpt, n, (Bool, RawImp)) ->
|
||||
(FC, RigCount, List FnOpt, n, (Bool, RawImp)) ->
|
||||
Core (n, RigCount, RawImp)
|
||||
getMethDecl {vars} env nest params mnames (fc, c, opts, n, (d, ty))
|
||||
= do ty_imp <- bindTypeNames (map fst params ++ mnames ++ vars) ty
|
||||
@ -83,7 +83,7 @@ getMethDecl {vars} env nest params mnames (fc, c, opts, n, (d, ty))
|
||||
then stripParams ps ret
|
||||
else IPi fc r p mn arg (stripParams ps ret)
|
||||
stripParams ps ty = ty
|
||||
|
||||
|
||||
-- bind the auto implicit for the interface - put it after all the other
|
||||
-- implicits
|
||||
bindIFace : FC -> RawImp -> RawImp -> RawImp
|
||||
@ -93,42 +93,42 @@ bindIFace _ ity (IPi fc rig AutoImplicit n ty sc)
|
||||
= IPi fc rig AutoImplicit n ty (bindIFace fc ity sc)
|
||||
bindIFace fc ity sc = IPi fc RigW AutoImplicit (Just (UN "__con")) ity sc
|
||||
|
||||
-- Get the top level function for implementing a method
|
||||
-- Get the top level function for implementing a method
|
||||
getMethToplevel : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars -> Visibility ->
|
||||
Env Term vars -> Visibility ->
|
||||
Name -> Name ->
|
||||
(constraints : List (Maybe Name)) ->
|
||||
(allmeths : List Name) ->
|
||||
(params : List Name) ->
|
||||
(FC, RigCount, List FnOpt, Name, (Bool, RawImp)) ->
|
||||
(FC, RigCount, List FnOpt, Name, (Bool, RawImp)) ->
|
||||
Core (List ImpDecl)
|
||||
getMethToplevel {vars} env vis iname cname constraints allmeths params
|
||||
getMethToplevel {vars} env vis iname cname constraints allmeths params
|
||||
(fc, c, opts, n, (d, ty))
|
||||
= do let ity = apply (IVar fc iname) (map (IVar fc) params)
|
||||
= do let ity = apply (IVar fc iname) (map (IVar fc) params)
|
||||
-- Make the constraint application explicit for any method names
|
||||
-- which appear in other method types
|
||||
let ty_constr = substNames vars (map applyCon allmeths) ty
|
||||
ty_imp <- bindTypeNames vars (bindIFace fc ity ty_constr)
|
||||
cn <- inCurrentNS n
|
||||
let tydecl = IClaim fc c vis (if d then [Inline, Invertible]
|
||||
else [Inline])
|
||||
(MkImpTy fc cn ty_imp)
|
||||
else [Inline])
|
||||
(MkImpTy fc cn ty_imp)
|
||||
let conapp = apply (IVar fc cname)
|
||||
(map (const (Implicit fc True)) constraints ++
|
||||
map (IBindVar fc) (map bindName allmeths))
|
||||
let argns = getExplicitArgs 0 ty
|
||||
-- eta expand the RHS so that we put implicits in the right place
|
||||
let fnclause = PatClause fc (IImplicitApp fc (IVar fc cn)
|
||||
let fnclause = PatClause fc (IImplicitApp fc (IVar fc cn)
|
||||
(Just (UN "__con"))
|
||||
conapp)
|
||||
(mkLam argns
|
||||
(mkLam argns
|
||||
(apply (IVar fc (methName n))
|
||||
(map (IVar fc) argns)))
|
||||
let fndef = IDef fc cn [fnclause]
|
||||
pure [tydecl, fndef]
|
||||
where
|
||||
applyCon : Name -> (Name, RawImp)
|
||||
applyCon n = (n, IImplicitApp fc (IVar fc n)
|
||||
applyCon n = (n, IImplicitApp fc (IVar fc n)
|
||||
(Just (UN "__con")) (IVar fc (UN "__con")))
|
||||
|
||||
getExplicitArgs : Int -> RawImp -> List Name
|
||||
@ -139,7 +139,7 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params
|
||||
|
||||
mkLam : List Name -> RawImp -> RawImp
|
||||
mkLam [] tm = tm
|
||||
mkLam (x :: xs) tm
|
||||
mkLam (x :: xs) tm
|
||||
= ILam fc RigW Explicit (Just x) (Implicit fc False) (mkLam xs tm)
|
||||
|
||||
bindName : Name -> String
|
||||
@ -153,7 +153,7 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params
|
||||
-- Get the function for chasing a constraint. This is one of the
|
||||
-- arguments to the record, appearing before the method arguments.
|
||||
getConstraintHint : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Env Term vars -> Visibility ->
|
||||
FC -> Env Term vars -> Visibility ->
|
||||
Name -> Name ->
|
||||
(constraints : List Name) ->
|
||||
(allmeths : List Name) ->
|
||||
@ -162,14 +162,14 @@ getConstraintHint : {auto c : Ref Ctxt Defs} ->
|
||||
getConstraintHint {vars} fc env vis iname cname constraints meths params (cn, con)
|
||||
= do let ity = apply (IVar fc iname) (map (IVar fc) params)
|
||||
let fty = IPi fc RigW Explicit Nothing ity con
|
||||
ty_imp <- bindTypeNames (meths ++ vars) fty
|
||||
ty_imp <- bindTypeNames (meths ++ vars) fty
|
||||
let hintname = DN ("Constraint " ++ show con)
|
||||
(UN ("__" ++ show iname ++ "_" ++ show con))
|
||||
let tydecl = IClaim fc RigW vis [Inline, Hint False]
|
||||
let tydecl = IClaim fc RigW vis [Inline, Hint False]
|
||||
(MkImpTy fc hintname ty_imp)
|
||||
let conapp = apply (IVar fc cname)
|
||||
(map (IBindVar fc) (map bindName constraints) ++
|
||||
map (const (Implicit fc True)) meths)
|
||||
map (const (Implicit fc True)) meths)
|
||||
let fnclause = PatClause fc (IApp fc (IVar fc hintname) conapp)
|
||||
(IVar fc (constName cn))
|
||||
let fndef = IDef fc hintname [fnclause]
|
||||
@ -193,9 +193,9 @@ getDefault (IDef fc n cs) = Just (fc, [], n, cs)
|
||||
getDefault _ = Nothing
|
||||
|
||||
mkCon : FC -> Name -> Name
|
||||
mkCon loc (NS ns (UN n))
|
||||
mkCon loc (NS ns (UN n))
|
||||
= NS ns (DN (n ++ " at " ++ show loc) (UN ("__mk" ++ n)))
|
||||
mkCon loc n
|
||||
mkCon loc n
|
||||
= DN (show n ++ " at " ++ show loc) (UN ("__mk" ++ show n))
|
||||
|
||||
updateIfaceSyn : {auto s : Ref Syn SyntaxInfo} ->
|
||||
@ -212,7 +212,7 @@ elabInterface : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
FC -> Visibility ->
|
||||
FC -> Visibility ->
|
||||
Env Term vars -> NestedNames vars ->
|
||||
(constraints : List (Maybe Name, RawImp)) ->
|
||||
Name ->
|
||||
@ -239,13 +239,13 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
|
||||
ns_meths <- traverse (\mt => do n <- inCurrentNS (fst mt)
|
||||
pure (n, snd mt)) meth_decls
|
||||
ns_iname <- inCurrentNS fullIName
|
||||
updateIfaceSyn ns_iname conName
|
||||
updateIfaceSyn ns_iname conName
|
||||
(map fst params) (map snd constraints)
|
||||
ns_meths ds
|
||||
where
|
||||
nameCons : Int -> List (Maybe Name, RawImp) -> List (Name, RawImp)
|
||||
nameCons i [] = []
|
||||
nameCons i ((_, ty) :: rest)
|
||||
nameCons i ((_, ty) :: rest)
|
||||
= (UN ("__con" ++ show i), ty) :: nameCons (i + 1) rest
|
||||
|
||||
-- Elaborate the data declaration part of the interface
|
||||
@ -257,13 +257,13 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
|
||||
-- signatures and constraint hints
|
||||
meths <- traverse (getMethDecl env nest params meth_names) meth_sigs
|
||||
log 5 $ "Method declarations: " ++ show meths
|
||||
|
||||
consts <- traverse (getMethDecl env nest params meth_names)
|
||||
|
||||
consts <- traverse (getMethDecl env nest params meth_names)
|
||||
(map (\c => (fc, Rig1, [], c))
|
||||
(map notData constraints))
|
||||
log 5 $ "Constraints: " ++ show consts
|
||||
|
||||
dt <- mkIfaceData fc vis env consts iname conName params
|
||||
dt <- mkIfaceData fc vis env consts iname conName params
|
||||
dets meths
|
||||
log 10 $ "Methods: " ++ show meths
|
||||
log 5 $ "Making interface data type " ++ show dt
|
||||
@ -273,7 +273,7 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
|
||||
notData : (n, t) -> (n, (Bool, t))
|
||||
notData (x, y) = (x, (False, y))
|
||||
|
||||
elabMethods : (conName : Name) -> List Name ->
|
||||
elabMethods : (conName : Name) -> List Name ->
|
||||
List (FC, RigCount, List FnOpt, Name, (Bool, RawImp)) ->
|
||||
Core ()
|
||||
elabMethods conName meth_names meth_sigs
|
||||
@ -293,14 +293,14 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
|
||||
-- we know it's okay, since we'll need to re-elaborate it for each
|
||||
-- instance, to specialise it
|
||||
elabDefault : List (Name, RigCount, (Bool, RawImp)) ->
|
||||
(FC, List FnOpt, Name, List ImpClause) ->
|
||||
(FC, List FnOpt, Name, List ImpClause) ->
|
||||
Core (Name, List ImpClause)
|
||||
elabDefault tydecls (fc, opts, n, cs)
|
||||
elabDefault tydecls (fc, opts, n, cs)
|
||||
= do -- orig <- branch
|
||||
let dn_in = UN ("Default implementation of " ++ show n)
|
||||
dn <- inCurrentNS dn_in
|
||||
|
||||
(rig, dty) <-
|
||||
(rig, dty) <-
|
||||
the (Core (RigCount, RawImp)) $
|
||||
case lookup n tydecls of
|
||||
Just (r, (_, t)) => pure (r, t)
|
||||
@ -310,9 +310,9 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
|
||||
|
||||
-- Substitute the method names with their top level function
|
||||
-- name, so they don't get implicitly bound in the name
|
||||
methNameMap <- traverse (\n =>
|
||||
methNameMap <- traverse (\n =>
|
||||
do cn <- inCurrentNS n
|
||||
pure (n, applyParams (IVar fc cn)
|
||||
pure (n, applyParams (IVar fc cn)
|
||||
(map fst params)))
|
||||
(map fst tydecls)
|
||||
let dty = substNames vars methNameMap dty
|
||||
@ -320,7 +320,7 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
|
||||
dty_imp <- bindTypeNames (map fst tydecls ++ vars)
|
||||
(bindIFace fc ity dty)
|
||||
log 5 $ "Default method " ++ show dn ++ " : " ++ show dty_imp
|
||||
let dtydecl = IClaim fc rig vis [] (MkImpTy fc dn dty_imp)
|
||||
let dtydecl = IClaim fc rig vis [] (MkImpTy fc dn dty_imp)
|
||||
processDecl [] nest env dtydecl
|
||||
|
||||
let cs' = map (changeName dn) cs
|
||||
@ -348,9 +348,9 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
|
||||
changeNameTerm dn tm = tm
|
||||
|
||||
changeName : Name -> ImpClause -> ImpClause
|
||||
changeName dn (PatClause fc lhs rhs)
|
||||
changeName dn (PatClause fc lhs rhs)
|
||||
= PatClause fc (changeNameTerm dn lhs) rhs
|
||||
changeName dn (ImpossibleClause fc lhs)
|
||||
changeName dn (ImpossibleClause fc lhs)
|
||||
= ImpossibleClause fc (changeNameTerm dn lhs)
|
||||
|
||||
elabConstraintHints : (conName : Name) -> List Name ->
|
||||
|
@ -16,7 +16,7 @@ import Parser.Support
|
||||
pshow : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
Env Term vars -> Term vars -> Core String
|
||||
pshow env tm
|
||||
pshow env tm
|
||||
= do defs <- get Ctxt
|
||||
itm <- resugar env !(normaliseHoles defs env tm)
|
||||
pure (show itm)
|
||||
@ -24,7 +24,7 @@ pshow env tm
|
||||
pshowNoNorm : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
Env Term vars -> Term vars -> Core String
|
||||
pshowNoNorm env tm
|
||||
pshowNoNorm env tm
|
||||
= do defs <- get Ctxt
|
||||
itm <- resugar env tm
|
||||
pure (show itm)
|
||||
@ -37,10 +37,10 @@ perror (Fatal err) = perror err
|
||||
perror (CantConvert _ env l r)
|
||||
= pure $ "Mismatch between:\n\t" ++ !(pshow env l) ++ "\nand\n\t" ++ !(pshow env r)
|
||||
perror (CantSolveEq _ env l r)
|
||||
= pure $ "Can't solve constraint between:\n\t" ++ !(pshow env l) ++
|
||||
= pure $ "Can't solve constraint between:\n\t" ++ !(pshow env l) ++
|
||||
"\nand\n\t" ++ !(pshow env r)
|
||||
perror (PatternVariableUnifies _ env n tm)
|
||||
= pure $ "Pattern variable " ++ showPVar n ++
|
||||
= pure $ "Pattern variable " ++ showPVar n ++
|
||||
" unifies with:\n\t" ++ !(pshow env tm)
|
||||
where
|
||||
showPVar : Name -> String
|
||||
@ -72,7 +72,7 @@ perror (NotCovering fc n (MissingCases cs))
|
||||
showSep "\n\t" !(traverse (pshow []) cs)
|
||||
perror (NotCovering fc n (NonCoveringCall ns))
|
||||
= pure $ show n ++ " is not covering:\n\t" ++
|
||||
"Calls non covering function"
|
||||
"Calls non covering function"
|
||||
++ case ns of
|
||||
[fn] => " " ++ show fn
|
||||
_ => "s: " ++ showSep ", " (map show ns)
|
||||
@ -94,11 +94,11 @@ perror (LinearMisuse fc n exp ctx)
|
||||
showRel Rig1 = "relevant"
|
||||
showRel RigW = "non-linear"
|
||||
perror (BorrowPartial fc env tm arg)
|
||||
= pure $ !(pshow env tm) ++
|
||||
" borrows argument " ++ !(pshow env arg) ++
|
||||
= pure $ !(pshow env tm) ++
|
||||
" borrows argument " ++ !(pshow env arg) ++
|
||||
" so must be fully applied"
|
||||
perror (BorrowPartialType fc env tm)
|
||||
= pure $ !(pshow env tm) ++
|
||||
= pure $ !(pshow env tm) ++
|
||||
" borrows, so must return a concrete type"
|
||||
perror (AmbiguousName fc ns) = pure $ "Ambiguous name " ++ show ns
|
||||
perror (AmbiguousElab fc env ts)
|
||||
@ -137,12 +137,12 @@ perror (NotRecordField fc fld (Just ty))
|
||||
perror (NotRecordType fc ty)
|
||||
= pure $ show ty ++ " is not a record type"
|
||||
perror (IncompatibleFieldUpdate fc flds)
|
||||
= pure $ "Field update " ++ showSep "->" flds ++
|
||||
= pure $ "Field update " ++ showSep "->" flds ++
|
||||
" not compatible with other updates"
|
||||
perror (InvalidImplicits _ env [Just n] tm)
|
||||
= pure $ show n ++ " is not a valid implicit argument in " ++ !(pshow env tm)
|
||||
perror (InvalidImplicits _ env ns tm)
|
||||
= pure $ showSep ", " (map show ns) ++
|
||||
= pure $ showSep ", " (map show ns) ++
|
||||
" are not valid implicit arguments in " ++ !(pshow env tm)
|
||||
perror (TryWithImplicits _ env imps)
|
||||
= pure $ "Need to bind implicits "
|
||||
@ -159,19 +159,19 @@ perror (CantSolveGoal _ env g)
|
||||
where
|
||||
-- For display, we don't want to see the full top level type; just the
|
||||
-- return type
|
||||
dropPis : Env Term vars -> Term vars ->
|
||||
dropPis : Env Term vars -> Term vars ->
|
||||
(ns ** (Env Term ns, Term ns))
|
||||
dropPis env (Bind _ n b@(Pi _ _ _) sc) = dropPis (b :: env) sc
|
||||
dropPis env (Bind _ n b@(Pi _ _ _) sc) = dropPis (b :: env) sc
|
||||
dropPis env tm = (_ ** (env, tm))
|
||||
|
||||
perror (DeterminingArg _ n i env g)
|
||||
= pure $ "Can't find an implementation for " ++ !(pshow env g) ++ "\n" ++
|
||||
"since I can't infer a value for argument " ++ show n
|
||||
perror (UnsolvedHoles hs)
|
||||
perror (UnsolvedHoles hs)
|
||||
= pure $ "Unsolved holes:\n" ++ showHoles hs
|
||||
where
|
||||
showHoles [] = ""
|
||||
showHoles ((fc, n) :: hs) = show n ++ " introduced at " ++ show fc ++ "\n"
|
||||
showHoles ((fc, n) :: hs) = show n ++ " introduced at " ++ show fc ++ "\n"
|
||||
++ showHoles hs
|
||||
perror (CantInferArgType _ env n h ty)
|
||||
= pure $ "Can't infer type for argument " ++ show n ++ "\n" ++
|
||||
@ -180,7 +180,7 @@ perror (SolvedNamedHole _ env h tm)
|
||||
= pure $ "Named hole " ++ show h ++ " has been solved by unification\n"
|
||||
++ "Result: " ++ !(pshow env tm)
|
||||
perror (VisibilityError fc vx x vy y)
|
||||
= pure $ show vx ++ " " ++ sugarName x ++
|
||||
= pure $ show vx ++ " " ++ sugarName x ++
|
||||
" cannot refer to " ++ show vy ++ " " ++ sugarName y
|
||||
perror (NonLinearPattern _ n) = pure $ "Non linear pattern " ++ sugarName n
|
||||
perror (BadPattern _ n) = pure $ "Pattern not allowed here: " ++ show n
|
||||
@ -189,7 +189,7 @@ perror (AlreadyDefined _ n) = pure $ show n ++ " is already defined"
|
||||
perror (NotFunctionType _ env tm)
|
||||
= pure $ !(pshow env tm) ++ " is not a function type"
|
||||
perror (RewriteNoChange _ env rule ty)
|
||||
= pure $ "Rewriting by " ++ !(pshow env rule) ++
|
||||
= pure $ "Rewriting by " ++ !(pshow env rule) ++
|
||||
" did not change type " ++ !(pshow env ty)
|
||||
perror (NotRewriteRule fc env rule)
|
||||
= pure $ !(pshow env rule) ++ " is not a rewrite rule type"
|
||||
@ -199,24 +199,24 @@ perror (CaseCompile _ n DifferingTypes)
|
||||
= pure $ "Patterns for " ++ show n ++ " require matching on different types"
|
||||
perror (CaseCompile _ n UnknownType)
|
||||
= pure $ "Can't infer type to match in " ++ show n
|
||||
perror (CaseCompile fc n (MatchErased (_ ** (env, tm))))
|
||||
= pure $ "Attempt to match on erased argument " ++ !(pshow env tm) ++
|
||||
perror (CaseCompile fc n (MatchErased (_ ** (env, tm))))
|
||||
= pure $ "Attempt to match on erased argument " ++ !(pshow env tm) ++
|
||||
" in " ++ show n
|
||||
perror (BadDotPattern _ env reason x y)
|
||||
= pure $ "Can't match on " ++ !(pshow env x) ++
|
||||
(if reason /= "" then " (" ++ reason ++ ")" else "") ++ "\n" ++
|
||||
"It elaborates to: " ++ !(pshow env y)
|
||||
perror (MatchTooSpecific _ env tm)
|
||||
= pure $ "Can't match on " ++ !(pshow env tm) ++
|
||||
= pure $ "Can't match on " ++ !(pshow env tm) ++
|
||||
" as it has a polymorphic type"
|
||||
perror (BadImplicit _ str)
|
||||
perror (BadImplicit _ str)
|
||||
= pure $ "Can't infer type for unbound implicit name " ++ str ++ "\n" ++
|
||||
"Try making it a bound implicit."
|
||||
perror (BadRunElab _ env script)
|
||||
= pure $ "Bad elaborator script " ++ !(pshow env script)
|
||||
perror (GenericMsg _ str) = pure str
|
||||
perror (TTCError msg) = pure $ "Error in TTC file: " ++ show msg
|
||||
perror (FileErr fname err)
|
||||
perror (FileErr fname err)
|
||||
= pure $ "File error in " ++ fname ++ ": " ++ show err
|
||||
perror (ParseFail _ err)
|
||||
= pure $ show err
|
||||
@ -231,22 +231,22 @@ perror ForceNeeded = pure "Internal error when resolving implicit laziness"
|
||||
perror (InternalError str) = pure $ "INTERNAL ERROR: " ++ str
|
||||
|
||||
perror (InType fc n err)
|
||||
= pure $ "While processing type of " ++ sugarName !(getFullName n) ++
|
||||
= pure $ "While processing type of " ++ sugarName !(getFullName n) ++
|
||||
" at " ++ show fc ++ ":\n" ++ !(perror err)
|
||||
perror (InCon fc n err)
|
||||
= pure $ "While processing constructor " ++ sugarName !(getFullName n) ++
|
||||
= pure $ "While processing constructor " ++ sugarName !(getFullName n) ++
|
||||
" at " ++ show fc ++ ":\n" ++ !(perror err)
|
||||
perror (InLHS fc n err)
|
||||
= pure $ "While processing left hand side of " ++ sugarName !(getFullName n) ++
|
||||
= pure $ "While processing left hand side of " ++ sugarName !(getFullName n) ++
|
||||
" at " ++ show fc ++ ":\n" ++ !(perror err)
|
||||
perror (InRHS fc n err)
|
||||
= pure $ "While processing right hand side of " ++ sugarName !(getFullName n) ++
|
||||
= pure $ "While processing right hand side of " ++ sugarName !(getFullName n) ++
|
||||
" at " ++ show fc ++ ":\n" ++ !(perror err)
|
||||
|
||||
export
|
||||
display : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
Error -> Core String
|
||||
display err
|
||||
display err
|
||||
= pure $ maybe "" (\f => show f ++ ":") (getErrorLoc err) ++
|
||||
!(perror err)
|
||||
|
@ -51,13 +51,13 @@ dump (Other str) = str
|
||||
data UPD : Type where
|
||||
|
||||
doUpdates : {auto u : Ref UPD (List String)} ->
|
||||
Defs -> List (String, String) -> List SourcePart ->
|
||||
Defs -> List (String, String) -> List SourcePart ->
|
||||
Core (List SourcePart)
|
||||
doUpdates defs ups [] = pure []
|
||||
doUpdates defs ups (LBrace :: xs)
|
||||
= case dropSpace xs of
|
||||
Name n :: RBrace :: rest =>
|
||||
pure (LBrace :: Name n ::
|
||||
pure (LBrace :: Name n ::
|
||||
Whitespace " " :: Equal :: Whitespace " " ::
|
||||
!(doUpdates defs ups (Name n :: RBrace :: rest)))
|
||||
Name n :: Equal :: rest =>
|
||||
@ -80,14 +80,14 @@ doUpdates defs ups (HoleName n :: xs)
|
||||
n' <- uniqueName defs used n
|
||||
put UPD (n' :: used)
|
||||
pure $ HoleName n' :: !(doUpdates defs ups xs)
|
||||
doUpdates defs ups (x :: xs)
|
||||
doUpdates defs ups (x :: xs)
|
||||
= pure $ x :: !(doUpdates defs ups xs)
|
||||
|
||||
-- State here is a list of new hole names we generated (so as not to reuse any).
|
||||
-- Update the token list with the string replacements for each match, and return
|
||||
-- Update the token list with the string replacements for each match, and return
|
||||
-- the newly generated strings.
|
||||
updateAll : {auto u : Ref UPD (List String)} ->
|
||||
Defs -> List SourcePart -> List (List (String, String)) ->
|
||||
Defs -> List SourcePart -> List (List (String, String)) ->
|
||||
Core (List String)
|
||||
updateAll defs l [] = pure []
|
||||
updateAll defs l (rs :: rss)
|
||||
@ -104,7 +104,7 @@ getReplaces : {auto c : Ref Ctxt Defs} ->
|
||||
getReplaces updates
|
||||
= do strups <- traverse toStrUpdate updates
|
||||
pure (concat strups)
|
||||
|
||||
|
||||
showImpossible : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
@ -131,7 +131,7 @@ updateCase splits line col
|
||||
let thisline = getLine (cast line) (lines file)
|
||||
case thisline of
|
||||
Nothing => throw (InternalError "File too short!")
|
||||
Just l =>
|
||||
Just l =>
|
||||
do let valid = mapMaybe getValid splits
|
||||
let bad = mapMaybe getBad splits
|
||||
if isNil valid
|
||||
@ -151,7 +151,7 @@ updateCase splits line col
|
||||
getBad _ = Nothing
|
||||
|
||||
fnName : Bool -> Name -> String
|
||||
fnName lhs (UN n)
|
||||
fnName lhs (UN n)
|
||||
= if any (not . identChar) (unpack n)
|
||||
then if lhs then "(" ++ n ++ ")"
|
||||
else "op"
|
||||
@ -171,7 +171,7 @@ getClause l n
|
||||
| Nothing => pure Nothing
|
||||
n <- getFullName nidx
|
||||
argns <- getEnvArgNames defs envlen !(nf defs [] ty)
|
||||
pure (Just (indent loc ++ fnName True n ++ concat (map (" " ++) argns) ++
|
||||
pure (Just (indent loc ++ fnName True n ++ concat (map (" " ++) argns) ++
|
||||
" = ?" ++ fnName False n ++ "_rhs"))
|
||||
where
|
||||
indent : FC -> String
|
||||
|
@ -6,7 +6,7 @@ import Parser.Lexer
|
||||
-- Implement make-with and make-case from the IDE mode
|
||||
|
||||
isLit : String -> (Bool, String)
|
||||
isLit str
|
||||
isLit str
|
||||
= assert_total $
|
||||
if length str > 0 && strHead str == '>'
|
||||
then (True, strTail str)
|
||||
@ -27,7 +27,7 @@ makeWith n srcline
|
||||
case span isSpace src of
|
||||
(spc, rest) => (length spc, rest)
|
||||
indent = fst isrc
|
||||
src = snd isrc
|
||||
src = snd isrc
|
||||
lhs = pack (readLHS 0 (unpack src)) in
|
||||
mkWithArg lit indent lhs ++ "\n" ++
|
||||
mkWithPat lit indent lhs ++ "\n"
|
||||
@ -43,15 +43,15 @@ makeWith n srcline
|
||||
|
||||
pref : Bool -> Nat -> String
|
||||
pref l ind
|
||||
= (if l then ">" else "") ++
|
||||
= (if l then ">" else "") ++
|
||||
pack (replicate ind ' ')
|
||||
|
||||
mkWithArg : Bool -> Nat -> String -> String
|
||||
mkWithArg lit indent lhs
|
||||
mkWithArg lit indent lhs
|
||||
= pref lit indent ++ lhs ++ "with (_)"
|
||||
|
||||
mkWithPat : Bool -> Nat -> String -> String
|
||||
mkWithPat lit indent lhs
|
||||
= pref lit (indent + 2) ++ lhs ++ "| with_pat = ?" ++
|
||||
mkWithPat lit indent lhs
|
||||
= pref lit (indent + 2) ++ lhs ++ "| with_pat = ?" ++
|
||||
showRHSName n ++ "_rhs"
|
||||
|
||||
|
@ -30,7 +30,7 @@ ident = pred startIdent <+> many (pred validIdent)
|
||||
validIdent x = isAlphaNum x
|
||||
|
||||
ideTokens : TokenMap Token
|
||||
ideTokens =
|
||||
ideTokens =
|
||||
map (\x => (exact x, Symbol)) symbols ++
|
||||
[(digits, \x => Literal (cast x)),
|
||||
(stringLit, \x => StrLit (stripQuotes x)),
|
||||
@ -42,11 +42,11 @@ ideTokens =
|
||||
stripQuotes = assert_total (strTail . reverse . strTail . reverse)
|
||||
|
||||
idelex : String -> Either (Int, Int, String) (List (TokenData Token))
|
||||
idelex str
|
||||
idelex str
|
||||
= case lex ideTokens str of
|
||||
-- Add the EndInput token so that we'll have a line and column
|
||||
-- number to read when storing spans in the file
|
||||
(tok, (l, c, "")) => Right (filter notComment tok ++
|
||||
(tok, (l, c, "")) => Right (filter notComment tok ++
|
||||
[MkToken l c EndInput])
|
||||
(_, fail) => Left fail
|
||||
where
|
||||
@ -65,7 +65,7 @@ sexp
|
||||
pure (IntegerAtom i)
|
||||
<|> do str <- strLit
|
||||
pure (StringAtom str)
|
||||
<|> do symbol ":"; x <- unqualifiedName
|
||||
<|> do symbol ":"; x <- unqualifiedName
|
||||
pure (SymbolAtom x)
|
||||
<|> do symbol "("
|
||||
xs <- many sexp
|
||||
@ -73,14 +73,14 @@ sexp
|
||||
pure (SExpList xs)
|
||||
|
||||
ideParser : String -> Grammar (TokenData Token) e ty -> Either ParseError ty
|
||||
ideParser str p
|
||||
ideParser str p
|
||||
= case idelex str of
|
||||
Left err => Left $ LexFail err
|
||||
Right toks =>
|
||||
Right toks =>
|
||||
case parse p toks of
|
||||
Left (Error err []) =>
|
||||
Left (Error err []) =>
|
||||
Left $ ParseFail err Nothing []
|
||||
Left (Error err (t :: ts)) =>
|
||||
Left (Error err (t :: ts)) =>
|
||||
Left $ ParseFail err (Just (line t, col t))
|
||||
(map tok (t :: ts))
|
||||
Right (val, _) => Right val
|
||||
|
@ -30,7 +30,7 @@ holeIdent : Lexer
|
||||
holeIdent = is '?' <+> ident
|
||||
|
||||
srcTokens : TokenMap SourcePart
|
||||
srcTokens =
|
||||
srcTokens =
|
||||
[(ident, Name),
|
||||
(holeIdent, \x => HoleName (assert_total (strTail x))),
|
||||
(space, Whitespace),
|
||||
@ -41,10 +41,10 @@ srcTokens =
|
||||
|
||||
export
|
||||
tokens : String -> List SourcePart
|
||||
tokens str
|
||||
tokens str
|
||||
= case lex srcTokens str of
|
||||
-- Add the EndInput token so that we'll have a line and column
|
||||
-- number to read when storing spans in the file
|
||||
(srctoks, (l, c, rest)) =>
|
||||
(srctoks, (l, c, rest)) =>
|
||||
map tok srctoks ++ (if rest == "" then [] else [Other rest])
|
||||
|
||||
|
@ -420,7 +420,9 @@ process (Eval itm)
|
||||
defs <- get Ctxt
|
||||
opts <- get ROpts
|
||||
let norm = nfun (evalMode opts)
|
||||
itm <- resugar [] !(norm defs [] tm)
|
||||
ntm <- norm defs [] tm
|
||||
itm <- resugar [] ntm
|
||||
logTermNF 5 "Normalised" [] ntm
|
||||
if showTypes opts
|
||||
then do ty <- getTerm gty
|
||||
ity <- resugar [] !(norm defs [] ty)
|
||||
@ -560,7 +562,7 @@ process (Editing cmd)
|
||||
processEdit cmd
|
||||
setPPrint ppopts
|
||||
pure True
|
||||
process Quit
|
||||
process Quit
|
||||
= pure False
|
||||
process NOP
|
||||
= pure True
|
||||
@ -653,8 +655,8 @@ repl
|
||||
repeat <- interpret inp
|
||||
end <- coreLift $ fEOF stdin
|
||||
if repeat && not end
|
||||
then repl
|
||||
else
|
||||
then repl
|
||||
else
|
||||
do iputStrLn "Bye for now!"
|
||||
pure ()
|
||||
|
||||
|
@ -22,9 +22,9 @@ iputStrLn msg
|
||||
REPL False => coreLift $ putStrLn msg
|
||||
REPL _ => pure ()
|
||||
IDEMode i _ f =>
|
||||
send f (SExpList [SymbolAtom "write-string",
|
||||
send f (SExpList [SymbolAtom "write-string",
|
||||
toSExp msg, toSExp i])
|
||||
|
||||
|
||||
|
||||
printWithStatus : {auto o : Ref ROpts REPLOpts} ->
|
||||
String -> String -> Core ()
|
||||
@ -33,7 +33,7 @@ printWithStatus status msg
|
||||
case idemode opts of
|
||||
REPL _ => coreLift $ putStrLn msg
|
||||
IDEMode i _ f =>
|
||||
do let m = SExpList [SymbolAtom status, toSExp msg,
|
||||
do let m = SExpList [SymbolAtom status, toSExp msg,
|
||||
-- highlighting; currently blank
|
||||
SExpList []]
|
||||
send f (SExpList [SymbolAtom "return", m, toSExp i])
|
||||
@ -58,7 +58,7 @@ emitError : {auto c : Ref Ctxt Defs} ->
|
||||
emitError err
|
||||
= do opts <- get ROpts
|
||||
case idemode opts of
|
||||
REPL _ =>
|
||||
REPL _ =>
|
||||
do msg <- display err
|
||||
coreLift $ putStrLn msg
|
||||
IDEMode i _ f =>
|
||||
@ -66,10 +66,10 @@ emitError err
|
||||
case getErrorLoc err of
|
||||
Nothing => iputStrLn msg
|
||||
Just fc =>
|
||||
send f (SExpList [SymbolAtom "warning",
|
||||
SExpList [toSExp (file fc),
|
||||
toSExp (addOne (startPos fc)),
|
||||
toSExp (addOne (endPos fc)),
|
||||
send f (SExpList [SymbolAtom "warning",
|
||||
SExpList [toSExp (file fc),
|
||||
toSExp (addOne (startPos fc)),
|
||||
toSExp (addOne (endPos fc)),
|
||||
toSExp msg,
|
||||
-- highlighting; currently blank
|
||||
SExpList []],
|
||||
|
@ -4,7 +4,7 @@ import Idris.Syntax
|
||||
import Idris.Socket
|
||||
|
||||
public export
|
||||
data OutputMode
|
||||
data OutputMode
|
||||
= IDEMode Integer File File
|
||||
| REPL Bool -- quiet flag (ignore iputStrLn)
|
||||
|
||||
@ -53,7 +53,7 @@ setSource src
|
||||
export
|
||||
getSource : {auto o : Ref ROpts REPLOpts} ->
|
||||
Core String
|
||||
getSource
|
||||
getSource
|
||||
= do opts <- get ROpts
|
||||
pure (source opts)
|
||||
|
||||
|
@ -131,7 +131,7 @@ mutual
|
||||
= pure (sugarApp (PRef fc (UN n)))
|
||||
toPTerm p (IVar loc (Nested _ n))
|
||||
= toPTerm p (IVar loc n)
|
||||
toPTerm p (IVar fc n)
|
||||
toPTerm p (IVar fc n)
|
||||
= do ns <- fullNamespace
|
||||
pure (sugarApp (PRef fc (if ns then n else dropNS n)))
|
||||
toPTerm p (IPi fc rig Implicit n arg ret)
|
||||
@ -168,7 +168,7 @@ mutual
|
||||
bracket p startPrec (mkIf (PCase fc sc' alts'))
|
||||
where
|
||||
mkIf : PTerm -> PTerm
|
||||
mkIf tm@(PCase loc sc [MkPatClause _ (PRef _ tval) t [],
|
||||
mkIf tm@(PCase loc sc [MkPatClause _ (PRef _ tval) t [],
|
||||
MkPatClause _ (PRef _ fval) f []])
|
||||
= if dropNS tval == UN "True" && dropNS fval == UN "False"
|
||||
then PIfThenElse loc sc t f
|
||||
@ -190,7 +190,7 @@ mutual
|
||||
= do arg' <- toPTerm startPrec arg
|
||||
fn' <- toPTerm startPrec fn
|
||||
bracket p appPrec (PWithApp fc fn' arg')
|
||||
toPTerm p (IImplicitApp fc fn n arg)
|
||||
toPTerm p (IImplicitApp fc fn n arg)
|
||||
= do arg' <- toPTerm startPrec arg
|
||||
app <- toPTermApp fn [(fc, Just n, arg')]
|
||||
imp <- showImplicits
|
||||
@ -199,7 +199,7 @@ mutual
|
||||
else mkOp app
|
||||
toPTerm p (ISearch fc d) = pure (PSearch fc d)
|
||||
toPTerm p (IAlternative fc _ _) = pure (PImplicit fc)
|
||||
toPTerm p (IRewrite fc rule tm)
|
||||
toPTerm p (IRewrite fc rule tm)
|
||||
= pure (PRewrite fc !(toPTerm startPrec rule)
|
||||
!(toPTerm startPrec tm))
|
||||
toPTerm p (ICoerced fc tm) = toPTerm p tm
|
||||
@ -221,7 +221,7 @@ mutual
|
||||
mkApp : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
PTerm -> List (FC, Maybe (Maybe Name), PTerm) -> Core PTerm
|
||||
mkApp fn [] = pure fn
|
||||
mkApp fn [] = pure fn
|
||||
mkApp fn ((fc, Nothing, arg) :: rest)
|
||||
= do let ap = sugarApp (PApp fc fn arg)
|
||||
mkApp ap rest
|
||||
@ -236,10 +236,10 @@ mutual
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
RawImp -> List (FC, Maybe (Maybe Name), PTerm) ->
|
||||
Core PTerm
|
||||
toPTermApp (IApp fc f a) args
|
||||
toPTermApp (IApp fc f a) args
|
||||
= do a' <- toPTerm argPrec a
|
||||
toPTermApp f ((fc, Nothing, a') :: args)
|
||||
toPTermApp (IImplicitApp fc f n a) args
|
||||
toPTermApp (IImplicitApp fc f n a) args
|
||||
= do a' <- toPTerm startPrec a
|
||||
toPTermApp f ((fc, Just n, a') :: args)
|
||||
toPTermApp fn@(IVar fc n) args
|
||||
@ -249,12 +249,12 @@ mutual
|
||||
mkApp fn' args
|
||||
Just def => do fn' <- toPTerm appPrec fn
|
||||
fenv <- showFullEnv
|
||||
let args'
|
||||
= if fenv
|
||||
let args'
|
||||
= if fenv
|
||||
then args
|
||||
else drop (length (vars def)) args
|
||||
mkApp fn' args'
|
||||
toPTermApp fn args
|
||||
toPTermApp fn args
|
||||
= do fn' <- toPTerm appPrec fn
|
||||
mkApp fn' args
|
||||
|
||||
@ -306,10 +306,10 @@ mutual
|
||||
|
||||
toPRecord : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
ImpRecord ->
|
||||
ImpRecord ->
|
||||
Core (Name, List (Name, PTerm), Maybe Name, List PField)
|
||||
toPRecord (MkImpRecord fc n ps con fs)
|
||||
= do ps' <- traverse (\ (n, ty) =>
|
||||
= do ps' <- traverse (\ (n, ty) =>
|
||||
do ty' <- toPTerm startPrec ty
|
||||
pure (n, ty')) ps
|
||||
fs' <- traverse toPField fs
|
||||
@ -319,7 +319,7 @@ mutual
|
||||
toPDecl : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
ImpDecl -> Core (Maybe PDecl)
|
||||
toPDecl (IClaim fc rig vis opts ty)
|
||||
toPDecl (IClaim fc rig vis opts ty)
|
||||
= pure (Just (PClaim fc rig vis opts !(toPTypeDecl ty)))
|
||||
toPDecl (IData fc vis d)
|
||||
= pure (Just (PData fc vis !(toPData d)))
|
||||
@ -327,7 +327,7 @@ mutual
|
||||
= pure (Just (PDef fc !(traverse toPClause cs)))
|
||||
toPDecl (IParameters fc ps ds)
|
||||
= do ds' <- traverse toPDecl ds
|
||||
pure (Just (PParameters fc
|
||||
pure (Just (PParameters fc
|
||||
!(traverse (\ntm => do tm' <- toPTerm startPrec (snd ntm)
|
||||
pure (fst ntm, tm')) ps)
|
||||
(mapMaybe id ds')))
|
||||
@ -337,6 +337,9 @@ mutual
|
||||
toPDecl (INamespace fc _ ns ds)
|
||||
= do ds' <- traverse toPDecl ds
|
||||
pure (Just (PNamespace fc ns (mapMaybe id ds')))
|
||||
toPDecl (ITransform fc lhs rhs)
|
||||
= pure (Just (PTransform fc !(toPTerm startPrec lhs)
|
||||
!(toPTerm startPrec rhs)))
|
||||
toPDecl (IPragma _) = pure Nothing
|
||||
toPDecl (ILog _) = pure Nothing
|
||||
|
||||
@ -347,7 +350,7 @@ resugar : {auto c : Ref Ctxt Defs} ->
|
||||
resugar env tm
|
||||
= do tti <- unelab env tm
|
||||
toPTerm startPrec tti
|
||||
|
||||
|
||||
export
|
||||
resugarNoPatvars : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
@ -355,7 +358,7 @@ resugarNoPatvars : {auto c : Ref Ctxt Defs} ->
|
||||
resugarNoPatvars env tm
|
||||
= do tti <- unelabNoPatvars env tm
|
||||
toPTerm startPrec tti
|
||||
|
||||
|
||||
export
|
||||
pterm : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
|
@ -217,6 +217,7 @@ mutual
|
||||
PMutual : FC -> List PDecl -> PDecl
|
||||
PFixity : FC -> Fixity -> Nat -> OpStr -> PDecl
|
||||
PNamespace : FC -> List String -> List PDecl -> PDecl
|
||||
PTransform : FC -> PTerm -> PTerm -> PDecl
|
||||
PDirective : FC -> Directive -> PDecl
|
||||
|
||||
definedInData : PDataDecl -> List Name
|
||||
|
@ -44,8 +44,8 @@ comment = is '-' <+> is '-' <+> many (isNot '\n')
|
||||
|
||||
toEndComment : (k : Nat) -> Recognise (k /= 0)
|
||||
toEndComment Z = empty
|
||||
toEndComment (S k)
|
||||
= some (pred (\c => c /= '-' && c /= '{'))
|
||||
toEndComment (S k)
|
||||
= some (pred (\c => c /= '-' && c /= '{'))
|
||||
<+> toEndComment (S k)
|
||||
<|> is '{' <+> is '-' <+> toEndComment (S (S k))
|
||||
<|> is '-' <+> is '}' <+> toEndComment k
|
||||
@ -54,7 +54,7 @@ toEndComment (S k)
|
||||
|
||||
blockComment : Lexer
|
||||
blockComment = is '{' <+> is '-' <+> toEndComment 1
|
||||
|
||||
|
||||
docComment : Lexer
|
||||
docComment = is '|' <+> is '|' <+> is '|' <+> many (isNot '\n')
|
||||
|
||||
@ -74,7 +74,7 @@ holeIdent : Lexer
|
||||
holeIdent = is '?' <+> ident
|
||||
|
||||
doubleLit : Lexer
|
||||
doubleLit
|
||||
doubleLit
|
||||
= digits <+> is '.' <+> digits <+> opt
|
||||
(is 'e' <+> opt (is '-' <|> is '+') <+> digits)
|
||||
|
||||
@ -82,11 +82,11 @@ doubleLit
|
||||
-- a specific back end
|
||||
cgDirective : Lexer
|
||||
cgDirective
|
||||
= exact "%cg" <+>
|
||||
((some space <+>
|
||||
= exact "%cg" <+>
|
||||
((some space <+>
|
||||
some (pred isAlphaNum) <+> many space <+>
|
||||
is '{' <+> many (isNot '}') <+>
|
||||
is '}')
|
||||
is '{' <+> many (isNot '}') <+>
|
||||
is '}')
|
||||
<|> many (isNot '\n'))
|
||||
|
||||
mkDirective : String -> Token
|
||||
@ -113,10 +113,10 @@ special = ["%lam", "%pi", "%imppi", "%let"]
|
||||
-- don't match 'validSymbol'
|
||||
export
|
||||
symbols : List String
|
||||
symbols
|
||||
symbols
|
||||
= [".(", -- for things such as Foo.Bar.(+)
|
||||
"@{",
|
||||
"(", ")", "{", "}", "[", "]", ",", ";", "_",
|
||||
"(", ")", "{", "}", "[", "]", ",", ";", "_",
|
||||
"`(", "`"]
|
||||
|
||||
export
|
||||
@ -130,14 +130,14 @@ validSymbol = some (oneOf opChars)
|
||||
export
|
||||
reservedSymbols : List String
|
||||
reservedSymbols
|
||||
= symbols ++
|
||||
= symbols ++
|
||||
["%", "\\", ":", "=", "|", "|||", "<-", "->", "=>", "?", "&", "**", ".."]
|
||||
|
||||
symbolChar : Char -> Bool
|
||||
symbolChar c = c `elem` unpack opChars
|
||||
|
||||
rawTokens : TokenMap Token
|
||||
rawTokens =
|
||||
rawTokens =
|
||||
[(comment, Comment),
|
||||
(blockComment, Comment),
|
||||
(docComment, DocComment),
|
||||
@ -160,11 +160,11 @@ rawTokens =
|
||||
export
|
||||
lexTo : (TokenData Token -> Bool) ->
|
||||
String -> Either (Int, Int, String) (List (TokenData Token))
|
||||
lexTo pred str
|
||||
lexTo pred str
|
||||
= case lexTo pred rawTokens str of
|
||||
-- Add the EndInput token so that we'll have a line and column
|
||||
-- number to read when storing spans in the file
|
||||
(tok, (l, c, "")) => Right (filter notComment tok ++
|
||||
(tok, (l, c, "")) => Right (filter notComment tok ++
|
||||
[MkToken l c EndInput])
|
||||
(_, fail) => Left fail
|
||||
where
|
||||
|
@ -26,15 +26,15 @@ export
|
||||
Show ParseError where
|
||||
show (ParseFail err loc toks)
|
||||
= "Parse error: " ++ err ++ " (next tokens: "
|
||||
++ show (take 10 toks) ++ ")"
|
||||
show (LexFail (c, l, str))
|
||||
++ show (take 10 toks) ++ ")"
|
||||
show (LexFail (c, l, str))
|
||||
= "Lex error at " ++ show (c, l) ++ " input: " ++ str
|
||||
show (FileFail err)
|
||||
= "File error: " ++ show err
|
||||
|
||||
export
|
||||
eoi : EmptyRule ()
|
||||
eoi
|
||||
eoi
|
||||
= do nextIs "Expected end of input" (isEOI . tok)
|
||||
pure ()
|
||||
where
|
||||
@ -45,14 +45,14 @@ eoi
|
||||
export
|
||||
runParserTo : (TokenData Token -> Bool) ->
|
||||
String -> Grammar (TokenData Token) e ty -> Either ParseError ty
|
||||
runParserTo pred str p
|
||||
runParserTo pred str p
|
||||
= case lexTo pred str of
|
||||
Left err => Left $ LexFail err
|
||||
Right toks =>
|
||||
Right toks =>
|
||||
case parse p toks of
|
||||
Left (Error err []) =>
|
||||
Left (Error err []) =>
|
||||
Left $ ParseFail err Nothing []
|
||||
Left (Error err (t :: ts)) =>
|
||||
Left (Error err (t :: ts)) =>
|
||||
Left $ ParseFail err (Just (line t, col t))
|
||||
(map tok (t :: ts))
|
||||
Right (val, _) => Right val
|
||||
@ -73,7 +73,7 @@ parseFile fn p
|
||||
|
||||
export
|
||||
location : EmptyRule (Int, Int)
|
||||
location
|
||||
location
|
||||
= do tok <- peek
|
||||
pure (line tok, col tok)
|
||||
|
||||
@ -176,29 +176,29 @@ escape' ('\\' :: 't' :: xs) = pure $ '\t' :: !(escape' xs)
|
||||
escape' ('\\' :: 'v' :: xs) = pure $ '\v' :: !(escape' xs)
|
||||
escape' ('\\' :: '\'' :: xs) = pure $ '\'' :: !(escape' xs)
|
||||
escape' ('\\' :: '\"' :: xs) = pure $ '\"' :: !(escape' xs)
|
||||
escape' ('\\' :: 'x' :: xs)
|
||||
escape' ('\\' :: 'x' :: xs)
|
||||
= case span isHexDigit xs of
|
||||
([], rest) => assert_total (escape' rest)
|
||||
(ds, rest) => pure $ cast !(toHex 1 (reverse ds)) ::
|
||||
(ds, rest) => pure $ cast !(toHex 1 (reverse ds)) ::
|
||||
!(assert_total (escape' rest))
|
||||
where
|
||||
where
|
||||
toHex : Int -> List Char -> Maybe Int
|
||||
toHex _ [] = Just 0
|
||||
toHex m (d :: ds)
|
||||
toHex m (d :: ds)
|
||||
= pure $ !(hex (toLower d)) * m + !(toHex (m*16) ds)
|
||||
escape' ('\\' :: 'o' :: xs)
|
||||
escape' ('\\' :: 'o' :: xs)
|
||||
= case span isOctDigit xs of
|
||||
([], rest) => assert_total (escape' rest)
|
||||
(ds, rest) => pure $ cast !(toOct 1 (reverse ds)) ::
|
||||
(ds, rest) => pure $ cast !(toOct 1 (reverse ds)) ::
|
||||
!(assert_total (escape' rest))
|
||||
where
|
||||
where
|
||||
toOct : Int -> List Char -> Maybe Int
|
||||
toOct _ [] = Just 0
|
||||
toOct m (d :: ds)
|
||||
toOct m (d :: ds)
|
||||
= pure $ !(oct (toLower d)) * m + !(toOct (m*8) ds)
|
||||
escape' ('\\' :: xs)
|
||||
escape' ('\\' :: xs)
|
||||
= case span isDigit xs of
|
||||
([], (a :: b :: c :: rest)) =>
|
||||
([], (a :: b :: c :: rest)) =>
|
||||
case getEsc (pack (the (List _) [a, b, c])) of
|
||||
Just v => Just (v :: !(assert_total (escape' rest)))
|
||||
Nothing => case getEsc (pack (the (List _) [a, b])) of
|
||||
@ -209,7 +209,7 @@ escape' ('\\' :: xs)
|
||||
Just v => Just (v :: [])
|
||||
Nothing => escape' xs
|
||||
([], rest) => assert_total (escape' rest)
|
||||
(ds, rest) => Just $ cast (cast {to=Int} (pack ds)) ::
|
||||
(ds, rest) => Just $ cast (cast {to=Int} (pack ds)) ::
|
||||
!(assert_total (escape' rest))
|
||||
escape' (x :: xs) = Just $ x :: !(escape' xs)
|
||||
|
||||
@ -227,7 +227,7 @@ getCharLit str
|
||||
|
||||
export
|
||||
constant : Rule Constant
|
||||
constant
|
||||
constant
|
||||
= terminal "Expected constant"
|
||||
(\x => case tok x of
|
||||
Literal i => Just (BI i)
|
||||
@ -247,7 +247,7 @@ constant
|
||||
|
||||
export
|
||||
intLit : Rule Integer
|
||||
intLit
|
||||
intLit
|
||||
= terminal "Expected integer literal"
|
||||
(\x => case tok x of
|
||||
Literal i => Just i
|
||||
@ -255,7 +255,7 @@ intLit
|
||||
|
||||
export
|
||||
strLit : Rule String
|
||||
strLit
|
||||
strLit
|
||||
= terminal "Expected string literal"
|
||||
(\x => case tok x of
|
||||
StrLit s => Just s
|
||||
@ -293,14 +293,14 @@ operator : Rule String
|
||||
operator
|
||||
= terminal "Expected operator"
|
||||
(\x => case tok x of
|
||||
Symbol s =>
|
||||
if s `elem` reservedSymbols
|
||||
Symbol s =>
|
||||
if s `elem` reservedSymbols
|
||||
then Nothing
|
||||
else Just s
|
||||
_ => Nothing)
|
||||
|
||||
identPart : Rule String
|
||||
identPart
|
||||
identPart
|
||||
= terminal "Expected name"
|
||||
(\x => case tok x of
|
||||
Ident str => Just str
|
||||
@ -308,13 +308,13 @@ identPart
|
||||
|
||||
export
|
||||
namespace_ : Rule (List String)
|
||||
namespace_
|
||||
namespace_
|
||||
= do ns <- sepBy1 (do col <- column
|
||||
symbol "."
|
||||
col' <- column
|
||||
if (col' - col == 1)
|
||||
then pure ()
|
||||
else fail "No whitepace allowed after namespace separator")
|
||||
else fail "No whitepace allowed after namespace separator")
|
||||
identPart
|
||||
pure (reverse ns) -- innermost first, so reverse
|
||||
|
||||
@ -324,7 +324,7 @@ unqualifiedName = identPart
|
||||
|
||||
export
|
||||
holeName : Rule String
|
||||
holeName
|
||||
holeName
|
||||
= terminal "Expected hole name"
|
||||
(\x => case tok x of
|
||||
HoleIdent str => Just str
|
||||
@ -332,8 +332,8 @@ holeName
|
||||
|
||||
export
|
||||
name : Rule Name
|
||||
name
|
||||
= do ns <- namespace_
|
||||
name
|
||||
= do ns <- namespace_
|
||||
(do symbol ".("
|
||||
op <- operator
|
||||
symbol ")"
|
||||
@ -374,12 +374,12 @@ continue = continueF (fail "Unexpected end of expression")
|
||||
-- As 'continue' but failing is fatal (i.e. entire parse fails)
|
||||
export
|
||||
mustContinue : (indent : IndentInfo) -> Maybe String -> EmptyRule ()
|
||||
mustContinue indent Nothing
|
||||
mustContinue indent Nothing
|
||||
= continueF (fatalError "Unexpected end of expression") indent
|
||||
mustContinue indent (Just req)
|
||||
mustContinue indent (Just req)
|
||||
= continueF (fatalError ("Expected '" ++ req ++ "'")) indent
|
||||
|
||||
data ValidIndent
|
||||
data ValidIndent
|
||||
= AnyIndent -- In {}, entries can begin in any column
|
||||
| AtPos Int -- Entry must begin in a specific column
|
||||
| AfterPos Int -- Entry can begin in this column or later
|
||||
@ -416,7 +416,7 @@ isTerminator _ = False
|
||||
|
||||
-- Check we're at the end of a block entry, given the start column
|
||||
-- of the block.
|
||||
-- It's the end if we have a terminating token, or the next token starts
|
||||
-- It's the end if we have a terminating token, or the next token starts
|
||||
-- in or before indent. Works by looking ahead but not consuming.
|
||||
export
|
||||
atEnd : (indent : IndentInfo) -> EmptyRule ()
|
||||
@ -452,8 +452,8 @@ terminator valid laststart
|
||||
afterDedent valid col
|
||||
<|> pure EndOfBlock
|
||||
where
|
||||
-- Expected indentation for the next token can either be anything (if
|
||||
-- we're inside a brace delimited block) or anywhere after the initial
|
||||
-- Expected indentation for the next token can either be anything (if
|
||||
-- we're inside a brace delimited block) or anywhere after the initial
|
||||
-- column (if we're inside an indentation delimited block)
|
||||
afterSemi : ValidIndent -> ValidIndent
|
||||
afterSemi AnyIndent = AnyIndent -- in braces, anything goes
|
||||
@ -461,8 +461,8 @@ terminator valid laststart
|
||||
afterSemi (AfterPos c) = AfterPos c
|
||||
afterSemi EndOfBlock = EndOfBlock
|
||||
|
||||
-- Expected indentation for the next token can either be anything (if
|
||||
-- we're inside a brace delimited block) or in exactly the initial column
|
||||
-- Expected indentation for the next token can either be anything (if
|
||||
-- we're inside a brace delimited block) or in exactly the initial column
|
||||
-- (if we're inside an indentation delimited block)
|
||||
afterDedent : ValidIndent -> Int -> EmptyRule ValidIndent
|
||||
afterDedent AnyIndent col
|
||||
@ -480,7 +480,7 @@ terminator valid laststart
|
||||
afterDedent EndOfBlock col = pure EndOfBlock
|
||||
|
||||
-- Parse an entry in a block
|
||||
blockEntry : ValidIndent -> (IndentInfo -> Rule ty) ->
|
||||
blockEntry : ValidIndent -> (IndentInfo -> Rule ty) ->
|
||||
Rule (ty, ValidIndent)
|
||||
blockEntry valid rule
|
||||
= do col <- column
|
||||
@ -535,7 +535,7 @@ blockWithOptHeaderAfter {ty} mincol header item
|
||||
else do hidt <- optional $ blockEntry (AtPos col) header
|
||||
ps <- blockEntries (AtPos col) item
|
||||
pure (map fst hidt, ps)
|
||||
where
|
||||
where
|
||||
restOfBlock : Maybe (hd, ValidIndent) -> Rule (Maybe hd, List ty)
|
||||
restOfBlock (Just (h, idt)) = do ps <- blockEntries idt item
|
||||
symbol "}"
|
||||
|
@ -13,13 +13,13 @@ import Control.Monad.State
|
||||
-- Rename the IBindVars in a term. Anything which appears in the list 'renames'
|
||||
-- should be renamed, to something which is *not* in the list 'used'
|
||||
export
|
||||
renameIBinds : (renames : List String) ->
|
||||
(used : List String) ->
|
||||
renameIBinds : (renames : List String) ->
|
||||
(used : List String) ->
|
||||
RawImp -> State (List (String, String)) RawImp
|
||||
renameIBinds rs us (IPi fc c p (Just (UN n)) ty sc)
|
||||
= if n `elem` rs
|
||||
= if n `elem` rs
|
||||
then let n' = getUnique (rs ++ us) n
|
||||
sc' = substNames (map UN (filter (/= n) us))
|
||||
sc' = substNames (map UN (filter (/= n) us))
|
||||
[(UN n, IVar fc (UN n'))] sc in
|
||||
do scr <- renameIBinds rs (n' :: us) sc'
|
||||
ty' <- renameIBinds rs us ty
|
||||
@ -141,6 +141,6 @@ piBindNames loc env tm
|
||||
where
|
||||
piBind : List String -> RawImp -> RawImp
|
||||
piBind [] ty = ty
|
||||
piBind (n :: ns) ty
|
||||
piBind (n :: ns) ty
|
||||
= IPi loc Rig0 Implicit (Just (UN n)) (Implicit loc False) (piBind ns ty)
|
||||
|
||||
|
@ -6,6 +6,7 @@ import Core.Env
|
||||
import Core.LinearCheck
|
||||
import Core.Metadata
|
||||
import Core.Normalise
|
||||
import Core.Transform
|
||||
import Core.UnifyState
|
||||
import Core.Unify
|
||||
|
||||
@ -23,7 +24,7 @@ findPLetRenames (Bind fc n (PLet c (Local {name = x@(MN _ _)} _ _ _ p) ty) sc)
|
||||
findPLetRenames (Bind fc n _ sc) = findPLetRenames sc
|
||||
findPLetRenames tm = []
|
||||
|
||||
doPLetRenames : List (Name, (RigCount, Name)) ->
|
||||
doPLetRenames : List (Name, (RigCount, Name)) ->
|
||||
List Name -> Term vars -> Term vars
|
||||
doPLetRenames ns drops (Bind fc n b@(PLet _ _ _) sc)
|
||||
= if n `elem` drops
|
||||
@ -31,7 +32,7 @@ doPLetRenames ns drops (Bind fc n b@(PLet _ _ _) sc)
|
||||
else Bind fc n b (doPLetRenames ns drops sc)
|
||||
doPLetRenames ns drops (Bind fc n b sc)
|
||||
= case lookup n ns of
|
||||
Just (c, n') =>
|
||||
Just (c, n') =>
|
||||
Bind fc n' (setMultiplicity b (max c (multiplicity b)))
|
||||
(doPLetRenames ns (n' :: drops) (renameTop n' sc))
|
||||
Nothing => Bind fc n b (doPLetRenames ns drops sc)
|
||||
@ -46,7 +47,7 @@ getRigNeeded _ = Rig1
|
||||
-- away (since solved holes don't get written to .tti)
|
||||
export
|
||||
normaliseHoleTypes : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
Core ()
|
||||
normaliseHoleTypes
|
||||
= do ust <- get UST
|
||||
@ -145,7 +146,7 @@ elabTermSub {vars} defining mode opts nest env env' sub tm ty
|
||||
-- need to check here.
|
||||
else pure chktm
|
||||
normaliseHoleTypes
|
||||
-- Put the current hole state back to what it was (minus anything
|
||||
-- Put the current hole state back to what it was (minus anything
|
||||
-- which has been solved in the meantime)
|
||||
when (not incase) $
|
||||
do hs <- getHoles
|
||||
@ -155,15 +156,15 @@ elabTermSub {vars} defining mode opts nest env env' sub tm ty
|
||||
-- were of the form x@_, where the _ is inferred to be a variable,
|
||||
-- to just x)
|
||||
case mode of
|
||||
InLHS _ =>
|
||||
InLHS _ =>
|
||||
do let vs = findPLetRenames chktm
|
||||
let ret = doPLetRenames vs [] chktm
|
||||
pure (ret, gnf env (doPLetRenames vs [] !(getTerm chkty)))
|
||||
_ => do dumpConstraints 2 False
|
||||
pure (chktm, chkty)
|
||||
where
|
||||
addHoles : (acc : IntMap (FC, Name)) ->
|
||||
(allHoles : IntMap (FC, Name)) ->
|
||||
addHoles : (acc : IntMap (FC, Name)) ->
|
||||
(allHoles : IntMap (FC, Name)) ->
|
||||
List (Int, (FC, Name)) ->
|
||||
IntMap (FC, Name)
|
||||
addHoles acc allhs [] = acc
|
||||
@ -189,8 +190,8 @@ checkTermSub : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
Int -> ElabMode -> List ElabOpt ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
Int -> ElabMode -> List ElabOpt ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
Env Term inner -> SubVars inner vars ->
|
||||
RawImp -> Glued vars ->
|
||||
Core (Term vars)
|
||||
@ -201,9 +202,9 @@ checkTermSub defining mode opts nest env env' sub tm ty
|
||||
_ => get Ctxt
|
||||
ust <- get UST
|
||||
mv <- get MD
|
||||
res <-
|
||||
res <-
|
||||
catch {t = Error}
|
||||
(elabTermSub defining mode opts nest
|
||||
(elabTermSub defining mode opts nest
|
||||
env env' sub tm (Just ty))
|
||||
(\err => case err of
|
||||
TryWithImplicits loc benv ns
|
||||
@ -217,7 +218,7 @@ checkTermSub defining mode opts nest env env' sub tm ty
|
||||
_ => throw err)
|
||||
pure (fst res)
|
||||
where
|
||||
bindImps : FC -> Env Term vs -> List (Name, Term vs) -> RawImp ->
|
||||
bindImps : FC -> Env Term vs -> List (Name, Term vs) -> RawImp ->
|
||||
Core RawImp
|
||||
bindImps loc env [] ty = pure ty
|
||||
bindImps loc env ((n, ty) :: ntys) sc
|
||||
@ -229,8 +230,8 @@ checkTerm : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
Int -> ElabMode -> List ElabOpt ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
Int -> ElabMode -> List ElabOpt ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RawImp -> Glued vars ->
|
||||
Core (Term vars)
|
||||
checkTerm defining mode opts nest env tm ty
|
||||
|
@ -18,8 +18,8 @@ import TTImp.TTImp
|
||||
export
|
||||
expandAmbigName : {auto c : Ref Ctxt Defs} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
ElabMode -> NestedNames vars -> Env Term vars -> RawImp ->
|
||||
List (FC, Maybe (Maybe Name), RawImp) ->
|
||||
ElabMode -> NestedNames vars -> Env Term vars -> RawImp ->
|
||||
List (FC, Maybe (Maybe Name), RawImp) ->
|
||||
RawImp -> Maybe (Glued vars) -> Core RawImp
|
||||
expandAmbigName (InLHS _) nest env orig args (IBindVar fc n) exp
|
||||
= do est <- get EST
|
||||
@ -32,11 +32,11 @@ expandAmbigName mode nest env orig args (IVar fc x) exp
|
||||
Nothing => do
|
||||
defs <- get Ctxt
|
||||
case defined x env of
|
||||
Just _ =>
|
||||
if isNil args || notLHS mode
|
||||
Just _ =>
|
||||
if isNil args || notLHS mode
|
||||
then pure $ orig
|
||||
else pure $ IMustUnify fc "Name applied to arguments" orig
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do est <- get EST
|
||||
fi <- fromIntegerName
|
||||
si <- fromStringName
|
||||
@ -51,41 +51,41 @@ expandAmbigName mode nest env orig args (IVar fc x) exp
|
||||
where
|
||||
-- If there's multiple alternatives and all else fails, resort to using
|
||||
-- the primitive directly
|
||||
uniqType : Maybe Name -> Maybe Name -> Maybe Name -> Name ->
|
||||
uniqType : Maybe Name -> Maybe Name -> Maybe Name -> Name ->
|
||||
List (FC, Maybe (Maybe Name), RawImp) -> AltType
|
||||
uniqType (Just fi) _ _ n [(_, _, IPrimVal fc (BI x))]
|
||||
uniqType (Just fi) _ _ n [(_, _, IPrimVal fc (BI x))]
|
||||
= UniqueDefault (IPrimVal fc (BI x))
|
||||
uniqType _ (Just si) _ n [(_, _, IPrimVal fc (Str x))]
|
||||
uniqType _ (Just si) _ n [(_, _, IPrimVal fc (Str x))]
|
||||
= UniqueDefault (IPrimVal fc (Str x))
|
||||
uniqType _ _ (Just ci) n [(_, _, IPrimVal fc (Ch x))]
|
||||
uniqType _ _ (Just ci) n [(_, _, IPrimVal fc (Ch x))]
|
||||
= UniqueDefault (IPrimVal fc (Ch x))
|
||||
uniqType _ _ _ _ _ = Unique
|
||||
|
||||
buildAlt : RawImp -> List (FC, Maybe (Maybe Name), RawImp) ->
|
||||
buildAlt : RawImp -> List (FC, Maybe (Maybe Name), RawImp) ->
|
||||
RawImp
|
||||
buildAlt f [] = f
|
||||
buildAlt f ((fc', Nothing, a) :: as)
|
||||
buildAlt f ((fc', Nothing, a) :: as)
|
||||
= buildAlt (IApp fc' f a) as
|
||||
buildAlt f ((fc', Just i, a) :: as)
|
||||
buildAlt f ((fc', Just i, a) :: as)
|
||||
= buildAlt (IImplicitApp fc' f i a) as
|
||||
|
||||
|
||||
isPrimName : List Name -> Name -> Bool
|
||||
isPrimName [] fn = False
|
||||
isPrimName (p :: ps) fn
|
||||
isPrimName (p :: ps) fn
|
||||
= dropNS fn == p || isPrimName ps fn
|
||||
|
||||
-- If it's not a constructor application, dot it
|
||||
wrapDot : Bool -> EState vars ->
|
||||
ElabMode -> Name -> List RawImp -> Def -> RawImp -> RawImp
|
||||
ElabMode -> Name -> List RawImp -> Def -> RawImp -> RawImp
|
||||
wrapDot _ _ _ _ _ (DCon _ _) tm = tm
|
||||
wrapDot _ _ _ _ _ (TCon _ _ _ _ _ _) tm = tm
|
||||
-- Leave primitive applications alone, because they'll be inlined
|
||||
-- before compiling the case tree
|
||||
wrapDot prim est (InLHS _) n' [arg] _ tm
|
||||
wrapDot prim est (InLHS _) n' [arg] _ tm
|
||||
= if n' == Resolved (defining est) || prim
|
||||
then tm
|
||||
else IMustUnify fc "Not a constructor application or primitive" tm
|
||||
wrapDot prim est (InLHS _) n' _ _ tm
|
||||
wrapDot prim est (InLHS _) n' _ _ tm
|
||||
= if n' == Resolved (defining est)
|
||||
then tm
|
||||
else IMustUnify fc "Not a constructor application or primitive" tm
|
||||
@ -93,12 +93,12 @@ expandAmbigName mode nest env orig args (IVar fc x) exp
|
||||
|
||||
|
||||
mkTerm : Bool -> EState vars -> Name -> GlobalDef -> RawImp
|
||||
mkTerm prim est n def
|
||||
mkTerm prim est n def
|
||||
= wrapDot prim est mode n (map (snd . snd) args)
|
||||
(definition def) (buildAlt (IVar fc n) args)
|
||||
|
||||
mkAlt : Bool -> EState vars -> (Name, Int, GlobalDef) -> RawImp
|
||||
mkAlt prim est (fullname, i, gdef)
|
||||
mkAlt prim est (fullname, i, gdef)
|
||||
= mkTerm prim est (Resolved i) gdef
|
||||
|
||||
notLHS : ElabMode -> Bool
|
||||
@ -106,10 +106,10 @@ expandAmbigName mode nest env orig args (IVar fc x) exp
|
||||
notLHS _ = True
|
||||
|
||||
expandAmbigName mode nest env orig args (IApp fc f a) exp
|
||||
= expandAmbigName mode nest env orig
|
||||
= expandAmbigName mode nest env orig
|
||||
((fc, Nothing, a) :: args) f exp
|
||||
expandAmbigName mode nest env orig args (IImplicitApp fc f n a) exp
|
||||
= expandAmbigName mode nest env orig
|
||||
= expandAmbigName mode nest env orig
|
||||
((fc, Just n, a) :: args) f exp
|
||||
expandAmbigName elabmode nest env orig args tm exp = pure orig
|
||||
|
||||
@ -126,11 +126,11 @@ Show TypeMatch where
|
||||
|
||||
mutual
|
||||
mightMatchD : Defs -> NF vars -> NF [] -> Core TypeMatch
|
||||
mightMatchD defs l r
|
||||
mightMatchD defs l r
|
||||
= mightMatch defs (stripDelay l) (stripDelay r)
|
||||
|
||||
mightMatchArg : Defs ->
|
||||
Closure vars -> Closure [] ->
|
||||
mightMatchArg : Defs ->
|
||||
Closure vars -> Closure [] ->
|
||||
Core Bool
|
||||
mightMatchArg defs l r
|
||||
= case !(mightMatchD defs !(evalClosure defs l) !(evalClosure defs r)) of
|
||||
@ -152,16 +152,16 @@ mutual
|
||||
mightMatch defs target (NBind fc n (Pi _ _ _) sc)
|
||||
= mightMatchD defs target !(sc defs (toClosure defaultOpts [] (Erased fc)))
|
||||
mightMatch defs (NTCon _ n t a args) (NTCon _ n' t' a' args')
|
||||
= if n == n'
|
||||
= if n == n'
|
||||
then do amatch <- mightMatchArgs defs args args'
|
||||
if amatch then pure Concrete else pure NoMatch
|
||||
else pure NoMatch
|
||||
mightMatch defs (NDCon _ n t a args) (NDCon _ n' t' a' args')
|
||||
= if t == t'
|
||||
= if t == t'
|
||||
then do amatch <- mightMatchArgs defs args args'
|
||||
if amatch then pure Concrete else pure NoMatch
|
||||
else pure NoMatch
|
||||
mightMatch defs (NPrimVal _ x) (NPrimVal _ y)
|
||||
mightMatch defs (NPrimVal _ x) (NPrimVal _ y)
|
||||
= if x == y then pure Concrete else pure NoMatch
|
||||
mightMatch defs (NType _) (NType _) = pure Concrete
|
||||
mightMatch defs (NApp _ _ _) _ = pure Poly
|
||||
@ -187,17 +187,17 @@ couldBeFn defs ty _ = pure Poly
|
||||
-- Just (True, app) if it's a match on concrete return type
|
||||
-- Just (False, app) if it might be a match due to being polymorphic
|
||||
couldBe : Defs -> NF vars -> RawImp -> Core (Maybe (Bool, RawImp))
|
||||
couldBe {vars} defs ty@(NTCon _ n _ _ _) app
|
||||
couldBe {vars} defs ty@(NTCon _ n _ _ _) app
|
||||
= case !(couldBeFn {vars} defs ty (getFn app)) of
|
||||
Concrete => pure $ Just (True, app)
|
||||
Poly => pure $ Just (False, app)
|
||||
NoMatch => pure Nothing
|
||||
couldBe {vars} defs ty@(NPrimVal _ _) app
|
||||
couldBe {vars} defs ty@(NPrimVal _ _) app
|
||||
= case !(couldBeFn {vars} defs ty (getFn app)) of
|
||||
Concrete => pure $ Just (True, app)
|
||||
Poly => pure $ Just (False, app)
|
||||
NoMatch => pure Nothing
|
||||
couldBe {vars} defs ty@(NType _) app
|
||||
couldBe {vars} defs ty@(NType _) app
|
||||
= case !(couldBeFn {vars} defs ty (getFn app)) of
|
||||
Concrete => pure $ Just (True, app)
|
||||
Poly => pure $ Just (False, app)
|
||||
@ -230,7 +230,7 @@ filterCore f (x :: xs)
|
||||
|
||||
pruneByType : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
Env Term vars -> NF vars -> List RawImp ->
|
||||
Env Term vars -> NF vars -> List RawImp ->
|
||||
Core (List RawImp)
|
||||
pruneByType env target alts
|
||||
= do defs <- get Ctxt
|
||||
@ -273,8 +273,8 @@ checkAlternative : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> AltType -> List RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkAlternative rig elabinfo nest env fc (UniqueDefault def) alts mexpected
|
||||
@ -285,35 +285,35 @@ checkAlternative rig elabinfo nest env fc (UniqueDefault def) alts mexpected
|
||||
let solvemode = case elabMode elabinfo of
|
||||
InLHS c => InLHS
|
||||
_ => InTerm
|
||||
delayOnFailure fc rig env expected ambiguous $
|
||||
delayOnFailure fc rig env expected ambiguous $
|
||||
\delayed =>
|
||||
do solveConstraints solvemode Normal
|
||||
defs <- get Ctxt
|
||||
exp <- getTerm expected
|
||||
|
||||
-- We can't just use the old NF on the second attempt,
|
||||
-- We can't just use the old NF on the second attempt,
|
||||
-- because we might know more now, so recalculate it
|
||||
let exp' = if delayed
|
||||
let exp' = if delayed
|
||||
then gnf env exp
|
||||
else expected
|
||||
|
||||
alts' <- pruneByType env !(getNF exp') alts
|
||||
|
||||
logGlueNF 5 ("Ambiguous elaboration " ++ show alts' ++
|
||||
logGlueNF 5 ("Ambiguous elaboration " ++ show alts' ++
|
||||
" at " ++ show fc ++
|
||||
"\nWith default. Target type ") env exp'
|
||||
if delayed -- use the default if there's still ambiguity
|
||||
then try
|
||||
(exactlyOne fc env
|
||||
(map (\t =>
|
||||
(getName t,
|
||||
checkImp rig elabinfo nest env t
|
||||
then try
|
||||
(exactlyOne fc env
|
||||
(map (\t =>
|
||||
(getName t,
|
||||
checkImp rig elabinfo nest env t
|
||||
(Just exp'))) alts'))
|
||||
(do log 5 "All failed, running default"
|
||||
checkImp rig elabinfo nest env def (Just exp'))
|
||||
else exactlyOne fc env
|
||||
(map (\t =>
|
||||
(getName t,
|
||||
(map (\t =>
|
||||
(getName t,
|
||||
checkImp rig elabinfo nest env t (Just exp')))
|
||||
alts')
|
||||
checkAlternative rig elabinfo nest env fc uniq alts mexpected
|
||||
@ -329,28 +329,28 @@ checkAlternative rig elabinfo nest env fc uniq alts mexpected
|
||||
let solvemode = case elabMode elabinfo of
|
||||
InLHS c => InLHS
|
||||
_ => InTerm
|
||||
delayOnFailure fc rig env expected ambiguous $
|
||||
delayOnFailure fc rig env expected ambiguous $
|
||||
\delayed =>
|
||||
do solveConstraints solvemode Normal
|
||||
defs <- get Ctxt
|
||||
exp <- getTerm expected
|
||||
|
||||
-- We can't just use the old NF on the second attempt,
|
||||
-- We can't just use the old NF on the second attempt,
|
||||
-- because we might know more now, so recalculate it
|
||||
let exp' = if delayed
|
||||
let exp' = if delayed
|
||||
then gnf env exp
|
||||
else expected
|
||||
|
||||
alts' <- pruneByType env !(getNF exp') alts
|
||||
|
||||
logGlueNF 5 ("Ambiguous elaboration " ++ show alts' ++
|
||||
logGlueNF 5 ("Ambiguous elaboration " ++ show alts' ++
|
||||
" at " ++ show fc ++
|
||||
"\nTarget type ") env exp'
|
||||
let tryall = case uniq of
|
||||
FirstSuccess => anyOne fc
|
||||
_ => exactlyOne fc env
|
||||
tryall (map (\t =>
|
||||
(getName t,
|
||||
tryall (map (\t =>
|
||||
(getName t,
|
||||
do res <- checkImp rig elabinfo nest env t (Just exp')
|
||||
-- Do it twice for interface resolution;
|
||||
-- first pass gets the determining argument
|
||||
|
@ -36,17 +36,17 @@ getNameType : {vars : _} ->
|
||||
Core (Term vars, Glued vars)
|
||||
getNameType rigc env fc x
|
||||
= case defined x env of
|
||||
Just (MkIsDefined rigb lv) =>
|
||||
Just (MkIsDefined rigb lv) =>
|
||||
do rigSafe rigb rigc
|
||||
let binder = getBinder lv env
|
||||
let bty = binderType binder
|
||||
addNameType fc x env bty
|
||||
when (isLinear rigb) $
|
||||
do est <- get EST
|
||||
put EST
|
||||
put EST
|
||||
(record { linearUsed $= ((MkVar lv) :: ) } est)
|
||||
pure (Local fc (Just (isLet binder)) _ lv, gnf env bty)
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do defs <- get Ctxt
|
||||
[(pname, i, def)] <- lookupCtxtName x (gamma defs)
|
||||
| [] => throw (UndefinedName fc x)
|
||||
@ -102,9 +102,9 @@ getVarType rigc nest env fc x
|
||||
where
|
||||
useVars : List (Term vars) -> Term vars -> Term vars
|
||||
useVars [] sc = sc
|
||||
useVars (a :: as) (Bind bfc n (Pi c _ ty) sc)
|
||||
useVars (a :: as) (Bind bfc n (Pi c _ ty) sc)
|
||||
= Bind bfc n (Let c a ty) (useVars (map weaken as) sc)
|
||||
useVars as (Bind bfc n (Let c v ty) sc)
|
||||
useVars as (Bind bfc n (Let c v ty) sc)
|
||||
= Bind bfc n (Let c v ty) (useVars (map weaken as) sc)
|
||||
useVars _ sc = sc -- Can't happen?
|
||||
|
||||
@ -130,9 +130,9 @@ mutual
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fntm : Term vars) ->
|
||||
RigCount -> RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fntm : Term vars) ->
|
||||
Name -> NF vars -> (Defs -> Closure vars -> Core (NF vars)) ->
|
||||
(expargs : List RawImp) ->
|
||||
(impargs : List (Maybe Name, RawImp)) ->
|
||||
@ -158,9 +158,9 @@ mutual
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fntm : Term vars) ->
|
||||
RigCount -> RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fntm : Term vars) ->
|
||||
Name -> NF vars -> (Defs -> Closure vars -> Core (NF vars)) ->
|
||||
(expargs : List RawImp) ->
|
||||
(impargs : List (Maybe Name, RawImp)) ->
|
||||
@ -171,7 +171,7 @@ mutual
|
||||
-- on the LHS, just treat it as an implicit pattern variable.
|
||||
-- on the RHS, add a searchable hole
|
||||
= case elabMode elabinfo of
|
||||
InLHS _ =>
|
||||
InLHS _ =>
|
||||
do defs <- get Ctxt
|
||||
nm <- genMVName x
|
||||
empty <- clearDefs defs
|
||||
@ -217,7 +217,7 @@ mutual
|
||||
needsDelayExpr True (ISearch _ _) = pure True
|
||||
needsDelayExpr True (IRewrite _ _ _) = pure True
|
||||
needsDelayExpr True _ = pure False
|
||||
|
||||
|
||||
-- On the LHS, for any concrete thing, we need to make sure we know
|
||||
-- its type before we proceed so that we can reject it if the type turns
|
||||
-- out to be polymorphic
|
||||
@ -270,8 +270,8 @@ mutual
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fntm : Term vars) -> Name ->
|
||||
(aty : NF vars) -> (sc : Defs -> Closure vars -> Core (NF vars)) ->
|
||||
(arg : RawImp) ->
|
||||
@ -308,7 +308,7 @@ mutual
|
||||
defs <- get Ctxt
|
||||
-- If we're on the LHS, reinstantiate it with 'argv' because it
|
||||
-- *may* have as patterns in it and we need to retain them.
|
||||
-- (As patterns are a bit of a hack but I don't yet see a
|
||||
-- (As patterns are a bit of a hack but I don't yet see a
|
||||
-- better way that leads to good code...)
|
||||
logTerm 5 ("Solving " ++ show metaval ++ " with") argv
|
||||
ok <- solveIfUndefined env metaval argv
|
||||
@ -351,25 +351,25 @@ mutual
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fntm : Term vars) -> (fnty : NF vars) ->
|
||||
(expargs : List RawImp) ->
|
||||
(impargs : List (Maybe Name, RawImp)) ->
|
||||
(knownret : Bool) -> -- Do we know what the return type is yet?
|
||||
-- if we do, we might be able to use it to work
|
||||
-- if we do, we might be able to use it to work
|
||||
-- out the types of arguments before elaborating them
|
||||
(expected : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
-- Ordinary explicit argument
|
||||
checkAppWith rig elabinfo nest env fc tm (NBind tfc x (Pi rigb Explicit aty) sc)
|
||||
(arg :: expargs) impargs kr expty
|
||||
(arg :: expargs) impargs kr expty
|
||||
= do let argRig = rigMult rig rigb
|
||||
checkRestApp rig argRig elabinfo nest env fc
|
||||
checkRestApp rig argRig elabinfo nest env fc
|
||||
tm x aty sc arg expargs impargs kr expty
|
||||
-- Function type is delayed, so force the term and continue
|
||||
checkAppWith rig elabinfo nest env fc tm (NDelayed dfc r ty@(NBind _ _ (Pi _ _ _) sc)) expargs impargs kr expty
|
||||
= checkAppWith rig elabinfo nest env fc (TForce dfc tm) ty expargs impargs kr expty
|
||||
= checkAppWith rig elabinfo nest env fc (TForce dfc r tm) ty expargs impargs kr expty
|
||||
-- If there's no more arguments given, and the plicities of the type and
|
||||
-- the expected type line up, stop
|
||||
checkAppWith rig elabinfo nest env fc tm ty@(NBind tfc x (Pi rigb Implicit aty) sc)
|
||||
@ -396,10 +396,10 @@ mutual
|
||||
expargs impargs kr expty
|
||||
= let argRig = rigMult rig rigb in
|
||||
case useAutoImp [] impargs of
|
||||
Nothing => makeAutoImplicit rig argRig elabinfo nest env fc tm
|
||||
Nothing => makeAutoImplicit rig argRig elabinfo nest env fc tm
|
||||
x aty sc expargs impargs kr expty
|
||||
Just (arg, impargs') =>
|
||||
checkRestApp rig argRig elabinfo nest env fc
|
||||
checkRestApp rig argRig elabinfo nest env fc
|
||||
tm x aty sc arg expargs impargs' kr expty
|
||||
where
|
||||
useAutoImp : List (Maybe Name, RawImp) -> List (Maybe Name, RawImp) ->
|
||||
@ -418,10 +418,10 @@ mutual
|
||||
expargs impargs kr expty
|
||||
= let argRig = rigMult rig rigb in
|
||||
case useImp [] impargs of
|
||||
Nothing => makeImplicit rig argRig elabinfo nest env fc tm
|
||||
Nothing => makeImplicit rig argRig elabinfo nest env fc tm
|
||||
x aty sc expargs impargs kr expty
|
||||
Just (arg, impargs') =>
|
||||
checkRestApp rig argRig elabinfo nest env fc
|
||||
checkRestApp rig argRig elabinfo nest env fc
|
||||
tm x aty sc arg expargs impargs' kr expty
|
||||
where
|
||||
useImp : List (Maybe Name, RawImp) -> List (Maybe Name, RawImp) ->
|
||||
@ -434,10 +434,10 @@ mutual
|
||||
useImp acc (ximp :: rest)
|
||||
= useImp (ximp :: acc) rest
|
||||
|
||||
checkAppWith rig elabinfo nest env fc tm ty [] [] kr expty
|
||||
checkAppWith rig elabinfo nest env fc tm ty [] [] kr expty
|
||||
= do defs <- get Ctxt
|
||||
checkExp rig elabinfo env fc tm (glueBack defs env ty) expty
|
||||
checkAppWith rig elabinfo nest env fc tm ty [] impargs kr expty
|
||||
checkAppWith rig elabinfo nest env fc tm ty [] impargs kr expty
|
||||
= case filter notInfer impargs of
|
||||
[] => checkAppWith rig elabinfo nest env fc tm ty [] [] kr expty
|
||||
is => throw (InvalidImplicits fc env (map fst is) tm)
|
||||
@ -446,7 +446,7 @@ mutual
|
||||
notInfer (_, Implicit _ _) = False
|
||||
notInfer (n, IAs _ _ _ i) = notInfer (n, i)
|
||||
notInfer _ = True
|
||||
checkAppWith {vars} rig elabinfo nest env fc tm ty (arg :: expargs) impargs kr expty
|
||||
checkAppWith {vars} rig elabinfo nest env fc tm ty (arg :: expargs) impargs kr expty
|
||||
= -- Invent a function type, and hope that we'll know enough to solve it
|
||||
-- later when we unify with expty
|
||||
do logNF 10 "Function type" env ty
|
||||
@ -456,7 +456,7 @@ mutual
|
||||
argTy <- metaVar fc Rig0 env argn (TType fc)
|
||||
let argTyG = gnf env argTy
|
||||
retTy <- metaVar -- {vars = argn :: vars}
|
||||
fc Rig0 env -- (Pi RigW Explicit argTy :: env)
|
||||
fc Rig0 env -- (Pi RigW Explicit argTy :: env)
|
||||
retn (TType fc)
|
||||
(argv, argt) <- check rig elabinfo
|
||||
nest env arg (Just argTyG)
|
||||
@ -480,10 +480,10 @@ checkApp : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fn : RawImp) ->
|
||||
(expargs : List RawImp) ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (fn : RawImp) ->
|
||||
(expargs : List RawImp) ->
|
||||
(impargs : List (Maybe Name, RawImp)) ->
|
||||
Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
@ -497,7 +497,7 @@ checkApp rig elabinfo nest env fc (IVar fc' n) expargs impargs exp
|
||||
elabinfo <- updateElabInfo (elabMode elabinfo) n expargs elabinfo
|
||||
logC 10 (do defs <- get Ctxt
|
||||
fnty <- quote defs env nty
|
||||
exptyt <- maybe (pure Nothing)
|
||||
exptyt <- maybe (pure Nothing)
|
||||
(\t => do ety <- getTerm t
|
||||
etynf <- normaliseHoles defs env ety
|
||||
pure (Just !(toFullNames etynf)))
|
||||
@ -510,15 +510,15 @@ checkApp rig elabinfo nest env fc (IVar fc' n) expargs impargs exp
|
||||
where
|
||||
isPrimName : List Name -> Name -> Bool
|
||||
isPrimName [] fn = False
|
||||
isPrimName (p :: ps) fn
|
||||
isPrimName (p :: ps) fn
|
||||
= dropNS fn == p || isPrimName ps fn
|
||||
|
||||
|
||||
updateElabInfo : ElabMode -> Name -> List RawImp -> ElabInfo -> Core ElabInfo
|
||||
-- If it's a primitive function applied to a constant on the LHS, treat it
|
||||
-- as an expression because we'll normalise the function away and match on
|
||||
-- the result
|
||||
updateElabInfo (InLHS _) n [IPrimVal fc c] elabinfo =
|
||||
do let prims = mapMaybe id
|
||||
do let prims = mapMaybe id
|
||||
[!fromIntegerName, !fromStringName, !fromCharName]
|
||||
if isPrimName prims !(getFullName n)
|
||||
then pure (record { elabMode = InExpr } elabinfo)
|
||||
|
@ -23,11 +23,11 @@ checkAs : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> UseSide -> Name -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkAs rig elabinfo nest env fc side n_in pat topexp
|
||||
checkAs rig elabinfo nest env fc side n_in pat topexp
|
||||
= do let elabmode = elabMode elabinfo
|
||||
let InLHS _ = elabmode
|
||||
| _ => throw (GenericMsg fc "@-patterns only allowed in pattern clauses")
|
||||
@ -40,7 +40,7 @@ checkAs rig elabinfo nest env fc side n_in pat topexp
|
||||
noteLHSPatVar elabmode str
|
||||
notePatVar n
|
||||
case lookup n (boundNames est) of
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do (pattm, patty) <- check rigPat elabinfo nest env pat topexp
|
||||
(tm, exp, bty) <- mkPatternHole fc rig n env
|
||||
(implicitMode elabinfo)
|
||||
@ -56,7 +56,7 @@ checkAs rig elabinfo nest env fc side n_in pat topexp
|
||||
(ntm, nty) <- checkExp rig elabinfo env fc tm (gnf env exp)
|
||||
(Just patty)
|
||||
pure (As fc ntm pattm, patty)
|
||||
Just bty => throw (NonLinearPattern fc n_in)
|
||||
Just bty => throw (NonLinearPattern fc n_in)
|
||||
where
|
||||
-- Only one side can be usable if it's linear! Normally we'd assume this
|
||||
-- to be the new variable (UseRight), but in generated case blocks it's
|
||||
|
@ -23,8 +23,8 @@ dropName n nest = record { names $= drop } nest
|
||||
where
|
||||
drop : List (Name, a) -> List (Name, a)
|
||||
drop [] = []
|
||||
drop ((x, y) :: xs)
|
||||
= if x == n then drop xs
|
||||
drop ((x, y) :: xs)
|
||||
= if x == n then drop xs
|
||||
else (x, y) :: drop xs
|
||||
|
||||
export
|
||||
@ -33,21 +33,21 @@ checkPi : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> PiInfo -> (n : Name) ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> PiInfo -> (n : Name) ->
|
||||
(argTy : RawImp) -> (retTy : RawImp) ->
|
||||
(expTy : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkPi rig elabinfo nest env fc rigf info n argTy retTy expTy
|
||||
= do let pirig = getRig (elabMode elabinfo)
|
||||
(tyv, tyt) <- check pirig elabinfo nest env argTy
|
||||
(tyv, tyt) <- check pirig elabinfo nest env argTy
|
||||
(Just (gType fc))
|
||||
let env' : Env Term (n :: _) = Pi rigf info tyv :: env
|
||||
let nest' = weaken (dropName n nest)
|
||||
(scopev, scopet) <-
|
||||
inScope fc env' (\e' =>
|
||||
(scopev, scopet) <-
|
||||
inScope fc env' (\e' =>
|
||||
check {e=e'} pirig elabinfo nest' env' retTy (Just (gType fc)))
|
||||
checkExp rig elabinfo env fc (Bind fc n (Pi rigf info tyv) scopev) (gType fc) expTy
|
||||
where
|
||||
@ -71,10 +71,10 @@ inferLambda : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> PiInfo -> (n : Name) ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> PiInfo -> (n : Name) ->
|
||||
(argTy : RawImp) -> (scope : RawImp) ->
|
||||
(expTy : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
@ -85,7 +85,7 @@ inferLambda rig elabinfo nest env fc rigl info n argTy scope expTy
|
||||
let env' : Env Term (n :: _) = Lam rigb info tyv :: env
|
||||
let nest' = weaken (dropName n nest)
|
||||
(scopev, scopet) <- inScope fc env' (\e' =>
|
||||
check {e=e'} rig elabinfo
|
||||
check {e=e'} rig elabinfo
|
||||
nest' env' scope Nothing)
|
||||
let lamty = gnf env (Bind fc n (Pi rigb info tyv) !(getTerm scopet))
|
||||
logGlue 5 "Inferred lambda type" env lamty
|
||||
@ -110,10 +110,10 @@ checkLambda : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> PiInfo -> (n : Name) ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> PiInfo -> (n : Name) ->
|
||||
(argTy : RawImp) -> (scope : RawImp) ->
|
||||
(expTy : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
@ -127,21 +127,21 @@ checkLambda rig_in elabinfo nest env fc rigl info n argTy scope (Just expty_in)
|
||||
defs <- get Ctxt
|
||||
case exptynf of
|
||||
Bind bfc bn (Pi c _ pty) psc =>
|
||||
do (tyv, tyt) <- check Rig0 elabinfo nest env
|
||||
do (tyv, tyt) <- check Rig0 elabinfo nest env
|
||||
argTy (Just (gType fc))
|
||||
let rigb = min rigl c
|
||||
let env' : Env Term (n :: _) = Lam rigb info tyv :: env
|
||||
convert fc elabinfo env (gnf env tyv) (gnf env pty)
|
||||
convert fc elabinfo env (gnf env tyv) (gnf env pty)
|
||||
let nest' = weaken (dropName n nest)
|
||||
(scopev, scopet) <-
|
||||
inScope fc env' (\e' =>
|
||||
check {e=e'} rig elabinfo nest' env' scope
|
||||
check {e=e'} rig elabinfo nest' env' scope
|
||||
(Just (gnf env' (renameTop n psc))))
|
||||
logTermNF 10 "Lambda type" env exptynf
|
||||
logGlueNF 10 "Got scope type" env' scopet
|
||||
checkExp rig elabinfo env fc
|
||||
checkExp rig elabinfo env fc
|
||||
(Bind fc n (Lam rigb info tyv) scopev)
|
||||
(gnf env
|
||||
(gnf env
|
||||
(Bind fc n (Pi rigb info tyv) !(getTerm scopet)))
|
||||
(Just (gnf env
|
||||
(Bind fc bn (Pi rigb info pty) psc)))
|
||||
@ -160,10 +160,10 @@ checkLet : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> (n : Name) ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC ->
|
||||
RigCount -> (n : Name) ->
|
||||
(nTy : RawImp) -> (nVal : RawImp) -> (scope : RawImp) ->
|
||||
(expTy : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
@ -179,15 +179,15 @@ checkLet rigc_in elabinfo nest env fc rigl n nTy nVal scope expty
|
||||
pure (fst c, snd c, rigMult rigl rigc))
|
||||
(\err => case err of
|
||||
LinearMisuse _ _ Rig1 _
|
||||
=> do c <- check Rig1 elabinfo
|
||||
=> do c <- check Rig1 elabinfo
|
||||
nest env nVal (Just (gnf env tyv))
|
||||
pure (fst c, snd c, Rig1)
|
||||
e => throw e)
|
||||
let env' : Env Term (n :: _) = Lam rigb Explicit tyv :: env
|
||||
let nest' = weaken (dropName n nest)
|
||||
expScope <- weakenExp env' expty
|
||||
(scopev, gscopet) <-
|
||||
inScope fc env' (\e' =>
|
||||
expScope <- weakenExp env' expty
|
||||
(scopev, gscopet) <-
|
||||
inScope fc env' (\e' =>
|
||||
check {e=e'} rigc elabinfo nest' env' scope expScope)
|
||||
scopet <- getTerm gscopet
|
||||
|
||||
|
@ -26,10 +26,10 @@ changeVar (MkVar {i=x} old) (MkVar new) (Local fc r idx p)
|
||||
else Local fc r _ p
|
||||
changeVar old new (Meta fc nm i args)
|
||||
= Meta fc nm i (map (changeVar old new) args)
|
||||
changeVar (MkVar old) (MkVar new) (Bind fc x b sc)
|
||||
= Bind fc x (assert_total (map (changeVar (MkVar old) (MkVar new)) b))
|
||||
changeVar (MkVar old) (MkVar new) (Bind fc x b sc)
|
||||
= Bind fc x (assert_total (map (changeVar (MkVar old) (MkVar new)) b))
|
||||
(changeVar (MkVar (Later old)) (MkVar (Later new)) sc)
|
||||
changeVar old new (App fc fn arg)
|
||||
changeVar old new (App fc fn arg)
|
||||
= App fc (changeVar old new fn) (changeVar old new arg)
|
||||
changeVar old new (As fc nm p)
|
||||
= As fc (changeVar old new nm) (changeVar old new p)
|
||||
@ -37,8 +37,8 @@ changeVar old new (TDelayed fc r p)
|
||||
= TDelayed fc r (changeVar old new p)
|
||||
changeVar old new (TDelay fc r t p)
|
||||
= TDelay fc r (changeVar old new t) (changeVar old new p)
|
||||
changeVar old new (TForce fc p)
|
||||
= TForce fc (changeVar old new p)
|
||||
changeVar old new (TForce fc r p)
|
||||
= TForce fc r (changeVar old new p)
|
||||
changeVar old new tm = tm
|
||||
|
||||
findLater : (x : Name) -> (newer : List Name) -> Var (newer ++ x :: older)
|
||||
@ -48,10 +48,10 @@ findLater {older} x (_ :: xs)
|
||||
MkVar (Later p)
|
||||
|
||||
-- For any variable *not* in vs', re-abstract over it in the term
|
||||
absOthers : FC -> Env Term vs -> SubVars vs' vs ->
|
||||
absOthers : FC -> Env Term vs -> SubVars vs' vs ->
|
||||
Term (done ++ vs) -> Term (done ++ vs)
|
||||
absOthers fc _ SubRefl tm = tm
|
||||
absOthers {done} {vs = x :: vars} fc (b :: env) (KeepCons sub) tm
|
||||
absOthers {done} {vs = x :: vars} fc (b :: env) (KeepCons sub) tm
|
||||
= rewrite appendAssociative done [x] vars in
|
||||
absOthers {done = done ++ [x]} fc env sub
|
||||
(rewrite sym (appendAssociative done [x] vars) in tm)
|
||||
@ -60,7 +60,7 @@ absOthers {done} {vs = x :: vars} fc (b :: env) (DropCons sub) tm
|
||||
= rewrite appendAssociative done [x] vars in
|
||||
map (weakenNs (done ++ [x])) b
|
||||
b' = Pi (multiplicity b) Explicit (binderType bindervs)
|
||||
btm = Bind fc x b'
|
||||
btm = Bind fc x b'
|
||||
(changeVar (findLater _ (x :: done)) (MkVar First) (weaken tm)) in
|
||||
rewrite appendAssociative done [x] vars in
|
||||
absOthers {done = done ++ [x]} fc env sub
|
||||
@ -70,43 +70,43 @@ absOthers {done} {vs = x :: vars} fc (b :: env) (DropCons sub) tm
|
||||
-- unbound implicits at the point in the environment where they were
|
||||
-- created (which is after the outer environment was bound)
|
||||
export
|
||||
abstractOver : FC -> Defs -> BindMode -> Env Term vs ->
|
||||
abstractOver : FC -> Defs -> BindMode -> Env Term vs ->
|
||||
Maybe (SubVars outer vs) -> List (Name, ImplBinding outer) ->
|
||||
(tm : Term vs) -> Core ClosedTerm
|
||||
abstractOver fc defs bindmode [] Nothing imps tm = pure tm
|
||||
abstractOver fc defs bindmode [] (Just SubRefl) imps tm
|
||||
= do tm' <- if isNil imps
|
||||
abstractOver fc defs bindmode [] (Just SubRefl) imps tm
|
||||
= do tm' <- if isNil imps
|
||||
then pure tm
|
||||
else normaliseHoles defs [] tm
|
||||
(bimptm, _) <- bindImplicits fc bindmode defs [] imps tm' (TType fc)
|
||||
pure bimptm
|
||||
abstractOver fc defs bindmode (b :: env) (Just SubRefl) imps tm
|
||||
abstractOver fc defs bindmode (b :: env) (Just SubRefl) imps tm
|
||||
= do let c : RigCount
|
||||
= case multiplicity b of
|
||||
Rig1 => Rig0
|
||||
r => r
|
||||
tm' <- if isNil imps
|
||||
tm' <- if isNil imps
|
||||
then pure tm
|
||||
else normaliseHoles defs (b :: env) tm
|
||||
(bimptm, _) <- bindImplicits fc bindmode defs (b :: env) imps
|
||||
tm' (TType fc)
|
||||
abstractOver fc defs bindmode env Nothing imps
|
||||
abstractOver fc defs bindmode env Nothing imps
|
||||
(Bind fc _ (Pi c Explicit (binderType b)) bimptm)
|
||||
abstractOver fc defs bindmode (b :: env) (Just (DropCons p)) imps tm
|
||||
abstractOver fc defs bindmode (b :: env) (Just (DropCons p)) imps tm
|
||||
= let c = case multiplicity b of
|
||||
Rig1 => Rig0
|
||||
r => r in
|
||||
abstractOver fc defs bindmode env (Just p) imps
|
||||
abstractOver fc defs bindmode env (Just p) imps
|
||||
(Bind fc _ (Pi c Explicit (binderType b)) tm)
|
||||
abstractOver fc defs bindmode (b :: env) _ imps tm
|
||||
abstractOver fc defs bindmode (b :: env) _ imps tm
|
||||
= let c = case multiplicity b of
|
||||
Rig1 => Rig0
|
||||
r => r in
|
||||
abstractOver fc defs bindmode env Nothing imps
|
||||
abstractOver fc defs bindmode env Nothing imps
|
||||
(Bind fc _ (Pi c Explicit (binderType b)) tm)
|
||||
|
||||
toRig1 : {idx : Nat} -> .(IsVar name idx vs) -> Env Term vs -> Env Term vs
|
||||
toRig1 First (b :: bs)
|
||||
toRig1 First (b :: bs)
|
||||
= if multiplicity b == Rig0
|
||||
then setMultiplicity b Rig1 :: bs
|
||||
else b :: bs
|
||||
@ -126,7 +126,7 @@ updateMults : List (Var vs) -> Env Term vs -> Env Term vs
|
||||
updateMults [] env = env
|
||||
updateMults (MkVar p :: us) env = updateMults us (toRig0 p env)
|
||||
|
||||
shrinkImp : SubVars outer vars ->
|
||||
shrinkImp : SubVars outer vars ->
|
||||
(Name, ImplBinding vars) -> Maybe (Name, ImplBinding outer)
|
||||
shrinkImp sub (n, NameBinding c p tm ty)
|
||||
= do tm' <- shrinkTerm tm sub
|
||||
@ -141,11 +141,11 @@ shrinkImp sub (n, AsBinding c p tm ty pat)
|
||||
findImpsIn : FC -> Env Term vars -> List (Name, Term vars) -> Term vars ->
|
||||
Core ()
|
||||
findImpsIn fc env ns (Bind _ n b@(Pi _ Implicit ty) sc)
|
||||
= findImpsIn fc (b :: env)
|
||||
= findImpsIn fc (b :: env)
|
||||
((n, weaken ty) :: map (\x => (fst x, weaken (snd x))) ns)
|
||||
sc
|
||||
findImpsIn fc env ns (Bind _ n b sc)
|
||||
= findImpsIn fc (b :: env)
|
||||
= findImpsIn fc (b :: env)
|
||||
(map (\x => (fst x, weaken (snd x))) ns)
|
||||
sc
|
||||
findImpsIn fc env ns ty
|
||||
@ -160,7 +160,7 @@ merge (v :: vs) xs
|
||||
|
||||
-- Extend the list of variables we need in the environment so far, removing
|
||||
-- duplicates
|
||||
extendNeeded : Binder (Term vs) ->
|
||||
extendNeeded : Binder (Term vs) ->
|
||||
Env Term vs -> List (Var vs) -> List (Var vs)
|
||||
extendNeeded (Let c ty val) env needed
|
||||
= merge (findUsedLocs env ty) (merge (findUsedLocs env val) needed)
|
||||
@ -183,16 +183,16 @@ isNeeded x (MkVar {i} _ :: xs) = x == i || isNeeded x xs
|
||||
-- starting from the 'outerEnv' which is the fragment of the environment
|
||||
-- used for the outer scope)
|
||||
shrinkEnv : Defs -> SubVars outer vs -> List (Var vs) ->
|
||||
(done : List Name) -> Env Term vs ->
|
||||
(done : List Name) -> Env Term vs ->
|
||||
Core (outer' ** SubVars outer' vs)
|
||||
shrinkEnv defs SubRefl needed done env
|
||||
shrinkEnv defs SubRefl needed done env
|
||||
= pure (_ ** SubRefl) -- keep them all
|
||||
-- usable name, so drop from the outer environment
|
||||
shrinkEnv {vs = UN n :: _} defs (DropCons p) needed done (b :: env)
|
||||
shrinkEnv {vs = UN n :: _} defs (DropCons p) needed done (b :: env)
|
||||
= do b' <- traverse (normaliseHoles defs env) b
|
||||
(_ ** p') <- shrinkEnv defs p
|
||||
(extendNeeded b'
|
||||
env (dropFirst needed))
|
||||
(_ ** p') <- shrinkEnv defs p
|
||||
(extendNeeded b'
|
||||
env (dropFirst needed))
|
||||
(UN n :: done) env
|
||||
-- if shadowed and not needed, keep in the outer env
|
||||
if (UN n `elem` done) && not (isNeeded 0 needed)
|
||||
@ -200,22 +200,22 @@ shrinkEnv {vs = UN n :: _} defs (DropCons p) needed done (b :: env)
|
||||
else pure (_ ** DropCons p')
|
||||
shrinkEnv {vs = n :: _} defs (DropCons p) needed done (b :: env)
|
||||
= do b' <- traverse (normaliseHoles defs env) b
|
||||
(_ ** p') <- shrinkEnv defs p
|
||||
(extendNeeded b'
|
||||
env (dropFirst needed))
|
||||
(_ ** p') <- shrinkEnv defs p
|
||||
(extendNeeded b'
|
||||
env (dropFirst needed))
|
||||
(n :: done) env
|
||||
if isNeeded 0 needed || notLam b
|
||||
then pure (_ ** DropCons p')
|
||||
then pure (_ ** DropCons p')
|
||||
else pure (_ ** KeepCons p')
|
||||
where
|
||||
notLam : Binder t -> Bool
|
||||
notLam (Lam _ _ _) = False
|
||||
notLam _ = True
|
||||
shrinkEnv {vs = n :: _} defs (KeepCons p) needed done (b :: env)
|
||||
shrinkEnv {vs = n :: _} defs (KeepCons p) needed done (b :: env)
|
||||
= do b' <- traverse (normaliseHoles defs env) b
|
||||
(_ ** p') <- shrinkEnv defs p
|
||||
(_ ** p') <- shrinkEnv defs p
|
||||
(extendNeeded b'
|
||||
env (dropFirst needed))
|
||||
env (dropFirst needed))
|
||||
(n :: done) env
|
||||
pure (_ ** KeepCons p') -- still keep it
|
||||
|
||||
@ -237,14 +237,14 @@ findScrutinee _ _ _ = Nothing
|
||||
|
||||
export
|
||||
caseBlock : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount ->
|
||||
ElabInfo -> FC ->
|
||||
NestedNames vars ->
|
||||
Env Term vars ->
|
||||
RigCount ->
|
||||
ElabInfo -> FC ->
|
||||
NestedNames vars ->
|
||||
Env Term vars ->
|
||||
RawImp -> -- original scrutinee
|
||||
Term vars -> -- checked scrutinee
|
||||
Term vars -> -- its type
|
||||
@ -261,7 +261,7 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
|
||||
(implicitMode elabinfo) env []
|
||||
let fullImps = mapMaybe (shrinkImp (subEnv est)) fullImps_in
|
||||
log 5 $ "Doing a case under unbound implicits " ++ show fullImps
|
||||
|
||||
|
||||
scrn <- genVarName "scr"
|
||||
casen <- genCaseName (defining est)
|
||||
|
||||
@ -287,15 +287,15 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
|
||||
-- if the scrutinee is ones of the arguments in 'env' we should
|
||||
-- split on that, rather than adding it as a new argument
|
||||
let splitOn = findScrutinee env smaller scr
|
||||
|
||||
|
||||
caseretty <- the (Core (Term vars)) $ case expected of
|
||||
Just ty => getTerm ty
|
||||
_ =>
|
||||
do nmty <- genName "caseTy"
|
||||
metaVar fc Rig0 env nmty (TType fc)
|
||||
|
||||
let envscope = absOthers {done = []} fc (allow splitOn env) smaller
|
||||
(maybe (Bind fc scrn (Pi caseRig Explicit scrty)
|
||||
let envscope = absOthers {done = []} fc (allow splitOn env) smaller
|
||||
(maybe (Bind fc scrn (Pi caseRig Explicit scrty)
|
||||
(weaken caseretty))
|
||||
(const caseretty) splitOn)
|
||||
casefnty <- abstractOver fc defs (implicitMode elabinfo)
|
||||
@ -317,20 +317,20 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
|
||||
|
||||
let alts' = map (updateClause casen splitOn env smaller) alts
|
||||
log 2 $ "Generated alts: " ++ show alts'
|
||||
let nest' = record { names $= ((Resolved cidx, (Nothing,
|
||||
(\fc, nt => applyToFull fc caseRef pre_env))) ::) }
|
||||
let nest' = record { names $= ((Resolved cidx, (Nothing,
|
||||
(\fc, nt => applyToFull fc caseRef pre_env))) ::) }
|
||||
nest
|
||||
processDecl [InCase] nest' pre_env (IDef fc casen alts')
|
||||
|
||||
let applyEnv = applyToOthers fc (applyToFull fc caseRef env) env smaller
|
||||
pure (maybe (App fc applyEnv scrtm)
|
||||
(const applyEnv) splitOn,
|
||||
pure (maybe (App fc applyEnv scrtm)
|
||||
(const applyEnv) splitOn,
|
||||
gnf env caseretty)
|
||||
|
||||
where
|
||||
mkLocalEnv : Env Term vs -> Env Term vs
|
||||
mkLocalEnv [] = []
|
||||
mkLocalEnv (b :: bs)
|
||||
mkLocalEnv (b :: bs)
|
||||
= let b' = if isLinear (multiplicity b)
|
||||
then setMultiplicity b Rig0
|
||||
else b in
|
||||
@ -341,40 +341,41 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
|
||||
= if n `elem` vs then Nothing else Just n
|
||||
canBindName _ vs = Nothing
|
||||
|
||||
addEnv : Env Term vs -> SubVars vs' vs -> List Name ->
|
||||
(List (Maybe RawImp), List Name)
|
||||
addEnv [] sub used = ([], used)
|
||||
addEnv : Env Term vs -> SubVars vs' vs -> List Name ->
|
||||
List (Maybe RawImp)
|
||||
addEnv [] sub used = []
|
||||
addEnv {vs = v :: vs} (b :: bs) SubRefl used
|
||||
= let (rest, used') = addEnv bs SubRefl used in
|
||||
(Nothing :: rest, used')
|
||||
= let rest = addEnv bs SubRefl used in
|
||||
Nothing :: rest
|
||||
addEnv (b :: bs) (KeepCons p) used
|
||||
= let (rest, used') = addEnv bs p used in
|
||||
(Nothing :: rest, used')
|
||||
= let rest = addEnv bs p used in
|
||||
Nothing :: rest
|
||||
addEnv {vs = v :: vs} (b :: bs) (DropCons p) used
|
||||
= let (rest, used') = addEnv bs p used in
|
||||
case canBindName v used' of
|
||||
Just n => (Just (IAs fc UseLeft n (Implicit fc True)) :: rest, n :: used')
|
||||
_ => (Just (Implicit fc True) :: rest, used')
|
||||
= case canBindName v used of
|
||||
Just n => let rest = addEnv bs p (n :: used) in
|
||||
Just (IAs fc UseLeft n (Implicit fc True)) :: rest
|
||||
_ => let rest = addEnv bs p used in
|
||||
Just (Implicit fc True) :: rest
|
||||
|
||||
-- Replace a variable in the argument list; if the reference is to
|
||||
-- a variable kept in the outer environment (therefore not an argument
|
||||
-- in the list) don't consume it
|
||||
replace : {idx : Nat} -> .(IsVar name idx vs) ->
|
||||
replace : {idx : Nat} -> .(IsVar name idx vs) ->
|
||||
RawImp -> List (Maybe RawImp) ->
|
||||
List RawImp
|
||||
replace First lhs (old :: xs)
|
||||
replace First lhs (old :: xs)
|
||||
= let lhs' = case old of
|
||||
Just (IAs loc' side n _) => IAs loc' side n lhs
|
||||
Just (IAs loc' side n _) => IAs loc' side n lhs
|
||||
_ => lhs in
|
||||
lhs' :: mapMaybe id xs
|
||||
replace (Later p) lhs (Nothing :: xs)
|
||||
replace (Later p) lhs (Nothing :: xs)
|
||||
= replace p lhs xs
|
||||
replace (Later p) lhs (Just x :: xs)
|
||||
replace (Later p) lhs (Just x :: xs)
|
||||
= x :: replace p lhs xs
|
||||
replace _ _ xs = mapMaybe id xs
|
||||
|
||||
mkSplit : Maybe (Var vs) ->
|
||||
RawImp -> List (Maybe RawImp) ->
|
||||
mkSplit : Maybe (Var vs) ->
|
||||
RawImp -> List (Maybe RawImp) ->
|
||||
List RawImp
|
||||
mkSplit Nothing lhs args = reverse (lhs :: mapMaybe id args)
|
||||
mkSplit (Just (MkVar prf)) lhs args
|
||||
@ -393,16 +394,16 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
|
||||
Env Term vars -> SubVars vs' vars ->
|
||||
ImpClause -> ImpClause
|
||||
updateClause casen splitOn env sub (PatClause loc' lhs rhs)
|
||||
= let args = fst (addEnv env sub (usedIn lhs))
|
||||
= let args = addEnv env sub (usedIn lhs)
|
||||
args' = mkSplit splitOn lhs args in
|
||||
PatClause loc' (apply (IVar loc' casen) args') rhs
|
||||
-- With isn't allowed in a case block but include for completeness
|
||||
updateClause casen splitOn env sub (WithClause loc' lhs wval cs)
|
||||
= let args = fst (addEnv env sub (usedIn lhs))
|
||||
= let args = addEnv env sub (usedIn lhs)
|
||||
args' = mkSplit splitOn lhs args in
|
||||
WithClause loc' (apply (IVar loc' casen) args') wval cs
|
||||
updateClause casen splitOn env sub (ImpossibleClause loc' lhs)
|
||||
= let args = fst (addEnv env sub (usedIn lhs))
|
||||
= let args = addEnv env sub (usedIn lhs)
|
||||
args' = mkSplit splitOn lhs args in
|
||||
ImpossibleClause loc' (apply (IVar loc' casen) args')
|
||||
|
||||
@ -414,15 +415,15 @@ checkCase : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (scr : RawImp) -> (ty : RawImp) -> List ImpClause ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> (scr : RawImp) -> (ty : RawImp) -> List ImpClause ->
|
||||
Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkCase rig elabinfo nest env fc scr scrty_exp alts exp
|
||||
= delayElab fc rig env exp $
|
||||
do (scrtyv, scrtyt) <- check Rig0 elabinfo nest env scrty_exp
|
||||
do (scrtyv, scrtyt) <- check Rig0 elabinfo nest env scrty_exp
|
||||
(Just (gType fc))
|
||||
|
||||
|
||||
logTerm 10 "Expected scrutinee type" scrtyv
|
||||
-- Try checking at the given multiplicity; if that doesn't work,
|
||||
-- try checking at Rig1 (meaning that we're using a linear variable
|
||||
@ -432,7 +433,7 @@ checkCase rig elabinfo nest env fc scr scrty_exp alts exp
|
||||
pure (fst c, snd c, RigW))
|
||||
(\err => case err of
|
||||
LinearMisuse _ _ Rig1 _
|
||||
=> do c <- check Rig1 elabinfo nest env scr
|
||||
=> do c <- check Rig1 elabinfo nest env scr
|
||||
(Just (gnf env scrtyv))
|
||||
pure (fst c, snd c, Rig1)
|
||||
e => throw e)
|
||||
|
@ -50,7 +50,7 @@ Show (ImplBinding vars) where
|
||||
export
|
||||
bindingMetas : ImplBinding vars -> NameMap Bool
|
||||
bindingMetas (NameBinding c p tm ty) = getMetas ty
|
||||
bindingMetas (AsBinding c p tm ty pat)
|
||||
bindingMetas (AsBinding c p tm ty pat)
|
||||
= insertAll (toList (getMetas ty)) (getMetas pat)
|
||||
where
|
||||
insertAll : List (Name, Bool) -> NameMap Bool -> NameMap Bool
|
||||
@ -92,7 +92,7 @@ record EState (vars : List Name) where
|
||||
outerEnv : Env Term outer
|
||||
subEnv : SubVars outer vars
|
||||
boundNames : List (Name, ImplBinding vars)
|
||||
-- implicit pattern/type variable bindings and the
|
||||
-- implicit pattern/type variable bindings and the
|
||||
-- term/type they elaborated to
|
||||
toBind : List (Name, ImplBinding vars)
|
||||
-- implicit pattern/type variables which haven't been
|
||||
@ -100,8 +100,8 @@ record EState (vars : List Name) where
|
||||
-- pattern vars need to be dealt with in with-application on
|
||||
-- the RHS)
|
||||
bindIfUnsolved : List (Name, RigCount, PiInfo,
|
||||
(vars' ** (Env Term vars', Term vars', Term vars',
|
||||
SubVars outer vars')))
|
||||
(vars' ** (Env Term vars', Term vars', Term vars',
|
||||
SubVars outer vars')))
|
||||
-- names to add as unbound implicits if they are still holes
|
||||
-- when unbound implicits are added
|
||||
lhsPatVars : List String
|
||||
@ -129,20 +129,20 @@ weakenedEState : {auto e : Ref EST (EState vars)} ->
|
||||
Core (Ref EST (EState (n :: vars)))
|
||||
weakenedEState {e}
|
||||
= do est <- get EST
|
||||
eref <- newRef EST
|
||||
(MkEState (defining est)
|
||||
eref <- newRef EST
|
||||
(MkEState (defining est)
|
||||
(outerEnv est)
|
||||
(DropCons (subEnv est))
|
||||
(map wknTms (boundNames est))
|
||||
(map wknTms (toBind est))
|
||||
(bindIfUnsolved est)
|
||||
(bindIfUnsolved est)
|
||||
(lhsPatVars est)
|
||||
(allPatVars est)
|
||||
(allowDelay est)
|
||||
(map weaken (linearUsed est)))
|
||||
pure eref
|
||||
where
|
||||
wknTms : (Name, ImplBinding vs) ->
|
||||
wknTms : (Name, ImplBinding vs) ->
|
||||
(Name, ImplBinding (n :: vs))
|
||||
wknTms (f, NameBinding c p x y) = (f, NameBinding c p (weaken x) (weaken y))
|
||||
wknTms (f, AsBinding c p x y z)
|
||||
@ -158,13 +158,13 @@ strengthenedEState {n} {vars} c e fc env
|
||||
svs <- dropSub (subEnv est)
|
||||
bns <- traverse (strTms defs) (boundNames est)
|
||||
todo <- traverse (strTms defs) (toBind est)
|
||||
|
||||
pure (MkEState (defining est)
|
||||
|
||||
pure (MkEState (defining est)
|
||||
(outerEnv est)
|
||||
svs
|
||||
bns
|
||||
bns
|
||||
todo
|
||||
(bindIfUnsolved est)
|
||||
(bindIfUnsolved est)
|
||||
(lhsPatVars est)
|
||||
(allPatVars est)
|
||||
(allowDelay est)
|
||||
@ -183,10 +183,10 @@ strengthenedEState {n} {vars} c e fc env
|
||||
-- in scope.
|
||||
removeArgVars : List (Term (n :: vs)) -> Maybe (List (Term vs))
|
||||
removeArgVars [] = pure []
|
||||
removeArgVars (Local fc r _ (Later p) :: args)
|
||||
removeArgVars (Local fc r _ (Later p) :: args)
|
||||
= do args' <- removeArgVars args
|
||||
pure (Local fc r _ p :: args')
|
||||
removeArgVars (Local fc r _ First :: args)
|
||||
removeArgVars (Local fc r _ First :: args)
|
||||
= removeArgVars args
|
||||
removeArgVars (a :: args)
|
||||
= do a' <- shrinkTerm a (DropCons SubRefl)
|
||||
@ -201,12 +201,12 @@ strengthenedEState {n} {vars} c e fc env
|
||||
f' <- shrinkTerm f (DropCons SubRefl)
|
||||
pure (apply (getLoc f) f' args')
|
||||
|
||||
strTms : Defs -> (Name, ImplBinding (n :: vars)) ->
|
||||
strTms : Defs -> (Name, ImplBinding (n :: vars)) ->
|
||||
Core (Name, ImplBinding vars)
|
||||
strTms defs (f, NameBinding c p x y)
|
||||
= do xnf <- normaliseHoles defs env x
|
||||
ynf <- normaliseHoles defs env y
|
||||
case (removeArg xnf,
|
||||
case (removeArg xnf,
|
||||
shrinkTerm ynf (DropCons SubRefl)) of
|
||||
(Just x', Just y') => pure (f, NameBinding c p x' y')
|
||||
_ => throw (BadUnboundImplicit fc env f y)
|
||||
@ -214,7 +214,7 @@ strengthenedEState {n} {vars} c e fc env
|
||||
= do xnf <- normaliseHoles defs env x
|
||||
ynf <- normaliseHoles defs env y
|
||||
znf <- normaliseHoles defs env z
|
||||
case (shrinkTerm xnf (DropCons SubRefl),
|
||||
case (shrinkTerm xnf (DropCons SubRefl),
|
||||
shrinkTerm ynf (DropCons SubRefl),
|
||||
shrinkTerm znf (DropCons SubRefl)) of
|
||||
(Just x', Just y', Just z') => pure (f, AsBinding c p x' y' z')
|
||||
@ -227,8 +227,8 @@ strengthenedEState {n} {vars} c e fc env
|
||||
export
|
||||
inScope : {auto c : Ref Ctxt Defs} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
FC -> Env Term (n :: vars) ->
|
||||
(Ref EST (EState (n :: vars)) -> Core a) ->
|
||||
FC -> Env Term (n :: vars) ->
|
||||
(Ref EST (EState (n :: vars)) -> Core a) ->
|
||||
Core a
|
||||
inScope {c} {e} fc env elab
|
||||
= do e' <- weakenedEState
|
||||
@ -238,7 +238,7 @@ inScope {c} {e} fc env elab
|
||||
pure res
|
||||
|
||||
export
|
||||
updateEnv : Env Term new -> SubVars new vars ->
|
||||
updateEnv : Env Term new -> SubVars new vars ->
|
||||
List (Name, RigCount, PiInfo, (vars' ** (Env Term vars', Term vars', Term vars', SubVars new vars'))) ->
|
||||
EState vars -> EState vars
|
||||
updateEnv env sub bif st
|
||||
@ -256,7 +256,7 @@ addBindIfUnsolved : Name -> RigCount -> PiInfo ->
|
||||
addBindIfUnsolved hn r p env tm ty st
|
||||
= MkEState (defining st)
|
||||
(outerEnv st) (subEnv st)
|
||||
(boundNames st) (toBind st)
|
||||
(boundNames st) (toBind st)
|
||||
((hn, r, p, (_ ** (env, tm, ty, subEnv st))) :: bindIfUnsolved st)
|
||||
(lhsPatVars st)
|
||||
(allPatVars st)
|
||||
@ -279,7 +279,7 @@ clearToBind : {auto e : Ref EST (EState vars)} ->
|
||||
(excepts : List Name) -> Core ()
|
||||
clearToBind excepts
|
||||
= do est <- get EST
|
||||
put EST (record { toBind $= filter (\x => fst x `elem` excepts) }
|
||||
put EST (record { toBind $= filter (\x => fst x `elem` excepts) }
|
||||
(clearBindIfUnsolved est))
|
||||
|
||||
export
|
||||
@ -403,7 +403,7 @@ successful : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
List (Maybe Name, Core a) ->
|
||||
Core (List (Either (Maybe Name, Error)
|
||||
Core (List (Either (Maybe Name, Error)
|
||||
(a, Defs, UState, EState vars)))
|
||||
successful [] = pure []
|
||||
successful ((tm, elab) :: elabs)
|
||||
@ -411,7 +411,7 @@ successful ((tm, elab) :: elabs)
|
||||
est <- get EST
|
||||
md <- get MD
|
||||
defs <- branch
|
||||
catch (do -- Run the elaborator
|
||||
catch (do -- Run the elaborator
|
||||
log 5 $ "Running " ++ show tm
|
||||
res <- elab
|
||||
-- Record post-elaborator state
|
||||
@ -440,14 +440,14 @@ exactlyOne : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
FC -> Env Term vars ->
|
||||
FC -> Env Term vars ->
|
||||
List (Maybe Name, Core (Term vars, Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
exactlyOne fc env [(tm, elab)] = elab
|
||||
exactlyOne {vars} fc env all
|
||||
= do elabs <- successful all
|
||||
case rights elabs of
|
||||
[(res, defs, ust, est)] =>
|
||||
[(res, defs, ust, est)] =>
|
||||
do put UST ust
|
||||
put EST est
|
||||
put Ctxt defs
|
||||
@ -460,7 +460,7 @@ exactlyOne {vars} fc env all
|
||||
|
||||
-- If they've all failed, collect all the errors
|
||||
-- If more than one succeeded, report the ambiguity
|
||||
altError : List (Maybe Name, Error) ->
|
||||
altError : List (Maybe Name, Error) ->
|
||||
List ((Term vars, Glued vars), st) ->
|
||||
Error
|
||||
altError ls [] = AllFailed ls
|
||||
@ -474,7 +474,7 @@ anyOne : {vars : _} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
FC -> List (Maybe Name, Core (Term vars, Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
anyOne fc [] = throw (GenericMsg fc "No elaborators provided")
|
||||
anyOne fc [] = throw (GenericMsg fc "No elaborators provided")
|
||||
anyOne fc [(tm, elab)] = elab
|
||||
anyOne fc ((tm, elab) :: es) = try elab (anyOne fc es)
|
||||
|
||||
@ -486,8 +486,8 @@ check : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars -> RawImp ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars -> RawImp ->
|
||||
Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
|
||||
@ -498,7 +498,7 @@ checkImp : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
|
||||
@ -508,9 +508,9 @@ processDecl : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
Env Term vars -> ImpDecl -> Core ()
|
||||
|
||||
|
||||
-- Check whether two terms are convertible. May solve metavariables (in Ctxt)
|
||||
-- in doing so.
|
||||
-- Returns a list of constraints which need to be solved for the conversion
|
||||
@ -528,7 +528,7 @@ convertWithLazy withLazy fc elabinfo env x y
|
||||
= case elabMode elabinfo of
|
||||
InLHS _ => InLHS
|
||||
_ => InTerm in
|
||||
catch
|
||||
catch
|
||||
(do logGlueNF 5 "Unifying" env x
|
||||
logGlueNF 5 "....with" env y
|
||||
vs <- if isFromTerm x && isFromTerm y
|
||||
@ -545,7 +545,7 @@ convertWithLazy withLazy fc elabinfo env x y
|
||||
when (holesSolved vs) $
|
||||
solveConstraints umode Normal
|
||||
pure vs)
|
||||
(\err =>
|
||||
(\err =>
|
||||
do defs <- get Ctxt
|
||||
xtm <- getTerm x
|
||||
ytm <- getTerm y
|
||||
@ -579,19 +579,19 @@ checkExp : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo -> Env Term vars -> FC ->
|
||||
(term : Term vars) ->
|
||||
(got : Glued vars) -> (expected : Maybe (Glued vars)) ->
|
||||
(term : Term vars) ->
|
||||
(got : Glued vars) -> (expected : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkExp rig elabinfo env fc tm got (Just exp)
|
||||
checkExp rig elabinfo env fc tm got (Just exp)
|
||||
= do vs <- convertWithLazy True fc elabinfo env got exp
|
||||
case (constraints vs) of
|
||||
[] => case addLazy vs of
|
||||
NoLazy => do logTerm 5 "Solved" tm
|
||||
pure (tm, got)
|
||||
AddForce => do logTerm 5 "Force" tm
|
||||
logGlue 5 "Got" env got
|
||||
logGlue 5 "Exp" env exp
|
||||
pure (TForce fc tm, exp)
|
||||
AddForce r => do logTerm 5 "Force" tm
|
||||
logGlue 5 "Got" env got
|
||||
logGlue 5 "Exp" env exp
|
||||
pure (TForce fc r tm, exp)
|
||||
AddDelay r => do ty <- getTerm got
|
||||
logTerm 5 "Delay" tm
|
||||
pure (TDelay fc r ty tm, exp)
|
||||
@ -603,7 +603,7 @@ checkExp rig elabinfo env fc tm got (Just exp)
|
||||
dumpConstraints 5 False
|
||||
case addLazy vs of
|
||||
NoLazy => pure (ctm, got)
|
||||
AddForce => pure (TForce fc tm, got)
|
||||
AddForce r => pure (TForce fc r tm, got)
|
||||
AddDelay r => do ty <- getTerm got
|
||||
pure (TDelay fc r ty tm, got)
|
||||
checkExp rig elabinfo env fc tm got Nothing = pure (tm, got)
|
||||
|
@ -18,15 +18,15 @@ import Data.IntMap
|
||||
%default covering
|
||||
|
||||
-- We run the elaborator in the given environment, but need to end up with a
|
||||
-- closed term.
|
||||
mkClosedElab : FC -> Env Term vars ->
|
||||
-- closed term.
|
||||
mkClosedElab : FC -> Env Term vars ->
|
||||
(Core (Term vars, Glued vars)) ->
|
||||
Core ClosedTerm
|
||||
mkClosedElab fc [] elab
|
||||
mkClosedElab fc [] elab
|
||||
= do (tm, _) <- elab
|
||||
pure tm
|
||||
mkClosedElab {vars = x :: vars} fc (b :: env) elab
|
||||
= mkClosedElab fc env
|
||||
= mkClosedElab fc env
|
||||
(do (sc', _) <- elab
|
||||
let b' = newBinder b
|
||||
pure (Bind fc x b' sc', gErased fc))
|
||||
@ -42,50 +42,50 @@ mkClosedElab {vars = x :: vars} fc (b :: env) elab
|
||||
-- predicate, make a hole and try it again later when more holes might
|
||||
-- have been resolved
|
||||
export
|
||||
delayOnFailure : {auto c : Ref Ctxt Defs} ->
|
||||
delayOnFailure : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
FC -> RigCount -> Env Term vars ->
|
||||
(expected : Glued vars) ->
|
||||
(Error -> Bool) ->
|
||||
(Bool -> Core (Term vars, Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
delayOnFailure fc rig env expected pred elab
|
||||
delayOnFailure fc rig env expected pred elab
|
||||
= handle (elab False)
|
||||
(\err =>
|
||||
(\err =>
|
||||
do est <- get EST
|
||||
if pred err && allowDelay est
|
||||
then
|
||||
then
|
||||
do nm <- genName "delayed"
|
||||
(ci, dtm) <- newDelayed fc Rig1 env nm !(getTerm expected)
|
||||
logGlueNF 5 ("Postponing elaborator " ++ show nm ++
|
||||
logGlueNF 5 ("Postponing elaborator " ++ show nm ++
|
||||
" at " ++ show fc ++
|
||||
" for") env expected
|
||||
log 10 ("Due to error " ++ show err)
|
||||
ust <- get UST
|
||||
put UST (record { delayedElab $=
|
||||
((ci, mkClosedElab fc env
|
||||
put UST (record { delayedElab $=
|
||||
((ci, mkClosedElab fc env
|
||||
(do est <- get EST
|
||||
put EST (record { allowDelay = False } est)
|
||||
tm <- elab True
|
||||
est <- get EST
|
||||
put EST (record { allowDelay = True } est)
|
||||
pure tm)) :: ) }
|
||||
pure tm)) :: ) }
|
||||
ust)
|
||||
pure (dtm, expected)
|
||||
else throw err)
|
||||
|
||||
export
|
||||
delayElab : {auto c : Ref Ctxt Defs} ->
|
||||
delayElab : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
FC -> RigCount -> Env Term vars ->
|
||||
(expected : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
delayElab {vars} fc rig env exp elab
|
||||
delayElab {vars} fc rig env exp elab
|
||||
= do est <- get EST
|
||||
if not (allowDelay est)
|
||||
then elab
|
||||
@ -93,17 +93,17 @@ delayElab {vars} fc rig env exp elab
|
||||
nm <- genName "delayed"
|
||||
expected <- mkExpected exp
|
||||
(ci, dtm) <- newDelayed fc Rig1 env nm !(getTerm expected)
|
||||
logGlueNF 5 ("Postponing elaborator " ++ show nm ++
|
||||
logGlueNF 5 ("Postponing elaborator " ++ show nm ++
|
||||
" for") env expected
|
||||
ust <- get UST
|
||||
put UST (record { delayedElab $=
|
||||
((ci, mkClosedElab fc env
|
||||
put UST (record { delayedElab $=
|
||||
((ci, mkClosedElab fc env
|
||||
(do est <- get EST
|
||||
put EST (record { allowDelay = False } est)
|
||||
tm <- elab
|
||||
est <- get EST
|
||||
put EST (record { allowDelay = True } est)
|
||||
pure tm)) :: ) }
|
||||
pure tm)) :: ) }
|
||||
ust)
|
||||
pure (dtm, expected)
|
||||
where
|
||||
@ -115,9 +115,9 @@ delayElab {vars} fc rig env exp elab
|
||||
pure (gnf env ty)
|
||||
|
||||
export
|
||||
retryDelayed : {auto c : Ref Ctxt Defs} ->
|
||||
retryDelayed : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
List (Int, Core ClosedTerm) ->
|
||||
Core ()
|
||||
retryDelayed [] = pure ()
|
||||
@ -127,7 +127,7 @@ retryDelayed ((i, elab) :: ds)
|
||||
| _ => retryDelayed ds
|
||||
log 5 ("Retrying delayed hole " ++ show !(getFullName (Resolved i)))
|
||||
tm <- elab
|
||||
updateDef (Resolved i) (const (Just
|
||||
updateDef (Resolved i) (const (Just
|
||||
(PMDef True [] (STerm tm) (STerm tm) [])))
|
||||
logTerm 5 ("Resolved delayed hole " ++ show i) tm
|
||||
logTermNF 5 ("Resolved delayed hole NF " ++ show i) [] tm
|
||||
|
@ -23,26 +23,26 @@ checkDot : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> String -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkDot rig elabinfo nest env fc reason tm Nothing
|
||||
= throw (GenericMsg fc ("Dot pattern not valid here (unknown type) "
|
||||
= throw (GenericMsg fc ("Dot pattern not valid here (unknown type) "
|
||||
++ show tm))
|
||||
checkDot rig elabinfo nest env fc reason tm (Just gexpty)
|
||||
= case elabMode elabinfo of
|
||||
InLHS _ =>
|
||||
do (wantedTm, wantedTy) <- checkImp rig
|
||||
do (wantedTm, wantedTy) <- checkImp rig
|
||||
(record { elabMode = InExpr }
|
||||
elabinfo)
|
||||
nest env
|
||||
nest env
|
||||
tm (Just gexpty)
|
||||
nm <- genName "dotTm"
|
||||
expty <- getTerm gexpty
|
||||
metaval <- metaVar fc rig env nm expty
|
||||
addDot fc env nm wantedTm reason metaval
|
||||
checkExp rig elabinfo env fc metaval wantedTy (Just gexpty)
|
||||
_ => throw (GenericMsg fc ("Dot pattern not valid here (Not LHS) "
|
||||
_ => throw (GenericMsg fc ("Dot pattern not valid here (Not LHS) "
|
||||
++ show tm))
|
||||
|
||||
|
@ -21,7 +21,7 @@ checkHole : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> String -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkHole rig elabinfo nest env fc n_in (Just gexpty)
|
||||
|
@ -19,11 +19,11 @@ import Data.NameMap
|
||||
|
||||
%default covering
|
||||
|
||||
varEmbedSub : SubVars small vars ->
|
||||
{idx : Nat} -> .(IsVar n idx small) ->
|
||||
varEmbedSub : SubVars small vars ->
|
||||
{idx : Nat} -> .(IsVar n idx small) ->
|
||||
Var vars
|
||||
varEmbedSub SubRefl y = MkVar y
|
||||
varEmbedSub (DropCons prf) y
|
||||
varEmbedSub (DropCons prf) y
|
||||
= let MkVar y' = varEmbedSub prf y in
|
||||
MkVar (Later y')
|
||||
varEmbedSub (KeepCons prf) First = MkVar First
|
||||
@ -33,21 +33,21 @@ varEmbedSub (KeepCons prf) (Later p)
|
||||
|
||||
mutual
|
||||
embedSub : SubVars small vars -> Term small -> Term vars
|
||||
embedSub sub (Local fc x idx y)
|
||||
embedSub sub (Local fc x idx y)
|
||||
= let MkVar y' = varEmbedSub sub y in Local fc x _ y'
|
||||
embedSub sub (Ref fc x name) = Ref fc x name
|
||||
embedSub sub (Meta fc x y xs)
|
||||
embedSub sub (Meta fc x y xs)
|
||||
= Meta fc x y (map (embedSub sub) xs)
|
||||
embedSub sub (Bind fc x b scope)
|
||||
embedSub sub (Bind fc x b scope)
|
||||
= Bind fc x (map (embedSub sub) b) (embedSub (KeepCons sub) scope)
|
||||
embedSub sub (App fc fn arg)
|
||||
embedSub sub (App fc fn arg)
|
||||
= App fc (embedSub sub fn) (embedSub sub arg)
|
||||
embedSub sub (As fc nm pat)
|
||||
embedSub sub (As fc nm pat)
|
||||
= As fc (embedSub sub nm) (embedSub sub pat)
|
||||
embedSub sub (TDelayed fc x y) = TDelayed fc x (embedSub sub y)
|
||||
embedSub sub (TDelay fc x t y)
|
||||
embedSub sub (TDelay fc x t y)
|
||||
= TDelay fc x (embedSub sub t) (embedSub sub y)
|
||||
embedSub sub (TForce fc x) = TForce fc (embedSub sub x)
|
||||
embedSub sub (TForce fc r x) = TForce fc r (embedSub sub x)
|
||||
embedSub sub (PrimVal fc c) = PrimVal fc c
|
||||
embedSub sub (Erased fc) = Erased fc
|
||||
embedSub sub (TType fc) = TType fc
|
||||
@ -67,7 +67,7 @@ mkOuterHole loc rig n topenv (Just expty_in)
|
||||
case shrinkTerm expected sub of
|
||||
-- Can't shrink so rely on unification with expected type later
|
||||
Nothing => mkOuterHole loc rig n topenv Nothing
|
||||
Just exp' =>
|
||||
Just exp' =>
|
||||
do let env = outerEnv est
|
||||
tm <- implBindVar loc rig env n exp'
|
||||
pure (embedSub sub tm, embedSub sub exp')
|
||||
@ -106,7 +106,7 @@ mkPatternHole {vars} loc rig n topenv imode (Just expty_in)
|
||||
Nothing => mkPatternHole loc rig n topenv imode Nothing
|
||||
Just exp' =>
|
||||
do tm <- implBindVar loc rig env n exp'
|
||||
pure (apply loc (embedSub sub tm) (mkArgs sub),
|
||||
pure (apply loc (embedSub sub tm) (mkArgs sub),
|
||||
expected,
|
||||
embedSub sub exp')
|
||||
where
|
||||
@ -118,7 +118,7 @@ mkPatternHole {vars} loc rig n topenv imode (Just expty_in)
|
||||
-- This is for the specific situation where we're pattern matching on
|
||||
-- function types, which is realistically the only time we'll legitimately
|
||||
-- encounter a type variable under a binder
|
||||
bindInner : Env Term vs -> Term vs -> SubVars newvars vs ->
|
||||
bindInner : Env Term vs -> Term vs -> SubVars newvars vs ->
|
||||
Maybe (Term newvars)
|
||||
bindInner env ty SubRefl = Just ty
|
||||
bindInner {vs = x :: _} (b :: env) ty (DropCons p)
|
||||
@ -136,7 +136,7 @@ bindUnsolved : {auto c : Ref Ctxt Defs} -> {auto e : Ref EST (EState vars)} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> ElabMode -> BindMode -> Core ()
|
||||
bindUnsolved fc elabmode NONE = pure ()
|
||||
bindUnsolved {vars} fc elabmode _
|
||||
bindUnsolved {vars} fc elabmode _
|
||||
= do est <- get EST
|
||||
defs <- get Ctxt
|
||||
let bifs = bindIfUnsolved est
|
||||
@ -150,7 +150,7 @@ bindUnsolved {vars} fc elabmode _
|
||||
= case shrinkTerm expected sub of
|
||||
Nothing => do tmn <- toFullNames expected
|
||||
throw (GenericMsg fc ("Can't bind implicit " ++ show n ++ " of type " ++ show tmn))
|
||||
Just exp' =>
|
||||
Just exp' =>
|
||||
do impn <- genVarName (nameRoot n)
|
||||
tm <- metaVar fc rig env impn exp'
|
||||
est <- get EST
|
||||
@ -160,8 +160,8 @@ bindUnsolved {vars} fc elabmode _
|
||||
pure (embedSub sub tm)
|
||||
|
||||
mkImplicit : Defs -> Env Term outer -> SubVars outer vars ->
|
||||
(Name, RigCount, PiInfo, (vars' **
|
||||
(Env Term vars', Term vars', Term vars', SubVars outer vars'))) ->
|
||||
(Name, RigCount, PiInfo, (vars' **
|
||||
(Env Term vars', Term vars', Term vars', SubVars outer vars'))) ->
|
||||
Core ()
|
||||
mkImplicit defs outerEnv subEnv (n, rig, p, (vs ** (env, tm, exp, sub)))
|
||||
= do Just (Hole _ _) <- lookupDefExact n (gamma defs)
|
||||
@ -176,33 +176,33 @@ bindUnsolved {vars} fc elabmode _
|
||||
fc env tm bindtm
|
||||
pure ()
|
||||
|
||||
swapIsVarH : {idx : Nat} -> .(IsVar name idx (x :: y :: xs)) ->
|
||||
swapIsVarH : {idx : Nat} -> .(IsVar name idx (x :: y :: xs)) ->
|
||||
Var (y :: x :: xs)
|
||||
swapIsVarH First = MkVar (Later First)
|
||||
swapIsVarH (Later First) = MkVar First
|
||||
swapIsVarH (Later (Later x)) = MkVar (Later (Later x))
|
||||
|
||||
swapIsVar : (vs : List Name) ->
|
||||
{idx : Nat} -> .(IsVar name idx (vs ++ x :: y :: xs)) ->
|
||||
{idx : Nat} -> .(IsVar name idx (vs ++ x :: y :: xs)) ->
|
||||
Var (vs ++ y :: x :: xs)
|
||||
swapIsVar [] prf = swapIsVarH prf
|
||||
swapIsVar (x :: xs) First = MkVar First
|
||||
swapIsVar (x :: xs) (Later p)
|
||||
swapIsVar (x :: xs) (Later p)
|
||||
= let MkVar p' = swapIsVar xs p in MkVar (Later p')
|
||||
|
||||
swapVars : {vs : List Name} ->
|
||||
Term (vs ++ x :: y :: ys) -> Term (vs ++ y :: x :: ys)
|
||||
swapVars (Local fc x idx p)
|
||||
swapVars (Local fc x idx p)
|
||||
= let MkVar p' = swapIsVar _ p in Local fc x _ p'
|
||||
swapVars (Ref fc x name) = Ref fc x name
|
||||
swapVars (Meta fc n i xs) = Meta fc n i (map swapVars xs)
|
||||
swapVars {vs} (Bind fc x b scope)
|
||||
swapVars {vs} (Bind fc x b scope)
|
||||
= Bind fc x (map swapVars b) (swapVars {vs = x :: vs} scope)
|
||||
swapVars (App fc fn arg) = App fc (swapVars fn) (swapVars arg)
|
||||
swapVars (As fc nm pat) = As fc (swapVars nm) (swapVars pat)
|
||||
swapVars (TDelayed fc x tm) = TDelayed fc x (swapVars tm)
|
||||
swapVars (TDelay fc x ty tm) = TDelay fc x (swapVars ty) (swapVars tm)
|
||||
swapVars (TForce fc tm) = TForce fc (swapVars tm)
|
||||
swapVars (TForce fc r tm) = TForce fc r (swapVars tm)
|
||||
swapVars (PrimVal fc c) = PrimVal fc c
|
||||
swapVars (Erased fc) = Erased fc
|
||||
swapVars (TType fc) = TType fc
|
||||
@ -215,7 +215,7 @@ push ofc n b tm@(Bind fc (PV x i) (Pi c Implicit ty) sc) -- only push past 'PV's
|
||||
= case shrinkTerm ty (DropCons SubRefl) of
|
||||
Nothing => -- needs explicit pi, do nothing
|
||||
Bind ofc n b tm
|
||||
Just ty' => Bind fc (PV x i) (Pi c Implicit ty')
|
||||
Just ty' => Bind fc (PV x i) (Pi c Implicit ty')
|
||||
(push ofc n (map weaken b) (swapVars {vs = []} sc))
|
||||
push ofc n b tm = Bind ofc n b tm
|
||||
|
||||
@ -227,7 +227,7 @@ liftImps : BindMode -> (Term vars, Term vars) -> (Term vars, Term vars)
|
||||
liftImps (PI _) (tm, TType) = (liftImps' tm, TType)
|
||||
where
|
||||
liftImps' : Term vars -> Term vars
|
||||
liftImps' (Bind fc (PV n i) (Pi c Implicit ty) sc)
|
||||
liftImps' (Bind fc (PV n i) (Pi c Implicit ty) sc)
|
||||
= Bind fc (PV n i) (Pi c Implicit ty) (liftImps' sc)
|
||||
liftImps' (Bind fc n (Pi c p ty) sc)
|
||||
= push fc n (Pi c p ty) (liftImps' sc)
|
||||
@ -241,7 +241,7 @@ bindImplVars : FC -> BindMode ->
|
||||
List (Name, ImplBinding vars) ->
|
||||
Term vars -> Term vars -> (Term vars, Term vars)
|
||||
bindImplVars fc NONE gam env imps_in scope scty = (scope, scty)
|
||||
bindImplVars {vars} fc mode gam env imps_in scope scty
|
||||
bindImplVars {vars} fc mode gam env imps_in scope scty
|
||||
= let imps = map (\ (x, bind) => (tidyName x, x, bind)) imps_in in
|
||||
getBinds imps None scope scty
|
||||
where
|
||||
@ -253,31 +253,31 @@ bindImplVars {vars} fc mode gam env imps_in scope scty
|
||||
tidyName (Nested n inner) = tidyName inner
|
||||
tidyName n = n
|
||||
|
||||
getBinds : (imps : List (Name, Name, ImplBinding vs)) ->
|
||||
getBinds : (imps : List (Name, Name, ImplBinding vs)) ->
|
||||
Bounds new -> (tm : Term vs) -> (ty : Term vs) ->
|
||||
(Term (new ++ vs), Term (new ++ vs))
|
||||
getBinds [] bs tm ty = (refsToLocals bs tm, refsToLocals bs ty)
|
||||
getBinds ((n, metan, NameBinding c p _ bty) :: imps) bs tm ty
|
||||
= let (tm', ty') = getBinds imps (Add n metan bs) tm ty
|
||||
= let (tm', ty') = getBinds imps (Add n metan bs) tm ty
|
||||
bty' = refsToLocals bs bty in
|
||||
case mode of
|
||||
PI c =>
|
||||
(Bind fc _ (Pi c Implicit bty') tm',
|
||||
(Bind fc _ (Pi c Implicit bty') tm',
|
||||
TType fc)
|
||||
_ =>
|
||||
(Bind fc _ (PVar c p bty') tm',
|
||||
(Bind fc _ (PVar c p bty') tm',
|
||||
Bind fc _ (PVTy c bty') ty')
|
||||
getBinds ((n, metan, AsBinding c _ _ bty bpat) :: imps) bs tm ty
|
||||
= let (tm', ty') = getBinds imps (Add n metan bs) tm ty
|
||||
= let (tm', ty') = getBinds imps (Add n metan bs) tm ty
|
||||
bty' = refsToLocals bs bty
|
||||
bpat' = refsToLocals bs bpat in
|
||||
(Bind fc _ (PLet c bpat' bty') tm',
|
||||
(Bind fc _ (PLet c bpat' bty') tm',
|
||||
Bind fc _ (PLet c bpat' bty') ty')
|
||||
|
||||
normaliseHolesScope : Defs -> Env Term vars -> Term vars -> Core (Term vars)
|
||||
normaliseHolesScope defs env (Bind fc n b sc)
|
||||
= pure $ Bind fc n b
|
||||
!(normaliseHolesScope defs
|
||||
normaliseHolesScope defs env (Bind fc n b sc)
|
||||
= pure $ Bind fc n b
|
||||
!(normaliseHolesScope defs
|
||||
-- use Lam because we don't want it reducing in the scope
|
||||
(Lam (multiplicity b) Explicit (binderType b) :: env) sc)
|
||||
normaliseHolesScope defs env tm = normaliseHoles defs env tm
|
||||
@ -288,7 +288,7 @@ bindImplicits : FC -> BindMode ->
|
||||
List (Name, ImplBinding vars) ->
|
||||
Term vars -> Term vars -> Core (Term vars, Term vars)
|
||||
bindImplicits fc NONE defs env hs tm ty = pure (tm, ty)
|
||||
bindImplicits {vars} fc mode defs env hs tm ty
|
||||
bindImplicits {vars} fc mode defs env hs tm ty
|
||||
= do hs' <- traverse nHoles hs
|
||||
pure $ liftImps mode $ bindImplVars fc mode defs env hs' tm ty
|
||||
where
|
||||
@ -302,7 +302,7 @@ export
|
||||
implicitBind : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
Name -> Core ()
|
||||
implicitBind n
|
||||
implicitBind n
|
||||
= do defs <- get Ctxt
|
||||
Just (Hole _ _) <- lookupDefExact n (gamma defs)
|
||||
| _ => pure ()
|
||||
@ -320,9 +320,9 @@ getToBind : {auto c : Ref Ctxt Defs} -> {auto e : Ref EST (EState vars)} ->
|
||||
FC -> ElabMode -> BindMode ->
|
||||
Env Term vars -> (excepts : List Name) ->
|
||||
Core (List (Name, ImplBinding vars))
|
||||
getToBind fc elabmode NONE env excepts
|
||||
getToBind fc elabmode NONE env excepts
|
||||
= pure [] -- We should probably never get here, but for completeness...
|
||||
getToBind {vars} fc elabmode impmode env excepts
|
||||
getToBind {vars} fc elabmode impmode env excepts
|
||||
= do solveConstraints (case elabmode of
|
||||
InLHS _ => InLHS
|
||||
_ => InTerm) Normal
|
||||
@ -337,7 +337,7 @@ getToBind {vars} fc elabmode impmode env excepts
|
||||
-- Make sure all the hole names are normalised in the implicitly
|
||||
-- bound types, because otherwise we'll bind them too
|
||||
res <- normImps defs [] tob
|
||||
let hnames = map fst res
|
||||
let hnames = map fst res
|
||||
-- Return then in dependency order
|
||||
let res' = depSort hnames res
|
||||
log 10 $ "Bound names: " ++ show res
|
||||
@ -348,14 +348,14 @@ getToBind {vars} fc elabmode impmode env excepts
|
||||
normBindingTy defs (NameBinding c p tm ty)
|
||||
= pure $ NameBinding c p tm !(normaliseHoles defs env ty)
|
||||
normBindingTy defs (AsBinding c p tm ty pat)
|
||||
= pure $ AsBinding c p tm !(normaliseHoles defs env ty)
|
||||
= pure $ AsBinding c p tm !(normaliseHoles defs env ty)
|
||||
!(normaliseHoles defs env pat)
|
||||
|
||||
normImps : Defs -> List Name -> List (Name, ImplBinding vars) ->
|
||||
normImps : Defs -> List Name -> List (Name, ImplBinding vars) ->
|
||||
Core (List (Name, ImplBinding vars))
|
||||
normImps defs ns [] = pure []
|
||||
normImps defs ns ((PV n i, bty) :: ts)
|
||||
= do logTermNF 10 ("Implicit pattern var " ++ show (PV n i)) env
|
||||
normImps defs ns ((PV n i, bty) :: ts)
|
||||
= do logTermNF 10 ("Implicit pattern var " ++ show (PV n i)) env
|
||||
(bindingType bty)
|
||||
if PV n i `elem` ns
|
||||
then normImps defs ns ts
|
||||
@ -379,8 +379,8 @@ getToBind {vars} fc elabmode impmode env excepts
|
||||
|
||||
-- Insert the hole/binding pair into the list before the first thing
|
||||
-- which refers to it
|
||||
insert : (Name, ImplBinding vars) -> List Name -> List Name ->
|
||||
List (Name, ImplBinding vars) ->
|
||||
insert : (Name, ImplBinding vars) -> List Name -> List Name ->
|
||||
List (Name, ImplBinding vars) ->
|
||||
List (Name, ImplBinding vars)
|
||||
insert h ns sofar [] = [h]
|
||||
insert (hn, bty) ns sofar ((hn', bty') :: rest)
|
||||
@ -389,14 +389,14 @@ getToBind {vars} fc elabmode impmode env excepts
|
||||
-- introduced in *this* expression (there may be others unresolved
|
||||
-- from elsewhere, for type inference purposes)
|
||||
if hn `elem` used
|
||||
then (hn, bty) ::
|
||||
then (hn, bty) ::
|
||||
(hn', bty') :: rest
|
||||
else (hn', bty') ::
|
||||
else (hn', bty') ::
|
||||
insert (hn, bty) ns (hn' :: sofar) rest
|
||||
|
||||
|
||||
-- Sort the list of implicits so that each binding is inserted *after*
|
||||
-- all the things it depends on (assumes no cycles)
|
||||
depSort : List Name -> List (Name, ImplBinding vars) ->
|
||||
depSort : List Name -> List (Name, ImplBinding vars) ->
|
||||
List (Name, ImplBinding vars)
|
||||
depSort hnames [] = []
|
||||
depSort hnames (h :: hs) = insert h hnames [] (depSort hnames hs)
|
||||
@ -407,9 +407,9 @@ checkBindVar : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> String -> -- string is base of the pattern name
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> String -> -- string is base of the pattern name
|
||||
Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkBindVar rig elabinfo nest env fc str topexp
|
||||
@ -426,7 +426,7 @@ checkBindVar rig elabinfo nest env fc str topexp
|
||||
notePatVar n
|
||||
est <- get EST
|
||||
case lookup n (boundNames est) of
|
||||
Nothing =>
|
||||
Nothing =>
|
||||
do (tm, exp, bty) <- mkPatternHole fc rig n env
|
||||
(implicitMode elabinfo)
|
||||
topexp
|
||||
@ -451,7 +451,7 @@ checkBindVar rig elabinfo nest env fc str topexp
|
||||
addNameType fc (UN str) env ty
|
||||
checkExp rig elabinfo env fc tm (gnf env ty) topexp
|
||||
where
|
||||
updateRig : Name -> RigCount -> List (Name, ImplBinding vars) ->
|
||||
updateRig : Name -> RigCount -> List (Name, ImplBinding vars) ->
|
||||
List (Name, ImplBinding vars)
|
||||
updateRig n c [] = []
|
||||
updateRig n c ((bn, r) :: bs)
|
||||
@ -467,7 +467,7 @@ checkBindVar rig elabinfo nest env fc str topexp
|
||||
combine n RigW Rig1 = throw (LinearUsed fc 2 n)
|
||||
combine n RigW RigW = pure ()
|
||||
combine n Rig0 c = pure ()
|
||||
combine n c Rig0
|
||||
combine n c Rig0
|
||||
-- It was 0, make it c
|
||||
= do est <- get EST
|
||||
put EST (record { boundNames $= updateRig n c,
|
||||
@ -479,8 +479,8 @@ checkBindHere : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> BindMode -> RawImp ->
|
||||
Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
@ -518,12 +518,12 @@ checkBindHere rig elabinfo nest env fc bindmode tm exp
|
||||
bindmode env dontbind
|
||||
clearToBind dontbind
|
||||
est <- get EST
|
||||
put EST (updateEnv oldenv oldsub oldbif
|
||||
put EST (updateEnv oldenv oldsub oldbif
|
||||
(record { boundNames = [] } est))
|
||||
ty <- getTerm tmt
|
||||
defs <- get Ctxt
|
||||
(bv, bt) <- bindImplicits fc bindmode
|
||||
defs env argImps
|
||||
defs env argImps
|
||||
!(normaliseHoles defs env tmv)
|
||||
!(normaliseHoles defs env ty)
|
||||
traverse_ implicitBind (map fst argImps)
|
||||
|
@ -22,7 +22,7 @@ checkDelayed : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> LazyReason -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkDelayed rig elabinfo nest env fc r tm exp
|
||||
@ -36,7 +36,7 @@ checkDelay : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkDelay rig elabinfo nest env fc tm mexpected
|
||||
@ -59,7 +59,7 @@ checkDelay rig elabinfo nest env fc tm mexpected
|
||||
(Just (glueBack defs env expnf))
|
||||
tynf <- getNF gty
|
||||
ty <- getTerm gty
|
||||
pure (TDelay fc r ty tm',
|
||||
pure (TDelay fc r ty tm',
|
||||
glueBack defs env (NDelayed fc r tynf))
|
||||
_ => throw (GenericMsg fc ("Can't infer delay type")))
|
||||
where
|
||||
@ -74,19 +74,19 @@ checkForce : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkForce rig elabinfo nest env fc tm exp
|
||||
= do defs <- get Ctxt
|
||||
expf <- maybe (pure Nothing)
|
||||
(\gty => do tynf <- getNF gty
|
||||
pure (Just (glueBack defs env
|
||||
pure (Just (glueBack defs env
|
||||
(NDelayed fc LUnknown tynf))))
|
||||
exp
|
||||
(tm', gty) <- check rig elabinfo nest env tm expf
|
||||
tynf <- getNF gty
|
||||
case tynf of
|
||||
NDelayed _ _ expnf =>
|
||||
pure (TForce fc tm', glueBack defs env expnf)
|
||||
NDelayed _ r expnf =>
|
||||
pure (TForce fc r tm', glueBack defs env expnf)
|
||||
_ => throw (GenericMsg fc "Forcing a non-delayed type")
|
||||
|
@ -20,16 +20,16 @@ checkLocal : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> List ImpDecl -> (scope : RawImp) ->
|
||||
(expTy : Maybe (Glued vars)) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkLocal {vars} rig elabinfo nest env fc nestdecls scope expty
|
||||
= do let defNames = definedInBlock [] nestdecls
|
||||
est <- get EST
|
||||
let f = defining est
|
||||
names' <- traverse (applyEnv f)
|
||||
est <- get EST
|
||||
let f = defining est
|
||||
names' <- traverse (applyEnv f)
|
||||
(nub defNames) -- binding names must be unique
|
||||
-- fixes bug #115
|
||||
let nest' = record { names $= (names' ++) } nest
|
||||
@ -43,31 +43,31 @@ checkLocal {vars} rig elabinfo nest env fc nestdecls scope expty
|
||||
-- ensuring the nested definition is used exactly once
|
||||
dropLinear : Env Term vs -> Env Term vs
|
||||
dropLinear [] = []
|
||||
dropLinear (b :: bs)
|
||||
dropLinear (b :: bs)
|
||||
= if isLinear (multiplicity b)
|
||||
then setMultiplicity b Rig0 :: dropLinear bs
|
||||
else b :: dropLinear bs
|
||||
|
||||
applyEnv : {auto c : Ref Ctxt Defs} -> Int -> Name ->
|
||||
applyEnv : {auto c : Ref Ctxt Defs} -> Int -> Name ->
|
||||
Core (Name, (Maybe Name, FC -> NameType -> Term vars))
|
||||
applyEnv outer inner
|
||||
applyEnv outer inner
|
||||
= do let nestedName = Nested outer inner
|
||||
n' <- addName nestedName
|
||||
pure (inner, (Just nestedName,
|
||||
\fc, nt => applyTo fc
|
||||
pure (inner, (Just nestedName,
|
||||
\fc, nt => applyTo fc
|
||||
(Ref fc nt (Resolved n')) env))
|
||||
|
||||
-- Update the names in the declarations to the new 'nested' names.
|
||||
-- When we encounter the names in elaboration, we'll update to an
|
||||
-- application of the nested name.
|
||||
newName : NestedNames vars -> Name -> Name
|
||||
newName nest n
|
||||
newName nest n
|
||||
= case lookup n (names nest) of
|
||||
Just (Just n', _) => n'
|
||||
_ => n
|
||||
|
||||
updateTyName : NestedNames vars -> ImpTy -> ImpTy
|
||||
updateTyName nest (MkImpTy loc' n ty)
|
||||
updateTyName nest (MkImpTy loc' n ty)
|
||||
= MkImpTy loc' (newName nest n) ty
|
||||
|
||||
updateDataName : NestedNames vars -> ImpData -> ImpData
|
||||
@ -78,11 +78,11 @@ checkLocal {vars} rig elabinfo nest env fc nestdecls scope expty
|
||||
= MkImpLater loc' (newName nest n) tycons
|
||||
|
||||
updateName : NestedNames vars -> ImpDecl -> ImpDecl
|
||||
updateName nest (IClaim loc' r vis fnopts ty)
|
||||
updateName nest (IClaim loc' r vis fnopts ty)
|
||||
= IClaim loc' r vis fnopts (updateTyName nest ty)
|
||||
updateName nest (IDef loc' n cs)
|
||||
updateName nest (IDef loc' n cs)
|
||||
= IDef loc' (newName nest n) cs
|
||||
updateName nest (IData loc' vis d)
|
||||
updateName nest (IData loc' vis d)
|
||||
= IData loc' vis (updateDataName nest d)
|
||||
updateName nest i = i
|
||||
|
||||
|
@ -63,25 +63,25 @@ genFieldName root
|
||||
|
||||
-- There's probably a generic version of this in the prelude isn't
|
||||
-- there?
|
||||
replace : String -> Rec ->
|
||||
replace : String -> Rec ->
|
||||
List (String, Rec) -> List (String, Rec)
|
||||
replace k v [] = []
|
||||
replace k v ((k', v') :: vs)
|
||||
= if k == k'
|
||||
= if k == k'
|
||||
then ((k, v) :: vs)
|
||||
else ((k', v') :: replace k v vs)
|
||||
|
||||
findPath : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> List String -> List String -> Maybe Name ->
|
||||
FC -> List String -> List String -> Maybe Name ->
|
||||
(String -> RawImp) ->
|
||||
Rec -> Core Rec
|
||||
findPath loc [] full tyn val (Field lhs _) = pure (Field lhs (val lhs))
|
||||
findPath loc [] full tyn val rec
|
||||
findPath loc [] full tyn val rec
|
||||
= throw (IncompatibleFieldUpdate loc full)
|
||||
findPath loc (p :: ps) full Nothing val (Field n v)
|
||||
findPath loc (p :: ps) full Nothing val (Field n v)
|
||||
= throw (NotRecordField loc p Nothing)
|
||||
findPath loc (p :: ps) full (Just tyn) val (Field n v)
|
||||
findPath loc (p :: ps) full (Just tyn) val (Field n v)
|
||||
= do defs <- get Ctxt
|
||||
Just con <- findConName defs tyn
|
||||
| Nothing => throw (NotRecordType loc tyn)
|
||||
@ -91,7 +91,7 @@ findPath loc (p :: ps) full (Just tyn) val (Field n v)
|
||||
let rec' = Constr con args
|
||||
findPath loc (p :: ps) full (Just tyn) val rec'
|
||||
where
|
||||
mkArgs : List (String, Maybe Name) ->
|
||||
mkArgs : List (String, Maybe Name) ->
|
||||
Core (List (String, Rec))
|
||||
mkArgs [] = pure []
|
||||
mkArgs ((p, _) :: ps)
|
||||
@ -99,7 +99,7 @@ findPath loc (p :: ps) full (Just tyn) val (Field n v)
|
||||
args' <- mkArgs ps
|
||||
pure ((p, Field fldn (IVar loc (UN fldn))) :: args')
|
||||
|
||||
findPath loc (p :: ps) full tyn val (Constr con args)
|
||||
findPath loc (p :: ps) full tyn val (Constr con args)
|
||||
= do let Just prec = lookup p args
|
||||
| Nothing => throw (NotRecordField loc p tyn)
|
||||
defs <- get Ctxt
|
||||
@ -112,13 +112,13 @@ findPath loc (p :: ps) full tyn val (Constr con args)
|
||||
|
||||
getSides : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> IFieldUpdate -> Name -> RawImp -> Rec ->
|
||||
FC -> IFieldUpdate -> Name -> RawImp -> Rec ->
|
||||
Core Rec
|
||||
getSides loc (ISetField path val) tyn orig rec
|
||||
getSides loc (ISetField path val) tyn orig rec
|
||||
-- update 'rec' so that 'path' is accessible on the lhs and rhs,
|
||||
-- then set the path on the rhs to 'val'
|
||||
= findPath loc path path (Just tyn) (const val) rec
|
||||
getSides loc (ISetFieldApp path val) tyn orig rec
|
||||
getSides loc (ISetFieldApp path val) tyn orig rec
|
||||
= findPath loc path path (Just tyn) (\n => apply val [IVar loc (UN n)]) rec
|
||||
where
|
||||
get : List String -> RawImp -> RawImp
|
||||
@ -127,23 +127,23 @@ getSides loc (ISetFieldApp path val) tyn orig rec
|
||||
|
||||
getAllSides : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> List IFieldUpdate -> Name ->
|
||||
RawImp -> Rec ->
|
||||
FC -> List IFieldUpdate -> Name ->
|
||||
RawImp -> Rec ->
|
||||
Core Rec
|
||||
getAllSides loc [] tyn orig rec = pure rec
|
||||
getAllSides loc (u :: upds) tyn orig rec
|
||||
getAllSides loc (u :: upds) tyn orig rec
|
||||
= getAllSides loc upds tyn orig !(getSides loc u tyn orig rec)
|
||||
|
||||
-- Convert the collection of high level field accesses into a case expression
|
||||
-- which does the updates all in one go
|
||||
export
|
||||
recUpdate : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo -> FC ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
List IFieldUpdate ->
|
||||
RigCount -> ElabInfo -> FC ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
List IFieldUpdate ->
|
||||
(rec : RawImp) -> (grecty : Glued vars) ->
|
||||
Core RawImp
|
||||
recUpdate rigc elabinfo loc nest env flds rec grecty
|
||||
@ -164,14 +164,14 @@ checkUpdate : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> List IFieldUpdate -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkUpdate rig elabinfo nest env fc upds rec expected
|
||||
= do recty <- case expected of
|
||||
Just ret => pure ret
|
||||
_ => do (_, ty) <- checkImp rig elabinfo
|
||||
_ => do (_, ty) <- checkImp rig elabinfo
|
||||
nest env rec Nothing
|
||||
pure ty
|
||||
rcase <- recUpdate rig elabinfo fc nest env upds rec recty
|
||||
|
@ -18,7 +18,7 @@ import TTImp.TTImp
|
||||
|
||||
-- TODO: Later, we'll get the name of the lemma from the type, if it's one
|
||||
-- that's generated for a dependent type. For now, always return the default
|
||||
findRewriteLemma : {auto c : Ref Ctxt Defs} ->
|
||||
findRewriteLemma : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> (rulety : Term vars) ->
|
||||
Core Name
|
||||
findRewriteLemma loc rulety
|
||||
@ -33,7 +33,7 @@ getRewriteTerms loc defs (NTCon nfc eq t a args) err
|
||||
= if !(isEqualTy eq)
|
||||
then case reverse args of
|
||||
(rhs :: lhs :: rhsty :: lhsty :: _) =>
|
||||
pure (!(evalClosure defs lhs),
|
||||
pure (!(evalClosure defs lhs),
|
||||
!(evalClosure defs rhs),
|
||||
!(evalClosure defs lhsty))
|
||||
_ => throw err
|
||||
@ -56,9 +56,9 @@ rewriteErr _ = False
|
||||
export
|
||||
elabRewrite : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Env Term vars ->
|
||||
(expected : Term vars) ->
|
||||
(expected : Term vars) ->
|
||||
(rulety : Term vars) ->
|
||||
Core (Name, Term vars, Term vars)
|
||||
elabRewrite loc env expected rulety
|
||||
@ -72,13 +72,13 @@ elabRewrite loc env expected rulety
|
||||
-- the metavariables might have been updated
|
||||
expnf <- nf defs env expected
|
||||
|
||||
logNF 5 "Rewriting" env lt
|
||||
logNF 5 "Rewriting" env lt
|
||||
logNF 5 "Rewriting in" env expnf
|
||||
rwexp_sc <- replace defs env lt (Ref loc Bound parg) expnf
|
||||
logTerm 5 "Rewritten to" rwexp_sc
|
||||
|
||||
empty <- clearDefs defs
|
||||
let pred = Bind loc parg (Lam RigW Explicit
|
||||
let pred = Bind loc parg (Lam RigW Explicit
|
||||
!(quote empty env lty))
|
||||
(refsToLocals (Add parg parg None) rwexp_sc)
|
||||
gpredty <- getType env pred
|
||||
@ -97,8 +97,8 @@ checkRewrite : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto e : Ref EST (EState vars)} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars ->
|
||||
FC -> RawImp -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkRewrite rigc elabinfo nest env fc rule tm Nothing
|
||||
@ -129,12 +129,12 @@ checkRewrite {vars} rigc elabinfo nest env fc rule tm (Just expected)
|
||||
(\e'' => check {e = e''} {vars = rname :: pname :: vars}
|
||||
rigc elabinfo (weaken (weaken nest)) env'
|
||||
(apply (IVar fc lemma) [IVar fc pname,
|
||||
IVar fc rname,
|
||||
tm])
|
||||
IVar fc rname,
|
||||
tm])
|
||||
(Just (gnf env'
|
||||
(weakenNs [rname, pname] expTy)))
|
||||
))
|
||||
rwty <- getTerm grwty
|
||||
pure (Bind fc pname pbind (Bind fc rname rbind rwtm),
|
||||
pure (Bind fc pname pbind (Bind fc rname rbind rwtm),
|
||||
gnf env (Bind fc pname pbind (Bind fc rname rbind rwty))))
|
||||
|
||||
|
@ -5,6 +5,7 @@ import Core.Core
|
||||
import Core.Env
|
||||
import Core.Metadata
|
||||
import Core.Normalise
|
||||
import Core.Reflect
|
||||
import Core.Unify
|
||||
import Core.TT
|
||||
import Core.Value
|
||||
@ -23,12 +24,13 @@ import TTImp.Elab.Local
|
||||
import TTImp.Elab.Prim
|
||||
import TTImp.Elab.Record
|
||||
import TTImp.Elab.Rewrite
|
||||
import TTImp.Reflect
|
||||
import TTImp.TTImp
|
||||
|
||||
%default covering
|
||||
|
||||
-- If the expected type has an implicit pi, elaborate with leading
|
||||
-- implicit lambdas if they aren't there already.
|
||||
-- implicit lambdas if they aren't there already.
|
||||
insertImpLam : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
Env Term vars ->
|
||||
@ -49,7 +51,7 @@ insertImpLam {vars} env tm (Just ty) = bindLam tm ty
|
||||
bindLamTm tm exp
|
||||
= case getFn exp of
|
||||
Ref _ Func _ => pure Nothing -- might still be implicit
|
||||
TForce _ _ => pure Nothing
|
||||
TForce _ _ _ => pure Nothing
|
||||
Bind _ _ (Lam _ _ _) _ => pure Nothing
|
||||
_ => pure $ Just tm
|
||||
|
||||
@ -84,22 +86,22 @@ checkTerm : {vars : _} ->
|
||||
RigCount -> ElabInfo ->
|
||||
NestedNames vars -> Env Term vars -> RawImp -> Maybe (Glued vars) ->
|
||||
Core (Term vars, Glued vars)
|
||||
checkTerm rig elabinfo nest env (IVar fc n) exp
|
||||
checkTerm rig elabinfo nest env (IVar fc n) exp
|
||||
= -- It may actually turn out to be an application, if the expected
|
||||
-- type is expecting an implicit argument, so check it as an
|
||||
-- application with no arguments
|
||||
checkApp rig elabinfo nest env fc (IVar fc n) [] [] exp
|
||||
checkTerm rig elabinfo nest env (IPi fc r p (Just n) argTy retTy) exp
|
||||
checkTerm rig elabinfo nest env (IPi fc r p (Just n) argTy retTy) exp
|
||||
= checkPi rig elabinfo nest env fc r p n argTy retTy exp
|
||||
checkTerm rig elabinfo nest env (IPi fc r p Nothing argTy retTy) exp
|
||||
checkTerm rig elabinfo nest env (IPi fc r p Nothing argTy retTy) exp
|
||||
= do n <- case p of
|
||||
Explicit => genVarName "arg"
|
||||
Implicit => genVarName "impArg"
|
||||
AutoImplicit => genVarName "conArg"
|
||||
checkPi rig elabinfo nest env fc r p n argTy retTy exp
|
||||
checkTerm rig elabinfo nest env (ILam fc r p (Just n) argTy scope) exp
|
||||
checkTerm rig elabinfo nest env (ILam fc r p (Just n) argTy scope) exp
|
||||
= checkLambda rig elabinfo nest env fc r p n argTy scope exp
|
||||
checkTerm rig elabinfo nest env (ILam fc r p Nothing argTy scope) exp
|
||||
checkTerm rig elabinfo nest env (ILam fc r p Nothing argTy scope) exp
|
||||
= do n <- genVarName "_"
|
||||
checkLambda rig elabinfo nest env fc r p n argTy scope exp
|
||||
checkTerm rig elabinfo nest env (ILet fc r n nTy nVal scope) exp
|
||||
@ -110,9 +112,9 @@ checkTerm rig elabinfo nest env (ILocal fc nested scope) exp
|
||||
= checkLocal rig elabinfo nest env fc nested scope exp
|
||||
checkTerm rig elabinfo nest env (IUpdate fc upds rec) exp
|
||||
= checkUpdate rig elabinfo nest env fc upds rec exp
|
||||
checkTerm rig elabinfo nest env (IApp fc fn arg) exp
|
||||
checkTerm rig elabinfo nest env (IApp fc fn arg) exp
|
||||
= checkApp rig elabinfo nest env fc fn [arg] [] exp
|
||||
checkTerm rig elabinfo nest env (IWithApp fc fn arg) exp
|
||||
checkTerm rig elabinfo nest env (IWithApp fc fn arg) exp
|
||||
= throw (GenericMsg fc "with application not implemented yet")
|
||||
checkTerm rig elabinfo nest env (IImplicitApp fc fn nm arg) exp
|
||||
= checkApp rig elabinfo nest env fc fn [] [(nm, arg)] exp
|
||||
@ -149,10 +151,10 @@ checkTerm rig elabinfo nest env (IDelay fc tm) exp
|
||||
= checkDelay rig elabinfo nest env fc tm exp
|
||||
checkTerm rig elabinfo nest env (IForce fc tm) exp
|
||||
= checkForce rig elabinfo nest env fc tm exp
|
||||
checkTerm {vars} rig elabinfo nest env (IPrimVal fc c) exp
|
||||
checkTerm {vars} rig elabinfo nest env (IPrimVal fc c) exp
|
||||
= do let (cval, cty) = checkPrim {vars} fc c
|
||||
checkExp rig elabinfo env fc cval (gnf env cty) exp
|
||||
checkTerm rig elabinfo nest env (IType fc) exp
|
||||
checkTerm rig elabinfo nest env (IType fc) exp
|
||||
= checkExp rig elabinfo env fc (TType fc) (gType fc) exp
|
||||
|
||||
checkTerm rig elabinfo nest env (IHole fc str) exp
|
||||
@ -185,7 +187,7 @@ checkTerm rig elabinfo nest env (Implicit fc b) Nothing
|
||||
-- {auto m : Ref MD Metadata} ->
|
||||
-- {auto u : Ref UST UState} ->
|
||||
-- {auto e : Ref EST (EState vars)} ->
|
||||
-- RigCount -> ElabInfo -> Env Term vars -> RawImp ->
|
||||
-- RigCount -> ElabInfo -> Env Term vars -> RawImp ->
|
||||
-- Maybe (Glued vars) ->
|
||||
-- Core (Term vars, Glued vars)
|
||||
-- If we've just inserted an implicit coercion (in practice, that's either
|
||||
@ -199,7 +201,7 @@ TTImp.Elab.Check.check rigc elabinfo nest env tm@(ILocal fc ds sc) exp
|
||||
= checkImp rigc elabinfo nest env tm exp
|
||||
TTImp.Elab.Check.check rigc elabinfo nest env tm@(IUpdate fc fs rec) exp
|
||||
= checkImp rigc elabinfo nest env tm exp
|
||||
TTImp.Elab.Check.check rigc elabinfo nest env tm_in exp
|
||||
TTImp.Elab.Check.check rigc elabinfo nest env tm_in exp
|
||||
= do defs <- get Ctxt
|
||||
est <- get EST
|
||||
tm <- expandAmbigName (elabMode elabinfo) nest env tm_in [] tm_in exp
|
||||
|
@ -23,7 +23,7 @@ findErasedFrom defs pos tm = pure []
|
||||
export
|
||||
findErased : {auto c : Ref Ctxt Defs} ->
|
||||
ClosedTerm -> Core (List Nat)
|
||||
findErased tm
|
||||
findErased tm
|
||||
= do defs <- get Ctxt
|
||||
tmnf <- nf defs [] tm
|
||||
findErasedFrom defs 0 tmnf
|
||||
|
@ -36,7 +36,7 @@ Show ClauseUpdate where
|
||||
show (Impossible lhs) = "Impossible: " ++ show lhs
|
||||
show Invalid = "Invalid"
|
||||
|
||||
public export
|
||||
public export
|
||||
data SplitError : Type where
|
||||
NoValidSplit : SplitError -- None of the splits either type check, or fail
|
||||
-- in a way which is valid as an 'impossible' case
|
||||
@ -60,7 +60,7 @@ Show a => Show (SplitResult a) where
|
||||
show (OK res) = "OK: " ++ show res
|
||||
|
||||
findTyName : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> Env Term vars -> Name -> Term vars ->
|
||||
Defs -> Env Term vars -> Name -> Term vars ->
|
||||
Core (Maybe Name)
|
||||
findTyName defs env n (Bind _ x (PVar c p ty) sc)
|
||||
-- Take the first one, which is the most recently bound
|
||||
@ -84,19 +84,19 @@ findCons : {auto c : Ref Ctxt Defs} ->
|
||||
Name -> Term [] -> Core (SplitResult (Name, Name, List Name))
|
||||
findCons n lhs
|
||||
= case getDefining lhs of
|
||||
Nothing => pure (SplitFail
|
||||
Nothing => pure (SplitFail
|
||||
(CantSplitThis n "Can't find function name on LHS"))
|
||||
Just fn =>
|
||||
do defs <- get Ctxt
|
||||
case !(findTyName defs [] n lhs) of
|
||||
Nothing => pure (SplitFail (CantSplitThis n
|
||||
Nothing => pure (SplitFail (CantSplitThis n
|
||||
("Can't find type of " ++ show n ++ " in LHS")))
|
||||
Just tyn =>
|
||||
do Just (TCon _ _ _ _ _ cons) <-
|
||||
lookupDefExact tyn (gamma defs)
|
||||
| res => pure (SplitFail
|
||||
(CantSplitThis n
|
||||
("Not a type constructor " ++
|
||||
| res => pure (SplitFail
|
||||
(CantSplitThis n
|
||||
("Not a type constructor " ++
|
||||
show res)))
|
||||
pure (OK (fn, tyn, cons))
|
||||
|
||||
@ -123,7 +123,7 @@ defaultNames = ["x", "y", "z", "w", "v", "s", "t", "u"]
|
||||
export
|
||||
getArgName : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> Name -> List Name -> NF vars -> Core String
|
||||
getArgName defs x allvars ty
|
||||
getArgName defs x allvars ty
|
||||
= do defnames <- findNames ty
|
||||
pure $ getName x defnames allvars
|
||||
where
|
||||
@ -148,9 +148,9 @@ getArgName defs x allvars ty
|
||||
|
||||
export
|
||||
getArgNames : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> List Name -> Env Term vars -> NF vars ->
|
||||
Defs -> List Name -> Env Term vars -> NF vars ->
|
||||
Core (List String)
|
||||
getArgNames defs allvars env (NBind fc x (Pi _ p ty) sc)
|
||||
getArgNames defs allvars env (NBind fc x (Pi _ p ty) sc)
|
||||
= do ns <- case p of
|
||||
Explicit => pure [!(getArgName defs x allvars ty)]
|
||||
_ => pure []
|
||||
@ -172,14 +172,14 @@ expandCon fc usedvars con
|
||||
= do defs <- get Ctxt
|
||||
Just ty <- lookupTyExact con (gamma defs)
|
||||
| Nothing => throw (UndefinedName fc con)
|
||||
pure (apply (IVar fc con)
|
||||
pure (apply (IVar fc con)
|
||||
(map (IBindVar fc)
|
||||
!(getArgNames defs usedvars []
|
||||
!(getArgNames defs usedvars []
|
||||
!(nf defs [] ty))))
|
||||
|
||||
updateArg : {auto c : Ref Ctxt Defs} ->
|
||||
List Name -> -- all the variable names
|
||||
(var : Name) -> (con : Name) ->
|
||||
(var : Name) -> (con : Name) ->
|
||||
RawImp -> Core RawImp
|
||||
updateArg allvars var con (IVar fc n)
|
||||
= if n `elem` allvars
|
||||
@ -188,22 +188,22 @@ updateArg allvars var con (IVar fc n)
|
||||
else pure $ Implicit fc True
|
||||
else pure $ IVar fc n
|
||||
updateArg allvars var con (IApp fc f a)
|
||||
= pure $ IApp fc !(updateArg allvars var con f)
|
||||
= pure $ IApp fc !(updateArg allvars var con f)
|
||||
!(updateArg allvars var con a)
|
||||
updateArg allvars var con (IImplicitApp fc f n a)
|
||||
= pure $ IImplicitApp fc !(updateArg allvars var con f) n
|
||||
= pure $ IImplicitApp fc !(updateArg allvars var con f) n
|
||||
!(updateArg allvars var con a)
|
||||
updateArg allvars var con (IAs fc s n p)
|
||||
= updateArg allvars var con p
|
||||
updateArg allvars var con tm = pure $ Implicit (getFC tm) True
|
||||
|
||||
data ArgType
|
||||
data ArgType
|
||||
= Explicit FC RawImp
|
||||
| Implicit FC (Maybe Name) RawImp
|
||||
|
||||
update : {auto c : Ref Ctxt Defs} ->
|
||||
List Name -> -- all the variable names
|
||||
(var : Name) -> (con : Name) ->
|
||||
(var : Name) -> (con : Name) ->
|
||||
ArgType -> Core ArgType
|
||||
update allvars var con (Explicit fc arg)
|
||||
= pure $ Explicit fc !(updateArg allvars var con arg)
|
||||
@ -226,14 +226,14 @@ apply f [] = f
|
||||
-- Also replace any variables with '_' to allow elaboration to
|
||||
-- expand them
|
||||
newLHS : {auto c : Ref Ctxt Defs} ->
|
||||
FC ->
|
||||
FC ->
|
||||
Nat -> -- previous environment length; leave these alone
|
||||
List Name -> -- all the variable names
|
||||
(var : Name) -> (con : Name) ->
|
||||
(var : Name) -> (con : Name) ->
|
||||
RawImp -> Core RawImp
|
||||
newLHS fc envlen allvars var con tm
|
||||
= do let (f, args) = getFnArgs tm []
|
||||
let keep = map (const (Explicit fc (Implicit fc True)))
|
||||
let keep = map (const (Explicit fc (Implicit fc True)))
|
||||
(take envlen args)
|
||||
let ups = drop envlen args
|
||||
ups' <- traverse (update allvars var con) ups
|
||||
@ -305,7 +305,7 @@ mkCase {c} {u} fn orig lhs_raw
|
||||
(\err => case err of
|
||||
WhenUnifying _ env l r err
|
||||
=> do defs <- get Ctxt
|
||||
if !(impossibleOK defs !(nf defs env l)
|
||||
if !(impossibleOK defs !(nf defs env l)
|
||||
!(nf defs env r))
|
||||
then pure (Impossible lhs_raw)
|
||||
else pure Invalid
|
||||
@ -328,7 +328,7 @@ export
|
||||
getSplitsLHS : {auto m : Ref MD Metadata} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Nat -> ClosedTerm -> Name ->
|
||||
FC -> Nat -> ClosedTerm -> Name ->
|
||||
Core (SplitResult (List ClauseUpdate))
|
||||
getSplitsLHS fc envlen lhs_in n
|
||||
= do let lhs = substLets lhs_in
|
||||
@ -337,7 +337,7 @@ getSplitsLHS fc envlen lhs_in n
|
||||
defs <- get Ctxt
|
||||
OK (fn, tyn, cons) <- findCons n lhs
|
||||
| SplitFail err => pure (SplitFail err)
|
||||
|
||||
|
||||
rawlhs <- unelabNoSugar [] lhs
|
||||
trycases <- traverse (\c => newLHS fc envlen usedns n c rawlhs) cons
|
||||
|
||||
@ -351,7 +351,7 @@ export
|
||||
getSplits : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
(FC -> ClosedTerm -> Bool) -> Name ->
|
||||
(FC -> ClosedTerm -> Bool) -> Name ->
|
||||
Core (SplitResult (List ClauseUpdate))
|
||||
getSplits p n
|
||||
= do Just (loc, envlen, lhs_in) <- findLHSAt p
|
||||
|
@ -43,30 +43,30 @@ search : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount ->
|
||||
SearchOpts -> Maybe RecData ->
|
||||
SearchOpts -> Maybe RecData ->
|
||||
ClosedTerm ->
|
||||
Name -> Core (List ClosedTerm)
|
||||
|
||||
getAllEnv : FC -> (done : List Name) ->
|
||||
Env Term vars ->
|
||||
getAllEnv : FC -> (done : List Name) ->
|
||||
Env Term vars ->
|
||||
List (Term (done ++ vars), Term (done ++ vars))
|
||||
getAllEnv fc done [] = []
|
||||
getAllEnv {vars = v :: vs} fc done (b :: env)
|
||||
= let rest = getAllEnv fc (done ++ [v]) env
|
||||
getAllEnv {vars = v :: vs} fc done (b :: env)
|
||||
= let rest = getAllEnv fc (done ++ [v]) env
|
||||
MkVar p = weakenVar done {inner = v :: vs} First in
|
||||
(Local fc Nothing _ p,
|
||||
rewrite appendAssociative done [v] vs in
|
||||
weakenNs (done ++ [v]) (binderType b)) ::
|
||||
rewrite appendAssociative done [v] vs in
|
||||
weakenNs (done ++ [v]) (binderType b)) ::
|
||||
rewrite appendAssociative done [v] vs in rest
|
||||
|
||||
-- Search recursively, but only if the given name wasn't solved by unification
|
||||
searchIfHole : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> SearchOpts -> Maybe RecData -> ClosedTerm ->
|
||||
Env Term vars -> ArgInfo vars ->
|
||||
FC -> SearchOpts -> Maybe RecData -> ClosedTerm ->
|
||||
Env Term vars -> ArgInfo vars ->
|
||||
Core (List (Term vars))
|
||||
searchIfHole fc opts defining topty env arg
|
||||
searchIfHole fc opts defining topty env arg
|
||||
= case depth opts of
|
||||
Z => pure []
|
||||
S k =>
|
||||
@ -86,7 +86,7 @@ searchIfHole fc opts defining topty env arg
|
||||
traverse (\tm => pure !(normaliseHoles defs env
|
||||
(applyTo fc (embed tm) env))) tms
|
||||
|
||||
mkCandidates : FC -> Term vars -> List (List (Term vars)) ->
|
||||
mkCandidates : FC -> Term vars -> List (List (Term vars)) ->
|
||||
List (Term vars)
|
||||
mkCandidates fc f [] = pure f
|
||||
mkCandidates fc f (args :: argss)
|
||||
@ -102,12 +102,12 @@ explicit ai = plicit ai == Explicit
|
||||
searchName : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount -> SearchOpts ->
|
||||
FC -> RigCount -> SearchOpts ->
|
||||
Env Term vars -> NF vars -> ClosedTerm ->
|
||||
Maybe RecData -> (Name, GlobalDef) -> Core (List (Term vars))
|
||||
searchName fc rigc opts env target topty defining (n, ndef)
|
||||
= do defs <- get Ctxt
|
||||
let True = visibleInAny (!getNS :: !getNestedNS)
|
||||
let True = visibleInAny (!getNS :: !getNestedNS)
|
||||
(fullname ndef) (visibility ndef)
|
||||
| _ => pure []
|
||||
let ty = type ndef
|
||||
@ -149,7 +149,7 @@ successful (elab :: elabs)
|
||||
= catch (do res <- elab
|
||||
rest <- successful elabs
|
||||
pure (Right res :: rest))
|
||||
(\err =>
|
||||
(\err =>
|
||||
do rest <- successful elabs
|
||||
pure (Left err :: rest))
|
||||
|
||||
@ -171,7 +171,7 @@ getSuccessful {vars} fc rig opts mkHole env ty topty defining all
|
||||
(\r => nameRoot (recname r) ++ "_rhs")
|
||||
defining
|
||||
hn <- uniqueName defs (map nameRoot vars) base
|
||||
(idx, tm) <- newMeta fc rig env (UN hn) ty
|
||||
(idx, tm) <- newMeta fc rig env (UN hn) ty
|
||||
(Hole (length env) False)
|
||||
False
|
||||
pure [tm]
|
||||
@ -197,7 +197,7 @@ searchNames fc rig opts env ty topty defining (n :: ns)
|
||||
getSuccessful fc rig opts False env ty topty defining
|
||||
(map (searchName fc rig opts env nfty topty defining) visns)
|
||||
where
|
||||
visible : Context -> List (List String) -> Name ->
|
||||
visible : Context -> List (List String) -> Name ->
|
||||
Core (Maybe (Name, GlobalDef))
|
||||
visible gam nspace n
|
||||
= do Just def <- lookupCtxtExact n gam
|
||||
@ -209,16 +209,16 @@ searchNames fc rig opts env ty topty defining (n :: ns)
|
||||
tryRecursive : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount -> SearchOpts ->
|
||||
FC -> RigCount -> SearchOpts ->
|
||||
Env Term vars -> Term vars -> ClosedTerm ->
|
||||
Maybe RecData -> Core (List (Term vars))
|
||||
tryRecursive fc rig opts env ty topty Nothing = pure []
|
||||
tryRecursive fc rig opts env ty topty (Just rdata)
|
||||
tryRecursive fc rig opts env ty topty (Just rdata)
|
||||
= do defs <- get Ctxt
|
||||
case !(lookupCtxtExact (recname rdata) (gamma defs)) of
|
||||
Nothing => pure []
|
||||
Just def =>
|
||||
do tms <- searchName fc rig opts env !(nf defs env ty)
|
||||
do tms <- searchName fc rig opts env !(nf defs env ty)
|
||||
topty Nothing
|
||||
(recname rdata, def)
|
||||
defs <- get Ctxt
|
||||
@ -235,7 +235,7 @@ tryRecursive fc rig opts env ty topty (Just rdata)
|
||||
argDiff (Ref _ _ fn) (Ref _ _ fn') = fn /= fn'
|
||||
argDiff (Bind _ _ _ _) _ = False
|
||||
argDiff _ (Bind _ _ _ _) = False
|
||||
argDiff (App _ f a) (App _ f' a')
|
||||
argDiff (App _ f a) (App _ f' a')
|
||||
= structDiff f f' || structDiff a a'
|
||||
argDiff (PrimVal _ c) (PrimVal _ c') = c /= c'
|
||||
argDiff (Erased _) _ = False
|
||||
@ -244,7 +244,7 @@ tryRecursive fc rig opts env ty topty (Just rdata)
|
||||
argDiff (As _ _ x) y = argDiff x y
|
||||
argDiff x (As _ _ y) = argDiff x y
|
||||
argDiff _ _ = True
|
||||
|
||||
|
||||
appsDiff : Term vs -> Term vs' -> List (Term vs) -> List (Term vs') ->
|
||||
Bool
|
||||
appsDiff (Ref _ (DataCon _ _) f) (Ref _ (DataCon _ _) f') args args'
|
||||
@ -277,9 +277,9 @@ usableLocal loc _ _ = True
|
||||
searchLocalWith : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount -> SearchOpts -> Env Term vars ->
|
||||
FC -> RigCount -> SearchOpts -> Env Term vars ->
|
||||
List (Term vars, Term vars) ->
|
||||
Term vars -> ClosedTerm -> Maybe RecData ->
|
||||
Term vars -> ClosedTerm -> Maybe RecData ->
|
||||
Core (List (Term vars))
|
||||
searchLocalWith fc rig opts env [] ty topty defining
|
||||
= pure []
|
||||
@ -291,7 +291,7 @@ searchLocalWith {vars} fc rig opts env ((p, pty) :: rest) ty topty defining
|
||||
searchLocalWith fc rig opts env rest ty
|
||||
topty defining]
|
||||
where
|
||||
findDirect : Defs -> Term vars ->
|
||||
findDirect : Defs -> Term vars ->
|
||||
(Term vars -> Term vars) ->
|
||||
NF vars -> -- local variable's type
|
||||
NF vars -> -- type we're looking for
|
||||
@ -301,7 +301,7 @@ searchLocalWith {vars} fc rig opts env ((p, pty) :: rest) ty topty defining
|
||||
-- We can only use a local variable if it's not an unsolved
|
||||
-- hole
|
||||
if usableLocal fc env nty
|
||||
then
|
||||
then
|
||||
tryUnify -- try with no arguments first
|
||||
(do ures <- unify InTerm fc env ty nty
|
||||
let [] = constraints ures
|
||||
@ -310,12 +310,12 @@ searchLocalWith {vars} fc rig opts env ((p, pty) :: rest) ty topty defining
|
||||
(do ures <- unify InTerm fc env ty appTy
|
||||
let [] = constraints ures
|
||||
| _ => pure []
|
||||
args' <- traverse (searchIfHole fc opts defining topty env)
|
||||
args' <- traverse (searchIfHole fc opts defining topty env)
|
||||
args
|
||||
pure (mkCandidates fc (f prf) args'))
|
||||
else pure []
|
||||
|
||||
findPos : Defs -> Term vars ->
|
||||
findPos : Defs -> Term vars ->
|
||||
(Term vars -> Term vars) ->
|
||||
NF vars -> NF vars -> Core (List (Term vars))
|
||||
findPos defs prf f x@(NTCon pfc pn _ _ [xty, yty]) target
|
||||
@ -340,7 +340,7 @@ searchLocalWith {vars} fc rig opts env ((p, pty) :: rest) ty topty defining
|
||||
f arg])
|
||||
xtynf target),
|
||||
(do ytynf <- evalClosure defs yty
|
||||
findPos defs prf
|
||||
findPos defs prf
|
||||
(\arg => apply fc (Ref fc Func sname)
|
||||
[xtytm,
|
||||
ytytm,
|
||||
@ -352,12 +352,12 @@ searchLocalWith {vars} fc rig opts env ((p, pty) :: rest) ty topty defining
|
||||
searchLocal : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount -> SearchOpts ->
|
||||
FC -> RigCount -> SearchOpts ->
|
||||
Env Term vars -> Term vars -> ClosedTerm ->
|
||||
Maybe RecData -> Core (List (Term vars))
|
||||
searchLocal fc rig opts env ty topty defining
|
||||
searchLocal fc rig opts env ty topty defining
|
||||
= do defs <- get Ctxt
|
||||
searchLocalWith fc rig opts env (getAllEnv fc [] env)
|
||||
searchLocalWith fc rig opts env (getAllEnv fc [] env)
|
||||
ty topty defining
|
||||
|
||||
searchType : {auto c : Ref Ctxt Defs} ->
|
||||
@ -399,12 +399,12 @@ searchType fc rig opts env defining topty _ ty
|
||||
([searchLocal fc rig opts env ty topty defining,
|
||||
searchNames fc rig opts env ty topty defining
|
||||
allHints]
|
||||
++ if recOK opts
|
||||
++ if recOK opts
|
||||
then [tryRecursive fc rig opts env ty topty defining]
|
||||
else [])
|
||||
else pure []
|
||||
_ => do logTerm 10 "Searching locals only at" ty
|
||||
getSuccessful fc rig opts True env ty topty defining
|
||||
getSuccessful fc rig opts True env ty topty defining
|
||||
([searchLocal fc rig opts env ty topty defining]
|
||||
++ if recOK opts
|
||||
then [tryRecursive fc rig opts env ty topty defining]
|
||||
@ -413,7 +413,7 @@ searchType fc rig opts env defining topty _ ty
|
||||
searchHole : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> RigCount -> SearchOpts -> Maybe RecData -> Name ->
|
||||
FC -> RigCount -> SearchOpts -> Maybe RecData -> Name ->
|
||||
Nat -> ClosedTerm ->
|
||||
Defs -> GlobalDef -> Core (List ClosedTerm)
|
||||
searchHole fc rig opts defining n locs topty defs glob
|
||||
@ -430,15 +430,15 @@ search fc rig opts defining topty n_in
|
||||
-- if it's arising from an auto implicit
|
||||
case definition gdef of
|
||||
Hole locs _ => searchHole fc rig opts defining n locs topty defs gdef
|
||||
BySearch _ _ _ => searchHole fc rig opts defining n
|
||||
!(getArity defs [] (type gdef))
|
||||
BySearch _ _ _ => searchHole fc rig opts defining n
|
||||
!(getArity defs [] (type gdef))
|
||||
topty defs gdef
|
||||
_ => do log 10 $ show n_in ++ " not a hole"
|
||||
throw (InternalError $ "Not a hole: " ++ show n ++ " in " ++ show (map recname defining))
|
||||
_ => do log 10 $ show n_in ++ " not found"
|
||||
throw (UndefinedName fc n_in)
|
||||
where
|
||||
lookupHoleName : Name -> Context ->
|
||||
lookupHoleName : Name -> Context ->
|
||||
Core (Maybe (Name, Int, GlobalDef))
|
||||
lookupHoleName n defs
|
||||
= case !(lookupCtxtExactI n defs) of
|
||||
@ -454,7 +454,7 @@ getLHSData defs (Just tm) = pure $ getLHS !(normaliseHoles defs [] tm)
|
||||
getLHS : Term vars -> Maybe RecData
|
||||
getLHS (Bind _ _ (PVar _ _ _) sc) = getLHS sc
|
||||
getLHS (Bind _ _ (PLet _ _ _) sc) = getLHS sc
|
||||
getLHS sc
|
||||
getLHS sc
|
||||
= case getFn sc of
|
||||
Ref _ _ n => Just (MkRecData n sc)
|
||||
_ => Nothing
|
||||
@ -465,7 +465,7 @@ dropLinearErrors : {auto c : Ref Ctxt Defs} ->
|
||||
Core (List ClosedTerm)
|
||||
dropLinearErrors fc [] = pure []
|
||||
dropLinearErrors fc (t :: ts)
|
||||
= tryUnify
|
||||
= tryUnify
|
||||
(do linearCheck fc Rig1 False [] t
|
||||
ts' <- dropLinearErrors fc ts
|
||||
pure (t :: ts'))
|
||||
|
@ -23,7 +23,7 @@ import TTImp.Utils
|
||||
%default covering
|
||||
|
||||
fnName : Bool -> Name -> String
|
||||
fnName lhs (UN n)
|
||||
fnName lhs (UN n)
|
||||
= if any (not . identChar) (unpack n)
|
||||
then if lhs then "(" ++ n ++ ")"
|
||||
else "op"
|
||||
@ -49,10 +49,10 @@ uniqueRHS c = pure c
|
||||
expandClause : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Int -> ImpClause ->
|
||||
FC -> Int -> ImpClause ->
|
||||
Core (List ImpClause)
|
||||
expandClause loc n c
|
||||
= do log 10 $ "Trying clause " ++ show c
|
||||
= do log 10 $ "Trying clause " ++ show c
|
||||
c <- uniqueRHS c
|
||||
Just clause <- checkClause Rig1 False n [] (MkNested []) [] c
|
||||
| Nothing => pure [] -- TODO: impossible clause, do something
|
||||
@ -71,8 +71,8 @@ expandClause loc n c
|
||||
rhsnf <- normaliseHoles defs [] rhs'
|
||||
let (_ ** (env', rhsenv)) = dropLams locs [] rhsnf
|
||||
|
||||
rhsraw <- unelab env' rhsenv
|
||||
logTermNF 5 "Got clause" env lhs
|
||||
rhsraw <- unelab env' rhsenv
|
||||
logTermNF 5 "Got clause" env lhs
|
||||
logTermNF 5 " = " env' rhsenv
|
||||
pure [updateRHS c rhsraw]
|
||||
where
|
||||
@ -82,10 +82,10 @@ expandClause loc n c
|
||||
updateRHS (WithClause fc lhs wval cs) rhs = WithClause fc lhs wval cs
|
||||
updateRHS (ImpossibleClause fc lhs) _ = ImpossibleClause fc lhs
|
||||
|
||||
dropLams : Nat -> Env Term vars -> Term vars ->
|
||||
dropLams : Nat -> Env Term vars -> Term vars ->
|
||||
(vars' ** (Env Term vars', Term vars'))
|
||||
dropLams Z env tm = (_ ** (env, tm))
|
||||
dropLams (S k) env (Bind _ _ b sc) = dropLams k (b :: env) sc
|
||||
dropLams (S k) env (Bind _ _ b sc) = dropLams k (b :: env) sc
|
||||
dropLams _ env tm = (_ ** (env, tm))
|
||||
|
||||
splittableNames : RawImp -> List Name
|
||||
@ -129,19 +129,19 @@ trySplit loc lhsraw lhs rhs n
|
||||
Nothing => IBindVar loc' n
|
||||
Just tm => fixNames tm
|
||||
updateLHS ups (IApp loc' f a) = IApp loc' (updateLHS ups f) (updateLHS ups a)
|
||||
updateLHS ups (IImplicitApp loc' f t a)
|
||||
updateLHS ups (IImplicitApp loc' f t a)
|
||||
= IImplicitApp loc' (updateLHS ups f) t (updateLHS ups a)
|
||||
updateLHS ups tm = tm
|
||||
|
||||
generateSplits : {auto m : Ref MD Metadata} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Int -> ImpClause ->
|
||||
FC -> Int -> ImpClause ->
|
||||
Core (List (Name, List ImpClause))
|
||||
generateSplits loc fn (ImpossibleClause fc lhs) = pure []
|
||||
generateSplits loc fn (WithClause fc lhs wval cs) = pure []
|
||||
generateSplits {c} {m} {u} loc fn (PatClause fc lhs rhs)
|
||||
= do (lhstm, _) <-
|
||||
generateSplits {c} {m} {u} loc fn (PatClause fc lhs rhs)
|
||||
= do (lhstm, _) <-
|
||||
elabTerm fn (InLHS Rig1) [] (MkNested []) []
|
||||
(IBindHere loc PATTERN lhs) Nothing
|
||||
traverse (trySplit fc lhs lhstm rhs) (splittableNames lhs)
|
||||
@ -151,10 +151,10 @@ mutual
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Int -> Error ->
|
||||
List (Name, List ImpClause) ->
|
||||
List (Name, List ImpClause) ->
|
||||
Core (List ImpClause)
|
||||
tryAllSplits loc n err [] = throw err
|
||||
tryAllSplits loc n err ((x, []) :: rest)
|
||||
tryAllSplits loc n err ((x, []) :: rest)
|
||||
= tryAllSplits loc n err rest
|
||||
tryAllSplits loc n err ((x, cs) :: rest)
|
||||
= do log 5 $ "Splitting on " ++ show x
|
||||
@ -165,7 +165,7 @@ mutual
|
||||
mkSplits : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
FC -> Int -> ImpClause ->
|
||||
FC -> Int -> ImpClause ->
|
||||
Core (List ImpClause)
|
||||
-- If the clause works, use it. Otherwise, split on one of the splittable
|
||||
-- variables and try all of the resulting clauses
|
||||
@ -180,9 +180,9 @@ export
|
||||
makeDef : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
(FC -> (Name, Nat, ClosedTerm) -> Bool) ->
|
||||
(FC -> (Name, Nat, ClosedTerm) -> Bool) ->
|
||||
Name -> Core (Maybe (FC, List ImpClause))
|
||||
makeDef p n
|
||||
makeDef p n
|
||||
= do Just (loc, nidx, envlen, ty) <- findTyDeclAt p
|
||||
| Nothing => pure Nothing
|
||||
n <- getFullName nidx
|
||||
@ -193,10 +193,10 @@ makeDef p n
|
||||
argns <- getEnvArgNames defs envlen !(nf defs [] ty)
|
||||
-- Need to add implicit patterns for the outer environment.
|
||||
-- We won't try splitting on these
|
||||
let pre_env = replicate envlen (Implicit loc True)
|
||||
let pre_env = replicate envlen (Implicit loc True)
|
||||
|
||||
rhshole <- uniqueName defs [] (fnName False n ++ "_rhs")
|
||||
let initcs = PatClause loc
|
||||
let initcs = PatClause loc
|
||||
(apply (IVar loc n) (pre_env ++ (map (IBindVar loc) argns)))
|
||||
(IHole loc rhshole)
|
||||
let Just nidx = getNameID n (gamma defs)
|
||||
|
@ -28,7 +28,7 @@ bindable p tm
|
||||
(Ref _ (DataCon _ _) _, args) => any (bindable p) args
|
||||
(TDelayed _ _ t, args) => any (bindable p) (t :: args)
|
||||
(TDelay _ _ _ t, args) => any (bindable p) (t :: args)
|
||||
(TForce _ t, args) => any (bindable p) (t :: args)
|
||||
(TForce _ _ t, args) => any (bindable p) (t :: args)
|
||||
(Local _ _ p' _, []) => p == p'
|
||||
_ => False
|
||||
|
||||
@ -38,7 +38,7 @@ bindableArg p (Bind _ _ (Pi _ _ ty) sc)
|
||||
bindableArg p _ = False
|
||||
|
||||
getArgs : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars -> Nat -> Term vars ->
|
||||
Env Term vars -> Nat -> Term vars ->
|
||||
Core (List (Name, Maybe Name, PiInfo, RigCount, RawImp), RawImp)
|
||||
getArgs {vars} env (S k) (Bind _ x (Pi c p ty) sc)
|
||||
= do ty' <- unelab env ty
|
||||
@ -59,8 +59,8 @@ getArgs env k ty
|
||||
= do ty' <- unelab env ty
|
||||
pure ([], ty')
|
||||
|
||||
mkType : FC -> List (Name, Maybe Name, PiInfo, RigCount, RawImp) ->
|
||||
RawImp -> RawImp
|
||||
mkType : FC -> List (Name, Maybe Name, PiInfo, RigCount, RawImp) ->
|
||||
RawImp -> RawImp
|
||||
mkType loc [] ret = ret
|
||||
mkType loc ((_, n, p, c, ty) :: rest) ret
|
||||
= IPi loc c p n ty (mkType loc rest ret)
|
||||
@ -70,7 +70,7 @@ mkApp : FC -> Name ->
|
||||
mkApp loc n args
|
||||
= apply (IVar loc n) (mapMaybe getArg args)
|
||||
where
|
||||
getArg : (Name, Maybe Name, PiInfo, RigCount, RawImp) ->
|
||||
getArg : (Name, Maybe Name, PiInfo, RigCount, RawImp) ->
|
||||
Maybe RawImp
|
||||
getArg (x, _, Explicit, _, _) = Just (IVar loc x)
|
||||
getArg _ = Nothing
|
||||
@ -80,7 +80,7 @@ mkApp loc n args
|
||||
export
|
||||
makeLemma : {auto m : Ref MD Metadata} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Nat -> ClosedTerm ->
|
||||
FC -> Name -> Nat -> ClosedTerm ->
|
||||
Core (RawImp, RawImp)
|
||||
makeLemma loc n nlocs ty
|
||||
= do (args, ret) <- getArgs [] nlocs ty
|
||||
|
@ -99,7 +99,7 @@ visOpt
|
||||
opt <- fnDirectOpt
|
||||
pure (Right opt)
|
||||
|
||||
getVisibility : Maybe Visibility -> List (Either Visibility FnOpt) ->
|
||||
getVisibility : Maybe Visibility -> List (Either Visibility FnOpt) ->
|
||||
EmptyRule Visibility
|
||||
getVisibility Nothing [] = pure Private
|
||||
getVisibility (Just vis) [] = pure vis
|
||||
@ -136,10 +136,10 @@ mutual
|
||||
applyExpImp start end f [] = f
|
||||
applyExpImp start end f (Left exp :: args)
|
||||
= applyExpImp start end (IApp (MkFC fname start end) f exp) args
|
||||
applyExpImp start end f (Right (n, imp) :: args)
|
||||
applyExpImp start end f (Right (n, imp) :: args)
|
||||
= applyExpImp start end (IImplicitApp (MkFC fname start end) f n imp) args
|
||||
|
||||
argExpr : FileName -> IndentInfo ->
|
||||
argExpr : FileName -> IndentInfo ->
|
||||
Rule (Either RawImp (Maybe Name, RawImp))
|
||||
argExpr fname indents
|
||||
= do continue indents
|
||||
@ -201,20 +201,20 @@ mutual
|
||||
getMult Nothing = pure RigW
|
||||
getMult _ = fatalError "Invalid multiplicity (must be 0 or 1)"
|
||||
|
||||
pibindAll : FC -> PiInfo -> List (RigCount, Maybe Name, RawImp) ->
|
||||
pibindAll : FC -> PiInfo -> List (RigCount, Maybe Name, RawImp) ->
|
||||
RawImp -> RawImp
|
||||
pibindAll fc p [] scope = scope
|
||||
pibindAll fc p ((rig, n, ty) :: rest) scope
|
||||
= IPi fc rig p n ty (pibindAll fc p rest scope)
|
||||
|
||||
bindList : FileName -> FilePos -> IndentInfo ->
|
||||
bindList : FileName -> FilePos -> IndentInfo ->
|
||||
Rule (List (RigCount, Name, RawImp))
|
||||
bindList fname start indents
|
||||
= sepBy1 (symbol ",")
|
||||
(do rigc <- multiplicity
|
||||
n <- unqualifiedName
|
||||
end <- location
|
||||
ty <- option
|
||||
ty <- option
|
||||
(Implicit (MkFC fname start end) False)
|
||||
(do symbol ":"
|
||||
appExpr fname indents)
|
||||
@ -222,7 +222,7 @@ mutual
|
||||
pure (rig, UN n, ty))
|
||||
|
||||
|
||||
pibindList : FileName -> FilePos -> IndentInfo ->
|
||||
pibindList : FileName -> FilePos -> IndentInfo ->
|
||||
Rule (List (RigCount, Maybe Name, RawImp))
|
||||
pibindList fname start indents
|
||||
= do rigc <- multiplicity
|
||||
@ -239,7 +239,7 @@ mutual
|
||||
ty <- expr fname indents
|
||||
rig <- getMult rigc
|
||||
pure (rig, Just n, ty))
|
||||
|
||||
|
||||
autoImplicitPi : FileName -> IndentInfo -> Rule RawImp
|
||||
autoImplicitPi fname indents
|
||||
= do start <- location
|
||||
@ -347,8 +347,8 @@ mutual
|
||||
= do start <- location
|
||||
lhs <- appExpr fname indents
|
||||
caseRHS fname indents start lhs
|
||||
|
||||
caseRHS : FileName -> IndentInfo -> (Int, Int) -> RawImp ->
|
||||
|
||||
caseRHS : FileName -> IndentInfo -> (Int, Int) -> RawImp ->
|
||||
Rule ImpClause
|
||||
caseRHS fname indents start lhs
|
||||
= do symbol "=>"
|
||||
@ -392,7 +392,7 @@ mutual
|
||||
tm <- expr fname indents
|
||||
end <- location
|
||||
pure (IRewrite (MkFC fname start end) rule tm)
|
||||
|
||||
|
||||
lazy : FileName -> IndentInfo -> Rule RawImp
|
||||
lazy fname indents
|
||||
= do start <- location
|
||||
@ -440,8 +440,8 @@ mutual
|
||||
where
|
||||
mkPi : FilePos -> FilePos -> RawImp -> List (PiInfo, RawImp) -> RawImp
|
||||
mkPi start end arg [] = arg
|
||||
mkPi start end arg ((exp, a) :: as)
|
||||
= IPi (MkFC fname start end) RigW exp Nothing arg
|
||||
mkPi start end arg ((exp, a) :: as)
|
||||
= IPi (MkFC fname start end) RigW exp Nothing arg
|
||||
(mkPi start end a as)
|
||||
|
||||
export
|
||||
@ -490,7 +490,7 @@ mutual
|
||||
getFn (IVar _ n) = pure n
|
||||
getFn (IApp _ f a) = getFn f
|
||||
getFn (IImplicitApp _ f _ a) = getFn f
|
||||
getFn _ = fail "Not a function application"
|
||||
getFn _ = fail "Not a function application"
|
||||
|
||||
clause : Nat -> FileName -> IndentInfo -> Rule (Name, ImpClause)
|
||||
clause withArgs fname indents
|
||||
@ -506,13 +506,13 @@ mutual
|
||||
applyArgs f ((fc, a) :: args) = applyArgs (IApp fc f a) args
|
||||
|
||||
parseWithArg : Rule (FC, RawImp)
|
||||
parseWithArg
|
||||
parseWithArg
|
||||
= do symbol "|"
|
||||
start <- location
|
||||
tm <- expr fname indents
|
||||
end <- location
|
||||
pure (MkFC fname start end, tm)
|
||||
|
||||
|
||||
definition : FileName -> IndentInfo -> Rule ImpDecl
|
||||
definition fname indents
|
||||
= do start <- location
|
||||
@ -666,7 +666,7 @@ collectDefs (IDef loc fn cs :: ds)
|
||||
(ys, zs) => (y ++ ys, zs)
|
||||
|
||||
isClause : Name -> ImpDecl -> Maybe (List ImpClause)
|
||||
isClause n (IDef _ n' cs)
|
||||
isClause n (IDef _ n' cs)
|
||||
= if n == n' then Just cs else Nothing
|
||||
isClause n _ = Nothing
|
||||
collectDefs (INamespace loc nest ns nds :: ds)
|
||||
|
@ -19,13 +19,13 @@ import Data.NameMap
|
||||
|
||||
processDataOpt : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> DataOpt -> Core ()
|
||||
processDataOpt fc n NoHints
|
||||
processDataOpt fc n NoHints
|
||||
= pure ()
|
||||
processDataOpt fc ndef (SearchBy dets)
|
||||
processDataOpt fc ndef (SearchBy dets)
|
||||
= setDetermining fc ndef dets
|
||||
|
||||
checkRetType : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars -> NF vars ->
|
||||
Env Term vars -> NF vars ->
|
||||
(NF vars -> Core ()) -> Core ()
|
||||
checkRetType env (NBind fc x (Pi _ _ ty) sc) chk
|
||||
= do defs <- get Ctxt
|
||||
@ -35,18 +35,18 @@ checkRetType env nf chk = chk nf
|
||||
checkIsType : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Env Term vars -> NF vars -> Core ()
|
||||
checkIsType loc n env nf
|
||||
= checkRetType env nf
|
||||
(\nf => case nf of
|
||||
= checkRetType env nf
|
||||
(\nf => case nf of
|
||||
NType _ => pure ()
|
||||
_ => throw (BadTypeConType loc n))
|
||||
|
||||
checkFamily : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> Name -> Env Term vars -> NF vars -> Core ()
|
||||
checkFamily loc cn tn env nf
|
||||
= checkRetType env nf
|
||||
(\nf => case nf of
|
||||
= checkRetType env nf
|
||||
(\nf => case nf of
|
||||
NType _ => throw (BadDataConType loc cn tn)
|
||||
NTCon _ n' _ _ _ =>
|
||||
NTCon _ n' _ _ _ =>
|
||||
if tn == n'
|
||||
then pure ()
|
||||
else throw (BadDataConType loc cn tn)
|
||||
@ -56,7 +56,7 @@ checkCon : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
Env Term vars -> Visibility -> Name ->
|
||||
ImpTy -> Core Constructor
|
||||
checkCon {vars} opts nest env vis tn (MkImpTy fc cn_in ty_raw)
|
||||
@ -67,9 +67,9 @@ checkCon {vars} opts nest env vis tn (MkImpTy fc cn_in ty_raw)
|
||||
-- Check 'cn' is undefined
|
||||
Nothing <- lookupCtxtExact cn (gamma defs)
|
||||
| Just gdef => throw (AlreadyDefined fc cn)
|
||||
ty <-
|
||||
ty <-
|
||||
wrapError (InCon fc cn) $
|
||||
checkTerm !(resolveName cn) InType opts nest env
|
||||
checkTerm !(resolveName cn) InType opts nest env
|
||||
(IBindHere fc (PI Rig0) ty_raw)
|
||||
(gType fc)
|
||||
|
||||
@ -96,10 +96,10 @@ processData : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
Env Term vars -> FC -> Visibility ->
|
||||
ImpData -> Core ()
|
||||
processData {vars} eopts nest env fc vis (MkImpLater dfc n_in ty_raw)
|
||||
processData {vars} eopts nest env fc vis (MkImpLater dfc n_in ty_raw)
|
||||
= do n <- inCurrentNS n_in
|
||||
ty_raw <- bindTypeNames vars ty_raw
|
||||
|
||||
@ -107,10 +107,10 @@ processData {vars} eopts nest env fc vis (MkImpLater dfc n_in ty_raw)
|
||||
-- Check 'n' is undefined
|
||||
Nothing <- lookupCtxtExact n (gamma defs)
|
||||
| Just gdef => throw (AlreadyDefined fc n)
|
||||
|
||||
(ty, _) <-
|
||||
|
||||
(ty, _) <-
|
||||
wrapError (InCon fc n) $
|
||||
elabTerm !(resolveName n) InType eopts nest env
|
||||
elabTerm !(resolveName n) InType eopts nest env
|
||||
(IBindHere fc (PI Rig0) ty_raw)
|
||||
(Just (gType dfc))
|
||||
let fullty = abstractEnvType dfc env ty
|
||||
@ -144,7 +144,7 @@ processData {vars} eopts nest env fc vis (MkImpData dfc n_in ty_raw opts cons_ra
|
||||
defs <- get Ctxt
|
||||
(ty, _) <-
|
||||
wrapError (InCon fc n) $
|
||||
elabTerm !(resolveName n) InType eopts nest env
|
||||
elabTerm !(resolveName n) InType eopts nest env
|
||||
(IBindHere fc (PI Rig0) ty_raw)
|
||||
(Just (gType dfc))
|
||||
let fullty = abstractEnvType dfc env ty
|
||||
|
@ -24,11 +24,11 @@ process : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
List ElabOpt ->
|
||||
NestedNames vars -> Env Term vars -> ImpDecl -> Core ()
|
||||
process eopts nest env (IClaim fc rig vis opts ty)
|
||||
process eopts nest env (IClaim fc rig vis opts ty)
|
||||
= processType eopts nest env fc rig vis opts ty
|
||||
process eopts nest env (IData fc vis ddef)
|
||||
process eopts nest env (IData fc vis ddef)
|
||||
= processData eopts nest env fc vis ddef
|
||||
process eopts nest env (IDef fc fname def)
|
||||
process eopts nest env (IDef fc fname def)
|
||||
= processDef eopts nest env fc fname def
|
||||
process eopts nest env (IParameters fc ps decls)
|
||||
= processParams nest env fc ps decls
|
||||
@ -42,10 +42,12 @@ process eopts nest env (INamespace fc nested ns decls)
|
||||
newns <- getNS
|
||||
traverse_ (processDecl eopts nest env) decls
|
||||
defs <- get Ctxt
|
||||
put Ctxt (record { currentNS = cns,
|
||||
put Ctxt (record { currentNS = cns,
|
||||
nestedNS = if nested
|
||||
then newns :: nns
|
||||
else nns } defs)
|
||||
process eopts nest env (ITransform fc lhs rhs)
|
||||
= throw (GenericMsg fc "%transform not yet implemented")
|
||||
process {c} eopts nest env (IPragma act)
|
||||
= act c nest env
|
||||
process eopts nest env (ILog n)
|
||||
|
@ -26,22 +26,22 @@ import Data.NameMap
|
||||
|
||||
mutual
|
||||
mismatchNF : Defs -> NF vars -> NF vars -> Core Bool
|
||||
mismatchNF defs (NTCon _ xn xt _ xargs) (NTCon _ yn yt _ yargs)
|
||||
mismatchNF defs (NTCon _ xn xt _ xargs) (NTCon _ yn yt _ yargs)
|
||||
= if xn /= yn
|
||||
then pure True
|
||||
else anyM (mismatch defs) (zip xargs yargs)
|
||||
mismatchNF defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs)
|
||||
else anyM (mismatch defs) (zip xargs yargs)
|
||||
mismatchNF defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs)
|
||||
= if xt /= yt
|
||||
then pure True
|
||||
else anyM (mismatch defs) (zip xargs yargs)
|
||||
else anyM (mismatch defs) (zip xargs yargs)
|
||||
mismatchNF defs (NPrimVal _ xc) (NPrimVal _ yc) = pure (xc /= yc)
|
||||
mismatchNF defs (NDelayed _ _ x) (NDelayed _ _ y) = mismatchNF defs x y
|
||||
mismatchNF defs (NDelay _ _ _ x) (NDelay _ _ _ y)
|
||||
mismatchNF defs (NDelay _ _ _ x) (NDelay _ _ _ y)
|
||||
= mismatchNF defs !(evalClosure defs x) !(evalClosure defs y)
|
||||
mismatchNF _ _ _ = pure False
|
||||
|
||||
mismatch : Defs -> (Closure vars, Closure vars) -> Core Bool
|
||||
mismatch defs (x, y)
|
||||
mismatch defs (x, y)
|
||||
= mismatchNF defs !(evalClosure defs x) !(evalClosure defs y)
|
||||
|
||||
-- If the terms have the same type constructor at the head, and one of
|
||||
@ -62,7 +62,7 @@ impossibleOK defs (NPrimVal _ x) (NPrimVal _ y) = pure (x /= y)
|
||||
impossibleOK defs (NDCon _ _ _ _ _) (NPrimVal _ _) = pure True
|
||||
impossibleOK defs (NPrimVal _ _) (NDCon _ _ _ _ _) = pure True
|
||||
impossibleOK defs x y = pure False
|
||||
|
||||
|
||||
export
|
||||
impossibleErrOK : {auto c : Ref Ctxt Defs} ->
|
||||
Defs -> Error -> Core Bool
|
||||
@ -91,10 +91,10 @@ impossibleErrOK defs _ = pure False
|
||||
-- so all the outer things are marked as 'DropCons'
|
||||
extendEnv : Env Term vars -> SubVars inner vars ->
|
||||
NestedNames vars ->
|
||||
Term vars -> Term vars ->
|
||||
Core (vars' **
|
||||
Term vars -> Term vars ->
|
||||
Core (vars' **
|
||||
(SubVars inner vars',
|
||||
Env Term vars', NestedNames vars',
|
||||
Env Term vars', NestedNames vars',
|
||||
Term vars', Term vars'))
|
||||
extendEnv env p nest (Bind _ n (PVar c pi tmty) sc) (Bind _ n' (PVTy _ _) tysc) with (nameEq n n')
|
||||
extendEnv env p nest (Bind _ n (PVar c pi tmty) sc) (Bind _ n' (PVTy _ _) tysc) | Nothing
|
||||
@ -107,7 +107,7 @@ extendEnv env p nest (Bind _ n (PLet c tmval tmty) sc) (Bind _ n' (PLet _ _ _) t
|
||||
-- PLet on the left becomes Let on the right, to give it computational force
|
||||
extendEnv env p nest (Bind _ n (PLet c tmval tmty) sc) (Bind _ n (PLet _ _ _) tysc) | (Just Refl)
|
||||
= extendEnv (Let c tmval tmty :: env) (DropCons p) (weaken nest) sc tysc
|
||||
extendEnv env p nest tm ty
|
||||
extendEnv env p nest tm ty
|
||||
= pure (_ ** (p, env, nest, tm, ty))
|
||||
|
||||
-- Find names which are applied to a function in a Rig1/Rig0 position,
|
||||
@ -117,9 +117,9 @@ extendEnv env p nest tm ty
|
||||
-- only ones we're checking linearity of (we may be shadowing names if this
|
||||
-- is a local definition, so we need to leave the earlier ones alone)
|
||||
findLinear : {auto c : Ref Ctxt Defs} ->
|
||||
Bool -> Nat -> RigCount -> Term vars ->
|
||||
Bool -> Nat -> RigCount -> Term vars ->
|
||||
Core (List (Name, RigCount))
|
||||
findLinear top bound rig (Bind fc n b sc)
|
||||
findLinear top bound rig (Bind fc n b sc)
|
||||
= findLinear top (S bound) rig sc
|
||||
findLinear top bound rig (As fc _ p)
|
||||
= findLinear top bound rig p
|
||||
@ -137,22 +137,22 @@ findLinear top bound rig tm
|
||||
accessible Func r = if top then r else Rig0
|
||||
accessible _ r = r
|
||||
|
||||
findLinArg : RigCount -> NF [] -> List (Term vars) ->
|
||||
findLinArg : RigCount -> NF [] -> List (Term vars) ->
|
||||
Core (List (Name, RigCount))
|
||||
findLinArg rig ty (As fc _ p :: as) = findLinArg rig ty (p :: as)
|
||||
findLinArg rig (NBind _ x (Pi c _ _) sc) (Local {name=a} fc _ idx prf :: as)
|
||||
findLinArg rig (NBind _ x (Pi c _ _) sc) (Local {name=a} fc _ idx prf :: as)
|
||||
= do defs <- get Ctxt
|
||||
if idx < bound
|
||||
then do sc' <- sc defs (toClosure defaultOpts [] (Ref fc Bound x))
|
||||
pure $ (a, rigMult c rig) ::
|
||||
pure $ (a, rigMult c rig) ::
|
||||
!(findLinArg rig sc' as)
|
||||
else do sc' <- sc defs (toClosure defaultOpts [] (Ref fc Bound x))
|
||||
findLinArg rig sc' as
|
||||
findLinArg rig (NBind fc x (Pi c _ _) sc) (a :: as)
|
||||
findLinArg rig (NBind fc x (Pi c _ _) sc) (a :: as)
|
||||
= do defs <- get Ctxt
|
||||
pure $ !(findLinear False bound (rigMult c rig) a) ++
|
||||
!(findLinArg rig !(sc defs (toClosure defaultOpts [] (Ref fc Bound x))) as)
|
||||
findLinArg rig ty (a :: as)
|
||||
findLinArg rig ty (a :: as)
|
||||
= pure $ !(findLinear False bound rig a) ++ !(findLinArg rig ty as)
|
||||
findLinArg _ _ [] = pure []
|
||||
|
||||
@ -178,7 +178,7 @@ combineLinear loc ((n, count) :: cs)
|
||||
= case lookupAll n cs of
|
||||
[] => pure $ (n, count) :: !(combineLinear loc cs)
|
||||
counts => do count' <- combineAll count counts
|
||||
pure $ (n, count') ::
|
||||
pure $ (n, count') ::
|
||||
!(combineLinear loc (filter notN cs))
|
||||
where
|
||||
notN : (Name, RigCount) -> Bool
|
||||
@ -186,7 +186,7 @@ combineLinear loc ((n, count) :: cs)
|
||||
|
||||
lookupAll : Name -> List (Name, RigCount) -> List RigCount
|
||||
lookupAll n [] = []
|
||||
lookupAll n ((n', c) :: cs)
|
||||
lookupAll n ((n', c) :: cs)
|
||||
= if n == n' then c :: lookupAll n cs else lookupAll n cs
|
||||
|
||||
combine : RigCount -> RigCount -> Core RigCount
|
||||
@ -209,14 +209,14 @@ checkLHS : {vars : _} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
(mult : RigCount) -> (hashit : Bool) ->
|
||||
Int -> List ElabOpt -> NestedNames vars -> Env Term vars ->
|
||||
FC -> RawImp ->
|
||||
FC -> RawImp ->
|
||||
Core (RawImp, -- checked LHS with implicits added
|
||||
(vars' ** (SubVars vars vars',
|
||||
Env Term vars', NestedNames vars',
|
||||
Env Term vars', NestedNames vars',
|
||||
Term vars', Term vars')))
|
||||
checkLHS {vars} mult hashit n opts nest env fc lhs_in
|
||||
= do defs <- get Ctxt
|
||||
lhs_raw <- lhsInCurrentNS nest lhs_in
|
||||
lhs_raw <- lhsInCurrentNS nest lhs_in
|
||||
autoimp <- isAutoImplicits
|
||||
autoImplicits True
|
||||
(_, lhs_bound) <- bindNames False lhs_raw
|
||||
@ -226,9 +226,9 @@ checkLHS {vars} mult hashit n opts nest env fc lhs_in
|
||||
log 5 $ "Checking LHS of " ++ show !(getFullName (Resolved n)) ++
|
||||
" " ++ show lhs
|
||||
logEnv 5 "In env" env
|
||||
(lhstm, lhstyg) <-
|
||||
(lhstm, lhstyg) <-
|
||||
wrapError (InLHS fc !(getFullName (Resolved n))) $
|
||||
elabTerm n (InLHS mult) opts nest env
|
||||
elabTerm n (InLHS mult) opts nest env
|
||||
(IBindHere fc PATTERN lhs) Nothing
|
||||
logTerm 10 "Checked LHS term" lhstm
|
||||
lhsty <- getTerm lhstyg
|
||||
@ -240,7 +240,7 @@ checkLHS {vars} mult hashit n opts nest env fc lhs_in
|
||||
lhsty <- normaliseHoles defs env lhsty
|
||||
linvars_in <- findLinear True 0 Rig1 lhstm
|
||||
logTerm 10 "Checked LHS term after normalise" lhstm
|
||||
log 5 $ "Linearity of names in " ++ show n ++ ": " ++
|
||||
log 5 $ "Linearity of names in " ++ show n ++ ": " ++
|
||||
show linvars_in
|
||||
|
||||
linvars <- combineLinear fc linvars_in
|
||||
@ -260,59 +260,59 @@ plicit (PVar _ p _) = p
|
||||
plicit _ = Explicit
|
||||
|
||||
bindNotReq : {vs : _} ->
|
||||
FC -> Int -> Env Term vs -> (sub : SubVars pre vs) ->
|
||||
FC -> Int -> Env Term vs -> (sub : SubVars pre vs) ->
|
||||
List (PiInfo, Name) ->
|
||||
Term vs -> (List (PiInfo, Name), Term pre)
|
||||
bindNotReq fc i [] SubRefl ns tm = (ns, embed tm)
|
||||
bindNotReq fc i (b :: env) SubRefl ns tm
|
||||
= let tmptm = subst (Ref fc Bound (MN "arg" i)) tm
|
||||
bindNotReq fc i (b :: env) SubRefl ns tm
|
||||
= let tmptm = subst (Ref fc Bound (MN "arg" i)) tm
|
||||
(ns', btm) = bindNotReq fc (1 + i) env SubRefl ns tmptm in
|
||||
(ns', refToLocal (MN "arg" i) _ btm)
|
||||
bindNotReq fc i (b :: env) (KeepCons p) ns tm
|
||||
= let tmptm = subst (Ref fc Bound (MN "arg" i)) tm
|
||||
bindNotReq fc i (b :: env) (KeepCons p) ns tm
|
||||
= let tmptm = subst (Ref fc Bound (MN "arg" i)) tm
|
||||
(ns', btm) = bindNotReq fc (1 + i) env p ns tmptm in
|
||||
(ns', refToLocal (MN "arg" i) _ btm)
|
||||
bindNotReq {vs = n :: _} fc i (b :: env) (DropCons p) ns tm
|
||||
bindNotReq {vs = n :: _} fc i (b :: env) (DropCons p) ns tm
|
||||
= bindNotReq fc i env p ((plicit b, n) :: ns)
|
||||
(Bind fc _ (Pi (multiplicity b) Explicit (binderType b)) tm)
|
||||
|
||||
bindReq : {vs : _} ->
|
||||
FC -> Env Term vs -> (sub : SubVars pre vs) ->
|
||||
FC -> Env Term vs -> (sub : SubVars pre vs) ->
|
||||
List (PiInfo, Name) ->
|
||||
Term pre -> Maybe (List (PiInfo, Name), List Name, ClosedTerm)
|
||||
bindReq {vs} fc env SubRefl ns tm
|
||||
bindReq {vs} fc env SubRefl ns tm
|
||||
= pure (ns, notLets [] _ env, abstractEnvType fc env tm)
|
||||
where
|
||||
notLets : List Name -> (vars : List Name) -> Env Term vars -> List Name
|
||||
notLets acc [] _ = acc
|
||||
notLets acc (v :: vs) (Let _ _ _ :: env) = notLets acc vs env
|
||||
notLets acc (v :: vs) (_ :: env) = notLets (v :: acc) vs env
|
||||
bindReq {vs = n :: _} fc (b :: env) (KeepCons p) ns tm
|
||||
bindReq {vs = n :: _} fc (b :: env) (KeepCons p) ns tm
|
||||
= do b' <- shrinkBinder b p
|
||||
bindReq fc env p ((plicit b, n) :: ns)
|
||||
(Bind fc _ (Pi (multiplicity b) Explicit (binderType b')) tm)
|
||||
bindReq fc (b :: env) (DropCons p) ns tm
|
||||
bindReq fc (b :: env) (DropCons p) ns tm
|
||||
= bindReq fc env p ns tm
|
||||
|
||||
-- Return whether any of the pattern variables are in a trivially empty
|
||||
-- type, where trivally empty means one of:
|
||||
-- * No constructors
|
||||
-- * Every constructor of the family has a return type which conflicts with
|
||||
-- * Every constructor of the family has a return type which conflicts with
|
||||
-- the given constructor's type
|
||||
hasEmptyPat : Defs -> Env Term vars -> Term vars -> Core Bool
|
||||
hasEmptyPat defs env (Bind fc x (PVar c p ty) sc)
|
||||
= pure $ !(isEmpty defs !(nf defs env ty))
|
||||
= pure $ !(isEmpty defs !(nf defs env ty))
|
||||
|| !(hasEmptyPat defs (PVar c p ty :: env) sc)
|
||||
hasEmptyPat defs env _ = pure False
|
||||
|
||||
|
||||
-- For checking with blocks as nested names
|
||||
applyEnv : {auto c : Ref Ctxt Defs} ->
|
||||
Env Term vars -> Name ->
|
||||
Env Term vars -> Name ->
|
||||
Core (Name, (Maybe Name, FC -> NameType -> Term vars))
|
||||
applyEnv env withname
|
||||
= do n' <- resolveName withname
|
||||
pure (withname, (Just withname,
|
||||
\fc, nt => applyTo fc
|
||||
pure (withname, (Just withname,
|
||||
\fc, nt => applyTo fc
|
||||
(Ref fc nt (Resolved n')) env))
|
||||
|
||||
-- Check a pattern clause, returning the component of the 'Case' expression it
|
||||
@ -335,15 +335,15 @@ checkClause mult hashit n opts nest env (ImpossibleClause fc lhs)
|
||||
|
||||
log 5 $ "Checking " ++ show lhs
|
||||
logEnv 5 "In env" env
|
||||
(lhstm, lhstyg) <-
|
||||
elabTerm n (InLHS mult) opts nest env
|
||||
(lhstm, lhstyg) <-
|
||||
elabTerm n (InLHS mult) opts nest env
|
||||
(IBindHere fc PATTERN lhs) Nothing
|
||||
defs <- get Ctxt
|
||||
lhs <- normaliseHoles defs env lhstm
|
||||
if !(hasEmptyPat defs env lhs)
|
||||
then pure Nothing
|
||||
else throw (ValidCase fc env (Left lhs)))
|
||||
(\err =>
|
||||
(\err =>
|
||||
case err of
|
||||
ValidCase _ _ _ => throw err
|
||||
_ => do defs <- get Ctxt
|
||||
@ -351,7 +351,7 @@ checkClause mult hashit n opts nest env (ImpossibleClause fc lhs)
|
||||
then pure Nothing
|
||||
else throw (ValidCase fc env (Right err)))
|
||||
checkClause {vars} mult hashit n opts nest env (PatClause fc lhs_in rhs)
|
||||
= do (_, (vars' ** (sub', env', nest', lhstm', lhsty'))) <-
|
||||
= do (_, (vars' ** (sub', env', nest', lhstm', lhsty'))) <-
|
||||
checkLHS mult hashit n opts nest env fc lhs_in
|
||||
let rhsMode = case mult of
|
||||
Rig0 => InType
|
||||
@ -362,21 +362,21 @@ checkClause {vars} mult hashit n opts nest env (PatClause fc lhs_in rhs)
|
||||
clearHoleLHS
|
||||
|
||||
logTerm 5 "RHS term" rhstm
|
||||
when hashit $
|
||||
when hashit $
|
||||
do addHash lhstm'
|
||||
addHash rhstm
|
||||
|
||||
-- If the rhs is a hole, record the lhs in the metadata because we
|
||||
-- If the rhs is a hole, record the lhs in the metadata because we
|
||||
-- might want to split it interactively
|
||||
case rhstm of
|
||||
Meta _ _ _ _ =>
|
||||
Meta _ _ _ _ =>
|
||||
addLHS (getFC lhs_in) (length env) env' lhstm'
|
||||
_ => pure ()
|
||||
|
||||
pure (Just (MkClause env' lhstm' rhstm))
|
||||
-- TODO: (to decide) With is complicated. Move this into its own module?
|
||||
checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs)
|
||||
= do (lhs, (vars' ** (sub', env', nest', lhspat, reqty))) <-
|
||||
= do (lhs, (vars' ** (sub', env', nest', lhspat, reqty))) <-
|
||||
checkLHS mult hashit n opts nest env fc lhs_in
|
||||
let wmode
|
||||
= case mult of
|
||||
@ -386,7 +386,7 @@ checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs
|
||||
(wval, gwvalTy) <- wrapError (InRHS fc !(getFullName (Resolved n))) $
|
||||
elabTermSub n wmode opts nest' env' env sub' wval_raw Nothing
|
||||
clearHoleLHS
|
||||
|
||||
|
||||
logTerm 5 "With value" wval
|
||||
logTerm 3 "Required type" reqty
|
||||
wvalTy <- getTerm gwvalTy
|
||||
@ -418,7 +418,7 @@ checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs
|
||||
|
||||
wtyScope <- replace defs scenv !(nf defs scenv (weaken wval))
|
||||
(Local fc (Just False) _ First)
|
||||
!(nf defs scenv
|
||||
!(nf defs scenv
|
||||
(weaken {n=wargn} notreqty))
|
||||
let bNotReq = Bind fc wargn (Pi RigW Explicit wvalTy) wtyScope
|
||||
|
||||
@ -428,11 +428,11 @@ checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs
|
||||
-- list of argument names - 'Just' means we need to match the name
|
||||
-- in the with clauses to find out what the pattern should be.
|
||||
-- 'Nothing' means it's the with pattern (so wargn)
|
||||
let wargNames
|
||||
= map Just reqns ++
|
||||
let wargNames
|
||||
= map Just reqns ++
|
||||
Nothing :: map Just notreqns
|
||||
|
||||
logTerm 3 "With function type" wtype
|
||||
logTerm 3 "With function type" wtype
|
||||
log 5 $ "Argument names " ++ show wargNames
|
||||
|
||||
wname <- genWithName n
|
||||
@ -443,7 +443,7 @@ checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs
|
||||
|
||||
log 3 $ "Applying to with argument " ++ show rhs_in
|
||||
rhs <- wrapError (InRHS fc !(getFullName (Resolved n))) $
|
||||
checkTermSub n wmode opts nest' env' env sub' rhs_in
|
||||
checkTermSub n wmode opts nest' env' env sub' rhs_in
|
||||
(gnf env' reqty)
|
||||
|
||||
-- Generate new clauses by rewriting the matched arguments
|
||||
@ -478,12 +478,12 @@ checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs
|
||||
keepOldEnv (KeepCons p) (KeepCons p')
|
||||
= let (_ ** rest) = keepOldEnv p p' in
|
||||
(_ ** KeepCons rest)
|
||||
|
||||
|
||||
-- Rewrite the clauses in the block to use an updated LHS.
|
||||
-- 'drop' is the number of additional with arguments we expect (i.e.
|
||||
-- the things to drop from the end before matching LHSs)
|
||||
mkClauseWith : (drop : Nat) -> Name -> List (Maybe (PiInfo, Name)) ->
|
||||
RawImp -> ImpClause ->
|
||||
RawImp -> ImpClause ->
|
||||
Core ImpClause
|
||||
mkClauseWith drop wname wargnames lhs (PatClause ploc patlhs rhs)
|
||||
= do newlhs <- getNewLHS ploc drop nest wname wargnames lhs patlhs
|
||||
@ -519,17 +519,17 @@ mkRunTime n
|
||||
let PMDef r cargs tree_ct _ pats = definition gdef
|
||||
| _ => pure () -- not a function definition
|
||||
let ty = type gdef
|
||||
(rargs ** tree_rt) <- getPMDef (location gdef) RunTime n ty
|
||||
(rargs ** tree_rt) <- getPMDef (location gdef) RunTime n ty
|
||||
!(traverse (toClause (location gdef)) pats)
|
||||
let Just Refl = nameListEq cargs rargs
|
||||
| Nothing => throw (InternalError "WAT")
|
||||
addDef n (record { definition = PMDef r cargs tree_ct tree_rt pats
|
||||
addDef n (record { definition = PMDef r cargs tree_ct tree_rt pats
|
||||
} gdef)
|
||||
pure ()
|
||||
where
|
||||
toClause : FC -> (vars ** (Env Term vars, Term vars, Term vars)) ->
|
||||
Core Clause
|
||||
toClause fc (_ ** (env, lhs, rhs))
|
||||
toClause fc (_ ** (env, lhs, rhs))
|
||||
= do rhs_erased <- linearCheck fc Rig1 True env rhs
|
||||
pure $ MkClause env lhs rhs_erased
|
||||
|
||||
@ -565,7 +565,7 @@ calcRefs at fn
|
||||
traverse_ (calcRefs at) (keys refs)
|
||||
|
||||
toPats : Clause -> (vs ** (Env Term vs, Term vs, Term vs))
|
||||
toPats (MkClause {vars} env lhs rhs)
|
||||
toPats (MkClause {vars} env lhs rhs)
|
||||
= (_ ** (env, lhs, rhs))
|
||||
|
||||
export
|
||||
@ -590,9 +590,9 @@ processDef opts nest env fc n_in cs_in
|
||||
cs <- traverse (checkClause mult hashit nidx opts nest env) cs_in
|
||||
let pats = map toPats (mapMaybe id cs)
|
||||
|
||||
(cargs ** tree_ct) <- getPMDef fc CompileTime n ty
|
||||
(cargs ** tree_ct) <- getPMDef fc CompileTime n ty
|
||||
(mapMaybe id cs)
|
||||
|
||||
|
||||
logC 5 (do t <- toFullNames tree_ct
|
||||
pure ("Case tree for " ++ show n ++ ": " ++ show t))
|
||||
|
||||
@ -645,10 +645,10 @@ processDef opts nest env fc n_in cs_in
|
||||
catchAll Nothing = False
|
||||
catchAll (Just (MkClause env lhs _))
|
||||
= all simplePat (getArgs lhs)
|
||||
|
||||
|
||||
-- Return 'Nothing' if the clause is impossible, otherwise return the
|
||||
-- original
|
||||
checkImpossible : Int -> RigCount -> ClosedTerm ->
|
||||
checkImpossible : Int -> RigCount -> ClosedTerm ->
|
||||
Core (Maybe ClosedTerm)
|
||||
checkImpossible n mult tm
|
||||
= do itm <- unelabNoPatvars [] tm
|
||||
@ -662,7 +662,7 @@ processDef opts nest env fc n_in cs_in
|
||||
(\err => case err of
|
||||
WhenUnifying _ env l r err
|
||||
=> do defs <- get Ctxt
|
||||
if !(impossibleOK defs !(nf defs env l)
|
||||
if !(impossibleOK defs !(nf defs env l)
|
||||
!(nf defs env r))
|
||||
then pure Nothing
|
||||
else pure (Just tm)
|
||||
@ -676,8 +676,8 @@ processDef opts nest env fc n_in cs_in
|
||||
pure []
|
||||
else getMissing fc (Resolved n) ctree
|
||||
logC 3 (do mc <- traverse toFullNames missCase
|
||||
pure ("Initially missing in " ++
|
||||
show !(getFullName (Resolved n)) ++ ":\n" ++
|
||||
pure ("Initially missing in " ++
|
||||
show !(getFullName (Resolved n)) ++ ":\n" ++
|
||||
showSep "\n" (map show mc)))
|
||||
missImp <- traverse (checkImpossible n mult) missCase
|
||||
let miss = mapMaybe id missImp
|
||||
|
@ -13,7 +13,7 @@ import TTImp.Elab.Check
|
||||
import TTImp.TTImp
|
||||
|
||||
%default covering
|
||||
|
||||
|
||||
extend : Env Term extvs -> SubVars vs extvs ->
|
||||
NestedNames extvs ->
|
||||
Term extvs ->
|
||||
@ -28,7 +28,7 @@ processParams : {vars : _} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
NestedNames vars ->
|
||||
Env Term vars ->
|
||||
Env Term vars ->
|
||||
FC -> List (Name, RawImp) -> List ImpDecl ->
|
||||
Core ()
|
||||
processParams {vars} {c} {m} {u} nest env fc ps ds
|
||||
@ -38,7 +38,7 @@ processParams {vars} {c} {m} {u} nest env fc ps ds
|
||||
let pty_raw = mkParamTy ps
|
||||
pty_imp <- bindTypeNames vars (IBindHere fc (PI Rig0) pty_raw)
|
||||
log 10 $ "Checking " ++ show pty_imp
|
||||
pty <- checkTerm (-1) InType []
|
||||
pty <- checkTerm (-1) InType []
|
||||
nest env pty_imp (gType fc)
|
||||
let (vs ** (prf, env', nest')) = extend env SubRefl nest pty
|
||||
|
||||
@ -53,13 +53,13 @@ processParams {vars} {c} {m} {u} nest env fc ps ds
|
||||
where
|
||||
mkParamTy : List (Name, RawImp) -> RawImp
|
||||
mkParamTy [] = IType fc
|
||||
mkParamTy ((n, ty) :: ps)
|
||||
mkParamTy ((n, ty) :: ps)
|
||||
= IPi fc RigW Explicit (Just n) ty (mkParamTy ps)
|
||||
|
||||
applyEnv : Env Term vs -> Name ->
|
||||
applyEnv : Env Term vs -> Name ->
|
||||
Core (Name, (Maybe Name, FC -> NameType -> Term vs))
|
||||
applyEnv env n
|
||||
= do n' <- resolveName n -- it'll be Resolved by expandAmbigName
|
||||
pure (Resolved n', (Nothing,
|
||||
\fc, nt => applyTo fc
|
||||
pure (Resolved n', (Nothing,
|
||||
\fc, nt => applyTo fc
|
||||
(Ref fc nt (Resolved n')) env))
|
||||
|
@ -17,7 +17,7 @@ import TTImp.Utils
|
||||
|
||||
mkDataTy : FC -> List (Name, RawImp) -> RawImp
|
||||
mkDataTy fc [] = IType fc
|
||||
mkDataTy fc ((n, ty) :: ps)
|
||||
mkDataTy fc ((n, ty) :: ps)
|
||||
= IPi fc RigW Explicit (Just n) ty (mkDataTy fc ps)
|
||||
|
||||
mkCon : Name -> Name
|
||||
@ -27,14 +27,14 @@ mkCon n = DN (show n) (MN ("__mk" ++ show n) 0)
|
||||
elabRecord : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
List ElabOpt -> FC -> Env Term vars ->
|
||||
List ElabOpt -> FC -> Env Term vars ->
|
||||
NestedNames vars -> Visibility ->
|
||||
Name ->
|
||||
(params : List (Name, RawImp)) ->
|
||||
(conName : Maybe Name) ->
|
||||
List IField ->
|
||||
Core ()
|
||||
elabRecord {vars} eopts fc env nest vis tn params rcon fields
|
||||
elabRecord {vars} eopts fc env nest vis tn params rcon fields
|
||||
= do let conName_in = maybe (mkCon tn) id rcon
|
||||
conName <- inCurrentNS conName_in
|
||||
elabAsData conName
|
||||
@ -50,18 +50,18 @@ elabRecord {vars} eopts fc env nest vis tn params rcon fields
|
||||
fname : IField -> Name
|
||||
fname (MkIField fc c p n ty) = n
|
||||
|
||||
farg : IField ->
|
||||
farg : IField ->
|
||||
(Maybe Name, RigCount, PiInfo, RawImp)
|
||||
farg (MkIField fc c p n ty) = (Just n, c, p, ty)
|
||||
|
||||
mkTy : List (Maybe Name, RigCount, PiInfo, RawImp) ->
|
||||
|
||||
mkTy : List (Maybe Name, RigCount, PiInfo, RawImp) ->
|
||||
RawImp -> RawImp
|
||||
mkTy [] ret = ret
|
||||
mkTy ((n, c, imp, argty) :: args) ret
|
||||
= IPi fc c imp n argty (mkTy args ret)
|
||||
|
||||
elabAsData : Name -> Core ()
|
||||
elabAsData cname
|
||||
elabAsData cname
|
||||
= do let retty = apply (IVar fc tn) (map (IVar fc) (map fst params))
|
||||
let conty = mkTy (map jname params) $
|
||||
mkTy (map farg fields) retty
|
||||
@ -79,7 +79,7 @@ elabRecord {vars} eopts fc env nest vis tn params rcon fields
|
||||
countExp _ = 0
|
||||
|
||||
-- Generate getters from the elaborated record constructor type
|
||||
elabGetters : Name -> RawImp ->
|
||||
elabGetters : Name -> RawImp ->
|
||||
(done : Nat) -> -- number of explicit fields processed
|
||||
List (Name, RawImp) -> -- names to update in types
|
||||
-- (for dependent records, where a field's type may depend
|
||||
@ -89,26 +89,26 @@ elabRecord {vars} eopts fc env nest vis tn params rcon fields
|
||||
elabGetters con recTy done upds tyenv (Bind bfc n b@(Pi rc imp ty_chk) sc)
|
||||
= if n `elem` map fst params
|
||||
then elabGetters con recTy
|
||||
(if imp == Explicit
|
||||
(if imp == Explicit
|
||||
then S done else done)
|
||||
upds (b :: tyenv) sc
|
||||
else
|
||||
else
|
||||
do let fldName = n
|
||||
gname <- inCurrentNS fldName
|
||||
ty <- unelab tyenv ty_chk
|
||||
let ty' = substNames vars upds ty
|
||||
let rname = MN "rec" 0
|
||||
gty <- bindTypeNames
|
||||
gty <- bindTypeNames
|
||||
(map fst params ++ map fname fields ++ vars) $
|
||||
mkTy (map jname params) $
|
||||
IPi fc RigW Explicit (Just rname) recTy ty'
|
||||
log 5 $ "Projection " ++ show gname ++ " : " ++ show gty
|
||||
let lhs_exp
|
||||
let lhs_exp
|
||||
= apply (IVar fc con)
|
||||
(replicate done (Implicit fc True) ++
|
||||
(if imp == Explicit
|
||||
then [IBindVar fc (nameRoot fldName)]
|
||||
else []) ++
|
||||
else []) ++
|
||||
(replicate (countExp sc) (Implicit fc True)))
|
||||
let lhs = IApp fc (IVar fc gname)
|
||||
(if imp == Explicit
|
||||
@ -123,7 +123,7 @@ elabRecord {vars} eopts fc env nest vis tn params rcon fields
|
||||
(IDef fc gname [PatClause fc lhs (IVar fc fldName)])
|
||||
let upds' = (n, IApp fc (IVar fc gname) (IVar fc rname)) :: upds
|
||||
elabGetters con recTy
|
||||
(if imp == Explicit
|
||||
(if imp == Explicit
|
||||
then S done else done)
|
||||
upds' (b :: tyenv) sc
|
||||
elabGetters con recTy done upds _ _ = pure ()
|
||||
@ -133,7 +133,7 @@ export
|
||||
processRecord : {auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
List ElabOpt -> NestedNames vars ->
|
||||
Env Term vars -> Visibility ->
|
||||
ImpRecord -> Core ()
|
||||
processRecord eopts nest env vis (MkImpRecord fc n ps cons fs)
|
||||
|
@ -25,12 +25,12 @@ getRetTy defs (NBind fc _ (Pi _ _ _) sc)
|
||||
= getRetTy defs !(sc defs (toClosure defaultOpts [] (Erased fc)))
|
||||
getRetTy defs (NTCon _ n _ _ _) = pure n
|
||||
getRetTy defs ty
|
||||
= throw (GenericMsg (getLoc ty)
|
||||
= throw (GenericMsg (getLoc ty)
|
||||
"Can only add hints for concrete return types")
|
||||
|
||||
processFnOpt : {auto c : Ref Ctxt Defs} ->
|
||||
FC -> Name -> FnOpt -> Core ()
|
||||
processFnOpt fc ndef Inline
|
||||
processFnOpt fc ndef Inline
|
||||
= setFlag fc ndef Inline
|
||||
processFnOpt fc ndef (Hint d)
|
||||
= do defs <- get Ctxt
|
||||
@ -59,7 +59,7 @@ processFnOpt fc ndef PartialOK
|
||||
-- Similarly set foreign definitions. Possibly combine these?
|
||||
initDef : {auto c : Ref Ctxt Defs} ->
|
||||
Name -> Env Term vars -> Term vars -> List FnOpt -> Core Def
|
||||
initDef n env ty []
|
||||
initDef n env ty []
|
||||
= do addUserHole n
|
||||
pure None
|
||||
initDef n env ty (ExternFn :: opts)
|
||||
@ -77,7 +77,7 @@ processType : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
List ElabOpt -> NestedNames vars -> Env Term vars ->
|
||||
List ElabOpt -> NestedNames vars -> Env Term vars ->
|
||||
FC -> RigCount -> Visibility ->
|
||||
List FnOpt -> ImpTy -> Core ()
|
||||
processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc n_in ty_raw)
|
||||
@ -86,17 +86,17 @@ processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc n_in ty_raw)
|
||||
|
||||
log 1 $ "Processing " ++ show n
|
||||
log 5 $ "Checking type decl " ++ show n ++ " : " ++ show ty_raw
|
||||
idx <- resolveName n
|
||||
idx <- resolveName n
|
||||
|
||||
-- Check 'n' is undefined
|
||||
defs <- get Ctxt
|
||||
Nothing <- lookupCtxtExact (Resolved idx) (gamma defs)
|
||||
| Just _ => throw (AlreadyDefined fc n)
|
||||
|
||||
ty <-
|
||||
|
||||
ty <-
|
||||
wrapError (InType fc n) $
|
||||
checkTerm idx InType (HolesOkay :: eopts) nest env
|
||||
(IBindHere fc (PI Rig0) ty_raw)
|
||||
checkTerm idx InType (HolesOkay :: eopts) nest env
|
||||
(IBindHere fc (PI Rig0) ty_raw)
|
||||
(gType fc)
|
||||
logTermNF 5 ("Type of " ++ show n) [] (abstractEnvType tfc env ty)
|
||||
-- TODO: Check name visibility
|
||||
@ -105,10 +105,10 @@ processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc n_in ty_raw)
|
||||
let fullty = abstractEnvType tfc env ty
|
||||
erased <- findErased fullty
|
||||
|
||||
addDef (Resolved idx)
|
||||
addDef (Resolved idx)
|
||||
(record { eraseArgs = erased }
|
||||
(newDef fc n rig vars fullty vis def))
|
||||
-- Flag it as checked, because we're going to check the clauses
|
||||
-- Flag it as checked, because we're going to check the clauses
|
||||
-- from the top level.
|
||||
-- But, if it's a case block, it'll be checked as part of the top
|
||||
-- level check so don't set the flag.
|
||||
|
623
src/TTImp/Reflect.idr
Normal file
623
src/TTImp/Reflect.idr
Normal file
@ -0,0 +1,623 @@
|
||||
module TTImp.Reflect
|
||||
|
||||
import Core.Context
|
||||
import Core.Env
|
||||
import Core.Normalise
|
||||
import Core.Reflect
|
||||
import Core.TT
|
||||
import Core.Value
|
||||
|
||||
import TTImp.TTImp
|
||||
|
||||
%default covering
|
||||
|
||||
export
|
||||
Reify BindMode where
|
||||
reify defs (NDCon _ (NS _ (UN "PI")) _ _ [c])
|
||||
= do c' <- reify defs !(evalClosure defs c)
|
||||
pure (PI c')
|
||||
reify defs (NDCon _ (NS _ (UN "PATTERN")) _ _ _)
|
||||
= pure PATTERN
|
||||
reify defs (NDCon _ (NS _ (UN "NONE")) _ _ _)
|
||||
= pure NONE
|
||||
reify deva val = cantReify val "BindMode"
|
||||
|
||||
export
|
||||
Reflect BindMode where
|
||||
reflect fc defs env (PI c)
|
||||
= do c' <- reflect fc defs env c
|
||||
appCon fc defs (reflection "PI") [c']
|
||||
reflect fc defs env PATTERN
|
||||
= getCon fc defs (reflection "PATTERN")
|
||||
reflect fc defs env NONE
|
||||
= getCon fc defs (reflection "NONE")
|
||||
|
||||
export
|
||||
Reify UseSide where
|
||||
reify defs (NDCon _ (NS _ (UN "UseLeft")) _ _ _)
|
||||
= pure UseLeft
|
||||
reify defs (NDCon _ (NS _ (UN "UseRight")) _ _ _)
|
||||
= pure UseRight
|
||||
reify deva val = cantReify val "UseSide"
|
||||
|
||||
export
|
||||
Reflect UseSide where
|
||||
reflect fc defs env UseLeft
|
||||
= getCon fc defs (reflection "UseLeft")
|
||||
reflect fc defs env UseRight
|
||||
= getCon fc defs (reflection "UseRight")
|
||||
|
||||
mutual
|
||||
export
|
||||
Reify RawImp where
|
||||
reify defs (NDCon _ (NS _ (UN "IVar")) _ _ [fc, n])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
n' <- reify defs !(evalClosure defs n)
|
||||
pure (IVar fc' n')
|
||||
reify defs (NDCon _ (NS _ (UN "IPi")) _ _ [fc, c, p, mn, aty, rty])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
c' <- reify defs !(evalClosure defs c)
|
||||
p' <- reify defs !(evalClosure defs p)
|
||||
mn' <- reify defs !(evalClosure defs mn)
|
||||
aty' <- reify defs !(evalClosure defs aty)
|
||||
rty' <- reify defs !(evalClosure defs rty)
|
||||
pure (IPi fc' c' p' mn' aty' rty')
|
||||
reify defs (NDCon _ (NS _ (UN "ILam")) _ _ [fc, c, p, mn, aty, lty])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
c' <- reify defs !(evalClosure defs c)
|
||||
p' <- reify defs !(evalClosure defs p)
|
||||
mn' <- reify defs !(evalClosure defs mn)
|
||||
aty' <- reify defs !(evalClosure defs aty)
|
||||
lty' <- reify defs !(evalClosure defs lty)
|
||||
pure (ILam fc' c' p' mn' aty' lty')
|
||||
reify defs (NDCon _ (NS _ (UN "ILet")) _ _ [fc, c, n, ty, val, sc])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
c' <- reify defs !(evalClosure defs c)
|
||||
n' <- reify defs !(evalClosure defs n)
|
||||
ty' <- reify defs !(evalClosure defs ty)
|
||||
val' <- reify defs !(evalClosure defs val)
|
||||
sc' <- reify defs !(evalClosure defs sc)
|
||||
pure (ILet fc' c' n' ty' val' sc')
|
||||
reify defs (NDCon _ (NS _ (UN "ICase")) _ _ [fc, sc, ty, cs])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
sc' <- reify defs !(evalClosure defs sc)
|
||||
ty' <- reify defs !(evalClosure defs ty)
|
||||
cs' <- reify defs !(evalClosure defs cs)
|
||||
pure (ICase fc' sc' ty' cs')
|
||||
reify defs (NDCon _ (NS _ (UN "ILocal")) _ _ [fc, ds, sc])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
ds' <- reify defs !(evalClosure defs ds)
|
||||
sc' <- reify defs !(evalClosure defs sc)
|
||||
pure (ILocal fc' ds' sc')
|
||||
reify defs (NDCon _ (NS _ (UN "IUpdate")) _ _ [fc, ds, sc])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
ds' <- reify defs !(evalClosure defs ds)
|
||||
sc' <- reify defs !(evalClosure defs sc)
|
||||
pure (IUpdate fc' ds' sc')
|
||||
reify defs (NDCon _ (NS _ (UN "IApp")) _ _ [fc, f, a])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
f' <- reify defs !(evalClosure defs f)
|
||||
a' <- reify defs !(evalClosure defs a)
|
||||
pure (IApp fc' f' a')
|
||||
reify defs (NDCon _ (NS _ (UN "IImplicitApp")) _ _ [fc, f, m, a])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
f' <- reify defs !(evalClosure defs f)
|
||||
m' <- reify defs !(evalClosure defs m)
|
||||
a' <- reify defs !(evalClosure defs a)
|
||||
pure (IImplicitApp fc' f' m' a')
|
||||
reify defs (NDCon _ (NS _ (UN "IWithApp")) _ _ [fc, f, a])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
f' <- reify defs !(evalClosure defs f)
|
||||
a' <- reify defs !(evalClosure defs a)
|
||||
pure (IWithApp fc' f' a')
|
||||
reify defs (NDCon _ (NS _ (UN "ISearch")) _ _ [fc, d])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
d' <- reify defs !(evalClosure defs d)
|
||||
pure (ISearch fc' d')
|
||||
reify defs (NDCon _ (NS _ (UN "IAlternative")) _ _ [fc, t, as])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
as' <- reify defs !(evalClosure defs as)
|
||||
pure (IAlternative fc' t' as')
|
||||
reify defs (NDCon _ (NS _ (UN "IRewrite")) _ _ [fc, t, sc])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
sc' <- reify defs !(evalClosure defs sc)
|
||||
pure (IRewrite fc' t' sc')
|
||||
reify defs (NDCon _ (NS _ (UN "IBindHere")) _ _ [fc, t, sc])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
sc' <- reify defs !(evalClosure defs sc)
|
||||
pure (IBindHere fc' t' sc')
|
||||
reify defs (NDCon _ (NS _ (UN "IBindVar")) _ _ [fc, n])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
n' <- reify defs !(evalClosure defs n)
|
||||
pure (IBindVar fc' n')
|
||||
reify defs (NDCon _ (NS _ (UN "IAs")) _ _ [fc, s, n, t])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
s' <- reify defs !(evalClosure defs s)
|
||||
n' <- reify defs !(evalClosure defs n)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
pure (IAs fc' s' n' t')
|
||||
reify defs (NDCon _ (NS _ (UN "IMustUnify")) _ _ [fc, r, t])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
r' <- reify defs !(evalClosure defs r)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
pure (IMustUnify fc' r' t')
|
||||
reify defs (NDCon _ (NS _ (UN "IDelayed")) _ _ [fc, r, t])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
r' <- reify defs !(evalClosure defs r)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
pure (IDelayed fc' r' t')
|
||||
reify defs (NDCon _ (NS _ (UN "IDelay")) _ _ [fc, t])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
pure (IDelay fc' t')
|
||||
reify defs (NDCon _ (NS _ (UN "IForce")) _ _ [fc, t])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
pure (IForce fc' t')
|
||||
reify defs (NDCon _ (NS _ (UN "IPrimVal")) _ _ [fc, t])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
t' <- reify defs !(evalClosure defs t)
|
||||
pure (IPrimVal fc' t')
|
||||
reify defs (NDCon _ (NS _ (UN "IType")) _ _ [fc])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
pure (IType fc')
|
||||
reify defs (NDCon _ (NS _ (UN "IHole")) _ _ [fc, n])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
n' <- reify defs !(evalClosure defs n)
|
||||
pure (IHole fc' n')
|
||||
reify defs (NDCon _ (NS _ (UN "Implicit")) _ _ [fc, n])
|
||||
= do fc' <- reify defs !(evalClosure defs fc)
|
||||
n' <- reify defs !(evalClosure defs n)
|
||||
pure (Implicit fc' n')
|
||||
reify defs val = cantReify val "TTImp"
|
||||
|
||||
export
|
||||
Reify IFieldUpdate where
|
||||
reify defs (NDCon _ (NS _ (UN "ISetField")) _ _ [x, y])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
pure (ISetField x' y')
|
||||
reify defs (NDCon _ (NS _ (UN "ISetFieldApp")) _ _ [x, y])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
pure (ISetFieldApp x' y')
|
||||
reify defs val = cantReify val "IFieldUpdate"
|
||||
|
||||
export
|
||||
Reify AltType where
|
||||
reify defs (NDCon _ (NS _ (UN "FirstSuccess")) _ _ _)
|
||||
= pure FirstSuccess
|
||||
reify defs (NDCon _ (NS _ (UN "Unique")) _ _ _)
|
||||
= pure Unique
|
||||
reify defs (NDCon _ (NS _ (UN "UniqueDefault")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (UniqueDefault x')
|
||||
reify defs val = cantReify val "AltType"
|
||||
|
||||
export
|
||||
Reify FnOpt where
|
||||
reify defs (NDCon _ (NS _ (UN "Inline")) _ _ _)
|
||||
= pure Inline
|
||||
reify defs (NDCon _ (NS _ (UN "Hint")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (Hint x')
|
||||
reify defs (NDCon _ (NS _ (UN "GlobalHint")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (GlobalHint x')
|
||||
reify defs (NDCon _ (NS _ (UN "ExternFn")) _ _ _)
|
||||
= pure ExternFn
|
||||
reify defs (NDCon _ (NS _ (UN "ForeignFn")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (ForeignFn x')
|
||||
reify defs (NDCon _ (NS _ (UN "Invertible")) _ _ _)
|
||||
= pure Invertible
|
||||
reify defs (NDCon _ (NS _ (UN "Total")) _ _ _)
|
||||
= pure Total
|
||||
reify defs (NDCon _ (NS _ (UN "Covering")) _ _ _)
|
||||
= pure Covering
|
||||
reify defs (NDCon _ (NS _ (UN "PartialOK")) _ _ _)
|
||||
= pure PartialOK
|
||||
reify defs val = cantReify val "FnOpt"
|
||||
|
||||
export
|
||||
Reify ImpTy where
|
||||
reify defs (NDCon _ (NS _ (UN "MkTy")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (MkImpTy x' y' z')
|
||||
reify defs val = cantReify val "ITy"
|
||||
|
||||
export
|
||||
Reify DataOpt where
|
||||
reify defs (NDCon _ (NS _ (UN "SearchBy")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (SearchBy x')
|
||||
reify defs (NDCon _ (NS _ (UN "NoHints")) _ _ _)
|
||||
= pure NoHints
|
||||
reify defs val = cantReify val "DataOpt"
|
||||
|
||||
export
|
||||
Reify ImpData where
|
||||
reify defs (NDCon _ (NS _ (UN "MkData")) _ _ [v,w,x,y,z])
|
||||
= do v' <- reify defs !(evalClosure defs v)
|
||||
w' <- reify defs !(evalClosure defs w)
|
||||
x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (MkImpData v' w' x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "MkLater")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (MkImpLater x' y' z')
|
||||
reify defs val = cantReify val "Data"
|
||||
|
||||
export
|
||||
Reify IField where
|
||||
reify defs (NDCon _ (NS _ (UN "MkIField")) _ _ [v,w,x,y,z])
|
||||
= do v' <- reify defs !(evalClosure defs v)
|
||||
w' <- reify defs !(evalClosure defs w)
|
||||
x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (MkIField v' w' x' y' z')
|
||||
reify defs val = cantReify val "IField"
|
||||
|
||||
export
|
||||
Reify ImpRecord where
|
||||
reify defs (NDCon _ (NS _ (UN "MkRecord")) _ _ [v,w,x,y,z])
|
||||
= do v' <- reify defs !(evalClosure defs v)
|
||||
w' <- reify defs !(evalClosure defs w)
|
||||
x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (MkImpRecord v' w' x' y' z')
|
||||
reify defs val = cantReify val "Record"
|
||||
|
||||
export
|
||||
Reify ImpClause where
|
||||
reify defs (NDCon _ (NS _ (UN "PatClause")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (PatClause x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "WithClause")) _ _ [w,x,y,z])
|
||||
= do w' <- reify defs !(evalClosure defs w)
|
||||
x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (WithClause w' x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "WithClause")) _ _ [x,y])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
pure (ImpossibleClause x' y')
|
||||
reify defs val = cantReify val "Clause"
|
||||
|
||||
export
|
||||
Reify ImpDecl where
|
||||
reify defs (NDCon _ (NS _ (UN "IClaim")) _ _ [v,w,x,y,z])
|
||||
= do v' <- reify defs !(evalClosure defs v)
|
||||
w' <- reify defs !(evalClosure defs w)
|
||||
x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (IClaim v' w' x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "IData")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (IData x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "IDef")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (IDef x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "IParameters")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (IParameters x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "IRecord")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (IRecord x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "INamespace")) _ _ [w,x,y,z])
|
||||
= do w' <- reify defs !(evalClosure defs w)
|
||||
x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (INamespace w' x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "ITransform")) _ _ [x,y,z])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
z' <- reify defs !(evalClosure defs z)
|
||||
pure (ITransform x' y' z')
|
||||
reify defs (NDCon _ (NS _ (UN "ILog")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (ILog x')
|
||||
reify defs val = cantReify val "Decl"
|
||||
|
||||
mutual
|
||||
export
|
||||
Reflect RawImp where
|
||||
reflect fc defs env (IVar tfc n)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
n' <- reflect fc defs env n
|
||||
appCon fc defs (reflection "IVar") [fc', n']
|
||||
reflect fc defs env (IPi tfc c p mn aty rty)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
c' <- reflect fc defs env c
|
||||
p' <- reflect fc defs env p
|
||||
mn' <- reflect fc defs env mn
|
||||
aty' <- reflect fc defs env aty
|
||||
rty' <- reflect fc defs env rty
|
||||
appCon fc defs (reflection "IPi") [fc', c', p', mn', aty', rty']
|
||||
reflect fc defs env (ILam tfc c p mn aty rty)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
c' <- reflect fc defs env c
|
||||
p' <- reflect fc defs env p
|
||||
mn' <- reflect fc defs env mn
|
||||
aty' <- reflect fc defs env aty
|
||||
rty' <- reflect fc defs env rty
|
||||
appCon fc defs (reflection "ILam") [fc', c', p', mn', aty', rty']
|
||||
reflect fc defs env (ILet tfc c n aty aval sc)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
c' <- reflect fc defs env c
|
||||
n' <- reflect fc defs env n
|
||||
aty' <- reflect fc defs env aty
|
||||
aval' <- reflect fc defs env aval
|
||||
sc' <- reflect fc defs env sc
|
||||
appCon fc defs (reflection "ILet") [fc', c', n', aty', aval', sc']
|
||||
reflect fc defs env (ICase tfc sc ty cs)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
sc' <- reflect fc defs env sc
|
||||
ty' <- reflect fc defs env ty
|
||||
cs' <- reflect fc defs env cs
|
||||
appCon fc defs (reflection "ICase") [fc', sc', ty', cs']
|
||||
reflect fc defs env (ILocal tfc ds sc)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
ds' <- reflect fc defs env ds
|
||||
sc' <- reflect fc defs env sc
|
||||
appCon fc defs (reflection "ILocal") [fc', ds', sc']
|
||||
reflect fc defs env (IUpdate tfc ds sc)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
ds' <- reflect fc defs env ds
|
||||
sc' <- reflect fc defs env sc
|
||||
appCon fc defs (reflection "IUpdate") [fc', ds', sc']
|
||||
reflect fc defs env (IApp tfc f a)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
f' <- reflect fc defs env f
|
||||
a' <- reflect fc defs env a
|
||||
appCon fc defs (reflection "IApp") [fc', f', a']
|
||||
reflect fc defs env (IImplicitApp tfc f m a)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
f' <- reflect fc defs env f
|
||||
m' <- reflect fc defs env m
|
||||
a' <- reflect fc defs env a
|
||||
appCon fc defs (reflection "IImplicitApp") [fc', f', m', a']
|
||||
reflect fc defs env (IWithApp tfc f a)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
f' <- reflect fc defs env f
|
||||
a' <- reflect fc defs env a
|
||||
appCon fc defs (reflection "IWithApp") [fc', f', a']
|
||||
reflect fc defs env (ISearch tfc d)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
d' <- reflect fc defs env d
|
||||
appCon fc defs (reflection "ISearch") [fc', d']
|
||||
reflect fc defs env (IAlternative tfc t as)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
t' <- reflect fc defs env t
|
||||
as' <- reflect fc defs env as
|
||||
appCon fc defs (reflection "IAlternative") [fc', t', as']
|
||||
reflect fc defs env (IRewrite tfc t sc)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
t' <- reflect fc defs env t
|
||||
sc' <- reflect fc defs env sc
|
||||
appCon fc defs (reflection "IRewrite") [fc', t', sc']
|
||||
reflect fc defs env (ICoerced tfc d) = reflect fc defs env d
|
||||
reflect fc defs env (IBindHere tfc n sc)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
n' <- reflect fc defs env n
|
||||
sc' <- reflect fc defs env sc
|
||||
appCon fc defs (reflection "IBindHere") [fc', n', sc']
|
||||
reflect fc defs env (IBindVar tfc n)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
n' <- reflect fc defs env n
|
||||
appCon fc defs (reflection "IBindVar") [fc', n']
|
||||
reflect fc defs env (IAs tfc s n t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
s' <- reflect fc defs env s
|
||||
n' <- reflect fc defs env n
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "IAs") [fc', s', n', t']
|
||||
reflect fc defs env (IMustUnify tfc r t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
r' <- reflect fc defs env r
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "IMustUnify") [fc', r', t']
|
||||
reflect fc defs env (IDelayed tfc r t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
r' <- reflect fc defs env r
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "IDelayed") [fc', r', t']
|
||||
reflect fc defs env (IDelay tfc t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "IDelay") [fc', t']
|
||||
reflect fc defs env (IForce tfc t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "IForce") [fc', t']
|
||||
reflect fc defs env (IPrimVal tfc t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "IPrimVal") [fc', t']
|
||||
reflect fc defs env (IType tfc)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
appCon fc defs (reflection "IType") [fc']
|
||||
reflect fc defs env (IHole tfc t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "IHole") [fc', t']
|
||||
reflect fc defs env (Implicit tfc t)
|
||||
= do fc' <- reflect fc defs env tfc
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "Implicit") [fc', t']
|
||||
|
||||
export
|
||||
Reflect IFieldUpdate where
|
||||
reflect fc defs env (ISetField p t)
|
||||
= do p' <- reflect fc defs env p
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "ISetField") [p', t']
|
||||
reflect fc defs env (ISetFieldApp p t)
|
||||
= do p' <- reflect fc defs env p
|
||||
t' <- reflect fc defs env t
|
||||
appCon fc defs (reflection "ISetFieldApp") [p', t']
|
||||
|
||||
export
|
||||
Reflect AltType where
|
||||
reflect fc defs env FirstSuccess = getCon fc defs (reflection "FirstSuccess")
|
||||
reflect fc defs env Unique = getCon fc defs (reflection "Unique")
|
||||
reflect fc defs env (UniqueDefault x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "UniqueDefault") [x']
|
||||
|
||||
export
|
||||
Reflect FnOpt where
|
||||
reflect fc defs env Inline = getCon fc defs (reflection "FirstSuccess")
|
||||
reflect fc defs env (Hint x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "Hint") [x']
|
||||
reflect fc defs env (GlobalHint x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "GlobalHint") [x']
|
||||
reflect fc defs env ExternFn = getCon fc defs (reflection "ExternFn")
|
||||
reflect fc defs env (ForeignFn x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "ForeignFn") [x']
|
||||
reflect fc defs env Invertible = getCon fc defs (reflection "Invertible")
|
||||
reflect fc defs env Total = getCon fc defs (reflection "Total")
|
||||
reflect fc defs env Covering = getCon fc defs (reflection "Covering")
|
||||
reflect fc defs env PartialOK = getCon fc defs (reflection "PartialOK")
|
||||
|
||||
export
|
||||
Reflect ImpTy where
|
||||
reflect fc defs env (MkImpTy x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "MkTy") [x', y', z']
|
||||
|
||||
export
|
||||
Reflect DataOpt where
|
||||
reflect fc defs env (SearchBy x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "SearchBy") [x']
|
||||
reflect fc defs env NoHints = getCon fc defs (reflection "NoHints")
|
||||
|
||||
export
|
||||
Reflect ImpData where
|
||||
reflect fc defs env (MkImpData v w x y z)
|
||||
= do v' <- reflect fc defs env v
|
||||
w' <- reflect fc defs env w
|
||||
x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "MkData") [v', w', x', y', z']
|
||||
reflect fc defs env (MkImpLater x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "MkLater") [x', y', z']
|
||||
|
||||
export
|
||||
Reflect IField where
|
||||
reflect fc defs env (MkIField v w x y z)
|
||||
= do v' <- reflect fc defs env v
|
||||
w' <- reflect fc defs env w
|
||||
x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "MkIField") [v', w', x', y', z']
|
||||
|
||||
export
|
||||
Reflect ImpRecord where
|
||||
reflect fc defs env (MkImpRecord v w x y z)
|
||||
= do v' <- reflect fc defs env v
|
||||
w' <- reflect fc defs env w
|
||||
x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "MkRecord") [v', w', x', y', z']
|
||||
|
||||
export
|
||||
Reflect ImpClause where
|
||||
reflect fc defs env (PatClause x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "PatClause") [x', y', z']
|
||||
reflect fc defs env (WithClause w x y z)
|
||||
= do w' <- reflect fc defs env x
|
||||
x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "WithClause") [w', x', y', z']
|
||||
reflect fc defs env (ImpossibleClause x y)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
appCon fc defs (reflection "ImpossibleClause") [x', y']
|
||||
|
||||
export
|
||||
Reflect ImpDecl where
|
||||
reflect fc defs env (IClaim v w x y z)
|
||||
= do v' <- reflect fc defs env v
|
||||
w' <- reflect fc defs env w
|
||||
x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "IClaim") [v', w', x', y', z']
|
||||
reflect fc defs env (IData x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "IData") [x', y', z']
|
||||
reflect fc defs env (IDef x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "IDef") [x', y', z']
|
||||
reflect fc defs env (IParameters x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "IParameters") [x', y', z']
|
||||
reflect fc defs env (IRecord x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "IRecord") [x', y', z']
|
||||
reflect fc defs env (INamespace w x y z)
|
||||
= do w' <- reflect fc defs env w
|
||||
x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "INamespace") [w', x', y', z']
|
||||
reflect fc defs env (ITransform x y z)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
z' <- reflect fc defs env z
|
||||
appCon fc defs (reflection "ITransform") [x', y', z']
|
||||
reflect fc defs env (IPragma x)
|
||||
= throw (GenericMsg fc "Can't reflect a pragma")
|
||||
reflect fc defs env (ILog x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "ILog") [x']
|
||||
|
||||
|
Some files were not shown because too many files have changed in this diff Show More
Loading…
Reference in New Issue
Block a user