[ fix #1959 ] use modern record update syntax (#2196)

This commit is contained in:
teggot 2021-12-16 20:23:18 +02:00 committed by GitHub
parent 516cb4a690
commit d3aed0404c
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
93 changed files with 693 additions and 644 deletions

View File

@ -539,7 +539,7 @@ Language changes:
be at least `covering`
+ That is, `%default covering` is the default status.
* Fields of records can be accessed (and updated) using the dot syntax,
such as `r.field1.field2` or `record { field1.field2 = 42 }`.
such as `r.field1.field2` or `{ field1.field2 := 42 }`.
For details, see [the "records" entry in the user manual](https://idris2.readthedocs.io/en/latest/reference/records.html)
* New function flag `%tcinline` which means that the function should be
inlined for the purposes of totality checking (but otherwise not inlined).

View File

@ -14,6 +14,7 @@ Arnaud Bailly
Brian Wignall
Bryn Keller
Christian Rasmussen
Danylo Lapirov
David Smith
Denis Buzdalov
Edwin Brady

View File

@ -176,14 +176,14 @@ Finally, the examples:
-- record update syntax uses dots now
-- prints 3.0
printLn $ (record { topLeft.x = 3 } rect).topLeft.x
printLn $ ({ topLeft.x := 3 } rect).topLeft.x
-- but for compatibility, we support the old syntax, too
printLn $ (record { topLeft->x = 3 } rect).topLeft.x
printLn $ ({ topLeft->x := 3 } rect).topLeft.x
-- prints 2.1
printLn $ (record { topLeft.x $= (+1) } rect).topLeft.x
printLn $ (record { topLeft->x $= (+1) } rect).topLeft.x
printLn $ ({ topLeft.x $= (+1) } rect).topLeft.x
printLn $ ({ topLeft->x $= (+1) } rect).topLeft.x
Parses but does not typecheck:

View File

@ -1074,13 +1074,13 @@ updated):
.. code-block:: bash
*Record> record { firstName = "Jim" } fred
*Record> { firstName := "Jim" } fred
MkPerson "Jim" "Joe" "Bloggs" 30 : Person
*Record> record { firstName = "Jim", age $= (+ 1) } fred
*Record> { firstName := "Jim", age $= (+ 1) } fred
MkPerson "Jim" "Joe" "Bloggs" 31 : Person
The syntax ``record { field = val, ... }`` generates a function which
updates the given fields in a record. ``=`` assigns a new value to a field,
The syntax ``{ field := val, ... }`` generates a function which
updates the given fields in a record. ``:=`` assigns a new value to a field,
and ``$=`` applies a function to update its value.
Each record is defined in its own namespace, which means that field names
@ -1103,7 +1103,7 @@ length because it will not affect the type of the record:
.. code-block:: idris
addStudent : Person -> Class -> Class
addStudent p c = record { students = p :: students c } c
addStudent p c = { students := p :: students c } c
::
@ -1115,7 +1115,7 @@ We could also use ``$=`` to define ``addStudent`` more concisely:
.. code-block:: idris
addStudent' : Person -> Class -> Class
addStudent' p c = record { students $= (p ::) } c
addStudent' p c = { students $= (p ::) } c
Nested record projection
~~~~~~~~~~~~~~~~~~~~~~~~
@ -1150,11 +1150,11 @@ syntax:
.. code-block:: idris
record { a.b.c = val } x
{ a.b.c := val } x
This returns a new record, with the field accessed by the path
``a.b.c`` set to ``val``. The syntax is first class, i.e. ``record {
a.b.c = val }`` itself has a function type.
``a.b.c`` set to ``val``. The syntax is first class, i.e. ``{
a.b.c := val }`` itself has a function type.
The ``$=`` notation is also valid for nested record updates.
@ -1191,7 +1191,7 @@ is added:
.. code-block:: idris
addStudent : Person -> SizedClass n -> SizedClass (S n)
addStudent p c = record { students = p :: students c } c
addStudent p c = { students := p :: students c } c
In fact, the dependent pair type we have just seen is, in practice, defined
as a record, with fields ``fst`` and ``snd`` which allow projecting values
@ -1211,8 +1211,8 @@ that all related fields are updated at once. For example:
cons : t -> (x : Nat ** Vect x t) -> (x : Nat ** Vect x t)
cons val xs
= record { fst = S (fst xs),
snd = (val :: snd xs) } xs
= { fst := S (fst xs),
snd := (val :: snd xs) } xs
Or even:
@ -1220,8 +1220,8 @@ Or even:
cons' : t -> (x : Nat ** Vect x t) -> (x : Nat ** Vect x t)
cons' val
= record { fst $= S,
snd $= (val ::) }
= { fst $= S,
snd $= (val ::) }
.. _sect-more-expr:

View File

@ -801,8 +801,8 @@ correspondingly:
.. code-block:: idris
addEntry : String -> WrapVect String -> WrapVect String
addEntry val = record { length $= S,
content $= (val :: ) }
addEntry val = { length $= S,
content $= (val :: ) }
Another novelty - new update syntax (previous one still functional):

View File

@ -130,7 +130,7 @@ covering
notListTerminating : (p : Subset Nat (HasLength xs)) -> P xs
notListTerminating p = case view p of
Z => PNil
S p => PCon (notListTerminating {xs = map id (tail xs)} (record { snd $= map id } p))
S p => PCon (notListTerminating {xs = map id (tail xs)} ({ snd $= map id } p))
natTerminating : (n : Nat) -> (0 p : HasLength xs n) -> P xs
natTerminating n p = case view n p of

View File

@ -65,7 +65,7 @@ record OptDescr a where
export
Functor OptDescr where
map f = record { argDescr $= map f }
map f = { argDescr $= map f }
-- kind of cmd line arg (internal use only):
data OptKind a
@ -160,7 +160,7 @@ emptyRes = MkResult [] [] [] []
export
Functor Result where
map f = record { options $= map f }
map f = { options $= map f }
OptFun : Type -> Type
OptFun a = List String -> List (OptDescr a) -> (OptKind a,List String)

View File

@ -292,9 +292,9 @@ append' left right =
if isAbsolute' right || isJust right.volume then
right
else if hasRoot right then
record { volume = left.volume } right
{ volume := left.volume } right
else
record { body = left.body ++ right.body, hasTrailSep = right.hasTrailSep } left
{ body := left.body ++ right.body, hasTrailSep := right.hasTrailSep } left
splitPath' : Path -> List Path
splitPath' path =
@ -319,7 +319,7 @@ splitParent' path =
[] => Nothing
(x::xs) =>
let
parent = record { body = init (x::xs), hasTrailSep = False } path
parent = { body := init (x::xs), hasTrailSep := False } path
child = MkPath Nothing False [last (x::xs)] path.hasTrailSep
in
Just (parent, child)

View File

@ -27,4 +27,4 @@ reflow = fillSep . words
||| Renders a document with a certain width.
export
putDocW : Nat -> Doc ann -> IO ()
putDocW w = renderIO . layoutPretty (record { layoutPageWidth = AvailablePerLine (cast w) 1 } defaultLayoutOptions)
putDocW w = renderIO . layoutPretty ({ layoutPageWidth := AvailablePerLine (cast w) 1 } defaultLayoutOptions)

View File

@ -155,17 +155,17 @@ options args = case args of
go : List String -> Maybe String -> Options -> Maybe (Maybe String, Options)
go rest only opts = case rest of
[] => pure (only, opts)
("--timing" :: xs) => go xs only (record { timing = True} opts)
("--interactive" :: xs) => go xs only (record { interactive = True } opts)
("--color" :: xs) => go xs only (record { color = True } opts)
("--colour" :: xs) => go xs only (record { color = True } opts)
("--no-color" :: xs) => go xs only (record { color = False } opts)
("--no-colour" :: xs) => go xs only (record { color = False } opts)
("--cg" :: cg :: xs) => go xs only (record { codegen = Just cg } opts)
("--timing" :: xs) => go xs only ({ timing := True} opts)
("--interactive" :: xs) => go xs only ({ interactive := True } opts)
("--color" :: xs) => go xs only ({ color := True } opts)
("--colour" :: xs) => go xs only ({ color := True } opts)
("--no-color" :: xs) => go xs only ({ color := False } opts)
("--no-colour" :: xs) => go xs only ({ color := False } opts)
("--cg" :: cg :: xs) => go xs only ({ codegen := Just cg } opts)
("--threads" :: n :: xs) => do let pos : Nat = !(parsePositive n)
go xs only (record { threads = pos } opts)
("--failure-file" :: p :: xs) => go xs only (record { failureFile = Just p } opts)
("--only" :: xs) => pure (only, record { onlyNames = xs } opts)
go xs only ({ threads := pos } opts)
("--failure-file" :: p :: xs) => go xs only ({ failureFile := Just p } opts)
("--only" :: xs) => pure (only, { onlyNames := xs } opts)
("--only-file" :: p :: xs) => go xs (Just p) opts
_ => Nothing
@ -178,7 +178,7 @@ options args = case args of
| Nothing => pure (Just opts)
Right only <- readFile fp
| Left err => fail (show err)
pure $ Just $ record { onlyNames $= ((lines only) ++) } opts
pure $ Just $ { onlyNames $= ((lines only) ++) } opts
||| Normalise strings between different OS.
|||
@ -501,7 +501,7 @@ runner tests
putStrLn usage
-- if no CG has been set, find a sensible default based on what is available
opts <- case codegen opts of
Nothing => pure $ record { codegen = !findCG } opts
Nothing => pure $ { codegen := !findCG } opts
Just _ => pure opts
-- run the tests
res <- concat <$> traverse (poolRunner opts) tests

View File

@ -23,13 +23,13 @@ filter p (x :: xs)
cons : t -> (x : Nat ** Vect x t) -> (x : Nat ** Vect x t)
cons val xs
= record { fst = S (fst xs),
snd = (val :: snd xs) } xs
= { fst := S (fst xs),
snd := (val :: snd xs) } xs
cons' : t -> (x : Nat ** Vect x t) -> (x : Nat ** Vect x t)
cons' val
= record { fst $= S,
snd $= (val ::) }
= { fst $= S,
snd $= (val ::) }
Show a => Show (Vect n a) where
show xs = "[" ++ show' xs ++ "]" where

View File

@ -16,10 +16,10 @@ record Class where
className : String
addStudent : Person -> Class -> Class
addStudent p c = record { students = p :: students c } c
addStudent p c = { students := p :: students c } c
addStudent' : Person -> SizedClass n -> SizedClass (S n)
addStudent' p c = SizedClassInfo (p :: students c) (className c)
addStudent'' : Person -> SizedClass n -> SizedClass (S n)
addStudent'' p c = record { students = p :: students c } c
addStudent'' p c = { students := p :: students c } c

View File

@ -116,14 +116,14 @@ record ESSt where
||| Map a local name to the given minimal expression
export
addLocal : { auto c : Ref ESs ESSt } -> Name -> Minimal -> Core ()
addLocal n v = update ESs $ record { locals $= insert n v }
addLocal n v = update ESs $ { locals $= insert n v }
||| Get and bump the local var index
export
nextLocal : { auto c : Ref ESs ESSt } -> Core Var
nextLocal = do
st <- get ESs
put ESs $ record { loc $= (+1) } st
put ESs $ { loc $= (+1) } st
pure $ VLoc st.loc
||| Register a `Name` as a local variable. The name is kept
@ -168,14 +168,14 @@ projections sc xs =
||| Map a toplevel function name to the given `Var`
export
addRef : { auto c : Ref ESs ESSt } -> Name -> Var -> Core ()
addRef n v = update ESs $ record { refs $= insert n v }
addRef n v = update ESs $ { refs $= insert n v }
||| Get and bump the local ref index
export
nextRef : { auto c : Ref ESs ESSt } -> Core Var
nextRef = do
st <- get ESs
put ESs $ record { ref $= (+1) } st
put ESs $ { ref $= (+1) } st
pure $ VRef st.ref
registerRef : {auto c : Ref ESs ESSt}
@ -216,7 +216,7 @@ addToPreamble : {auto c : Ref ESs ESSt}
addToPreamble name def = do
s <- get ESs
case lookup name (preamble s) of
Nothing => put ESs $ record { preamble $= insert name def } s
Nothing => put ESs $ { preamble $= insert name def } s
Just x =>
unless (x == def) $ do
errorConcat
@ -243,4 +243,4 @@ init mode isArg isFun ccs noMangle =
||| function.
export
reset : {auto c : Ref ESs ESSt} -> Core ()
reset = update ESs $ record { loc = 0, locals = empty }
reset = update ESs $ { loc := 0, locals := empty }

View File

@ -512,7 +512,7 @@ updateCallGraph n
Just def <- lookupCtxtExact n (gamma defs) | Nothing => pure ()
let Just cexpr = compexpr def | Nothing => pure ()
let refs = getRefs cexpr
ignore $ addDef n (record { refersToRuntimeM = Just refs } def)
ignore $ addDef n ({ refersToRuntimeM := Just refs } def)
export
fixArityDef : {auto c : Ref Ctxt Defs} ->

View File

@ -88,7 +88,7 @@ getReg stk RVal = do
getReg stk Discard = pure Null
setReg : Ref State InterpState => Stack -> Reg -> Object -> Core ()
setReg stk RVal obj = update State $ record { returnObj = Just obj }
setReg stk RVal obj = update State $ { returnObj := Just obj }
setReg stk (Loc i) obj = do
ls <- locals <$> get State
when (i >= max ls) $ interpError stk $ "Attempt to set register: " ++ show i ++ ", size of locals: " ++ show (max ls)
@ -162,7 +162,7 @@ beginFunction args (DECLARE _ :: is) maxLoc = beginFunction args is maxLoc
beginFunction args (START :: is) maxLoc = do
locals <- coreLift $ newArray (maxLoc + 1)
ignore $ traverse (\(idx, arg) => coreLift $ writeArray locals idx arg) args
update State $ record { locals = locals, returnObj = Nothing }
update State $ { locals := locals, returnObj := Nothing }
pure is
beginFunction args is maxLoc = pure is

View File

@ -146,7 +146,7 @@ genName : {auto l : Ref Lifts LDefs} ->
genName
= do ldefs <- get Lifts
let i = nextName ldefs
put Lifts (record { nextName = i + 1 } ldefs)
put Lifts ({ nextName := i + 1 } ldefs)
pure $ mkName (basename ldefs) i
where
mkName : Name -> Int -> Name
@ -235,7 +235,7 @@ mutual
scl' = dropUnused {outer=bound} unused scl
n <- genName
ldefs <- get Lifts
put Lifts (record { defs $= ((n, MkLFun (dropped vars unused) bound scl') ::) } ldefs)
put Lifts ({ defs $= ((n, MkLFun (dropped vars unused) bound scl') ::) } ldefs)
pure $ LUnderApp fc n (length bound) (allVars fc vars unused)
where
allPrfs : (vs : List Name) -> (unused : Vect (length vs) Bool) -> List (Var vs)

View File

@ -355,37 +355,37 @@ addAutoHint (hintn_in, d)
= do defs <- get Ctxt
hintn <- toResolvedNames hintn_in
put Ctxt (record { autoHints $= insert hintn d } defs)
put Ctxt ({ autoHints $= insert hintn d } defs)
export
updatePair : {auto c : Ref Ctxt Defs} ->
Maybe PairNames -> Core ()
updatePair p
= do defs <- get Ctxt
put Ctxt (record { options->pairnames $= (p <+>) } defs)
put Ctxt ({ options->pairnames $= (p <+>) } defs)
export
updateRewrite : {auto c : Ref Ctxt Defs} ->
Maybe RewriteNames -> Core ()
updateRewrite r
= do defs <- get Ctxt
put Ctxt (record { options->rewritenames $= (r <+>) } defs)
put Ctxt ({ options->rewritenames $= (r <+>) } defs)
export
updatePrimNames : PrimNames -> PrimNames -> PrimNames
updatePrimNames p
= record { fromIntegerName $= ((fromIntegerName p) <+>),
fromStringName $= ((fromStringName p) <+>),
fromCharName $= ((fromCharName p) <+>),
fromDoubleName $= ((fromDoubleName p) <+>)
}
= { fromIntegerName $= ((fromIntegerName p) <+>),
fromStringName $= ((fromStringName p) <+>),
fromCharName $= ((fromCharName p) <+>),
fromDoubleName $= ((fromDoubleName p) <+>)
}
export
updatePrims : {auto c : Ref Ctxt Defs} ->
PrimNames -> Core ()
updatePrims p
= do defs <- get Ctxt
put Ctxt (record { options->primnames $= updatePrimNames p } defs)
put Ctxt ({ options->primnames $= updatePrimNames p } defs)
export
updateNameDirectives : {auto c : Ref Ctxt Defs} ->
@ -393,7 +393,7 @@ updateNameDirectives : {auto c : Ref Ctxt Defs} ->
updateNameDirectives [] = pure ()
updateNameDirectives ((t, ns) :: nds)
= do defs <- get Ctxt
put Ctxt (record { namedirectives $= insert t ns } defs)
put Ctxt ({ namedirectives $= insert t ns } defs)
updateNameDirectives nds
export
@ -402,7 +402,7 @@ updateCGDirectives : {auto c : Ref Ctxt Defs} ->
updateCGDirectives cgs
= do defs <- get Ctxt
let cgs' = nub (cgs ++ cgdirectives defs)
put Ctxt (record { cgdirectives = cgs' } defs)
put Ctxt ({ cgdirectives := cgs' } defs)
export
updateTransforms : {auto c : Ref Ctxt Defs} ->
@ -417,9 +417,9 @@ updateTransforms ((n, t) :: ts)
= do defs <- get Ctxt
case lookup n (transforms defs) of
Nothing =>
put Ctxt (record { transforms $= insert n [t] } defs)
put Ctxt ({ transforms $= insert n [t] } defs)
Just ts =>
put Ctxt (record { transforms $= insert n (t :: ts) } defs)
put Ctxt ({ transforms $= insert n (t :: ts) } defs)
getNSas : (String, (ModuleIdent, Bool, Namespace)) ->
@ -450,7 +450,7 @@ readFromTTC nestedns loc reexp fname modNS importAs
-- this time, because we need to reexport the dependencies.)
let False = (modNS, reexp, importAs) `elem` map snd (allImported defs)
| True => pure Nothing
put Ctxt (record { allImported $= ((fname, (modNS, reexp, importAs)) :: ) } defs)
put Ctxt ({ allImported $= ((fname, (modNS, reexp, importAs)) :: ) } defs)
Right buffer <- coreLift $ readFromFile fname
| Left err => throw (InternalError (fname ++ ": " ++ show err))
@ -495,7 +495,7 @@ readFromTTC nestedns loc reexp fname modNS importAs
-- Finally, update the unification state with the holes from the
-- ttc
ust <- get UST
put UST (record { nextName = nextVar ttc } ust)
put UST ({ nextName := nextVar ttc } ust)
pure (Just (ex, ifaceHash ttc, imported ttc))
where
alreadyDone : ModuleIdent -> Namespace ->

View File

@ -114,7 +114,7 @@ updatePats {todo = pvar :: ns} env (NBind fc _ (Pi _ c _ farg) fsc) (p :: ps)
Unknown =>
do defs <- get Ctxt
empty <- clearDefs defs
pure (record { argType = Known c !(quote empty env farg) } p
pure ({ argType := Known c !(quote empty env farg) } p
:: !(updatePats env !(fsc defs (toClosure defaultOpts env (Ref fc Bound pvar))) ps))
_ => pure (p :: ps)
updatePats env nf (p :: ps)
@ -122,7 +122,7 @@ updatePats env nf (p :: ps)
Unknown =>
do defs <- get Ctxt
empty <- clearDefs defs
pure (record { argType = Stuck !(quote empty env nf) } p :: ps)
pure ({ argType := Stuck !(quote empty env nf) } p :: ps)
_ => pure (p :: ps)
substInPatInfo : {pvar, vars, todo : _} ->
@ -137,7 +137,7 @@ substInPatInfo {pvar} {vars} fc n tm p ps
tynf <- nf defs (mkEnv fc _) ty
case tynf of
NApp _ _ _ =>
pure (record { argType = Known c (substName n tm ty) } p, ps)
pure ({ argType := Known c (substName n tm ty) } p, ps)
-- Got a concrete type, and that's all we need, so stop
_ => pure (p, ps)
Stuck fty =>
@ -146,7 +146,7 @@ substInPatInfo {pvar} {vars} fc n tm p ps
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,
pure ({ argType := Known c !(quote empty env farg) } p,
!(updatePats env
!(fsc defs (toClosure defaultOpts env
(Ref pfc Bound pvar))) ps))
@ -472,7 +472,7 @@ newPats : (pargs : List Pat) -> LengthMatch pargs ns ->
NamedPats vars ns
newPats [] NilMatch rest = []
newPats (newpat :: xs) (ConsMatch w) (pi :: rest)
= record { pat = newpat} pi :: newPats xs w rest
= { pat := newpat} pi :: newPats xs w rest
updateNames : List (Name, Pat) -> List (Name, Name)
updateNames = mapMaybe update
@ -484,7 +484,7 @@ updateNames = mapMaybe update
updatePatNames : List (Name, Name) -> NamedPats vars todo -> NamedPats vars todo
updatePatNames _ [] = []
updatePatNames ns (pi :: ps)
= record { pat $= update } pi :: updatePatNames ns ps
= { pat $= update } pi :: updatePatNames ns ps
where
update : Pat -> Pat
update (PAs fc n p)

View File

@ -104,10 +104,10 @@ newEntry n ctxt
when (idx >= max arr) $
do arr' <- coreLift $ newArrayCopy (max arr + Grow) arr
put Arr arr'
pure (idx, record { nextEntry = idx + 1,
resolvedAs $= insert n idx,
possibles $= addPossible n idx
} ctxt)
pure (idx, { nextEntry := idx + 1,
resolvedAs $= insert n idx,
possibles $= addPossible n idx
} ctxt)
-- Get the position of the next entry in the context array, growing the
-- array if it's out of bounds.
@ -124,7 +124,7 @@ getPosition n ctxt
newAlias : Name -> Name -> Context -> Core Context
newAlias alias full ctxt
= do (idx, ctxt) <- getPosition full ctxt
pure $ record { possibles $= addAlias alias full idx } ctxt
pure $ { possibles $= addAlias alias full idx } ctxt
export
getNameID : Name -> Context -> Maybe Int
@ -144,7 +144,7 @@ addCtxt n val ctxt_in
coreLift $ writeArray arr idx (Decoded val)
pure (idx, ctxt)
else do (idx, ctxt) <- getPosition n ctxt_in
pure (idx, record { staging $= insert idx (Decoded val) } ctxt)
pure (idx, { staging $= insert idx (Decoded val) } ctxt)
export
addEntry : Name -> ContextEntry -> Context -> Core (Int, Context)
@ -156,7 +156,7 @@ addEntry n entry ctxt_in
coreLift $ writeArray arr idx entry
pure (idx, ctxt)
else do (idx, ctxt) <- getPosition n ctxt_in
pure (idx, record { staging $= insert idx entry } ctxt)
pure (idx, { staging $= insert idx entry } ctxt)
returnDef : Bool -> Int -> GlobalDef -> Maybe (Int, GlobalDef)
returnDef False idx def = Just (idx, def)
@ -282,13 +282,13 @@ lookupHiddenCtxtName : Name -> Context -> Core (List (Name, Int, GlobalDef))
lookupHiddenCtxtName = lookupCtxtName' True
hideName : Name -> Context -> Context
hideName n ctxt = record { hidden $= insert n () } ctxt
hideName n ctxt = { hidden $= insert n () } ctxt
unhideName : Name -> Context -> Context
unhideName n ctxt = record { hidden $= delete n } ctxt
unhideName n ctxt = { hidden $= delete n } ctxt
branchCtxt : Context -> Core Context
branchCtxt ctxt = pure (record { branchDepth $= S } ctxt)
branchCtxt ctxt = pure ({ branchDepth $= S } ctxt)
commitCtxt : Context -> Core Context
commitCtxt ctxt
@ -298,9 +298,9 @@ commitCtxt ctxt
do let a = content ctxt
arr <- get Arr
coreLift $ commitStaged (toList (staging ctxt)) arr
pure (record { staging = empty,
branchDepth = Z } ctxt)
S k => pure (record { branchDepth = k } ctxt)
pure ({ staging := empty,
branchDepth := Z } ctxt)
S k => pure ({ branchDepth := k } ctxt)
where
-- We know the array must be big enough, because it will have been resized
-- if necessary in the branch to fit the index we've been given here
@ -616,8 +616,8 @@ StripNamespace Def where
export
StripNamespace GlobalDef where
trimNS ns def = record { definition $= trimNS ns } def
restoreNS ns def = record { definition $= restoreNS ns } def
trimNS ns def = { definition $= trimNS ns } def
restoreNS ns def = { definition $= restoreNS ns } def
HasNames (NameMap a) where
full gam nmap
@ -675,8 +675,8 @@ HasNames Totality where
export
HasNames SCCall where
full gam sc = pure $ record { fnCall = !(full gam (fnCall sc)) } sc
resolved gam sc = pure $ record { fnCall = !(resolved gam (fnCall sc)) } sc
full gam sc = pure $ { fnCall := !(full gam (fnCall sc)) } sc
resolved gam sc = pure $ { fnCall := !(resolved gam (fnCall sc)) } sc
export
HasNames a => HasNames (Maybe a) where
@ -693,21 +693,21 @@ HasNames GlobalDef where
-- d <- full gam (definition def)
-- coreLift $ printLn (fullname def, ts)
-- coreLift $ printLn (fullname def, d)
pure $ record { type = !(full gam (type def)),
definition = !(full gam (definition def)),
totality = !(full gam (totality def)),
refersToM = !(full gam (refersToM def)),
refersToRuntimeM = !(full gam (refersToRuntimeM def)),
sizeChange = !(traverse (full gam) (sizeChange def))
} def
pure $ { type := !(full gam (type def)),
definition := !(full gam (definition def)),
totality := !(full gam (totality def)),
refersToM := !(full gam (refersToM def)),
refersToRuntimeM := !(full gam (refersToRuntimeM def)),
sizeChange := !(traverse (full gam) (sizeChange def))
} def
resolved gam def
= pure $ record { type = !(resolved gam (type def)),
definition = !(resolved gam (definition def)),
totality = !(resolved gam (totality def)),
refersToM = !(resolved gam (refersToM def)),
refersToRuntimeM = !(resolved gam (refersToRuntimeM def)),
sizeChange = !(traverse (resolved gam) (sizeChange def))
} def
= pure $ { type := !(resolved gam (type def)),
definition := !(resolved gam (definition def)),
totality := !(resolved gam (totality def)),
refersToM := !(resolved gam (refersToM def)),
refersToRuntimeM := !(resolved gam (refersToRuntimeM def)),
sizeChange := !(traverse (resolved gam) (sizeChange def))
} def
export
HasNames Transform where
@ -806,7 +806,7 @@ data Ctxt : Type where
export
clearDefs : Defs -> Core Defs
clearDefs defs
= pure (record { gamma->inlineOnly = True } defs)
= pure ({ gamma->inlineOnly := True } defs)
export
initDefs : Core Defs
@ -853,13 +853,13 @@ clearCtxt : {auto c : Ref Ctxt Defs} ->
Core ()
clearCtxt
= do defs <- get Ctxt
put Ctxt (record { options = resetElab (options defs),
timings = timings defs } !initDefs)
put Ctxt ({ options := resetElab (options defs),
timings := timings defs } !initDefs)
where
resetElab : Options -> Options
resetElab opts =
let tot = totalReq (session opts) in
record { elabDirectives = record { totality = tot } defaultElab } opts
{ elabDirectives := { totality := tot } defaultElab } opts
export
getFieldNames : Context -> Namespace -> List Name
@ -1000,14 +1000,14 @@ addHash : {auto c : Ref Ctxt Defs} ->
Hashable a => a -> Core ()
addHash x
= do defs <- get Ctxt
put Ctxt (record { ifaceHash = hashWithSalt (ifaceHash defs) x } defs)
put Ctxt ({ ifaceHash := hashWithSalt (ifaceHash defs) x } defs)
export
initHash : {auto c : Ref Ctxt Defs} ->
Core ()
initHash
= do defs <- get Ctxt
put Ctxt (record { ifaceHash = 5381 } defs)
put Ctxt ({ ifaceHash := 5381 } defs)
export
addUserHole : {auto c : Ref Ctxt Defs} ->
@ -1016,14 +1016,14 @@ addUserHole : {auto c : Ref Ctxt Defs} ->
Core ()
addUserHole ext n
= do defs <- get Ctxt
put Ctxt (record { userHoles $= insert n ext } defs)
put Ctxt ({ userHoles $= insert n ext } defs)
export
clearUserHole : {auto c : Ref Ctxt Defs} ->
Name -> Core ()
clearUserHole n
= do defs <- get Ctxt
put Ctxt (record { userHoles $= delete n } defs)
put Ctxt ({ userHoles $= delete n } defs)
export
getUserHoles : {auto c : Ref Ctxt Defs} ->
@ -1052,7 +1052,7 @@ addDef : {auto c : Ref Ctxt Defs} ->
addDef n def
= do defs <- get Ctxt
(idx, gam') <- addCtxt n def (gamma defs)
put Ctxt (record { gamma = gam' } defs)
put Ctxt ({ gamma := gam' } defs)
case definition def of
None => pure ()
Hole _ _ => pure ()
@ -1065,7 +1065,7 @@ addContextEntry : {auto c : Ref Ctxt Defs} ->
addContextEntry ns n def
= do defs <- get Ctxt
(idx, gam') <- addEntry n (Coded ns def) (gamma defs)
put Ctxt (record { gamma = gam' } defs)
put Ctxt ({ gamma := gam' } defs)
pure idx
export
@ -1076,7 +1076,7 @@ addContextAlias alias full
Nothing <- lookupCtxtExact alias (gamma defs)
| _ => pure () -- Don't add the alias if the name exists already
gam' <- newAlias alias full (gamma defs)
put Ctxt (record { gamma = gam' } defs)
put Ctxt ({ gamma := gam' } defs)
export
addBuiltin : {arity : _} ->
@ -1119,8 +1119,8 @@ updateDef n fdef
| Nothing => pure ()
case fdef (definition gdef) of
Nothing => pure ()
Just def' => ignore $ addDef n (record { definition = def',
schemeExpr = Nothing } gdef)
Just def' => ignore $ addDef n ({ definition := def',
schemeExpr := Nothing } gdef)
export
updateTy : {auto c : Ref Ctxt Defs} ->
@ -1129,7 +1129,7 @@ updateTy i ty
= do defs <- get Ctxt
Just gdef <- lookupCtxtExact (Resolved i) (gamma defs)
| Nothing => pure ()
ignore $ addDef (Resolved i) (record { type = ty } gdef)
ignore $ addDef (Resolved i) ({ type := ty } gdef)
export
setCompiled : {auto c : Ref Ctxt Defs} ->
@ -1138,7 +1138,7 @@ setCompiled n cexp
= do defs <- get Ctxt
Just gdef <- lookupCtxtExact n (gamma defs)
| Nothing => pure ()
ignore $ addDef n (record { compexpr = Just cexp } gdef)
ignore $ addDef n ({ compexpr := Just cexp } gdef)
-- Record that the name has been linearity checked so we don't need to do
-- it again
@ -1149,13 +1149,13 @@ setLinearCheck i chk
= do defs <- get Ctxt
Just gdef <- lookupCtxtExact (Resolved i) (gamma defs)
| Nothing => pure ()
ignore $ addDef (Resolved i) (record { linearChecked = chk } gdef)
ignore $ addDef (Resolved i) ({ linearChecked := chk } gdef)
export
setCtxt : {auto c : Ref Ctxt Defs} -> Context -> Core ()
setCtxt gam
= do defs <- get Ctxt
put Ctxt (record { gamma = gam } defs)
put Ctxt ({ gamma := gam } defs)
export
resolveName : {auto c : Ref Ctxt Defs} ->
@ -1222,9 +1222,9 @@ addToSave : {auto c : Ref Ctxt Defs} ->
addToSave n_in
= do defs <- get Ctxt
n <- full (gamma defs) n_in
put Ctxt (record { toSave $= insert n (),
toIR $= insert n ()
} defs)
put Ctxt ({ toSave $= insert n (),
toIR $= insert n ()
} defs)
-- Specific lookup functions
export
@ -1330,7 +1330,7 @@ setFlag fc n fl
Just def <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName fc n
let flags' = fl :: filter (/= fl) (flags def)
ignore $ addDef n (record { flags = flags' } def)
ignore $ addDef n ({ flags := flags' } def)
export
setNameFlag : {auto c : Ref Ctxt Defs} ->
@ -1340,7 +1340,7 @@ setNameFlag fc n fl
[(n', i, def)] <- lookupCtxtName n (gamma defs)
| res => ambiguousName fc n (map fst res)
let flags' = fl :: filter (/= fl) (flags def)
ignore $ addDef (Resolved i) (record { flags = flags' } def)
ignore $ addDef (Resolved i) ({ flags := flags' } def)
export
unsetFlag : {auto c : Ref Ctxt Defs} ->
@ -1350,7 +1350,7 @@ unsetFlag fc n fl
Just def <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName fc n
let flags' = filter (/= fl) (flags def)
ignore $ addDef n (record { flags = flags' } def)
ignore $ addDef n ({ flags := flags' } def)
export
hasFlag : {auto c : Ref Ctxt Defs} ->
@ -1368,7 +1368,7 @@ setSizeChange loc n sc
= do defs <- get Ctxt
Just def <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName loc n
ignore $ addDef n (record { sizeChange = sc } def)
ignore $ addDef n ({ sizeChange := sc } def)
export
setTotality : {auto c : Ref Ctxt Defs} ->
@ -1377,7 +1377,7 @@ setTotality loc n tot
= do defs <- get Ctxt
Just def <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName loc n
ignore $ addDef n (record { totality = tot } def)
ignore $ addDef n ({ totality := tot } def)
export
setCovering : {auto c : Ref Ctxt Defs} ->
@ -1386,7 +1386,7 @@ setCovering loc n tot
= do defs <- get Ctxt
Just def <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName loc n
ignore $ addDef n (record { totality->isCovering = tot } def)
ignore $ addDef n ({ totality->isCovering := tot } def)
export
setTerminating : {auto c : Ref Ctxt Defs} ->
@ -1395,7 +1395,7 @@ setTerminating loc n tot
= do defs <- get Ctxt
Just def <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName loc n
ignore $ addDef n (record { totality->isTerminating = tot } def)
ignore $ addDef n ({ totality->isTerminating := tot } def)
export
getTotality : {auto c : Ref Ctxt Defs} ->
@ -1422,7 +1422,7 @@ setVisibility fc n vis
= do defs <- get Ctxt
Just def <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName fc n
ignore $ addDef n (record { visibility = vis } def)
ignore $ addDef n ({ visibility := vis } def)
public export
record SearchData where
@ -1502,14 +1502,14 @@ addMutData : {auto c : Ref Ctxt Defs} ->
Name -> Core ()
addMutData n
= do defs <- get Ctxt
put Ctxt (record { mutData $= (n ::) } defs)
put Ctxt ({ mutData $= (n ::) } defs)
export
dropMutData : {auto c : Ref Ctxt Defs} ->
Name -> Core ()
dropMutData n
= do defs <- get Ctxt
put Ctxt (record { mutData $= filter (/= n) } defs)
put Ctxt ({ mutData $= filter (/= n) } defs)
export
setDetermining : {auto c : Ref Ctxt Defs} ->
@ -1555,7 +1555,7 @@ setUniqueSearch fc tyn u
| _ => undefinedName fc tyn
let TCon t a ps ds fl cons ms det = definition g
| _ => throw (GenericMsg fc (show (fullname g) ++ " is not a type constructor [setDetermining]"))
let fl' = record { uniqueAuto = u } fl
let fl' = { uniqueAuto := u } fl
updateDef tyn (const (Just (TCon t a ps ds fl' cons ms det)))
export
@ -1567,7 +1567,7 @@ setExternal fc tyn u
| _ => undefinedName fc tyn
let TCon t a ps ds fl cons ms det = definition g
| _ => throw (GenericMsg fc (show (fullname g) ++ " is not a type constructor [setDetermining]"))
let fl' = record { external = u } fl
let fl' = { external := u } fl
updateDef tyn (const (Just (TCon t a ps ds fl' cons ms det)))
export
@ -1586,12 +1586,12 @@ addHintFor fc tyn_in hintn_in direct loading
Nothing => []
if loading
then put Ctxt
(record { typeHints $= insert tyn ((hintn, direct) :: hs)
} defs)
({ typeHints $= insert tyn ((hintn, direct) :: hs)
} defs)
else put Ctxt
(record { typeHints $= insert tyn ((hintn, direct) :: hs),
saveTypeHints $= ((tyn, hintn, direct) :: )
} defs)
({ typeHints $= insert tyn ((hintn, direct) :: hs),
saveTypeHints $= ((tyn, hintn, direct) :: )
} defs)
export
addGlobalHint : {auto c : Ref Ctxt Defs} ->
@ -1600,8 +1600,8 @@ addGlobalHint hintn_in isdef
= do defs <- get Ctxt
hintn <- toResolvedNames hintn_in
put Ctxt (record { autoHints $= insert hintn isdef,
saveAutoHints $= ((hintn, isdef) ::) } defs)
put Ctxt ({ autoHints $= insert hintn isdef,
saveAutoHints $= ((hintn, isdef) ::) } defs)
export
addLocalHint : {auto c : Ref Ctxt Defs} ->
@ -1609,27 +1609,27 @@ addLocalHint : {auto c : Ref Ctxt Defs} ->
addLocalHint hintn_in
= do defs <- get Ctxt
hintn <- toResolvedNames hintn_in
put Ctxt (record { localHints $= insert hintn () } defs)
put Ctxt ({ localHints $= insert hintn () } defs)
export
addOpenHint : {auto c : Ref Ctxt Defs} -> Name -> Core ()
addOpenHint hintn_in
= do defs <- get Ctxt
hintn <- toResolvedNames hintn_in
put Ctxt (record { openHints $= insert hintn () } defs)
put Ctxt ({ openHints $= insert hintn () } defs)
export
dropOpenHint : {auto c : Ref Ctxt Defs} -> Name -> Core ()
dropOpenHint hintn_in
= do defs <- get Ctxt
hintn <- toResolvedNames hintn_in
put Ctxt (record { openHints $= delete hintn } defs)
put Ctxt ({ openHints $= delete hintn } defs)
export
setOpenHints : {auto c : Ref Ctxt Defs} -> NameMap () -> Core ()
setOpenHints hs
= do d <- get Ctxt
put Ctxt (record { openHints = hs } d)
put Ctxt ({ openHints := hs } d)
export
addTransform : {auto c : Ref Ctxt Defs} ->
@ -1645,18 +1645,18 @@ addTransform fc t_in
t_full <- toFullNames t_in
case lookup fn (transforms defs) of
Nothing =>
put Ctxt (record { transforms $= insert fn [t],
saveTransforms $= ((fn_full, t_full) ::) } defs)
put Ctxt ({ transforms $= insert fn [t],
saveTransforms $= ((fn_full, t_full) ::) } defs)
Just ts =>
put Ctxt (record { transforms $= insert fn (t :: ts),
saveTransforms $= ((fn_full, t_full) ::) } defs)
put Ctxt ({ transforms $= insert fn (t :: ts),
saveTransforms $= ((fn_full, t_full) ::) } defs)
export
clearSavedHints : {auto c : Ref Ctxt Defs} -> Core ()
clearSavedHints
= do defs <- get Ctxt
put Ctxt (record { saveTypeHints = [],
saveAutoHints = [] } defs)
put Ctxt ({ saveTypeHints := [],
saveAutoHints := [] } defs)
-- Set the default namespace for new definitions
export
@ -1664,7 +1664,7 @@ setNS : {auto c : Ref Ctxt Defs} ->
Namespace -> Core ()
setNS ns
= do defs <- get Ctxt
put Ctxt (record { currentNS = ns } defs)
put Ctxt ({ currentNS := ns } defs)
-- Set the nested namespaces we're allowed to look inside
export
@ -1672,7 +1672,7 @@ setNestedNS : {auto c : Ref Ctxt Defs} ->
List Namespace -> Core ()
setNestedNS ns
= do defs <- get Ctxt
put Ctxt (record { nestedNS = ns } defs)
put Ctxt ({ nestedNS := ns } defs)
-- Get the default namespace for new definitions
export
@ -1699,7 +1699,7 @@ addImported : {auto c : Ref Ctxt Defs} ->
(ModuleIdent, Bool, Namespace) -> Core ()
addImported mod
= do defs <- get Ctxt
put Ctxt (record { imported $= (mod ::) } defs)
put Ctxt ({ imported $= (mod ::) } defs)
export
getImported : {auto c : Ref Ctxt Defs} ->
@ -1717,7 +1717,7 @@ addDirective c str
Nothing => -- warn, rather than fail, because the CG may exist
-- but be unknown to this particular instance
coreLift $ putStrLn $ "Unknown code generator " ++ c
Just cg => put Ctxt (record { cgdirectives $= ((cg, str) ::) } defs)
Just cg => put Ctxt ({ cgdirectives $= ((cg, str) ::) } defs)
export
getDirectives : {auto c : Ref Ctxt Defs} ->
@ -1735,7 +1735,7 @@ getNextTypeTag : {auto c : Ref Ctxt Defs} ->
Core Int
getNextTypeTag
= do defs <- get Ctxt
put Ctxt (record { nextTag $= (+1) } defs)
put Ctxt ({ nextTag $= (+1) } defs)
pure (nextTag defs)
-- Add a new nested namespace to the current namespace for new definitions
@ -1747,7 +1747,7 @@ extendNS : {auto c : Ref Ctxt Defs} ->
Namespace -> Core ()
extendNS ns
= do defs <- get Ctxt
put Ctxt (record { currentNS $= (<.> ns) } defs)
put Ctxt ({ currentNS $= (<.> ns) } defs)
export
withExtendedNS : {auto c : Ref Ctxt Defs} ->
@ -1755,10 +1755,10 @@ withExtendedNS : {auto c : Ref Ctxt Defs} ->
withExtendedNS ns act
= do defs <- get Ctxt
let cns = currentNS defs
put Ctxt (record { currentNS = cns <.> ns } defs)
put Ctxt ({ currentNS := cns <.> ns } defs)
ma <- catch (Right <$> act) (pure . Left)
defs <- get Ctxt
put Ctxt (record { currentNS = cns } defs)
put Ctxt ({ currentNS := cns } defs)
case ma of
Left err => throw err
Right a => pure a
@ -1794,7 +1794,7 @@ setVisible : {auto c : Ref Ctxt Defs} ->
Namespace -> Core ()
setVisible nspace
= do defs <- get Ctxt
put Ctxt (record { gamma->visibleNS $= (nspace ::) } defs)
put Ctxt ({ gamma->visibleNS $= (nspace ::) } defs)
export
getVisible : {auto c : Ref Ctxt Defs} ->
@ -1811,7 +1811,7 @@ setAllPublic : {auto c : Ref Ctxt Defs} ->
(pub : Bool) -> Core ()
setAllPublic pub
= do defs <- get Ctxt
put Ctxt (record { gamma->allPublic = pub } defs)
put Ctxt ({ gamma->allPublic := pub } defs)
export
isAllPublic : {auto c : Ref Ctxt Defs} ->
@ -1851,7 +1851,7 @@ setNextEntry : {auto c : Ref Ctxt Defs} ->
Int -> Core ()
setNextEntry i
= do defs <- get Ctxt
put Ctxt (record { gamma->nextEntry = i } defs)
put Ctxt ({ gamma->nextEntry := i } defs)
-- Set the 'first entry' index (i.e. the first entry in the current file)
-- to the place we currently are in the context
@ -1860,7 +1860,7 @@ resetFirstEntry : {auto c : Ref Ctxt Defs} ->
Core ()
resetFirstEntry
= do defs <- get Ctxt
put Ctxt (record { gamma->firstEntry = nextEntry (gamma defs) } defs)
put Ctxt ({ gamma->firstEntry := nextEntry (gamma defs) } defs)
export
getFullName : {auto c : Ref Ctxt Defs} ->
@ -1886,14 +1886,14 @@ setPPrint : {auto c : Ref Ctxt Defs} ->
PPrinter -> Core ()
setPPrint ppopts
= do defs <- get Ctxt
put Ctxt (record { options->printing = ppopts } defs)
put Ctxt ({ options->printing := ppopts } defs)
export
setCG : {auto c : Ref Ctxt Defs} ->
CG -> Core ()
setCG cg
= do defs <- get Ctxt
put Ctxt (record { options->session->codegen = cg } defs)
put Ctxt ({ options->session->codegen := cg } defs)
export
getDirs : {auto c : Ref Ctxt Defs} -> Core Dirs
@ -1905,49 +1905,49 @@ export
addExtraDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
addExtraDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->extra_dirs $= (++ [dir]) } defs)
put Ctxt ({ options->dirs->extra_dirs $= (++ [dir]) } defs)
export
addPackageDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
addPackageDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->package_dirs $= (++ [dir]) } defs)
put Ctxt ({ options->dirs->package_dirs $= (++ [dir]) } defs)
export
addDataDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
addDataDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->data_dirs $= (++ [dir]) } defs)
put Ctxt ({ options->dirs->data_dirs $= (++ [dir]) } defs)
export
addLibDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
addLibDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->lib_dirs $= (++ [dir]) } defs)
put Ctxt ({ options->dirs->lib_dirs $= (++ [dir]) } defs)
export
setBuildDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
setBuildDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->build_dir = dir } defs)
put Ctxt ({ options->dirs->build_dir := dir } defs)
export
setDependsDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
setDependsDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->depends_dir = dir } defs)
put Ctxt ({ options->dirs->depends_dir := dir } defs)
export
setOutputDir : {auto c : Ref Ctxt Defs} -> Maybe String -> Core ()
setOutputDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->output_dir = dir } defs)
put Ctxt ({ options->dirs->output_dir := dir } defs)
export
setSourceDir : {auto c : Ref Ctxt Defs} -> Maybe String -> Core ()
setSourceDir mdir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->source_dir = mdir } defs)
put Ctxt ({ options->dirs->source_dir := mdir } defs)
export
setWorkingDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
@ -1956,7 +1956,7 @@ setWorkingDir dir
coreLift_ $ changeDir dir
Just cdir <- coreLift $ currentDir
| Nothing => throw (InternalError "Can't get current directory")
put Ctxt (record { options->dirs->working_dir = cdir } defs)
put Ctxt ({ options->dirs->working_dir := cdir } defs)
export
getWorkingDir : Core String
@ -1977,13 +1977,13 @@ export
setPrefix : {auto c : Ref Ctxt Defs} -> String -> Core ()
setPrefix dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->prefix_dir = dir } defs)
put Ctxt ({ options->dirs->prefix_dir := dir } defs)
export
setExtension : {auto c : Ref Ctxt Defs} -> LangExt -> Core ()
setExtension e
= do defs <- get Ctxt
put Ctxt (record { options $= setExtension e } defs)
put Ctxt ({ options $= setExtension e } defs)
export
isExtension : LangExt -> Defs -> Bool
@ -2003,55 +2003,55 @@ lazyActive : {auto c : Ref Ctxt Defs} ->
Bool -> Core ()
lazyActive a
= do defs <- get Ctxt
put Ctxt (record { options->elabDirectives->lazyActive = a } defs)
put Ctxt ({ options->elabDirectives->lazyActive := a } defs)
export
setUnboundImplicits : {auto c : Ref Ctxt Defs} ->
Bool -> Core ()
setUnboundImplicits a
= do defs <- get Ctxt
put Ctxt (record { options->elabDirectives->unboundImplicits = a } defs)
put Ctxt ({ options->elabDirectives->unboundImplicits := a } defs)
export
setPrefixRecordProjections : {auto c : Ref Ctxt Defs} -> Bool -> Core ()
setPrefixRecordProjections b = do
defs <- get Ctxt
put Ctxt (record { options->elabDirectives->prefixRecordProjections = b } defs)
put Ctxt ({ options->elabDirectives->prefixRecordProjections := b } defs)
export
setDefaultTotalityOption : {auto c : Ref Ctxt Defs} ->
TotalReq -> Core ()
setDefaultTotalityOption tot
= do defs <- get Ctxt
put Ctxt (record { options->elabDirectives->totality = tot } defs)
put Ctxt ({ options->elabDirectives->totality := tot } defs)
export
setAmbigLimit : {auto c : Ref Ctxt Defs} ->
Nat -> Core ()
setAmbigLimit max
= do defs <- get Ctxt
put Ctxt (record { options->elabDirectives->ambigLimit = max } defs)
put Ctxt ({ options->elabDirectives->ambigLimit := max } defs)
export
setAutoImplicitLimit : {auto c : Ref Ctxt Defs} ->
Nat -> Core ()
setAutoImplicitLimit max
= do defs <- get Ctxt
put Ctxt (record { options->elabDirectives->autoImplicitLimit = max } defs)
put Ctxt ({ options->elabDirectives->autoImplicitLimit := max } defs)
export
setNFThreshold : {auto c : Ref Ctxt Defs} ->
Nat -> Core ()
setNFThreshold max
= do defs <- get Ctxt
put Ctxt (record { options->elabDirectives->nfThreshold = max } defs)
put Ctxt ({ options->elabDirectives->nfThreshold := max } defs)
export
setSearchTimeout : {auto c : Ref Ctxt Defs} ->
Integer -> Core ()
setSearchTimeout t
= do defs <- get Ctxt
put Ctxt (record { options->session->searchTimeout = t } defs)
put Ctxt ({ options->session->searchTimeout := t } defs)
export
isLazyActive : {auto c : Ref Ctxt Defs} ->
@ -2102,7 +2102,7 @@ setPair fc ty f s
ty' <- checkUnambig fc ty
f' <- checkUnambig fc f
s' <- checkUnambig fc s
put Ctxt (record { options $= setPair ty' f' s' } defs)
put Ctxt ({ options $= setPair ty' f' s' } defs)
export
setRewrite : {auto c : Ref Ctxt Defs} ->
@ -2111,7 +2111,7 @@ setRewrite fc eq rw
= do defs <- get Ctxt
rw' <- checkUnambig fc rw
eq' <- checkUnambig fc eq
put Ctxt (record { options $= setRewrite eq' rw' } defs)
put Ctxt ({ options $= setRewrite eq' rw' } defs)
-- Don't check for ambiguity here; they're all meant to be overloadable
export
@ -2119,28 +2119,28 @@ setFromInteger : {auto c : Ref Ctxt Defs} ->
Name -> Core ()
setFromInteger n
= do defs <- get Ctxt
put Ctxt (record { options $= setFromInteger n } defs)
put Ctxt ({ options $= setFromInteger n } defs)
export
setFromString : {auto c : Ref Ctxt Defs} ->
Name -> Core ()
setFromString n
= do defs <- get Ctxt
put Ctxt (record { options $= setFromString n } defs)
put Ctxt ({ options $= setFromString n } defs)
export
setFromChar : {auto c : Ref Ctxt Defs} ->
Name -> Core ()
setFromChar n
= do defs <- get Ctxt
put Ctxt (record { options $= setFromChar n } defs)
put Ctxt ({ options $= setFromChar n } defs)
export
setFromDouble : {auto c : Ref Ctxt Defs} ->
Name -> Core ()
setFromDouble n
= do defs <- get Ctxt
put Ctxt (record { options $= setFromDouble n } defs)
put Ctxt ({ options $= setFromDouble n } defs)
export
addNameDirective : {auto c : Ref Ctxt Defs} ->
@ -2148,7 +2148,7 @@ addNameDirective : {auto c : Ref Ctxt Defs} ->
addNameDirective fc n ns
= do defs <- get Ctxt
n' <- checkUnambig fc n
put Ctxt (record { namedirectives $= insert n' ns } defs)
put Ctxt ({ namedirectives $= insert n' ns } defs)
-- Checking special names from Options
@ -2244,10 +2244,10 @@ addLogLevel : {auto c : Ref Ctxt Defs} ->
addLogLevel lvl
= do defs <- get Ctxt
case lvl of
Nothing => put Ctxt (record { options->session->logEnabled = True,
options->session->logLevel = defaultLogLevel } defs)
Just l => put Ctxt (record { options->session->logEnabled = True,
options->session->logLevel $= insertLogLevel l } defs)
Nothing => put Ctxt ({ options->session->logEnabled := True,
options->session->logLevel := defaultLogLevel } defs)
Just l => put Ctxt ({ options->session->logEnabled := True,
options->session->logLevel $= insertLogLevel l } defs)
export
withLogLevel : {auto c : Ref Ctxt Defs} ->
@ -2255,10 +2255,10 @@ withLogLevel : {auto c : Ref Ctxt Defs} ->
withLogLevel l comp = do
defs <- get Ctxt
let logs = logLevel (session (options defs))
put Ctxt (record { options->session->logLevel = insertLogLevel l logs } defs)
put Ctxt ({ options->session->logLevel := insertLogLevel l logs } defs)
r <- comp
defs <- get Ctxt
put Ctxt (record { options->session->logLevel = logs } defs)
put Ctxt ({ options->session->logLevel := logs } defs)
pure r
export
@ -2266,14 +2266,14 @@ setLogTimings : {auto c : Ref Ctxt Defs} ->
Bool -> Core ()
setLogTimings b
= do defs <- get Ctxt
put Ctxt (record { options->session->logTimings = b } defs)
put Ctxt ({ options->session->logTimings := b } defs)
export
setDebugElabCheck : {auto c : Ref Ctxt Defs} ->
Bool -> Core ()
setDebugElabCheck b
= do defs <- get Ctxt
put Ctxt (record { options->session->debugElabCheck = b } defs)
put Ctxt ({ options->session->debugElabCheck := b } defs)
export
getSession : {auto c : Ref Ctxt Defs} ->
@ -2287,7 +2287,7 @@ setSession : {auto c : Ref Ctxt Defs} ->
Session -> Core ()
setSession sopts
= do defs <- get Ctxt
put Ctxt (record { options->session = sopts } defs)
put Ctxt ({ options->session := sopts } defs)
%inline
export
@ -2300,7 +2300,7 @@ recordWarning : {auto c : Ref Ctxt Defs} -> Warning -> Core ()
recordWarning w
= do defs <- get Ctxt
session <- getSession
put Ctxt $ record { warnings $= (w ::) } defs
put Ctxt $ { warnings $= (w ::) } defs
export
getTime : Core Integer
@ -2325,14 +2325,14 @@ startTimer : {auto c : Ref Ctxt Defs} ->
startTimer tmax action
= do t <- getTime
defs <- get Ctxt
put Ctxt $ record { timer = Just (t + tmax * 1000000, action) } defs
put Ctxt $ { timer := Just (t + tmax * 1000000, action) } defs
||| Clear the timer
export
clearTimer : {auto c : Ref Ctxt Defs} -> Core ()
clearTimer
= do defs <- get Ctxt
put Ctxt $ record { timer = Nothing } defs
put Ctxt $ { timer := Nothing } defs
||| If the timer was started more than t milliseconds ago, throw an exception
export
@ -2363,9 +2363,9 @@ addImportedInc modNS inc
-- compile incrementally
do recordWarning (GenericWarn ("No incremental compile data for " ++ show modNS))
defs <- get Ctxt
put Ctxt (record { allIncData $= drop cg } defs)
put Ctxt ({ allIncData $= drop cg } defs)
Just (mods, extra) =>
put Ctxt (record { allIncData $= addMod cg (mods, extra) }
put Ctxt ({ allIncData $= addMod cg (mods, extra) }
defs)
where
addMod : CG -> (String, List String) ->
@ -2389,7 +2389,7 @@ setIncData : {auto c : Ref Ctxt Defs} ->
CG -> (String, List String) -> Core ()
setIncData cg res
= do defs <- get Ctxt
put Ctxt (record { incData $= ((cg, res) :: )} defs)
put Ctxt ({ incData $= ((cg, res) :: )} defs)
-- Set a name as Private that was previously visible (and, if 'everywhere' is
-- set, hide in any modules imported by this one)
@ -2400,7 +2400,7 @@ hide fc n
= do defs <- get Ctxt
[(nsn, _)] <- lookupCtxtName n (gamma defs)
| res => ambiguousName fc n (map fst res)
put Ctxt (record { gamma $= hideName nsn } defs)
put Ctxt ({ gamma $= hideName nsn } defs)
-- Set a name as Public that was previously hidden
-- Note: this is here at the bottom only becuase `recordWarning` is defined just above.
@ -2411,7 +2411,7 @@ unhide fc n
= do defs <- get Ctxt
[(nsn, _)] <- lookupHiddenCtxtName n (gamma defs)
| res => ambiguousName fc n (map fst res)
put Ctxt (record { gamma $= unhideName nsn } defs)
put Ctxt ({ gamma $= unhideName nsn } defs)
unless (isHidden nsn (gamma defs)) $ do
recordWarning $ GenericWarn $
"Trying to %unhide `" ++ show nsn ++ "`, which was not hidden in the first place"

View File

@ -109,7 +109,7 @@ addData vars vis tidx (MkData (MkCon dfc tyn arity tycon) datacons)
defaultFlags [] (map name datacons) Nothing)
(idx, gam') <- addCtxt tyn tydef (gamma defs)
gam'' <- addDataConstructors 0 datacons gam'
put Ctxt (record { gamma = gam'' } defs)
put Ctxt ({ gamma := gam'' } defs)
pure idx
where
allDet : Nat -> List Nat

View File

@ -162,7 +162,7 @@ logTimeRecord' key act
let tot = case lookup key (timings defs) of
Nothing => 0
Just (_, t) => t
put Ctxt (record { timings $= insert key (False, tot + time) } defs)
put Ctxt ({ timings $= insert key (False, tot + time) } defs)
pure res
-- for ad-hoc profiling, record the time the action takes and add it
@ -176,7 +176,7 @@ logTimeRecord key act
case lookup key (timings defs) of
Just (True, t) => act
Just (False, t)
=> do put Ctxt (record { timings $= insert key (True, t) } defs)
=> do put Ctxt ({ timings $= insert key (True, t) } defs)
logTimeRecord' key act
Nothing
=> logTimeRecord' key act

View File

@ -206,7 +206,7 @@ addLHS loc outerenvlen env tm
tm' <- toFullNames (bindEnv loc (toPat env) tm)
-- Put the lhs on the metadata if it's not empty
whenJust (isNonEmptyFC loc) $ \ neloc =>
put MD $ record { lhsApps $= ((neloc, outerenvlen, tm') ::) } meta
put MD $ { lhsApps $= ((neloc, outerenvlen, tm') ::) } meta
where
toPat : Env Term vs -> Env Term vs
@ -239,7 +239,7 @@ addNameType loc n env tm
-- Add the name to the metadata if the file context is not empty
whenJust (isConcreteFC loc) $ \ neloc => do
put MD $ record { names $= ((neloc, (n', 0, substEnv loc env tm)) ::) } meta
put MD $ { names $= ((neloc, (n', 0, substEnv loc env tm)) ::) } meta
log "metadata.names" 7 $ show n' ++ " at line " ++ show (1 + startLine neloc)
export
@ -253,7 +253,7 @@ addTyDecl loc n env tm
-- Add the type declaration to the metadata if the file context is not empty
whenJust (isNonEmptyFC loc) $ \ neloc =>
put MD $ record { tydecls $= ( (neloc, (n', length env, bindEnv loc env tm)) ::) } meta
put MD $ { tydecls $= ( (neloc, (n', length env, bindEnv loc env tm)) ::) } meta
export
addNameLoc : {auto m : Ref MD Metadata} ->
@ -263,21 +263,21 @@ addNameLoc loc n
= do meta <- get MD
n' <- getFullName n
whenJust (isConcreteFC loc) $ \neloc =>
put MD $ record { nameLocMap $= insert (neloc, n') } meta
put MD $ { nameLocMap $= insert (neloc, n') } meta
export
setHoleLHS : {auto m : Ref MD Metadata} ->
ClosedTerm -> Core ()
setHoleLHS tm
= do meta <- get MD
put MD (record { currentLHS = Just tm } meta)
put MD ({ currentLHS := Just tm } meta)
export
clearHoleLHS : {auto m : Ref MD Metadata} ->
Core ()
clearHoleLHS
= do meta <- get MD
put MD (record { currentLHS = Nothing } meta)
put MD ({ currentLHS := Nothing } meta)
export
withCurrentLHS : {auto c : Ref Ctxt Defs} ->
@ -287,7 +287,7 @@ withCurrentLHS n
= do meta <- get MD
n' <- getFullName n
maybe (pure ())
(\lhs => put MD (record { holeLHS $= ((n', lhs) ::) } meta))
(\lhs => put MD ({ holeLHS $= ((n', lhs) ::) } meta))
(currentLHS meta)
findEntryWith : (NonEmptyFC -> a -> Bool) -> List (NonEmptyFC, a) -> Maybe (NonEmptyFC, a)
@ -355,8 +355,8 @@ addSemanticDecorations decors
unless (isNil droppedDecors)
$ log "ide-mode.highlight" 19 $ "ignored adding decorations to "
++ show meta.sourceIdent ++ ": " ++ show droppedDecors
put MD $ record {semanticHighlighting
= (fromList newDecors) `union` posmap} meta
put MD $ {semanticHighlighting
:= (fromList newDecors) `union` posmap} meta
-- Normalise all the types of the names, since they might have had holes
@ -368,7 +368,7 @@ normaliseTypes
= do meta <- get MD
defs <- get Ctxt
ns' <- traverse (nfType defs) (names meta)
put MD (record { names = ns' } meta)
put MD ({ names := ns' } meta)
where
nfType : Defs -> (NonEmptyFC, (Name, Nat, ClosedTerm)) ->
Core (NonEmptyFC, (Name, Nat, ClosedTerm))
@ -396,13 +396,13 @@ TTC TTMFile where
HasNames Metadata where
full gam md
= pure $ record { lhsApps = !(traverse fullLHS $ md.lhsApps)
, names = !(traverse fullTy $ md.names)
, tydecls = !(traverse fullTy $ md.tydecls)
, currentLHS = Nothing
, holeLHS = !(traverse fullHLHS $ md.holeLHS)
, nameLocMap = fromList !(traverse fullDecls (toList $ md.nameLocMap))
} md
= pure $ { lhsApps := !(traverse fullLHS $ md.lhsApps)
, names := !(traverse fullTy $ md.names)
, tydecls := !(traverse fullTy $ md.tydecls)
, currentLHS := Nothing
, holeLHS := !(traverse fullHLHS $ md.holeLHS)
, nameLocMap := fromList !(traverse fullDecls (toList $ md.nameLocMap))
} md
where
fullLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm))
fullLHS (fc, (i, tm)) = pure (fc, (i, !(full gam tm)))

View File

@ -66,7 +66,7 @@ updateLimit Func n opts
Nothing => Nothing
Just Z => Nothing
Just (S k) =>
Just (record { reduceLimit $= set n k } opts)
Just ({ reduceLimit $= set n k } opts)
where
set : Name -> Nat -> List (Name, Nat) -> List (Name, Nat)
set n v [] = []
@ -410,7 +410,7 @@ parameters (defs : Defs, topopts : EvalOpts)
pure (Result res)
Just Z => pure GotStuck
Just (S k) =>
do let opts' = record { fuel = Just k } opts
do let opts' = { fuel := Just k } opts
res <- evalWithOpts defs opts' env loc (embed tm) stk
pure (Result res)
evalTree env loc opts fc stk _ = pure GotStuck

View File

@ -183,11 +183,11 @@ mutual
opts' <- case sizeLimit opts of
Nothing => pure opts
Just Z => throw (InternalError "Size limit exceeded")
Just (S k) => pure (record { sizeLimit = Just k } opts)
Just (S k) => pure ({ sizeLimit := Just k } opts)
args' <- if patterns opts && not (topLevel opts) && isRef f
then do empty <- clearDefs defs
quoteArgsWithFC q opts' empty bound env args
else quoteArgsWithFC q (record { topLevel = False } opts')
else quoteArgsWithFC q ({ topLevel := False } opts')
defs bound env args
pure $ applyWithFC f' args'
where
@ -216,8 +216,8 @@ mutual
where
toHolesOnly : Closure vs -> Closure vs
toHolesOnly (MkClosure opts locs env tm)
= MkClosure (record { holesOnly = True,
argHolesOnly = True } opts)
= MkClosure ({ holesOnly := True,
argHolesOnly := True } opts)
locs env tm
toHolesOnly c = c
quoteGenNF q opts defs bound env (NForce fc r arg args)

View File

@ -264,40 +264,40 @@ defaults
-- Reset the options which are set by source files
export
clearNames : Options -> Options
clearNames = record { pairnames = Nothing,
rewritenames = Nothing,
primnames = MkPrimNs Nothing Nothing Nothing Nothing,
extensions = []
}
clearNames = { pairnames := Nothing,
rewritenames := Nothing,
primnames := MkPrimNs Nothing Nothing Nothing Nothing,
extensions := []
}
export
setPair : (pairType : Name) -> (fstn : Name) -> (sndn : Name) ->
Options -> Options
setPair ty f s = record { pairnames = Just (MkPairNs ty f s) }
setPair ty f s = { pairnames := Just (MkPairNs ty f s) }
export
setRewrite : (eq : Name) -> (rwlemma : Name) -> Options -> Options
setRewrite eq rw = record { rewritenames = Just (MkRewriteNs eq rw) }
setRewrite eq rw = { rewritenames := Just (MkRewriteNs eq rw) }
export
setFromInteger : Name -> Options -> Options
setFromInteger n = record { primnames->fromIntegerName = Just n }
setFromInteger n = { primnames->fromIntegerName := Just n }
export
setFromString : Name -> Options -> Options
setFromString n = record { primnames->fromStringName = Just n }
setFromString n = { primnames->fromStringName := Just n }
export
setFromChar : Name -> Options -> Options
setFromChar n = record { primnames->fromCharName = Just n }
setFromChar n = { primnames->fromCharName := Just n }
export
setFromDouble : Name -> Options -> Options
setFromDouble n = record { primnames->fromDoubleName = Just n }
setFromDouble n = { primnames->fromDoubleName := Just n }
export
setExtension : LangExt -> Options -> Options
setExtension e = record { extensions $= (e ::) }
setExtension e = { extensions $= (e ::) }
export
isExtension : LangExt -> Options -> Bool
@ -305,4 +305,4 @@ isExtension e opts = e `elem` extensions opts
export
addCG : (String, CG) -> Options -> Options
addCG cg = record { additionalCGs $= (cg::) }
addCG cg = { additionalCGs $= (cg::) }

View File

@ -567,7 +567,7 @@ compileDef mode n_in
| Nothing => throw (InternalError ("Compiling " ++ show n ++ " failed"))
-- Record that this one is done
ignore $ addDef n (record { schemeExpr = Just (mode, schdef) } def)
ignore $ addDef n ({ schemeExpr := Just (mode, schdef) } def)
initEvalWith : {auto c : Ref Ctxt Defs} ->
String -> Core Bool
@ -579,7 +579,7 @@ initEvalWith "chez"
catch (do f <- readDataFile "chez/ct-support.ss"
Just _ <- coreLift $ evalSchemeStr $ "(begin " ++ f ++ ")"
| Nothing => pure False
put Ctxt (record { schemeEvalLoaded = True } defs)
put Ctxt ({ schemeEvalLoaded := True } defs)
pure True)
(\err => pure False)
initEvalWith "racket"
@ -590,7 +590,7 @@ initEvalWith "racket"
catch (do f <- readDataFile "racket/ct-support.rkt"
Just _ <- coreLift $ evalSchemeStr $ "(begin " ++ f ++ ")"
| Nothing => pure False
put Ctxt (record { schemeEvalLoaded = True } defs)
put Ctxt ({ schemeEvalLoaded := True } defs)
pure True)
(\err => do coreLift $ printLn err
pure False)

View File

@ -51,7 +51,7 @@ inSearch : UnifyInfo
inSearch = MkUnifyInfo True InSearch
lower : UnifyInfo -> UnifyInfo
lower = record { atTop = False }
lower = { atTop := False }
Eq UnifyMode where
InLHS == InLHS = True
@ -485,9 +485,9 @@ tryInstantiate {newvars} loc mode env mname mref num mdef locs otm tm
let simpleDef = MkPMDefInfo (SolvedHole num)
(not (isUserName mname) && isSimple rhs)
False
let newdef = record { definition =
PMDef simpleDef [] (STerm 0 rhs) (STerm 0 rhs) []
} mdef
let newdef = { definition :=
PMDef simpleDef [] (STerm 0 rhs) (STerm 0 rhs) []
} mdef
ignore $ addDef (Resolved mref) newdef
removeHole mref
pure True
@ -1290,10 +1290,10 @@ mutual
-- what the expected type turns out to be
then postpone loc mode "Postponing in lazy" env x tmy
else do vs <- unify (lower mode) loc env tmx tmy
pure (record { addLazy = AddForce r } vs)
pure ({ addLazy := AddForce r } vs)
unifyWithLazyD _ _ mode loc env tmx (NDelayed _ r tmy)
= do vs <- unify (lower mode) loc env tmx tmy
pure (record { addLazy = AddDelay r } vs)
pure ({ addLazy := AddDelay r } vs)
unifyWithLazyD _ _ mode loc env tmx tmy
= unify mode loc env tmx tmy
@ -1366,7 +1366,7 @@ setInvertible fc n
= do defs <- get Ctxt
Just gdef <- lookupCtxtExact n (gamma defs)
| Nothing => undefinedName fc n
ignore $ addDef n (record { invertible = True } gdef)
ignore $ addDef n ({ invertible := True } gdef)
public export
data SolveMode = Normal -- during elaboration: unifies and searches
@ -1470,7 +1470,7 @@ retryGuess mode smode (hid, (loc, hname))
handleUnify
(do tm <- search loc rig (smode == Defaults) depth defining
(type def) []
let gdef = record { definition = PMDef defaultPI [] (STerm 0 tm) (STerm 0 tm) [] } def
let gdef = { definition := PMDef defaultPI [] (STerm 0 tm) (STerm 0 tm) [] } def
logTermNF "unify.retry" 5 ("Solved " ++ show hname) [] tm
ignore $ addDef (Resolved hid) gdef
removeGuess hid
@ -1505,8 +1505,8 @@ retryGuess mode smode (hid, (loc, hname))
do ty <- getType [] tm
logTerm "unify.retry" 5 "Retry Delay" tm
pure $ delayMeta r envb !(getTerm ty) tm
let gdef = record { definition = PMDef (MkPMDefInfo NotHole True False)
[] (STerm 0 tm') (STerm 0 tm') [] } def
let gdef = { definition := PMDef (MkPMDefInfo NotHole True False)
[] (STerm 0 tm') (STerm 0 tm') [] } def
logTerm "unify.retry" 5 ("Resolved " ++ show hname) tm'
ignore $ addDef (Resolved hid) gdef
removeGuess hid
@ -1518,7 +1518,7 @@ retryGuess mode smode (hid, (loc, hname))
do ty <- getType [] tm
logTerm "unify.retry" 5 "Retry Delay (constrained)" tm
pure $ delayMeta r envb !(getTerm ty) tm
let gdef = record { definition = Guess tm' envb newcs } def
let gdef = { definition := Guess tm' envb newcs } def
ignore $ addDef (Resolved hid) gdef
pure False
Guess tm envb constrs =>
@ -1531,13 +1531,13 @@ retryGuess mode smode (hid, (loc, hname))
-- All constraints resolved, so turn into a
-- proper definition and remove it from the
-- hole list
[] => do let gdef = record { definition = PMDef (MkPMDefInfo NotHole True False)
[] (STerm 0 tm) (STerm 0 tm) [] } def
[] => do let gdef = { definition := PMDef (MkPMDefInfo NotHole True False)
[] (STerm 0 tm) (STerm 0 tm) [] } def
logTerm "unify.retry" 5 ("Resolved " ++ show hname) tm
ignore $ addDef (Resolved hid) gdef
removeGuess hid
pure (holesSolved csAll)
newcs => do let gdef = record { definition = Guess tm envb newcs } def
newcs => do let gdef = { definition := Guess tm envb newcs } def
ignore $ addDef (Resolved hid) gdef
pure False
_ => pure False
@ -1624,7 +1624,7 @@ checkDots
traverse_ checkConstraint (reverse (dotConstraints ust))
hs <- getCurrentHoles
ust <- get UST
put UST (record { dotConstraints = [] } ust)
put UST ({ dotConstraints := [] } ust)
where
getHoleName : Term [] -> Core (Maybe Name)
getHoleName tm
@ -1686,11 +1686,11 @@ checkDots
logTermNF "unify.constraint" 5 "Dot type" [] dty
-- Clear constraints so we don't report again
-- later
put UST (record { dotConstraints = [] } ust)
put UST ({ dotConstraints := [] } ust)
empty <- clearDefs defs
throw (BadDotPattern fc env reason
!(quote empty env x)
!(quote empty env y))
_ => do put UST (record { dotConstraints = [] } ust)
_ => do put UST ({ dotConstraints := [] } ust)
throw err)
checkConstraint _ = pure ()

View File

@ -145,7 +145,7 @@ resetNextVar : {auto u : Ref UST UState} ->
Core ()
resetNextVar
= do ust <- get UST
put UST (record { nextName = 0 } ust)
put UST ({ nextName := 0 } ust)
-- Generate a global name based on the given root, in the current namespace
export
@ -154,7 +154,7 @@ genName : {auto c : Ref Ctxt Defs} ->
String -> Core Name
genName str
= do ust <- get UST
put UST (record { nextName $= (+1) } ust)
put UST ({ nextName $= (+1) } ust)
n <- inCurrentNS (MN str (nextName ust))
pure n
@ -167,7 +167,7 @@ genMVName (UN str) = genName (displayUserName str)
genMVName (MN str _) = genName str
genMVName n
= do ust <- get UST
put UST (record { nextName $= (+1) } ust)
put UST ({ nextName $= (+1) } ust)
mn <- inCurrentNS (MN (show n) (nextName ust))
pure mn
@ -178,7 +178,7 @@ genVarName : {auto c : Ref Ctxt Defs} ->
String -> Core Name
genVarName str
= do ust <- get UST
put UST (record { nextName $= (+1) } ust)
put UST ({ nextName $= (+1) } ust)
pure (MN str (nextName ust))
-- Again, for case names
@ -188,7 +188,7 @@ genCaseName : {auto c : Ref Ctxt Defs} ->
String -> Core Name
genCaseName root
= do ust <- get UST
put UST (record { nextName $= (+1) } ust)
put UST ({ nextName $= (+1) } ust)
inCurrentNS (CaseBlock root (nextName ust))
export
@ -197,30 +197,30 @@ genWithName : {auto c : Ref Ctxt Defs} ->
String -> Core Name
genWithName root
= do ust <- get UST
put UST (record { nextName $= (+1) } ust)
put UST ({ nextName $= (+1) } ust)
inCurrentNS (WithBlock root (nextName ust))
addHoleName : {auto u : Ref UST UState} ->
FC -> Name -> Int -> Core ()
addHoleName fc n i
= do ust <- get UST
put UST (record { holes $= insert i (fc, n),
currentHoles $= insert i (fc, n) } ust)
put UST ({ holes $= insert i (fc, n),
currentHoles $= insert i (fc, n) } ust)
addGuessName : {auto u : Ref UST UState} ->
FC -> Name -> Int -> Core ()
addGuessName fc n i
= do ust <- get UST
put UST (record { guesses $= insert i (fc, n) } ust)
put UST ({ guesses $= insert i (fc, n) } ust)
export
removeHole : {auto u : Ref UST UState} ->
Int -> Core ()
removeHole n
= do ust <- get UST
put UST (record { holes $= delete n,
currentHoles $= delete n,
delayedHoles $= delete n } ust)
put UST ({ holes $= delete n,
currentHoles $= delete n,
delayedHoles $= delete n } ust)
export
removeHoleName : {auto c : Ref Ctxt Defs} ->
@ -237,21 +237,21 @@ addNoSolve : {auto u : Ref UST UState} ->
Int -> Core ()
addNoSolve i
= do ust <- get UST
put UST (record { noSolve $= insert i () } ust)
put UST ({ noSolve $= insert i () } ust)
export
removeNoSolve : {auto u : Ref UST UState} ->
Int -> Core ()
removeNoSolve i
= do ust <- get UST
put UST (record { noSolve $= delete i } ust)
put UST ({ noSolve $= delete i } ust)
export
saveHoles : {auto u : Ref UST UState} ->
Core (IntMap (FC, Name))
saveHoles
= do ust <- get UST
put UST (record { currentHoles = empty } ust)
put UST ({ currentHoles := empty } ust)
pure (currentHoles ust)
export
@ -259,14 +259,14 @@ restoreHoles : {auto u : Ref UST UState} ->
IntMap (FC, Name) -> Core ()
restoreHoles hs
= do ust <- get UST
put UST (record { currentHoles = hs } ust)
put UST ({ currentHoles := hs } ust)
export
removeGuess : {auto u : Ref UST UState} ->
Int -> Core ()
removeGuess n
= do ust <- get UST
put UST (record { guesses $= delete n } ust)
put UST ({ guesses $= delete n } ust)
-- Get all of the hole data
export
@ -312,14 +312,14 @@ setConstraint : {auto u : Ref UST UState} ->
Int -> Constraint -> Core ()
setConstraint cid c
= do ust <- get UST
put UST (record { constraints $= insert cid c } ust)
put UST ({ constraints $= insert cid c } ust)
export
deleteConstraint : {auto u : Ref UST UState} ->
Int -> Core ()
deleteConstraint cid
= do ust <- get UST
put UST (record { constraints $= delete cid } ust)
put UST ({ constraints $= delete cid } ust)
export
addConstraint : {auto u : Ref UST UState} ->
@ -328,8 +328,8 @@ addConstraint : {auto u : Ref UST UState} ->
addConstraint constr
= do ust <- get UST
let cid = nextConstraint ust
put UST (record { constraints $= insert cid constr,
nextConstraint = cid+1 } ust)
put UST ({ constraints $= insert cid constr,
nextConstraint := cid+1 } ust)
pure cid
export
@ -343,9 +343,9 @@ addDot fc env dotarg x reason y
defs <- get Ctxt
xnf <- nf defs env x
ynf <- nf defs env y
put UST (record { dotConstraints $=
((dotarg, reason, MkConstraint fc False env xnf ynf) ::)
} ust)
put UST ({ dotConstraints $=
((dotarg, reason, MkConstraint fc False env xnf ynf) ::)
} ust)
export
addPolyConstraint : {vars : _} ->
@ -354,7 +354,7 @@ addPolyConstraint : {vars : _} ->
Core ()
addPolyConstraint fc env arg x@(NApp _ (NMeta _ _ _) _) y
= do ust <- get UST
put UST (record { polyConstraints $= ((MkPolyConstraint fc env arg x y) ::) } ust)
put UST ({ polyConstraints $= ((MkPolyConstraint fc env arg x y) ::) } ust)
addPolyConstraint fc env arg x y
= pure ()
@ -460,7 +460,7 @@ newMetaLets : {vars : _} ->
newMetaLets {vars} fc rig env n ty def nocyc lets
= do let hty = if lets then abstractFullEnvType fc env ty
else abstractEnvType fc env ty
let hole = record { noCycles = nocyc }
let hole = { noCycles := nocyc }
(newDef fc n rig [] hty Public def)
log "unify.meta" 5 $ "Adding new meta " ++ show (n, fc, rig)
logTerm "unify.meta" 10 ("New meta type " ++ show n) hty
@ -573,7 +573,7 @@ tryErrorUnify elab
pure (Right res))
(\err => do put UST ust
defs' <- get Ctxt
put Ctxt (record { timings = timings defs' } defs)
put Ctxt ({ timings := timings defs' } defs)
pure (Left err))
export
@ -601,7 +601,7 @@ addDelayedHoleName : {auto u : Ref UST UState} ->
(Int, (FC, Name)) -> Core ()
addDelayedHoleName (idx, h)
= do ust <- get UST
put UST (record { delayedHoles $= insert idx h } ust)
put UST ({ delayedHoles $= insert idx h } ust)
export
checkDelayedHoles : {auto u : Ref UST UState} ->
@ -639,13 +639,13 @@ checkValidHole base (idx, (fc, n))
| Nothing => pure ()
case c of
MkConstraint fc l env x y =>
do put UST (record { guesses = empty } ust)
do put UST ({ guesses := empty } ust)
empty <- clearDefs defs
xnf <- quote empty env x
ynf <- quote empty env y
throw (CantSolveEq fc (gamma defs) env xnf ynf)
MkSeqConstraint fc env (x :: _) (y :: _) =>
do put UST (record { guesses = empty } ust)
do put UST ({ guesses := empty } ust)
empty <- clearDefs defs
xnf <- quote empty env x
ynf <- quote empty env y

View File

@ -41,11 +41,11 @@ withArgHoles = MkEvalOpts False True False False False Nothing [] CBN
export
tcOnly : EvalOpts
tcOnly = record { tcInline = True } withArgHoles
tcOnly = { tcInline := True } withArgHoles
export
onLHS : EvalOpts
onLHS = record { removeAs = False } defaultOpts
onLHS = { removeAs := False } defaultOpts
export
cbn : EvalOpts
@ -53,7 +53,7 @@ cbn = defaultOpts
export
cbv : EvalOpts
cbv = record { strategy = CBV } defaultOpts
cbv = { strategy := CBV } defaultOpts
mutual
public export

View File

@ -82,12 +82,12 @@ extendSyn newsyn
, "New (" ++ unwords (map show $ saveMod newsyn) ++ "): "
++ show (modDocstrings newsyn)
]
put Syn (record { infixes $= mergeLeft (infixes newsyn),
prefixes $= mergeLeft (prefixes newsyn),
ifaces $= merge (ifaces newsyn),
modDocstrings $= mergeLeft (modDocstrings newsyn),
defDocstrings $= merge (defDocstrings newsyn),
bracketholes $= ((bracketholes newsyn) ++) }
put Syn ({ infixes $= mergeLeft (infixes newsyn),
prefixes $= mergeLeft (prefixes newsyn),
ifaces $= merge (ifaces newsyn),
modDocstrings $= mergeLeft (modDocstrings newsyn),
defDocstrings $= merge (defDocstrings newsyn),
bracketholes $= ((bracketholes newsyn) ++) }
syn)
mkPrec : Fixity -> Nat -> OpPrec
@ -323,7 +323,7 @@ mutual
desugarB side ps (PHole fc br holename)
= do when br $
do syn <- get Syn
put Syn (record { bracketholes $= ((UN (Basic holename)) ::) } syn)
put Syn ({ bracketholes $= ((UN (Basic holename)) ::) } syn)
pure $ IHole fc holename
desugarB side ps (PType fc) = pure $ IType fc
desugarB side ps (PAs fc nameFC vname pattern)
@ -355,9 +355,9 @@ mutual
= do itm <- desugarB side ps term
bs <- get Bang
let bn = MN "bind" (nextName bs)
put Bang (record { nextName $= (+1),
bangNames $= ((bn, fc, itm) ::)
} bs)
put Bang ({ nextName $= (+1),
bangNames $= ((bn, fc, itm) ::)
} bs)
pure (IVar EmptyFC bn)
desugarB side ps (PIdiom fc term)
= do itm <- desugarB side ps term
@ -880,10 +880,10 @@ mutual
uimpls' <- traverse (\ ntm => do tm' <- desugar AnyExpr ps (snd ntm)
btm <- bindTypeNames fc oldu ps tm'
pure (fst ntm, btm)) uimpls
put Syn (record { usingImpl = uimpls' ++ oldu } syn)
put Syn ({ usingImpl := uimpls' ++ oldu } syn)
uds' <- traverse (desugarDecl ps) uds
syn <- get Syn
put Syn (record { usingImpl = oldu } syn)
put Syn ({ usingImpl := oldu } syn)
pure (concat uds')
desugarDecl ps (PReflect fc tm)
= throw (GenericMsg fc "Reflection not implemented yet")
@ -1017,11 +1017,11 @@ mutual
desugarDecl ps (PFixity fc Prefix prec (UN (Basic n)))
= do syn <- get Syn
put Syn (record { prefixes $= insert n prec } syn)
put Syn ({ prefixes $= insert n prec } syn)
pure []
desugarDecl ps (PFixity fc fix prec (UN (Basic n)))
= do syn <- get Syn
put Syn (record { infixes $= insert n (fix, prec) } syn)
put Syn ({ infixes $= insert n (fix, prec) } syn)
pure []
desugarDecl ps (PFixity fc _ _ _)
= throw (GenericMsg fc "Fixity declarations must be for unqualified names")

View File

@ -46,8 +46,8 @@ addModDocString : {auto s : Ref Syn SyntaxInfo} ->
Core ()
addModDocString mi doc
= do syn <- get Syn
put Syn (record { saveMod $= (mi ::)
, modDocstrings $= insert mi doc } syn)
put Syn ({ saveMod $= (mi ::)
, modDocstrings $= insert mi doc } syn)
-- Add a doc string for a name in the current namespace
export
@ -60,8 +60,8 @@ addDocString n_in doc
log "doc.record" 50 $
"Adding doc for " ++ show n_in ++ " (aka " ++ show n ++ " in current NS)"
syn <- get Syn
put Syn (record { defDocstrings $= addName n doc,
saveDocstrings $= insert n () } syn)
put Syn ({ defDocstrings $= addName n doc,
saveDocstrings $= insert n () } syn)
-- Add a doc string for a name, in an extended namespace (e.g. for
-- record getters)
@ -76,8 +76,8 @@ addDocStringNS ns n_in doc
NS old root => NS (old <.> ns) root
root => NS ns root
syn <- get Syn
put Syn (record { defDocstrings $= addName n' doc,
saveDocstrings $= insert n' () } syn)
put Syn ({ defDocstrings $= addName n' doc,
saveDocstrings $= insert n' () } syn)
prettyTerm : IPTerm -> Doc IdrisDocAnn
prettyTerm = reAnnotate Syntax . Idris.Pretty.prettyTerm

View File

@ -138,14 +138,14 @@ stMain cgs opts
| True => pure ()
defs <- initDefs
let updated = foldl (\o, (s, _) => addCG (s, Other s) o) (options defs) cgs
c <- newRef Ctxt (record { options = updated } defs)
c <- newRef Ctxt ({ options := updated } defs)
s <- newRef Syn initSyntax
setCG {c} $ maybe Chez (Other . fst) (head' cgs)
addPrimitives
setWorkingDir "."
when (ignoreMissingIpkg opts) $
setSession (record { ignoreMissingPkg = True } !getSession)
setSession ({ ignoreMissingPkg := True } !getSession)
let ide = ideMode opts
let ideSocket = ideModeSocket opts

View File

@ -246,7 +246,7 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i
-- Update nested names so we elaborate the body in the right
-- environment
names' <- traverse applyEnv (impName :: mtops)
let nest' = record { names $= (names' ++) } nest
let nest' = { names $= (names' ++) } nest
traverse_ (processDecl [] nest' env) [impFn]
unsetFlag vfc impName BlockedHint

View File

@ -299,8 +299,8 @@ updateIfaceSyn iname cn impps ps cs ms ds
= do syn <- get Syn
ms' <- traverse totMeth ms
let info = MkIFaceInfo cn impps ps cs ms' ds
put Syn (record { ifaces $= addName iname info,
saveIFaces $= (iname :: ) } syn)
put Syn ({ ifaces $= addName iname info,
saveIFaces $= (iname :: ) } syn)
where
findSetTotal : List FnOpt -> Maybe TotalReq
findSetTotal [] = Nothing
@ -355,7 +355,7 @@ elabInterface {vars} ifc vis env nest constraints iname params dets mcon body
ds <- traverse (elabDefault meth_decls) defaults
ns_meths <- traverse (\mt => do n <- inCurrentNS mt.name
pure (record { name = n } mt)) meth_decls
pure ({ name := n } mt)) meth_decls
defs <- get Ctxt
Just ty <- lookupTyExact ns_iname (gamma defs)
| Nothing => undefinedName ifc iname

View File

@ -310,7 +310,7 @@ perror (AmbiguousName fc ns)
<+> line <+> !(ploc fc)
perror (AmbiguousElab fc env ts_in)
= do pp <- getPPrint
setPPrint (record { fullNamespace = True } pp)
setPPrint ({ fullNamespace := True } pp)
ts_show <- traverse (\ (gam, t) =>
do defs <- get Ctxt
setCtxt gam

View File

@ -100,7 +100,7 @@ extractHoleData defs env fn (S args) (Bind fc x b sc)
log "idemode.hole" 10 $ "Showing name: " ++ show x
ity <- resugar env !(normalise defs env (binderType b))
let premise = MkHolePremise x ity (multiplicity b) (isImplicit b)
pure $ record { context $= (premise ::) } rest
pure $ { context $= (premise ::) } rest
extractHoleData defs env fn args ty
= do nty <- normalise defs env ty
ity <- resugar env nty
@ -123,7 +123,7 @@ holeData gam env fn args ty
pp <- getPPrint
pure $ if showImplicits pp
then hdata
else record { context $= dropShadows } hdata
else { context $= dropShadows } hdata
where
dropShadows : List HolePremise -> List HolePremise
dropShadows [] = []

View File

@ -160,13 +160,13 @@ field fname
mkBound : List Bound -> PkgVersionBounds -> EmptyRule PkgVersionBounds
mkBound (LT b i :: bs) pkgbs
= maybe (mkBound bs (record { upperBound = Just b,
upperInclusive = i } pkgbs))
= maybe (mkBound bs ({ upperBound := Just b,
upperInclusive := i } pkgbs))
(\_ => fail "Dependency already has an upper bound")
pkgbs.upperBound
mkBound (GT b i :: bs) pkgbs
= maybe (mkBound bs (record { lowerBound = Just b,
lowerInclusive = i } pkgbs))
= maybe (mkBound bs ({ lowerBound := Just b,
lowerInclusive := i } pkgbs))
(\_ => fail "Dependency already has a lower bound")
pkgbs.lowerBound
mkBound [] pkgbs = pure pkgbs
@ -204,36 +204,36 @@ addField : {auto c : Ref Ctxt Defs} ->
{auto p : Ref ParsedMods (List (FC, ModuleIdent))} ->
{auto m : Ref MainMod (Maybe (FC, ModuleIdent))} ->
DescField -> PkgDesc -> Core PkgDesc
addField (PVersion fc n) pkg = pure $ record { version = Just n } pkg
addField (PVersion fc n) pkg = pure $ { version := Just n } pkg
addField (PVersionDep fc n) pkg
= do emitWarning (Deprecated "version numbers must now be of the form x.y.z" Nothing)
pure pkg
addField (PAuthors fc a) pkg = pure $ record { authors = Just a } pkg
addField (PMaintainers fc a) pkg = pure $ record { maintainers = Just a } pkg
addField (PLicense fc a) pkg = pure $ record { license = Just a } pkg
addField (PBrief fc a) pkg = pure $ record { brief = Just a } pkg
addField (PReadMe fc a) pkg = pure $ record { readme = Just a } pkg
addField (PHomePage fc a) pkg = pure $ record { homepage = Just a } pkg
addField (PSourceLoc fc a) pkg = pure $ record { sourceloc = Just a } pkg
addField (PBugTracker fc a) pkg = pure $ record { bugtracker = Just a } pkg
addField (PDepends ds) pkg = pure $ record { depends = ds } pkg
addField (PAuthors fc a) pkg = pure $ { authors := Just a } pkg
addField (PMaintainers fc a) pkg = pure $ { maintainers := Just a } pkg
addField (PLicense fc a) pkg = pure $ { license := Just a } pkg
addField (PBrief fc a) pkg = pure $ { brief := Just a } pkg
addField (PReadMe fc a) pkg = pure $ { readme := Just a } pkg
addField (PHomePage fc a) pkg = pure $ { homepage := Just a } pkg
addField (PSourceLoc fc a) pkg = pure $ { sourceloc := Just a } pkg
addField (PBugTracker fc a) pkg = pure $ { bugtracker := Just a } pkg
addField (PDepends ds) pkg = pure $ { depends := ds } pkg
-- we can't resolve source files for modules without knowing the source directory,
-- so we save them for the second pass
addField (PModules ms) pkg = do put ParsedMods ms
pure pkg
addField (PMainMod loc n) pkg = do put MainMod (Just (loc, n))
pure pkg
addField (PExec e) pkg = pure $ record { executable = Just e } pkg
addField (POpts fc e) pkg = pure $ record { options = Just (fc, e) } pkg
addField (PSourceDir fc a) pkg = pure $ record { sourcedir = Just a } pkg
addField (PBuildDir fc a) pkg = pure $ record { builddir = Just a } pkg
addField (POutputDir fc a) pkg = pure $ record { outputdir = Just a } pkg
addField (PPrebuild fc e) pkg = pure $ record { prebuild = Just (fc, e) } pkg
addField (PPostbuild fc e) pkg = pure $ record { postbuild = Just (fc, e) } pkg
addField (PPreinstall fc e) pkg = pure $ record { preinstall = Just (fc, e) } pkg
addField (PPostinstall fc e) pkg = pure $ record { postinstall = Just (fc, e) } pkg
addField (PPreclean fc e) pkg = pure $ record { preclean = Just (fc, e) } pkg
addField (PPostclean fc e) pkg = pure $ record { postclean = Just (fc, e) } pkg
addField (PExec e) pkg = pure $ { executable := Just e } pkg
addField (POpts fc e) pkg = pure $ { options := Just (fc, e) } pkg
addField (PSourceDir fc a) pkg = pure $ { sourcedir := Just a } pkg
addField (PBuildDir fc a) pkg = pure $ { builddir := Just a } pkg
addField (POutputDir fc a) pkg = pure $ { outputdir := Just a } pkg
addField (PPrebuild fc e) pkg = pure $ { prebuild := Just (fc, e) } pkg
addField (PPostbuild fc e) pkg = pure $ { postbuild := Just (fc, e) } pkg
addField (PPreinstall fc e) pkg = pure $ { preinstall := Just (fc, e) } pkg
addField (PPostinstall fc e) pkg = pure $ { postinstall := Just (fc, e) } pkg
addField (PPreclean fc e) pkg = pure $ { preclean := Just (fc, e) } pkg
addField (PPostclean fc e) pkg = pure $ { postclean := Just (fc, e) } pkg
addFields : {auto c : Ref Ctxt Defs} ->
{auto s : Ref Syn SyntaxInfo} ->
@ -245,9 +245,9 @@ addFields xs desc = do p <- newRef ParsedMods []
setSourceDir (sourcedir added)
ms <- get ParsedMods
mmod <- get MainMod
pure $ record { modules = !(traverse toSource ms)
, mainmod = !(traverseOpt toSource mmod)
} added
pure $ { modules := !(traverse toSource ms)
, mainmod := !(traverseOpt toSource mmod)
} added
where
toSource : (FC, ModuleIdent) -> Core (ModuleIdent, String)
toSource (loc, ns) = pure (ns, !(nsToSource loc ns))
@ -769,10 +769,10 @@ partitionOpts opts = foldr pOptUpdate (MkPFR [] [] False) opts
pOptUpdate : CLOpt -> (PackageOpts -> PackageOpts)
pOptUpdate opt with (optType opt)
pOptUpdate opt | (PPackage cmd f) = record {pkgDetails $= ((cmd, f)::)}
pOptUpdate opt | POpt = record {oopts $= (opt::)}
pOptUpdate opt | (PPackage cmd f) = {pkgDetails $= ((cmd, f)::)}
pOptUpdate opt | POpt = {oopts $= (opt::)}
pOptUpdate opt | PIgnore = id
pOptUpdate opt | PErr = record {hasError = True}
pOptUpdate opt | PErr = {hasError := True}
errorMsg : String
errorMsg = unlines
@ -835,7 +835,7 @@ findIpkg fname
do let src' = up </> srcpath
setSource src'
opts <- get ROpts
put ROpts (record { mainfile = Just src' } opts)
put ROpts ({ mainfile := Just src' } opts)
pure (Just src')
where
dropHead : String -> List String -> List String

View File

@ -87,10 +87,10 @@ record ParseOpts where
withOK : Bool -- = with applications are parseable
peq : ParseOpts -> ParseOpts
peq = record { eqOK = True }
peq = { eqOK := True }
pnoeq : ParseOpts -> ParseOpts
pnoeq = record { eqOK = False }
pnoeq = { eqOK := False }
export
pdef : ParseOpts
@ -223,7 +223,7 @@ mutual
<|> if withOK q
then do continue indents
decoratedSymbol fname "|"
arg <- expr (record {withOK = False} q) fname indents
arg <- expr ({withOK := False} q) fname indents
pure [WithArg arg]
else fail "| not allowed here"
where
@ -793,15 +793,28 @@ mutual
record_ : OriginDesc -> IndentInfo -> Rule PTerm
record_ fname indents
= do b <- bounds (do kw <- option False
(decoratedKeyword fname "record"
$> True) -- TODO deprecated
decoratedSymbol fname "{"
commit
fs <- sepBy1 (decoratedSymbol fname ",") (field kw fname indents)
decoratedSymbol fname "}"
pure $ forget fs)
= do
b <- (
withWarning oldSyntaxWarning (
bounds (do
decoratedKeyword fname "record"
commit
body True
))
<|>
bounds (body False))
pure (PUpdate (boundToFC fname b) b.val)
where
oldSyntaxWarning : String
oldSyntaxWarning = "DEPRECATED: old record update syntax. Use \"{ f := v } p\" instead of \"record { f = v } p\""
body : Bool -> Rule (List PFieldUpdate)
body kw = do
decoratedSymbol fname "{"
commit
fs <- sepBy1 (decoratedSymbol fname ",") (field kw fname indents)
decoratedSymbol fname "}"
pure $ forget fs
field : Bool -> OriginDesc -> IndentInfo -> Rule PFieldUpdate
field kw fname indents

View File

@ -186,7 +186,7 @@ readAsMain fname
-- TODO: Maybe we should record this per namespace, since this is
-- a little bit of a hack? Or maybe that will have too much overhead.
ust <- get UST
put UST (record { nextName = nextName ustm } ust)
put UST ({ nextName := nextName ustm } ust)
setNS replNS
setNestedNS replNestedNS
@ -370,7 +370,7 @@ processMod sourceFileName ttcFileName msg sourcecode origin
-- If they haven't changed next time, and the source
-- file hasn't changed, no need to rebuild.
defs <- get Ctxt
put Ctxt (record { importHashes = importInterfaceHashes } defs)
put Ctxt ({ importHashes := importInterfaceHashes } defs)
pure (Just errs))
(\err => pure (Just [err]))

View File

@ -165,19 +165,19 @@ setOpt : {auto c : Ref Ctxt Defs} ->
REPLOpt -> Core ()
setOpt (ShowImplicits t)
= do pp <- getPPrint
setPPrint (record { showImplicits = t } pp)
setPPrint ({ showImplicits := t } pp)
setOpt (ShowNamespace t)
= do pp <- getPPrint
setPPrint (record { fullNamespace = t } pp)
setPPrint ({ fullNamespace := t } pp)
setOpt (ShowTypes t)
= do opts <- get ROpts
put ROpts (record { showTypes = t } opts)
put ROpts ({ showTypes := t } opts)
setOpt (EvalMode m)
= do opts <- get ROpts
put ROpts (record { evalMode = m } opts)
put ROpts ({ evalMode := m } opts)
setOpt (Editor e)
= do opts <- get ROpts
put ROpts (record { editor = e } opts)
put ROpts ({ editor := e } opts)
setOpt (CG e)
= do defs <- get Ctxt
case getCG (options defs) e of
@ -185,7 +185,7 @@ setOpt (CG e)
Nothing => iputStrLn (reflow "No such code generator available")
setOpt (Profile t)
= do pp <- getSession
setSession (record { profile = t } pp)
setSession ({ profile := t } pp)
setOpt (EvalTiming t)
= setEvalTiming t
@ -310,9 +310,9 @@ nextProofSearch
| Nothing => pure Nothing
Just (res, next) <- nextResult res
| Nothing =>
do put ROpts (record { psResult = Nothing } opts)
do put ROpts ({ psResult := Nothing } opts)
pure Nothing
put ROpts (record { psResult = Just (n, next) } opts)
put ROpts ({ psResult := Just (n, next) } opts)
pure (Just (n, res))
nextGenDef : {auto c : Ref Ctxt Defs} ->
@ -326,9 +326,9 @@ nextGenDef reject
| Nothing => pure Nothing
Just (res, next) <- nextResult res
| Nothing =>
do put ROpts (record { gdResult = Nothing } opts)
do put ROpts ({ gdResult := Nothing } opts)
pure Nothing
put ROpts (record { gdResult = Just (line, next) } opts)
put ROpts ({ gdResult := Just (line, next) } opts)
case reject of
Z => pure (Just (line, res))
S k => nextGenDef k
@ -449,7 +449,7 @@ processEdit (ExprSearch upd line name hints)
[(n, nidx, Hole locs _)] =>
do let searchtm = exprSearch replFC name hints
ropts <- get ROpts
put ROpts (record { psResult = Just (name, searchtm) } ropts)
put ROpts ({ psResult := Just (name, searchtm) } ropts)
defs <- get Ctxt
Just (_, restm) <- nextProofSearch
| Nothing => pure $ EditError "No search results"
@ -493,7 +493,7 @@ processEdit (GenerateDef upd line name rej)
do let searchdef = makeDefSort (\p, n => onLine (line - 1) p)
16 mostUsed n'
ropts <- get ROpts
put ROpts (record { gdResult = Just (line, searchdef) } ropts)
put ROpts ({ gdResult := Just (line, searchdef) } ropts)
Just (_, (fc, cs)) <- nextGenDef rej
| Nothing => pure (EditError "No search results")
@ -667,7 +667,7 @@ loadMainFile : {auto c : Ref Ctxt Defs} ->
String -> Core REPLResult
loadMainFile f
= do opts <- get ROpts
put ROpts (record { evalResultName = Nothing } opts)
put ROpts ({ evalResultName := Nothing } opts)
modIdent <- ctxtPathToNS f
resetContext (PhysicalIdrSrc modIdent)
Right res <- coreLift (readFile f)
@ -687,7 +687,7 @@ loadMainFile f
replEval : {auto c : Ref Ctxt Defs} ->
{vs : _} ->
REPLEval -> Defs -> Env Term vs -> Term vs -> Core (Term vs)
replEval NormaliseAll = normaliseOpts (record { strategy = CBV } withAll)
replEval NormaliseAll = normaliseOpts ({ strategy := CBV } withAll)
replEval _ = normalise
record TermWithType where
@ -780,7 +780,7 @@ process (Eval itm)
$ newDef replFC evalResultName top [] ty Private
$ PMDef defaultPI [] (STerm 0 ntm) (STerm 0 ntm) []
addToSave evalResultName
put ROpts (record { evalResultName = Just evalResultName } opts)
put ROpts ({ evalResultName := Just evalResultName } opts)
if showTypes opts
then do ity <- resugar [] !(norm defs [] ty)
pure (Evaluated itm (Just ity))
@ -822,7 +822,7 @@ process Reload
Just f => loadMainFile f
process (Load f)
= do opts <- get ROpts
put ROpts (record { mainfile = Just f } opts)
put ROpts ({ mainfile := Just f } opts)
-- Clear the context and load again
loadMainFile f
process (ImportMod m)
@ -930,12 +930,12 @@ process (Editing cmd)
-- Since we're working in a local environment, don't do the usual
-- thing of printing out the full environment for parameterised
-- calls or calls in where blocks
setPPrint (record { showFullEnv = False } ppopts)
setPPrint ({ showFullEnv := False } ppopts)
res <- processEdit cmd
setPPrint ppopts
pure $ Edited res
process (CGDirective str)
= do setSession (record { directives $= (str::) } !getSession)
= do setSession ({ directives $= (str::) } !getSession)
pure Done
process (RunShellCommand cmd)
= do coreLift_ (system cmd)

View File

@ -184,10 +184,10 @@ updateErrorLine : {auto o : Ref ROpts REPLOpts} ->
List Error -> Core ()
updateErrorLine []
= do opts <- get ROpts
put ROpts (record { errorLine = Nothing } opts)
put ROpts ({ errorLine := Nothing } opts)
updateErrorLine (e :: _)
= do opts <- get ROpts
put ROpts (record { errorLine = (getErrorLoc e) >>= getFCLine } opts)
put ROpts ({ errorLine := (getErrorLoc e) >>= getFCLine } opts)
export
resetContext : {auto c : Ref Ctxt Defs} ->
@ -198,7 +198,7 @@ resetContext : {auto c : Ref Ctxt Defs} ->
Core ()
resetContext origin
= do defs <- get Ctxt
put Ctxt (record { options = clearNames (options defs) } !initDefs)
put Ctxt ({ options := clearNames (options defs) } !initDefs)
addPrimitives
put UST initUState
put Syn initSyntax

View File

@ -104,7 +104,7 @@ setOutput : {auto o : Ref ROpts REPLOpts} ->
OutputMode -> Core ()
setOutput m
= do opts <- get ROpts
put ROpts (record { idemode = m } opts)
put ROpts ({ idemode := m } opts)
export
getOutput : {auto o : Ref ROpts REPLOpts} -> Core OutputMode
@ -116,23 +116,23 @@ setMainFile : {auto o : Ref ROpts REPLOpts} ->
Maybe String -> Core ()
setMainFile src
= do opts <- get ROpts
put ROpts (record { mainfile = src,
literateStyle = litStyle src } opts)
put ROpts ({ mainfile := src,
literateStyle := litStyle src } opts)
export
resetProofState : {auto o : Ref ROpts REPLOpts} ->
Core ()
resetProofState
= do opts <- get ROpts
put ROpts (record { psResult = Nothing,
gdResult = Nothing } opts)
put ROpts ({ psResult := Nothing,
gdResult := Nothing } opts)
export
setSource : {auto o : Ref ROpts REPLOpts} ->
String -> Core ()
setSource src
= do opts <- get ROpts
put ROpts (record { source = src } opts)
put ROpts ({ source := src } opts)
export
getSource : {auto o : Ref ROpts REPLOpts} ->
@ -160,7 +160,7 @@ setCurrentElabSource : {auto o : Ref ROpts REPLOpts} ->
String -> Core ()
setCurrentElabSource src
= do opts <- get ROpts
put ROpts (record { currentElabSource = src } opts)
put ROpts ({ currentElabSource := src } opts)
export
getCurrentElabSource : {auto o : Ref ROpts REPLOpts} ->
@ -172,7 +172,7 @@ getCurrentElabSource
addCodegen : {auto o : Ref ROpts REPLOpts} ->
String -> Codegen -> Core ()
addCodegen s cg = do opts <- get ROpts
put ROpts (record { extraCodegens $= ((s,cg)::) } opts)
put ROpts ({ extraCodegens $= ((s,cg)::) } opts)
export
getCodegen : {auto o : Ref ROpts REPLOpts} ->
@ -188,7 +188,7 @@ getConsoleWidth = do opts <- get ROpts
export
setConsoleWidth : {auto o : Ref ROpts REPLOpts} -> Maybe Nat -> Core ()
setConsoleWidth n = do opts <- get ROpts
put ROpts (record { consoleWidth = n } opts)
put ROpts ({ consoleWidth := n } opts)
export
getColor : {auto o : Ref ROpts REPLOpts} -> Core Bool
@ -198,7 +198,7 @@ getColor = do opts <- get ROpts
export
setColor : {auto o : Ref ROpts REPLOpts} -> Bool -> Core ()
setColor b = do opts <- get ROpts
put ROpts (record { color = b } opts)
put ROpts ({ color := b } opts)
export
getSynHighlightOn : {auto o : Ref ROpts REPLOpts} -> Core Bool
@ -208,7 +208,7 @@ getSynHighlightOn = do opts <- get ROpts
export
setSynHighlightOn : {auto o : Ref ROpts REPLOpts} -> Bool -> Core ()
setSynHighlightOn b = do opts <- get ROpts
put ROpts (record { synHighlightOn = b } opts)
put ROpts ({ synHighlightOn := b } opts)
export
getEvalTiming : {auto o : Ref ROpts REPLOpts} -> Core Bool
@ -220,4 +220,4 @@ export
setEvalTiming : {auto o : Ref ROpts REPLOpts} -> Bool -> Core ()
setEvalTiming b
= do opts <- get ROpts
put ROpts (record { evalTiming = b } opts)
put ROpts ({ evalTiming := b } opts)

View File

@ -274,7 +274,7 @@ setIncrementalCG failOnError cgn
then do coreLift $ putStrLn $ cgn ++ " does not support incremental builds"
coreLift $ exitWith (ExitFailure 1)
else pure ()
setSession (record { incrementalCGs $= (cg :: )} !getSession)
setSession ({ incrementalCGs $= (cg :: )} !getSession)
Nothing =>
if failOnError
then do coreLift $ putStrLn "No such code generator"
@ -290,32 +290,32 @@ preOptions : {auto c : Ref Ctxt Defs} ->
List CLOpt -> Core Bool
preOptions [] = pure True
preOptions (NoBanner :: opts)
= do setSession (record { nobanner = True } !getSession)
= do setSession ({ nobanner := True } !getSession)
preOptions opts
-- These things are processed later, but imply nobanner too
preOptions (OutputFile _ :: opts)
= do setSession (record { nobanner = True } !getSession)
= do setSession ({ nobanner := True } !getSession)
preOptions opts
preOptions (ExecFn _ :: opts)
= do setSession (record { nobanner = True } !getSession)
= do setSession ({ nobanner := True } !getSession)
preOptions opts
preOptions (IdeMode :: opts)
= do setSession (record { nobanner = True } !getSession)
= do setSession ({ nobanner := True } !getSession)
preOptions opts
preOptions (IdeModeSocket _ :: opts)
= do setSession (record { nobanner = True } !getSession)
= do setSession ({ nobanner := True } !getSession)
preOptions opts
preOptions (CheckOnly :: opts)
= do setSession (record { nobanner = True } !getSession)
= do setSession ({ nobanner := True } !getSession)
preOptions opts
preOptions (Profile :: opts)
= do setSession (record { profile = True } !getSession)
= do setSession ({ profile := True } !getSession)
preOptions opts
preOptions (Quiet :: opts)
= do setOutput (REPL VerbosityLvl.ErrorLvl)
preOptions opts
preOptions (NoPrelude :: opts)
= do setSession (record { noprelude = True } !getSession)
= do setSession ({ noprelude := True } !getSession)
preOptions opts
preOptions (SetCG e :: opts)
= do defs <- get Ctxt
@ -328,7 +328,7 @@ preOptions (SetCG e :: opts)
showSep ", " (map fst (availableCGs (options defs)))
coreLift $ exitWith (ExitFailure 1)
preOptions (Directive d :: opts)
= do setSession (record { directives $= (d::) } !getSession)
= do setSession ({ directives $= (d::) } !getSession)
preOptions opts
preOptions (PkgPath p :: opts)
= do addPkgDir p anyBounds
@ -356,33 +356,33 @@ preOptions (DebugElabCheck :: opts)
= do setDebugElabCheck True
preOptions opts
preOptions (AltErrorCount c :: opts)
= do setSession (record { logErrorCount = c } !getSession)
= do setSession ({ logErrorCount := c } !getSession)
preOptions opts
preOptions (RunREPL _ :: opts)
= do setOutput (REPL VerbosityLvl.ErrorLvl)
setSession (record { nobanner = True } !getSession)
setSession ({ nobanner := True } !getSession)
preOptions opts
preOptions (FindIPKG :: opts)
= do setSession (record { findipkg = True } !getSession)
= do setSession ({ findipkg := True } !getSession)
preOptions opts
preOptions (IgnoreMissingIPKG :: opts)
= do setSession (record { ignoreMissingPkg = True } !getSession)
= do setSession ({ ignoreMissingPkg := True } !getSession)
preOptions opts
preOptions (DumpCases f :: opts)
= do setSession (record { dumpcases = Just f } !getSession)
= do setSession ({ dumpcases := Just f } !getSession)
preOptions opts
preOptions (DumpLifted f :: opts)
= do setSession (record { dumplifted = Just f } !getSession)
= do setSession ({ dumplifted := Just f } !getSession)
preOptions opts
preOptions (DumpANF f :: opts)
= do setSession (record { dumpanf = Just f } !getSession)
= do setSession ({ dumpanf := Just f } !getSession)
preOptions opts
preOptions (DumpVMCode f :: opts)
= do setSession (record { dumpvmcode = Just f } !getSession)
= do setSession ({ dumpvmcode := Just f } !getSession)
preOptions opts
preOptions (Logging n :: opts)
= do setSession (record { logEnabled = True,
logLevel $= insertLogLevel n } !getSession)
= do setSession ({ logEnabled := True,
logLevel $= insertLogLevel n } !getSession)
preOptions opts
preOptions (ConsoleWidth n :: opts)
= do setConsoleWidth n
@ -391,24 +391,24 @@ preOptions (Color b :: opts)
= do setColor b
preOptions opts
preOptions (WarningsAsErrors :: opts)
= do updateSession (record { warningsAsErrors = True })
= do updateSession ({ warningsAsErrors := True })
preOptions opts
preOptions (IgnoreShadowingWarnings :: opts)
= do updateSession (record { showShadowingWarning = False })
= do updateSession ({ showShadowingWarning := False })
preOptions opts
preOptions (HashesInsteadOfModTime :: opts)
= do throw (InternalError "-Xcheck-hashes disabled (see issue #1935)")
updateSession (record { checkHashesInsteadOfModTime = True })
updateSession ({ checkHashesInsteadOfModTime := True })
preOptions opts
preOptions (CaseTreeHeuristics :: opts)
= do updateSession (record { caseTreeHeuristics = True })
= do updateSession ({ caseTreeHeuristics := True })
preOptions opts
preOptions (IncrementalCG e :: opts)
= do defs <- get Ctxt
setIncrementalCG True e
preOptions opts
preOptions (WholeProgram :: opts)
= do updateSession (record { wholeProgram = True })
= do updateSession ({ wholeProgram := True })
preOptions opts
preOptions (BashCompletion a b :: _)
= do os <- opts a b
@ -418,7 +418,7 @@ preOptions (BashCompletionScript fun :: _)
= do coreLift $ putStrLn $ completionScript fun
pure False
preOptions (Total :: opts)
= do updateSession (record { totalReq = Total })
= do updateSession ({ totalReq := Total })
preOptions opts
preOptions (_ :: opts) = preOptions opts

View File

@ -40,7 +40,7 @@ tarjan {cuid} deps = loop initial (SortedMap.keys deps)
Nothing => ts' -- no edges
Just edgeSet => loop ts' (SortedSet.toList edgeSet)
in case SortedMap.lookup v ts''.vertices of
Nothing => record { impossibleHappened = True } ts''
Nothing => { impossibleHappened := True } ts''
Just vtv =>
if vtv.index == vtv.lowlink
then createComponent ts'' v []
@ -49,14 +49,14 @@ tarjan {cuid} deps = loop initial (SortedMap.keys deps)
createComponent : TarjanState cuid -> cuid -> List cuid -> TarjanState cuid
createComponent ts v acc =
case ts.stack of
[] => record { impossibleHappened = True } ts
[] => { impossibleHappened := True } ts
w :: ws =>
let ts' : TarjanState cuid = record {
vertices $= SortedMap.adjust w record{ inStack = False },
stack = ws
let ts' : TarjanState cuid = {
vertices $= SortedMap.adjust w { inStack := False },
stack := ws
} ts
in if w == v
then record { components $= ((v ::: acc) ::) } ts' -- that's it
then { components $= ((v ::: acc) ::) } ts' -- that's it
else createComponent ts' v (w :: acc)
loop : TarjanState cuid -> List cuid -> TarjanState cuid
@ -66,16 +66,16 @@ tarjan {cuid} deps = loop initial (SortedMap.keys deps)
case SortedMap.lookup w ts.vertices of
Nothing => let ts' = strongConnect ts w in
case SortedMap.lookup w ts'.vertices of
Nothing => record { impossibleHappened = True } ts'
Just wtv => record { vertices $= SortedMap.adjust v record{ lowlink $= min wtv.lowlink } } ts'
Nothing => { impossibleHappened := True } ts'
Just wtv => { vertices $= SortedMap.adjust v { lowlink $= min wtv.lowlink } } ts'
Just wtv => case wtv.inStack of
False => ts -- nothing to do
True => record { vertices $= SortedMap.adjust v record{ lowlink $= min wtv.index } } ts
True => { vertices $= SortedMap.adjust v { lowlink $= min wtv.index } } ts
) ws
ts' : TarjanState cuid
ts' = record {
ts' = {
vertices $= SortedMap.insert v (TV ts.nextIndex ts.nextIndex True),
stack $= (v ::),
nextIndex $= (1+)

View File

@ -32,4 +32,4 @@ reflow = fillSep . words
||| Renders a document with a certain width.
export
putDocW : Nat -> Doc ann -> IO ()
putDocW w = renderIO . layoutPretty (record { layoutPageWidth = AvailablePerLine (cast w) 1 } defaultLayoutOptions)
putDocW w = renderIO . layoutPretty ({ layoutPageWidth := AvailablePerLine (cast w) 1 } defaultLayoutOptions)

View File

@ -41,7 +41,7 @@ appended i (MkBin b loc s used) = MkBin b (loc+i) s (used + i)
export
incLoc : Int -> Binary -> Binary
incLoc i c = record { loc $= (+i) } c
incLoc i c = { loc $= (+i) } c
export
dumpBin : Binary -> IO ()

View File

@ -298,9 +298,9 @@ append' left right =
if isAbsolute' right || isJust right.volume then
right
else if hasRoot right then
record { volume = left.volume } right
{ volume := left.volume } right
else
record { body = left.body ++ right.body, hasTrailSep = right.hasTrailSep } left
{ body := left.body ++ right.body, hasTrailSep := right.hasTrailSep } left
splitPath' : Path -> List Path
splitPath' path =
@ -325,7 +325,7 @@ splitParent' path =
[] => Nothing
(x::xs) =>
let
parent = record { body = init (x::xs), hasTrailSep = False } path
parent = { body := init (x::xs), hasTrailSep := False } path
child = MkPath Nothing False [last (x::xs)] path.hasTrailSep
in
Just (parent, child)

View File

@ -70,7 +70,7 @@ normaliseHoleTypes
updateType defs i def
= do ty' <- catch (tryNormaliseSizeLimit defs 10 [] (type def))
(\err => normaliseHoles defs [] (type def))
ignore $ addDef (Resolved i) (record { type = ty' } def)
ignore $ addDef (Resolved i) ({ type := ty' } def)
normaliseH : Defs -> Int -> Core ()
normaliseH defs i
@ -113,7 +113,7 @@ elabTermSub {vars} defining mode opts nest env env' sub tm ty
else pure empty
ust <- get UST
let olddelayed = delayedElab ust
put UST (record { delayedElab = [] } ust)
put UST ({ delayedElab := [] } ust)
constart <- getNextEntry
defs <- get Ctxt
@ -136,10 +136,10 @@ elabTermSub {vars} defining mode opts nest env env' sub tm ty
(delayedElab ust)))
(\err =>
do ust <- get UST
put UST (record { delayedElab = olddelayed } ust)
put UST ({ delayedElab := olddelayed } ust)
throw err)
ust <- get UST
put UST (record { delayedElab = olddelayed } ust)
put UST ({ delayedElab := olddelayed } ust)
solveConstraintsAfter constart solvemode MatchArgs
-- As long as we're not in the RHS of a case block,

View File

@ -329,7 +329,7 @@ addAmbig : List alts -> Maybe Name -> ElabInfo -> ElabInfo
addAmbig _ Nothing = id
addAmbig [] _ = id
addAmbig [_] _ = id
addAmbig _ (Just n) = record { ambigTries $= (n ::) }
addAmbig _ (Just n) = { ambigTries $= (n ::) }
export
checkAlternative : {vars : _} ->

View File

@ -59,7 +59,7 @@ getNameType elabMode rigc env fc x
when (isLinear rigb) $
do est <- get EST
put EST
(record { linearUsed $= ((MkVar lv) :: ) } est)
({ linearUsed $= ((MkVar lv) :: ) } est)
log "ide-mode.highlight" 8
$ "getNameType is trying to add Bound: "
++ show x ++ " (" ++ show fc ++ ")"
@ -482,7 +482,7 @@ mutual
| Nothing => pure ()
when (isErased (multiplicity gdef)) $ addNoSolve i
_ => pure ()
res <- check argRig (record { topLevel = False } elabinfo) nest env arg (Just $ glueBack defs env aty')
res <- check argRig ({ topLevel := False } elabinfo) nest env arg (Just $ glueBack defs env aty')
when (onLHS (elabMode elabinfo)) $
case aty' of
NApp _ (NMeta _ i _) _ => removeNoSolve i
@ -533,7 +533,7 @@ mutual
(\t => pure (Just !(toFullNames!(getTerm t))))
expty
pure ("Overall expected type: " ++ show ety))
res <- check argRig (record { topLevel = False } elabinfo)
res <- check argRig ({ topLevel := False } elabinfo)
nest env arg (Just (glueClosure defs env aty))
(argv, argt) <-
if not (onLHS (elabMode elabinfo))
@ -852,7 +852,7 @@ checkApp rig elabinfo nest env fc (IVar fc' n) expargs autoargs namedargs exp
-- the result
updateElabInfo prims (InLHS _) n [IPrimVal fc c] elabinfo =
do if isPrimName prims !(getFullName n)
then pure (record { elabMode = InExpr } elabinfo)
then pure ({ elabMode := InExpr } elabinfo)
else pure elabinfo
updateElabInfo _ _ _ _ info = pure info

View File

@ -49,9 +49,9 @@ checkAs rig elabinfo nest env fc nameFC side n_in pat topexp
defs <- get Ctxt
est <- get EST
put EST
(record { boundNames $= ((n, AsBinding rigAs Explicit tm exp pattm) :: ),
toBind $= ((n, AsBinding rigAs Explicit tm bty pattm) ::) }
est)
({ boundNames $= ((n, AsBinding rigAs Explicit tm exp pattm) :: ),
toBind $= ((n, AsBinding rigAs Explicit tm bty pattm) ::) }
est)
(ntm, nty) <- checkExp rig elabinfo env nameFC tm (gnf env exp)
(Just patty)

View File

@ -23,7 +23,7 @@ import TTImp.TTImp
-- bound name.
export
dropName : Name -> NestedNames vars -> NestedNames vars
dropName n nest = record { names $= drop } nest
dropName n nest = { names $= drop } nest
where
drop : List (Name, a) -> List (Name, a)
drop [] = []
@ -224,7 +224,7 @@ checkLet rigc_in elabinfo nest env fc lhsFC rigl n nTy nVal scope expty {vars}
-- Also elaborate any case blocks in the value via runDelays
(valv, valt, rigb) <- handle
(do c <- runDelays (==CaseBlock) $ check (rigl |*| rigc)
(record { preciseInf = True } elabinfo)
({ preciseInf := True } elabinfo)
nest env nVal (Just (gnf env tyv))
pure (fst c, snd c, rigl |*| rigc))
(\err => case linearErr err of

View File

@ -229,7 +229,7 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
-- actually bound! This is rather hacky, but a lot less fiddly than
-- the alternative of fixing up the environment
when (not (isNil fullImps)) $ findImpsIn fc [] [] casefnty
cidx <- addDef casen (record { eraseArgs = erasedargs }
cidx <- addDef casen ({ eraseArgs := erasedargs }
(newDef fc casen (if isErased rigc then erased else top)
[] casefnty vis None))
@ -261,7 +261,7 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
-- case block, because they're not going to make progress until
-- we come out again, so save them
let olddelayed = delayedElab ust
put UST (record { delayedElab = [] } ust)
put UST ({ delayedElab := [] } ust)
processDecl [InCase] nest' [] (IDef fc casen alts')
-- If there's no duplication of the scrutinee in the block,
@ -275,7 +275,7 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
when inlineOK $ setFlag fc casen Inline
ust <- get UST
put UST (record { delayedElab = olddelayed } ust)
put UST ({ delayedElab := olddelayed } ust)
pure (appTm, gnf env caseretty)
where

View File

@ -184,7 +184,7 @@ saveHole : {auto e : Ref EST (EState vars)} ->
Name -> Core ()
saveHole n
= do est <- get EST
put EST (record { saveHoles $= insert n () } est)
put EST ({ saveHoles $= insert n () } est)
weakenedEState : {n, vars : _} ->
{auto e : Ref EST (EState vars)} ->
@ -192,12 +192,12 @@ weakenedEState : {n, vars : _} ->
weakenedEState {e}
= do est <- get EST
eref <- newRef EST $
record { subEnv $= DropCons
, boundNames $= map wknTms
, toBind $= map wknTms
, linearUsed $= map weaken
, polyMetavars = [] -- no binders on LHS
} est
{ subEnv $= DropCons
, boundNames $= map wknTms
, toBind $= map wknTms
, linearUsed $= map weaken
, polyMetavars := [] -- no binders on LHS
} est
pure eref
where
wknTms : (Name, ImplBinding vs) ->
@ -218,12 +218,12 @@ strengthenedEState {n} {vars} c e fc env
svs <- dropSub (subEnv est)
bns <- traverse (strTms defs) (boundNames est)
todo <- traverse (strTms defs) (toBind est)
pure $ record { subEnv = svs
, boundNames = bns
, toBind = todo
, linearUsed $= mapMaybe dropTop
, polyMetavars = [] -- no binders on LHS
} est
pure $ { subEnv := svs
, boundNames := bns
, toBind := todo
, linearUsed $= mapMaybe dropTop
, polyMetavars := [] -- no binders on LHS
} est
where
dropSub : SubVars xs (y :: ys) -> Core (SubVars xs ys)
@ -304,7 +304,7 @@ mustBePoly : {auto e : Ref EST (EState vars)} ->
Term vars -> Term vars -> Core ()
mustBePoly fc env tm ty
= do est <- get EST
put EST (record { polyMetavars $= ((fc, env, tm, ty) :: ) } est)
put EST ({ polyMetavars $= ((fc, env, tm, ty) :: ) } est)
-- Return whether we already know the return type of the given function
-- type. If we know this, we can possibly infer some argument types before
@ -328,10 +328,10 @@ updateEnv : {new : _} ->
Term vars', Term vars', SubVars new vars'))) ->
EState vars -> EState vars
updateEnv env sub bif st
= record { outerEnv = env
, subEnv = sub
, bindIfUnsolved = bif
} st
= { outerEnv := env
, subEnv := sub
, bindIfUnsolved := bif
} st
export
addBindIfUnsolved : {vars : _} ->
@ -339,11 +339,11 @@ addBindIfUnsolved : {vars : _} ->
Env Term vars -> Term vars -> Term vars ->
EState vars -> EState vars
addBindIfUnsolved hn r p env tm ty st
= record { bindIfUnsolved $=
((hn, r, (_ ** (env, p, tm, ty, subEnv st))) ::)} st
= { bindIfUnsolved $=
((hn, r, (_ ** (env, p, tm, ty, subEnv st))) ::)} st
clearBindIfUnsolved : EState vars -> EState vars
clearBindIfUnsolved = record { bindIfUnsolved = [] }
clearBindIfUnsolved = { bindIfUnsolved := [] }
-- Clear the 'toBind' list, except for the names given
export
@ -351,15 +351,15 @@ 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) }
(clearBindIfUnsolved est))
put EST ({ toBind $= filter (\x => fst x `elem` excepts) }
(clearBindIfUnsolved est))
export
noteLHSPatVar : {auto e : Ref EST (EState vars)} ->
ElabMode -> Name -> Core ()
noteLHSPatVar (InLHS _) n
= do est <- get EST
put EST (record { lhsPatVars $= (n ::) } est)
put EST ({ lhsPatVars $= (n ::) } est)
noteLHSPatVar _ _ = pure ()
export
@ -367,7 +367,7 @@ notePatVar : {auto e : Ref EST (EState vars)} ->
Name -> Core ()
notePatVar n
= do est <- get EST
put EST (record { allPatVars $= (n ::) } est)
put EST ({ allPatVars $= (n ::) } est)
export
metaVar : {vars : _} ->
@ -400,7 +400,7 @@ metaVarI fc rig env n ty
tynf <- nf defs env ty
let hinf = case tynf of
NApp _ (NMeta _ _ _) _ =>
record { precisetype = True } (holeInit False)
{ precisetype := True } (holeInit False)
_ => holeInit False
newMeta fc rig env n ty (Hole (length env) hinf) True
@ -513,7 +513,7 @@ tryError elab
put EST est
put MD md
defs' <- get Ctxt
put Ctxt (record { timings = timings defs' } defs)
put Ctxt ({ timings := timings defs' } defs)
pure (Left err))
export

View File

@ -46,10 +46,10 @@ deeper : {auto e : Ref EST (EState vars)} ->
deeper elab
= do est <- get EST
let d = delayDepth est
put EST (record { delayDepth = 1 + d } est)
put EST ({ delayDepth := 1 + d } est)
res <- elab
est <- get EST
put EST (record { delayDepth = d } est)
put EST ({ delayDepth := d } est)
pure res
-- Try the given elaborator; if it fails, and the error matches the
@ -85,16 +85,16 @@ delayOnFailure fc rig env exp pred pri elab
log "elab.delay" 10 ("Due to error " ++ show err)
ust <- get UST
defs <- get Ctxt
put UST (record { delayedElab $=
put UST ({ delayedElab $=
((pri, ci, localHints defs,
mkClosedElab fc env
(deeper
(do ust <- get UST
let nos' = noSolve ust
put UST (record { noSolve = nos } ust)
put UST ({ noSolve := nos } ust)
res <- elab True
ust <- get UST
put UST (record { noSolve = nos' } ust)
put UST ({ noSolve := nos' } ust)
pure res))) :: ) }
ust)
pure (dtm, expected)
@ -130,14 +130,14 @@ delayElab {vars} fc rig env exp pri elab
" for") env expected
ust <- get UST
defs <- get Ctxt
put UST (record { delayedElab $=
put UST ({ delayedElab $=
((pri, ci, localHints defs, mkClosedElab fc env
(do ust <- get UST
let nos' = noSolve ust
put UST (record { noSolve = nos } ust)
put UST ({ noSolve := nos } ust)
res <- elab
ust <- get UST
put UST (record { noSolve = nos' } ust)
put UST ({ noSolve := nos' } ust)
pure res)) :: ) }
ust)
pure (dtm, expected)
@ -210,11 +210,11 @@ recoverable : {auto c : Ref Ctxt Defs} ->
Error -> Core Bool
recoverable (CantConvert _ gam env l r)
= do defs <- get Ctxt
let defs = record { gamma = gam } defs
let defs = { gamma := gam } defs
pure $ not !(contra defs !(nf defs env l) !(nf defs env r))
recoverable (CantSolveEq _ gam env l r)
= do defs <- get Ctxt
let defs = record { gamma = gam } defs
let defs = { gamma := gam } defs
pure $ not !(contra defs !(nf defs env l) !(nf defs env r))
recoverable (UndefinedName _ _) = pure False
recoverable (LinearMisuse _ _ _ _) = pure False
@ -256,9 +256,9 @@ retryDelayed' errmode p acc (d@(_, i, hints, elab) :: ds)
log "elab.retry" 5 (show (delayDepth est) ++ ": Retrying delayed hole " ++ show !(getFullName (Resolved i)))
-- elab itself might have delays internally, so keep track of them
ust <- get UST
put UST (record { delayedElab = [] } ust)
put UST ({ delayedElab := [] } ust)
defs <- get Ctxt
put Ctxt (record { localHints = hints } defs)
put Ctxt ({ localHints := hints } defs)
tm <- elab
ust <- get UST
@ -317,16 +317,16 @@ runDelays : {vars : _} ->
runDelays pri elab
= do ust <- get UST
let olddelayed = delayedElab ust
put UST (record { delayedElab = [] } ust)
put UST ({ delayedElab := [] } ust)
tm <- elab
ust <- get UST
log "elab.delay" 2 $ "Rerunning delayed in elaborator"
handle (do ignore $ retryDelayed' AllErrors False []
(reverse (filter hasPri (delayedElab ust))))
(\err => do put UST (record { delayedElab = olddelayed } ust)
(\err => do put UST ({ delayedElab := olddelayed } ust)
throw err)
ust <- get UST
put UST (record { delayedElab $= (++ olddelayed) } ust)
put UST ({ delayedElab $= (++ olddelayed) } ust)
pure tm
where
hasPri : (DelayReason, d) -> Bool

View File

@ -50,7 +50,7 @@ checkDot rig elabinfo nest env fc reason tm (Just gexpty)
= case elabMode elabinfo of
InLHS _ =>
do (wantedTm, wantedTy) <- check rig
(record { elabMode = InExpr }
({ elabMode := InExpr }
elabinfo)
nest env
tm (Just gexpty)

View File

@ -23,7 +23,7 @@ mkPrecise : {auto c : Ref Ctxt Defs} ->
mkPrecise (NApp _ (NMeta n i _) _)
= updateDef (Resolved i)
(\case
Hole i p => Just (Hole i (record { precisetype = True} p))
Hole i p => Just (Hole i ({ precisetype := True} p))
d => Nothing)
mkPrecise _ = pure ()

View File

@ -146,9 +146,9 @@ bindUnsolved {vars} fc elabmode _
tm <- metaVar fc rig env impn exp'
est <- get EST
let p' : PiInfo (Term vars) = forgetDef p
put EST (record { toBind $= ((impn, NameBinding rig p'
(embedSub subvars tm)
(embedSub subvars exp')) ::) } est)
put EST ({ toBind $= ((impn, NameBinding rig p'
(embedSub subvars tm)
(embedSub subvars exp')) ::) } est)
pure (embedSub sub tm)
mkImplicit : {outer : _} ->
@ -453,8 +453,8 @@ checkBindVar rig elabinfo nest env fc str topexp
_ => pure ()
log "elab.implicits" 5 $ "Added Bound implicit " ++ show (n, (rig, tm, exp, bty))
est <- get EST
put EST (record { boundNames $= ((n, NameBinding rig Explicit tm exp) ::),
toBind $= ((n, NameBinding rig Explicit tm bty) :: ) } est)
put EST ({ boundNames $= ((n, NameBinding rig Explicit tm exp) ::),
toBind $= ((n, NameBinding rig Explicit tm bty) :: ) } est)
log "metadata.names" 7 $ "checkBindVar is adding ↓"
addNameType fc (UN str) env exp
@ -548,9 +548,9 @@ checkBindHere rig elabinfo nest env fc bindmode tm exp
-- implicits should have access to whatever is in scope here
put EST (updateEnv env SubRefl [] est)
constart <- getNextEntry
(tmv, tmt) <- check rig (record { implicitMode = bindmode,
bindingVars = True }
elabinfo)
(tmv, tmt) <- check rig ({ implicitMode := bindmode,
bindingVars := True }
elabinfo)
nest env tm exp
let solvemode = case elabMode elabinfo of
InLHS c => inLHS
@ -561,14 +561,14 @@ checkBindHere rig elabinfo nest env fc bindmode tm exp
catch (retryDelayed solvemode (delayedElab ust))
(\err =>
do ust <- get UST
put UST (record { delayedElab = [] } ust)
put UST ({ delayedElab := [] } ust)
throw err)
-- Check all the patterns standing for polymorphic variables are
-- indeed polymorphic
ust <- get UST
let cons = polyConstraints ust
put UST (record { polyConstraints = [] } ust)
put UST ({ polyConstraints := [] } ust)
traverse_ solvePolyConstraint cons
traverse_ checkPolyConstraint cons
@ -586,7 +586,7 @@ checkBindHere rig elabinfo nest env fc bindmode tm exp
clearToBind dontbind
est <- get EST
put EST (updateEnv oldenv oldsub oldbif
(record { boundNames = [] } est))
({ boundNames := [] } est))
ty <- getTerm tmt
defs <- get Ctxt
(bv, bt) <- bindImplicits fc bindmode

View File

@ -55,14 +55,14 @@ localHelper {vars} nest env nestdecls_in func
names' <- traverse (applyEnv f)
(nub defNames) -- binding names must be unique
-- fixes bug #115
let nest' = record { names $= (names' ++) } nest
let nest' = { names $= (names' ++) } nest
let env' = dropLinear env
-- We don't want to keep rechecking delayed elaborators in the
-- locals block, because they're not going to make progress until
-- we come out again, so save them
ust <- get UST
let olddelayed = delayedElab ust
put UST (record { delayedElab = [] } ust)
put UST ({ delayedElab := [] } ust)
defs <- get Ctxt
-- store the local hints, so we can reset them after we've elaborated
-- everything
@ -73,11 +73,11 @@ localHelper {vars} nest env nestdecls_in func
traverse_ (processDecl [] nest' env') nestdecls
ust <- get UST
put UST (record { delayedElab = olddelayed } ust)
put UST ({ delayedElab := olddelayed } ust)
defs <- get Ctxt
res <- func nest'
defs <- get Ctxt
put Ctxt (record { localHints = oldhints } defs)
put Ctxt ({ localHints := oldhints } defs)
pure res
where
-- For the local definitions, don't allow access to linear things
@ -95,7 +95,7 @@ localHelper {vars} nest env nestdecls_in func
Core (Name, (Maybe Name, List (Var vars), FC -> NameType -> Term vars))
applyEnv outer inner
= do ust <- get UST
put UST (record { nextName $= (+1) } ust)
put UST ({ nextName $= (+1) } ust)
let nestedName_in = Nested (outer, nextName ust) inner
nestedName <- inCurrentNS nestedName_in
n' <- addName nestedName
@ -199,7 +199,7 @@ checkCaseLocal {vars} rig elabinfo nest env fc uname iname args sc expty
(app, args) <- getLocalTerm fc env name args
log "elab.local" 5 $ "Updating case local " ++ show uname ++ " " ++ show args
logTermNF "elab.local" 5 "To" env app
let nest' = record { names $= ((uname, (Just iname, args,
(\fc, nt => app))) :: ) }
nest
let nest' = { names $= ((uname, (Just iname, args,
(\fc, nt => app))) :: ) }
nest
check rig elabinfo nest' env sc expty

View File

@ -85,7 +85,7 @@ genFieldName : {auto u : Ref UST UState} ->
String -> Core String
genFieldName root
= do ust <- get UST
put UST (record { nextName $= (+1) } ust)
put UST ({ nextName $= (+1) } ust)
pure (root ++ show (nextName ust))
-- There's probably a generic version of this in the prelude isn't

View File

@ -241,14 +241,14 @@ checkTerm rig elabinfo nest env (IWithUnambigNames fc ns rhs) exp
= do -- enter the scope -> add unambiguous names
est <- get EST
rns <- resolveNames fc ns
put EST $ record { unambiguousNames = mergeLeft rns (unambiguousNames est) } est
put EST $ { unambiguousNames := mergeLeft rns (unambiguousNames est) } est
-- inside the scope -> check the RHS
result <- check rig elabinfo nest env rhs exp
-- exit the scope -> restore unambiguous names
newEST <- get EST
put EST $ record { unambiguousNames = unambiguousNames est } newEST
put EST $ { unambiguousNames := unambiguousNames est } newEST
pure result
where

View File

@ -65,9 +65,9 @@ updateErasable n
Just gdef <- lookupCtxtExact n (gamma defs)
| Nothing => pure ()
(es, dtes) <- findErased (type gdef)
ignore $ addDef n $ record
{ eraseArgs = es,
safeErase = dtes } gdef
ignore $ addDef n $
{ eraseArgs := es,
safeErase := dtes } gdef
export
wrapErrorC : List ElabOpt -> (Error -> Error) -> Core a -> Core a

View File

@ -273,7 +273,7 @@ recordUpdate : {auto u : Ref UPD Updates} ->
recordUpdate fc n tm
= do u <- get UPD
let nupdates = map (\x => (fst x, IVar fc (snd x))) (namemap u)
put UPD (record { updates $= ((n, substNames [] nupdates tm) ::) } u)
put UPD ({ updates $= ((n, substNames [] nupdates tm) ::) } u)
findUpdates : {auto u : Ref UPD Updates} ->
Defs -> RawImp -> RawImp -> Core ()
@ -283,8 +283,8 @@ findUpdates defs (IVar fc n) (IVar _ n')
Nothing =>
do u <- get UPD
case lookup n' (namemap u) of
Nothing => put UPD (record { namemap $= ((n', n) ::) } u)
Just nm => put UPD (record { updates $= ((n, IVar fc nm) ::) } u)
Nothing => put UPD ({ namemap $= ((n', n) ::) } u)
Just nm => put UPD ({ updates $= ((n, IVar fc nm) ::) } u)
findUpdates defs (IVar fc n) tm = recordUpdate fc n tm
findUpdates defs (IApp _ f a) (IApp _ f' a')
= do findUpdates defs f f'
@ -349,7 +349,7 @@ mkCase {c} {u} fn orig lhs_raw
put UST ust
case err of
WhenUnifying _ gam env l r err
=> do let defs = record { gamma = gam } defs
=> do let defs = { gamma := gam } defs
if !(impossibleOK defs !(nf defs env l)
!(nf defs env r))
then pure (Impossible lhs_raw)

View File

@ -214,8 +214,8 @@ searchIfHole fc opts hints topty env arg
let Hole _ _ = definition gdef
| _ => one (!(normaliseHoles defs env (metaApp arg)), [])
-- already solved
res <- search fc rig (record { depth = k,
inArg = True } opts) hints
res <- search fc rig ({ depth := k,
inArg := True } opts) hints
topty (Resolved hole)
-- When we solve an argument, we're also building a lambda
-- expression for its environment, so we need to apply it to
@ -412,7 +412,7 @@ tryRecursive fc rig opts hints env ty topty rdata
case !(lookupCtxtExact (recname rdata) (gamma defs)) of
Nothing => noResult
Just def =>
do res <- searchName fc rig (record { recData = Nothing } opts) hints
do res <- searchName fc rig ({ recData := Nothing } opts) hints
env !(nf defs env ty)
topty (recname rdata, def)
defs <- get Ctxt
@ -606,11 +606,11 @@ makeHelper fc rig opts env letty targetty (Result (locapp, ds) next)
-- added first.
-- There must be at least one case split.
((helper :: _), nextdef) <-
searchN 1 $ gendef (record { getRecData = False,
inUnwrap = True,
depth = depth',
ltor = False,
mustSplit = True } opts)
searchN 1 $ gendef ({ getRecData := False,
inUnwrap := True,
depth := depth',
ltor := False,
mustSplit := True } opts)
helpern 0 ty
| _ => do log "interaction.search" 10 "No results"
noResult
@ -665,7 +665,7 @@ tryIntermediateWith fc rig opts hints env ((p, pty) :: rest) ty topty
intnty <- genVarName "cty"
u <- uniVar fc
letty <- metaVar fc' erased env intnty (TType fc u)
let opts' = record { inUnwrap = True } opts
let opts' = { inUnwrap := True } opts
locsearch <- searchLocalWith fc True rig opts' hints env [(p, pty)]
letty topty
makeHelper fc rig opts env letty targetty locsearch
@ -707,8 +707,8 @@ tryIntermediateRec fc rig opts hints env ty topty (Just rd)
intnty <- genVarName "cty"
u <- uniVar fc
letty <- metaVar fc erased env intnty (TType fc u)
let opts' = record { inUnwrap = True,
recData = Nothing } opts
let opts' = { inUnwrap := True,
recData := Nothing } opts
logTerm "interaction.search" 10 "Trying recursive search for" ty
log "interaction.search" 10 $ show !(toFullNames (recname rd))
logTerm "interaction.search" 10 "LHS" !(toFullNames (lhsapp rd))
@ -894,7 +894,7 @@ exprSearchOpts opts fc n_in hints
log "interaction.search" 10 $ "LHS hole data " ++ show (n, lhs)
opts' <- if getRecData opts
then do d <- getLHSData defs lhs
pure (record { recData = d } opts)
pure ({ recData := d } opts)
else pure opts
validHints <- concat <$> for hints (\hint => do
defs <- get Ctxt

View File

@ -197,8 +197,8 @@ mutual
else expandClause loc opts n c)
(do cs <- generateSplits loc opts n c
log "interaction.generate" 5 $ "Splits: " ++ show cs
tryAllSplits loc (record { mustSplit = False,
doneSplit = True } opts) n cs)
tryAllSplits loc ({ mustSplit := False,
doneSplit := True } opts) n cs)
export
makeDefFromType : {auto c : Ref Ctxt Defs} ->
@ -247,7 +247,7 @@ makeDef p n
| Nothing => noResult
n <- getFullName nidx
logTerm "interaction.generate" 5 ("Searching for " ++ show n) ty
let opts = record { genExpr = Just (makeDefFromType (justFC loc)) }
let opts = { genExpr := Just (makeDefFromType (justFC loc)) }
(initSearchOpts True 5)
makeDefFromType (justFC loc) opts n envlen ty

View File

@ -290,7 +290,7 @@ mkSpecDef {vars} fc gdef pename sargs fn stk
do log "specialise" 1 $ "Partial evaluation of " ++ show !(toFullNames fn) ++ " failed" ++
"\n" ++ show err
defs <- get Ctxt
put Ctxt (record { peFailures $= insert pename () } defs)
put Ctxt ({ peFailures $= insert pename () } defs)
pure (applyWithFC (Ref fc Func fn) stk))
where
getAllRefs : NameMap Bool -> List ArgMode -> NameMap Bool

View File

@ -134,11 +134,11 @@ export
impossibleErrOK : {auto c : Ref Ctxt Defs} ->
Defs -> Error -> Core Bool
impossibleErrOK defs (CantConvert fc gam env l r)
= do let defs = record { gamma = gam } defs
= do let defs = { gamma := gam } defs
impossibleOK defs !(nf defs env l)
!(nf defs env r)
impossibleErrOK defs (CantSolveEq fc gam env l r)
= do let defs = record { gamma = gam } defs
= do let defs = { gamma := gam } defs
impossibleOK defs !(nf defs env l)
!(nf defs env r)
impossibleErrOK defs (BadDotPattern _ _ ErasedArg _ _) = pure True
@ -205,7 +205,7 @@ export
recoverableErr : {auto c : Ref Ctxt Defs} ->
Defs -> Error -> Core Bool
recoverableErr defs (CantConvert fc gam env l r)
= do let defs = record { gamma = gam } defs
= do let defs = { gamma := gam } defs
l <- nf defs env l
r <- nf defs env r
log "coverage.recover" 10 $ unlines
@ -217,7 +217,7 @@ recoverableErr defs (CantConvert fc gam env l r)
recoverable defs l r
recoverableErr defs (CantSolveEq fc gam env l r)
= do let defs = record { gamma = gam } defs
= do let defs = { gamma := gam } defs
recoverable defs !(nf defs env l)
!(nf defs env r)
recoverableErr defs (BadDotPattern _ _ ErasedArg _ _) = pure True
@ -582,7 +582,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
log "declare.def.clause.with" 5 $ "Argument names " ++ show wargNames
wname <- genWithName !(prettyName !(toFullNames (Resolved n)))
widx <- addDef wname (record {flags $= (SetTotal totreq ::)}
widx <- addDef wname ({flags $= (SetTotal totreq ::)}
(newDef vfc wname (if isErased mult then erased else top)
vars wtype vis None))
@ -611,7 +611,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
-- Elaborate the new definition here
nestname <- applyEnv env wname
let nest'' = record { names $= (nestname ::) } nest
let nest'' = { names $= (nestname ::) } nest
let wdef = IDef ifc wname cs'
processDecl [] nest'' env wdef
@ -762,8 +762,8 @@ calcRefs rt at fn
(dropErased (keys refs_all) refs_all)
(pure refs_all)
ignore $ ifThenElse rt
(addDef fn (record { refersToRuntimeM = Just refs } gdef))
(addDef fn (record { refersToM = Just refs } gdef))
(addDef fn ({ refersToRuntimeM := Just refs } gdef))
(addDef fn ({ refersToM := Just refs } gdef))
traverse_ (calcRefs rt at) (keys refs)
where
dropErased : List Name -> NameMap Bool -> Core (NameMap Bool)
@ -815,8 +815,8 @@ mkRunTime fc n
let Just Refl = nameListEq cargs rargs
| Nothing => throw (InternalError "WAT")
ignore $ addDef n $
record { definition = PMDef r rargs tree_ct tree_rt pats
} gdef
{ definition := PMDef r rargs tree_ct tree_rt pats
} gdef
-- If it's a case block, and not already set as inlinable,
-- check if it's safe to inline
when (caseName !(toFullNames n) && noInline (flags gdef)) $
@ -884,7 +884,7 @@ compileRunTime fc atotal
traverse_ (calcRefs True atotal) (toCompileCase defs)
defs <- get Ctxt
put Ctxt (record { toCompileCase = [] } defs)
put Ctxt ({ toCompileCase := [] } defs)
toPats : Clause -> (vs ** (Env Term vs, Term vs, Term vs))
toPats (MkClause {vars} env lhs rhs)
@ -1001,13 +1001,13 @@ processDef opts nest env fc n_in cs_in
defs <- get Ctxt
let pi = case lookup n (userHoles defs) of
Nothing => defaultPI
Just e => record { externalDecl = e } defaultPI
Just e => { externalDecl := e } defaultPI
-- Add compile time tree as a placeholder for the runtime tree,
-- but we'll rebuild that in a later pass once all the case
-- blocks etc are resolved
ignore $ addDef (Resolved nidx)
(record { definition = PMDef pi cargs tree_ct tree_ct pats
} gdef)
({ definition := PMDef pi cargs tree_ct tree_ct pats
} gdef)
when (visibility gdef == Public) $
do let rmetas = getMetas tree_ct
@ -1020,7 +1020,7 @@ processDef opts nest env fc n_in cs_in
-- Flag this name as one which needs compiling
defs <- get Ctxt
put Ctxt (record { toCompileCase $= (n ::) } defs)
put Ctxt ({ toCompileCase $= (n ::) } defs)
atotal <- toResolvedNames (NS builtinNS (UN $ Basic "assert_total"))
logTime ("+++ Building size change graphs " ++ show n) $

View File

@ -54,7 +54,7 @@ processParams {vars} {c} {m} {u} nest env fc ps ds
defs <- get Ctxt
let defNames = definedInBlock (currentNS defs) ds
names' <- traverse (applyEnv env') defNames
let nestBlock = record { names $= (names' ++) } nest'
let nestBlock = { names $= (names' ++) } nest'
traverse_ (processDecl [] nestBlock env') ds
where
mkParamTy : List (Name, RigCount, PiInfo RawImp, RawImp) -> RawImp

View File

@ -67,8 +67,8 @@ elabRecord {vars} eopts fc env nest newns vis mbtot tn_in params conName_in fiel
defs <- get Ctxt
-- Record that the current namespace is allowed to look
-- at private names in the nested namespace
put Ctxt (record { currentNS = cns,
nestedNS = newns :: nns } defs)
put Ctxt ({ currentNS := cns,
nestedNS := newns :: nns } defs)
where
paramTelescope : List (FC, Maybe Name, RigCount, PiInfo RawImp, RawImp)

View File

@ -93,7 +93,7 @@ processFnOpt fc _ ndef (SpecArgs ns)
ps <- getNamePos 0 nty
ddeps <- collectDDeps nty
specs <- collectSpec [] ddeps ps nty
ignore $ addDef ndef (record { specArgs = specs } gdef)
ignore $ addDef ndef ({ specArgs := specs } gdef)
where
insertDeps : List Nat -> List (Name, Nat) -> List Name -> List Nat
insertDeps acc ps [] = acc
@ -319,10 +319,10 @@ processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc nameFC n_in ty_ra
infargs <- findInferrable empty !(nf defs [] fullty)
ignore $ addDef (Resolved idx)
(record { eraseArgs = erased,
safeErase = dterased,
inferrable = infargs }
(newDef fc n rig vars fullty vis def))
({ eraseArgs := erased,
safeErase := dterased,
inferrable := infargs }
(newDef fc n rig vars fullty vis def))
-- 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

View File

@ -62,7 +62,7 @@ idrisTestsCasetree = MkTestPool "Case tree building" [] Nothing
idrisTestsWarning : TestPool
idrisTestsWarning = MkTestPool "Warnings" [] Nothing
["warning001", "warning002"]
["warning001", "warning002", "warning003"]
idrisTestsError : TestPool
idrisTestsError = MkTestPool "Error messages" [] Nothing
@ -357,4 +357,4 @@ main = runner $
where
testPaths : String -> TestPool -> TestPool
testPaths dir = record { testCases $= map ((dir ++ "/") ++) }
testPaths dir = { testCases $= map ((dir ++ "/") ++) }

View File

@ -5,9 +5,9 @@ record Attributes where
size : Nat
bigMono : Attributes -> Attributes
bigMono = record { font $= (++ "Mono")
, size = 20
}
bigMono = { font $= (++ "Mono")
, size := 20
}
smallMono : Attributes -> Attributes
smallMono = { size := 5 } . bigMono

View File

@ -20,19 +20,18 @@
0000e2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 25) (:end 7 34)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
0000e1(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 1) (:end 8 7)) ((:name "bigMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 9) (:end 8 9)) ((:decor :keyword)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 11) (:end 8 16)) ((:decor :keyword)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 18) (:end 8 18)) ((:decor :keyword)))))) 1)
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 20) (:end 8 23)) ((:decor :function)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 25) (:end 8 26)) ((:decor :keyword)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 28) (:end 8 28)) ((:decor :keyword)))))) 1)
0000e6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 29) (:end 8 30)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 32) (:end 8 37)) ((:decor :data)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 38) (:end 8 38)) ((:decor :keyword)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 18) (:end 9 18)) ((:decor :keyword)))))) 1)
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 20) (:end 9 23)) ((:decor :function)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 25) (:end 9 25)) ((:decor :keyword)))))) 1)
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 27) (:end 9 28)) ((:decor :data)))))) 1)
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 10 18) (:end 10 18)) ((:decor :keyword)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 11) (:end 8 11)) ((:decor :keyword)))))) 1)
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 13) (:end 8 16)) ((:decor :function)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 18) (:end 8 19)) ((:decor :keyword)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 21) (:end 8 21)) ((:decor :keyword)))))) 1)
0000e6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 22) (:end 8 23)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 25) (:end 8 30)) ((:decor :data)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 31) (:end 8 31)) ((:decor :keyword)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 11) (:end 9 11)) ((:decor :keyword)))))) 1)
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 13) (:end 9 16)) ((:decor :function)))))) 1)
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 18) (:end 9 19)) ((:decor :keyword)))))) 1)
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 21) (:end 9 22)) ((:decor :data)))))) 1)
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 10 11) (:end 10 11)) ((:decor :keyword)))))) 1)
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 1) (:end 12 9)) ((:decor :function)))))) 1)
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 11) (:end 12 11)) ((:decor :keyword)))))) 1)
0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 13) (:end 12 22)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)

View File

@ -142,7 +142,7 @@ mapSetMap : (a -> a') -> b' -> (c -> c') -> Three a b c -> Three a' b' c'
mapSetMap f y' g = {x $= f, y := y', z $= g}
setNameOld : String -> OrdinaryDog -> OrdinaryDog
setNameOld name' = record {name = name'}
setNameOld name' = {name := name'}
-------------------------------------

View File

@ -10,7 +10,7 @@ record Person animal where
-- Record update syntax with the record keyword
singleFieldUpdate : Person a -> Person a
singleFieldUpdate p = record { age $= (+one) } p
singleFieldUpdate p = { age $= (+one) } p
-- Not making it point free to keep a var here ^
where
-- A variable to use in the update
@ -18,7 +18,7 @@ singleFieldUpdate p = record { age $= (+one) } p
one = 1
multiFieldUpdate : Person a -> Person a
multiFieldUpdate p = record { name = emptyString, age = zero } p
multiFieldUpdate p = { name := emptyString, age := zero } p
-- Not point free just like above, in order to have a var here ^
where
-- Variables to use in the update

View File

@ -2,9 +2,9 @@
:typeat 7 3 age
:typeat 8 3 pet
:typeat 8 9 animal
:typeat 13 41 one
:typeat 13 48 p
:typeat 21 38 emptyString
:typeat 21 57 zero
:typeat 21 64 p
:typeat 13 33 one
:typeat 13 40 p
:typeat 21 32 emptyString
:typeat 21 51 zero
:typeat 21 58 p
:q

View File

@ -4,7 +4,7 @@ Main> Couldn't parse any alternatives:
(Interactive):1:4--1:5
1 | :t (3 : Nat)
^
... (52 others)
... (53 others)
Main> Expected string begin.
(Interactive):1:5--1:7

View File

@ -6,7 +6,7 @@ StrError1:2:24--2:25
1 | foo : String
2 | foo = "empty interp: \{}"
^
... (40 others)
... (41 others)
1/1: Building StrError2 (StrError2.idr)
Error: Couldn't parse any alternatives:
1: Expected 'case', 'if', 'do', application or operator expression.
@ -15,7 +15,7 @@ StrError2:2:19--2:20
1 | foo : String
2 | foo = "nested: \{ " \{ 1 + } " }"
^
... (28 others)
... (29 others)
1/1: Building StrError3 (StrError3.idr)
Error: Couldn't parse any alternatives:
1: Expected 'case', 'if', 'do', application or operator expression.
@ -24,7 +24,7 @@ StrError3:2:7--2:8
1 | foo : String
2 | foo = "incomplete: \{ a <+> }"
^
... (29 others)
... (30 others)
1/1: Building StrError4 (StrError4.idr)
Error: Bracket is not properly closed.

View File

@ -83,15 +83,15 @@ newQueue : RawMsgQueue
newQueue = MkRawMsgQueue [] []
enqueue : (1 val : a) -> RawMsgQueue -> RawMsgQueue
enqueue item q = record { inputStack $= (Entry item ::) } q
enqueue item q = { inputStack $= (Entry item ::) } q
dequeue : RawMsgQueue -> Maybe (RawMsgQueue, QueueEntry)
dequeue q
= case outputStack q of
[] => case reverse (inputStack q) of
[] => Nothing
(x :: xs) => Just (record { outputStack = xs, inputStack = [] } q, x)
(x :: xs) => Just (record { outputStack = xs } q, x)
(x :: xs) => Just ({ outputStack := xs, inputStack := [] } q, x)
(x :: xs) => Just ({ outputStack := xs } q, x)
public export
data Channel : {p : Protocol b} -> Actions a -> Type where

View File

@ -9,13 +9,13 @@ record MyDPair a p where
cons : t -> MyDPair Nat (\n => Vect n t) -> MyDPair Nat (\n => Vect n t)
cons val xs
= record { dfst = S (dfst xs),
dsnd = val :: dsnd xs } xs
= { dfst := S (dfst xs),
dsnd := val :: dsnd xs } xs
cons' : t -> MyDPair Nat (\n => Vect n t) -> MyDPair Nat (\n => Vect n t)
cons' val
= record { dfst $= S,
dsnd $= (val ::) }
= { dfst $= S,
dsnd $= (val ::) }
record Stats where
constructor MkStats
@ -35,4 +35,4 @@ testPerson : Person
testPerson = MkPerson "Fred" 1337 10 (MkStats 10 10)
grow : Person -> Person
grow = record { more.height $= (+1) }
grow = { more.height $= (+1) }

View File

@ -4,4 +4,7 @@ Main> MkPerson "Fred" 1337 10 (MkStats 12 10)
Main> MkMyDPair 6 [10, 1, 2, 3, 4, 5]
Main> MkMyDPair 6 [10, 1, 2, 3, 4, 5]
Main> MkPerson "Fred" 1337 10 (MkStats 10 94)
Main> MkPerson "Fred" 1337 10 (MkStats 10 11)
Main> MkPerson "Fred" 1337 10 (MkStats 10 94)
Main> MkPerson "Fred" 1337 10 (MkStats 10 11)
Main> Bye for now!

View File

@ -3,4 +3,7 @@ grow (grow testPerson)
cons 10 testPair
cons' 10 testPair
record { more.weight = 94 } testPerson
record { more.weight $= (+ 1) } testPerson
{ more.weight := 94 } testPerson
{ more.weight $= (+ 1) } testPerson
:q

View File

@ -16,8 +16,8 @@ Main.Point.x pt
Point.x pt
(x) pt
x pt
(record { topLeft.x = 3 } rect).topLeft.x
(record { topLeft.x $= (+1) } rect).topLeft.x
(record { topLeft->x = 3 } rect).topLeft.x
(record { topLeft->x $= (+1) } rect).topLeft.x
({ topLeft.x := 3 } rect).topLeft.x
({ topLeft.x $= (+1) } rect).topLeft.x
({ topLeft->x := 3 } rect).topLeft.x
({ topLeft->x $= (+1) } rect).topLeft.x
:q

View File

@ -6,4 +6,4 @@ AFoo : Foo
AFoo = MkFoo {Gnu = 'c'}
Bar : Foo
Bar = record { Gnu = '?' } AFoo
Bar = { Gnu := '?' } AFoo

View File

@ -11,9 +11,9 @@ foo1 = MkFoo
}
foo2 : Foo
foo2 = record
{ a = 3
, a = 4
, b = 2
, b = 1
foo2 =
{ a := 3
, a := 4
, b := 2
, b := 1
} foo1

View File

@ -13,11 +13,10 @@ Error: While processing right hand side of foo2. Duplicated record update paths:
b
Issue2095:14:8--19:9
14 | foo2 = record
15 | { a = 3
16 | , a = 4
17 | , b = 2
18 | , b = 1
Issue2095:15:3--19:9
15 | { a := 3
16 | , a := 4
17 | , b := 2
18 | , b := 1
19 | } foo1

View File

@ -3,4 +3,4 @@ record R (a : Type) where
x : a
rmap : (a -> b) -> R a -> R b
rmap f = record { x $= f }
rmap f = { x $= f }

View File

@ -0,0 +1,11 @@
module Main
record TestRecord where
constructor MkTestRecord
recField : Int
testRec : TestRecord
testRec = MkTestRecord 0
updatedRec : TestRecord
updatedRec = record { recField = 1 } testRec

View File

@ -0,0 +1,5 @@
package deprecated
modules = Foo
, Main

View File

@ -0,0 +1,11 @@
1/1: Building Main (Main.idr)
Warning: DEPRECATED: old record update syntax. Use "{ f := v } p" instead of "record { f = v } p"
Main:11:14--11:37
07 | testRec : TestRecord
08 | testRec = MkTestRecord 0
09 |
10 | updatedRec : TestRecord
11 | updatedRec = record { recField = 1 } testRec
^^^^^^^^^^^^^^^^^^^^^^^

4
tests/idris2/warning003/run Executable file
View File

@ -0,0 +1,4 @@
rm -rf ./build
$1 --no-color --console-width 0 --check Main.idr

View File

@ -40,5 +40,5 @@ Main> Couldn't parse any alternatives:
(Interactive):1:4--1:8
1 | :t with [] 4
^^^^
... (44 others)
... (45 others)
Main> Bye for now!

View File

@ -24,14 +24,14 @@ initState : GameState
initState = MkGameState (MkScore 0 0) 12
addWrong : GameState -> GameState
addWrong = record { score.attempted $= (+1) }
addWrong = { score.attempted $= (+1) }
addCorrect : GameState -> GameState
addCorrect = record { score.correct $= (+1),
score.attempted $= (+1) }
addCorrect = { score.correct $= (+1),
score.attempted $= (+1) }
setDifficulty : Int -> GameState -> GameState
setDifficulty newDiff state = record { difficulty = newDiff } state
setDifficulty newDiff state = { difficulty := newDiff } state
export
data Command : Type -> Type where