Properly erase 0 quantity arguments

Don't just have a placeholder. While this doesn't have a huge effect (if
any) on performance, it does generate smaller output for Chez to
process, and is tidier. Perhaps it's good for other back ends too, ones
that don't optimise as much as Chez does.

Only doing named functions, not higher order functions. HOFs may be
worth doing too, if we can, since this could remove lambdas and make
fewer closures.

The increment in TTC Version is necessary because otherwise there could
be inconsistencies between libraries and clients erasure properties.
This commit is contained in:
Edwin Brady 2021-05-03 13:37:38 +01:00
parent 5bbcb29a38
commit 73d374e435
8 changed files with 130 additions and 109 deletions

View File

@ -57,9 +57,9 @@ Ord UsePhase where
where
tag : UsePhase -> Int
tag Cases = 0
tag Lifted = 0
tag ANF = 0
tag VMCode = 0
tag Lifted = 1
tag ANF = 2
tag VMCode = 3
public export
record CompileData where

View File

@ -31,13 +31,22 @@ numArgs defs (Ref _ _ n)
case definition gdef of
DCon _ arity Nothing => pure (EraseArgs arity (eraseArgs gdef))
DCon _ arity (Just (_, pos)) => pure (NewTypeBy arity pos)
PMDef _ args _ _ _ => pure (Arity (length args))
PMDef _ args _ _ _ => pure (EraseArgs (length args) (eraseArgs gdef))
ExternDef arity => pure (Arity arity)
ForeignDef arity _ => pure (Arity arity)
Builtin {arity} f => pure (Arity arity)
_ => pure (Arity 0)
numArgs _ tm = pure (Arity 0)
mkSub : Nat -> (ns : List Name) -> List Nat -> (ns' ** SubVars ns' ns)
mkSub i _ [] = (_ ** SubRefl)
mkSub i [] ns = (_ ** SubRefl)
mkSub i (x :: xs) es
= let (ns' ** p) = mkSub (S i) xs es in
if i `elem` es
then (ns' ** DropCons p)
else (x :: ns' ** KeepCons p)
weakenVar : Var ns -> Var (a :: ns)
weakenVar (MkVar p) = (MkVar (Later p))
@ -98,23 +107,28 @@ applyNewType arity pos fn args
keepArg (CCon fc _ _ args) = keep 0 args
keepArg tm = CErased (getFC fn)
dropFrom : List Nat -> Nat -> List (CExp vs) -> List (CExp vs)
dropFrom epos i [] = []
dropFrom epos i (x :: xs)
= if i `elem` epos
then dropFrom epos (1 + i) xs
else x :: dropFrom epos (1 + i) xs
dropPos : List Nat -> CExp vs -> CExp vs
dropPos epos (CLam fc x sc) = CLam fc x (dropPos epos sc)
dropPos epos (CCon fc c a args) = CCon fc c a (drop 0 args)
where
drop : Nat -> List (CExp vs) -> List (CExp vs)
drop i [] = []
drop i (x :: xs)
= if i `elem` epos
then drop (1 + i) xs
else x :: drop (1 + i) xs
dropPos epos (CApp fc tm@(CApp _ _ _) args')
= CApp fc (dropPos epos tm) args'
dropPos epos (CApp fc f args) = CApp fc f (dropFrom epos 0 args)
dropPos epos (CCon fc c a args) = CCon fc c a (dropFrom epos 0 args)
dropPos epos tm = tm
eraseConArgs : {vars : _} ->
Nat -> List Nat -> CExp vars -> List (CExp vars) -> CExp vars
eraseConArgs arity epos fn args
= let fn' = expandToArity arity fn args in
dropPos epos fn' -- fn' might be lambdas, after eta expansion
if not (isNil epos)
then dropPos epos fn' -- fn' might be lambdas, after eta expansion
else fn'
mkDropSubst : Nat -> List Nat ->
(rest : List Name) ->
@ -638,13 +652,17 @@ getCFTypes args t
= pure (reverse args, !(nfToCFType (getLoc t) False t))
toCDef : {auto c : Ref Ctxt Defs} ->
Name -> ClosedTerm -> Def ->
Name -> ClosedTerm -> List Nat -> Def ->
Core CDef
toCDef n ty None
toCDef n ty _ None
= pure $ MkError $ CCrash emptyFC ("Encountered undefined name " ++ show !(getFullName n))
toCDef n ty (PMDef _ args _ tree _)
= pure $ MkFun _ !(toCExpTree n tree)
toCDef n ty (ExternDef arity)
toCDef n ty erased (PMDef _ args _ tree _)
= do let (args' ** p) = mkSub 0 args erased
comptree <- toCExpTree n tree
if isNil erased
then pure $ MkFun args comptree
else pure $ MkFun args' (shrinkCExp p comptree)
toCDef n ty _ (ExternDef arity)
= let (ns ** args) = mkArgList 0 arity in
pure $ MkFun _ (CExtPrim emptyFC !(getFullName n) (map toArgExp (getVars args)))
where
@ -654,11 +672,11 @@ toCDef n ty (ExternDef arity)
getVars : ArgList k ns -> List (Var ns)
getVars NoArgs = []
getVars (ConsArg a rest) = MkVar First :: map weakenVar (getVars rest)
toCDef n ty (ForeignDef arity cs)
toCDef n ty _ (ForeignDef arity cs)
= do defs <- get Ctxt
(atys, retty) <- getCFTypes [] !(nf defs [] ty)
pure $ MkForeign cs atys retty
toCDef n ty (Builtin {arity} op)
toCDef n ty _ (Builtin {arity} op)
= let (ns ** args) = mkArgList 0 arity in
pure $ MkFun _ (COp emptyFC op (map toArgExp (getVars args)))
where
@ -668,7 +686,7 @@ toCDef n ty (Builtin {arity} op)
getVars : ArgList k ns -> Vect k (Var ns)
getVars NoArgs = []
getVars (ConsArg a rest) = MkVar First :: map weakenVar (getVars rest)
toCDef n _ (DCon tag arity pos)
toCDef n _ _ (DCon tag arity pos)
= do let nt = snd <$> pos
defs <- get Ctxt
args <- numArgs {vars = []} defs (Ref EmptyFC (DataCon tag arity) n)
@ -677,20 +695,20 @@ toCDef n _ (DCon tag arity pos)
EraseArgs ar erased => ar `minus` length erased
Arity ar => ar
pure $ MkCon (Just tag) arity' nt
toCDef n _ (TCon tag arity _ _ _ _ _ _)
toCDef n _ _ (TCon tag arity _ _ _ _ _ _)
= pure $ MkCon Nothing arity Nothing
-- We do want to be able to compile these, but also report an error at run time
-- (and, TODO: warn at compile time)
toCDef n ty (Hole _ _)
toCDef n ty _ (Hole _ _)
= pure $ MkError $ CCrash emptyFC ("Encountered unimplemented hole " ++
show !(getFullName n))
toCDef n ty (Guess _ _ _)
toCDef n ty _ (Guess _ _ _)
= pure $ MkError $ CCrash emptyFC ("Encountered constrained hole " ++
show !(getFullName n))
toCDef n ty (BySearch _ _ _)
toCDef n ty _ (BySearch _ _ _)
= pure $ MkError $ CCrash emptyFC ("Encountered incomplete proof search " ++
show !(getFullName n))
toCDef n ty def
toCDef n ty _ def
= pure $ MkError $ CCrash emptyFC ("Encountered uncompilable name " ++
show (!(getFullName n), def))
@ -709,7 +727,7 @@ compileDef n
= do defs <- get Ctxt
Just gdef <- lookupCtxtExact n (gamma defs)
| Nothing => throw (InternalError ("Trying to compile unknown name " ++ show n))
ce <- toCDef n (type gdef)
ce <- toCDef n (type gdef) (eraseArgs gdef)
!(toFullNames (definition gdef))
setCompiled n ce

View File

@ -150,7 +150,7 @@ mutual
-- boost by removing unnecessary lambdas that we'll keep the special case.
eval rec env stk (CRef fc n)
= case (n == NS primIONS (UN "io_bind"), stk) of
(True, _ :: _ :: act :: cont :: world :: stk) =>
(True, act :: cont :: world :: stk) =>
do xn <- genName "act"
sc <- eval rec [] [] (CApp fc cont [CRef fc xn, world])
pure $ unload stk $

View File

@ -332,7 +332,6 @@ mutual
usedConstAlt : {default Nothing lazy : Maybe LazyReason} ->
Used vars -> LiftedConstAlt vars -> Used vars
usedConstAlt used (MkLConstAlt c sc) = usedVars used sc
usedVars used (LPrimVal _ _) = used
usedVars used (LErased _) = used
usedVars used (LCrash _ _) = used

View File

@ -31,7 +31,7 @@ import Data.Buffer
-- TTC files can only be compatible if the version number is the same
export
ttcVersion : Int
ttcVersion = 48
ttcVersion = 49
export
checkTTCVersion : String -> Int -> Int -> Core ()

View File

@ -13,6 +13,7 @@ import Core.Value
import TTImp.Elab.Check
import TTImp.Elab.Delayed
import TTImp.Elab.ImplicitBind
import TTImp.Elab.Utils
import TTImp.TTImp
import TTImp.Utils
@ -205,6 +206,7 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
(maybe (Bind fc scrn (Pi fc caseRig Explicit scrty)
(weaken caseretty))
(const caseretty) splitOn)
(erasedargs, _) <- findErased casefnty
logEnv "elab.case" 10 "Case env" env
logTermNF "elab.case" 2 ("Case function type: " ++ show casen) [] casefnty
@ -215,8 +217,10 @@ 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 (newDef fc casen (if isErased rigc then erased else top)
[] casefnty vis None)
cidx <- addDef casen (record { eraseArgs = erasedargs }
(newDef fc casen (if isErased rigc then erased else top)
[] casefnty vis None))
-- don't worry about totality of the case block; it'll be handled
-- by the totality of the parent function
setFlag fc (Resolved cidx) (SetTotal PartialOK)

View File

@ -11,26 +11,26 @@ Main.S = Constructor tag Just 0 arity 1
Prelude.Basics.intToBool = [{arg:N}]: (%case !{arg:N} [(%constcase 0 1)] Just 0)
Prelude.Basics.True = Constructor tag Just 0 arity 0
Prelude.Basics.False = Constructor tag Just 1 arity 0
Builtin.believe_me = [{arg:N}, {arg:N}, {ext:N}]: (believe_me [___, ___, !{ext:N}])
Prelude.Types.case block in prim__integerToNat = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Builtin.believe_me [___, ___, !{arg:N}])), (%constcase 1 0)] Nothing)
Builtin.believe_me = [{ext:N}]: (believe_me [___, ___, !{ext:N}])
Prelude.Types.case block in prim__integerToNat = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Builtin.believe_me [!{arg:N}])), (%constcase 1 0)] Nothing)
Prelude.Types.prim__integerToNat = [{arg:N}]: (Prelude.Types.case block in prim__integerToNat [!{arg:N}, (%case (<=Integer [0, !{arg:N}]) [(%constcase 0 1)] Just 0)])
Prelude.Types.Z = Constructor tag Just 0 arity 0
Prelude.Interfaces.Functor at Prelude/Interfaces.idr:L:C--L:C = Constructor tag Just 0 arity 2 (newtype by 1)
Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C = Constructor tag Just 0 arity 3
Prelude.Interfaces.pure = [{arg:N}, {arg:N}, {arg:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{ext:N}]))] Nothing)
PrimIO.case block in unsafePerformIO = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (PrimIO.unsafeDestroyWorld [___, ___, !{arg:N}])
PrimIO.case block in case block in io_bind = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (!{arg:N} [!{arg:N}])
PrimIO.case block in io_bind = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (PrimIO.case block in case block in io_bind [___, ___, ___, ___, ___, !{arg:N}, ___, (!{arg:N} [!{arg:N}])])
PrimIO.unsafePerformIO = [{arg:N}, {arg:N}]: (PrimIO.unsafeCreateWorld [___, (%lam w (PrimIO.case block in unsafePerformIO [___, !{arg:N}, ___, (!{arg:N} [!w])]))])
PrimIO.unsafeDestroyWorld = [{arg:N}, {arg:N}, {arg:N}]: !{arg:N}
PrimIO.unsafeCreateWorld = [{arg:N}, {arg:N}]: (!{arg:N} [%MkWorld])
PrimIO.io_pure = [{arg:N}, {arg:N}, {ext:N}]: !{arg:N}
PrimIO.io_bind = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}]: (PrimIO.case block in io_bind [___, ___, ___, !{arg:N}, ___, (!{arg:N} [!{ext:N}])])
Prelude.Interfaces.pure = [{arg:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{ext:N}]))] Nothing)
PrimIO.case block in unsafePerformIO = [{arg:N}, {arg:N}]: (PrimIO.unsafeDestroyWorld [___, !{arg:N}])
PrimIO.case block in case block in io_bind = [{arg:N}, {arg:N}, {arg:N}]: (!{arg:N} [!{arg:N}])
PrimIO.case block in io_bind = [{arg:N}, {arg:N}]: (PrimIO.case block in case block in io_bind [!{arg:N}, ___, (!{arg:N} [!{arg:N}])])
PrimIO.unsafePerformIO = [{arg:N}]: (PrimIO.unsafeCreateWorld [(%lam w (PrimIO.case block in unsafePerformIO [!{arg:N}, (!{arg:N} [!w])]))])
PrimIO.unsafeDestroyWorld = [{arg:N}, {arg:N}]: !{arg:N}
PrimIO.unsafeCreateWorld = [{arg:N}]: (!{arg:N} [%MkWorld])
PrimIO.io_pure = [{arg:N}, {ext:N}]: !{arg:N}
PrimIO.io_bind = [{arg:N}, {arg:N}, {ext:N}]: (PrimIO.case block in io_bind [!{arg:N}, (!{arg:N} [!{ext:N}])])
PrimIO.MkIORes = Constructor tag Just 0 arity 3 (newtype by 1)
PrimIO.MkIO = Constructor tag Just 0 arity 2 (newtype by 1)
Prelude.IO.pure = [{arg:N}, {arg:N}, {ext:N}]: !{arg:N}
Prelude.IO.map = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (!{arg:N} [!{act:N}]))
Prelude.IO.Functor implementation at Prelude/IO.idr:L:C--L:C = [{ext:N}, {ext:N}, {ext:N}, {ext:N}, {ext:N}]: (Prelude.IO.map [___, ___, !{ext:N}, !{ext:N}, !{ext:N}])
Prelude.IO.Applicative implementation at Prelude/IO.idr:L:C--L:C = []: (%con Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam b (%lam a (%lam func (%lam {arg:N} (%lam {eta:N} (Prelude.IO.map [___, ___, !func, !{arg:N}, !{eta:N}])))))), (%lam a (%lam {arg:N} (%lam {eta:N} !{arg:N}))), (%lam b (%lam a (%lam {arg:N} (%lam {arg:N} (%lam {eta:N} (%let {act:N} (!{arg:N} [!{eta:N}]) (%let {act:N} (!{arg:N} [!{eta:N}]) (!{act:N} [!{act:N}]))))))))])
Prelude.IO.<*> = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (%let {act:N} (!{arg:N} [!{ext:N}]) (!{act:N} [!{act:N}])))
Prelude.IO.pure = [{arg:N}, {ext:N}]: !{arg:N}
Prelude.IO.map = [{arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (!{arg:N} [!{act:N}]))
Prelude.IO.Functor implementation at Prelude/IO.idr:L:C--L:C = [{ext:N}, {ext:N}, {ext:N}, {ext:N}, {ext:N}]: (Prelude.IO.map [!{ext:N}, !{ext:N}, !{ext:N}])
Prelude.IO.Applicative implementation at Prelude/IO.idr:L:C--L:C = []: (%con Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam b (%lam a (%lam func (%lam {arg:N} (%lam {eta:N} (Prelude.IO.map [!func, !{arg:N}, !{eta:N}])))))), (%lam a (%lam {arg:N} (%lam {eta:N} !{arg:N}))), (%lam b (%lam a (%lam {arg:N} (%lam {arg:N} (%lam {eta:N} (%let {act:N} (!{arg:N} [!{eta:N}]) (%let {act:N} (!{arg:N} [!{eta:N}]) (!{act:N} [!{act:N}]))))))))])
Prelude.IO.<*> = [{arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (%let {act:N} (!{arg:N} [!{ext:N}]) (!{act:N} [!{act:N}])))

View File

@ -16,52 +16,52 @@ prim__gt_Int = [{arg:N}, {arg:N}]: (>Int [!{arg:N}, !{arg:N}])
prim__believe_me = [{arg:N}, {arg:N}, {arg:N}]: (believe_me [!{arg:N}, !{arg:N}, !{arg:N}])
prim__crash = [{arg:N}, {arg:N}]: (crash [!{arg:N}, !{arg:N}])
prim__cast_IntegerInt = [{arg:N}]: (cast-Integer-Int [!{arg:N}])
Main.main = [{ext:N}]: (Prelude.Interfaces.sum [___, ___, (%con Builtin.MkPair Just 0 [(%con Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [___, ___, !func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [___, ___, !func, !init, !input])))))), (%lam elem (%lam {arg:N} (Prelude.Types.null [___, !{arg:N}]))), (%lam elem (%lam acc (%lam m (%lam {i_con:N} (%lam funcM (%lam init (%lam input (Prelude.Types.foldlM [___, ___, ___, !{i_con:N}, !funcM, !init, !input]))))))))]), (%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))])]), (Prelude.Types.rangeFromTo [___, (%con Builtin.MkPair Just 0 [(%con Prelude.Num.Integral at Prelude/Num.idr:L:C--L:C Just 0 [(%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.div [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.mod [!{arg:N}, !{arg:N}])))]), (%con Builtin.MkPair Just 0 [(%con Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [(%con Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd./= [!{arg:N}, !{arg:N}])))]), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.compare [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.> [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.<= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.>= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.max [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.min [!{arg:N}, !{arg:N}])))]), (%con Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [(%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (Prelude.Num.negate [!{arg:N}])), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.- [!{arg:N}, !{arg:N}])))])])]), (cast-Integer-Int [1]), (cast-Integer-Int [100])])])
Main.main = [{ext:N}]: (Prelude.Interfaces.sum [(%con Builtin.MkPair Just 0 [(%con Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [!func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [!func, !init, !input])))))), (%lam elem (%lam {arg:N} (Prelude.Types.null [!{arg:N}]))), (%lam elem (%lam acc (%lam m (%lam {i_con:N} (%lam funcM (%lam init (%lam input (Prelude.Types.foldlM [!{i_con:N}, !funcM, !init, !input]))))))))]), (%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))])]), (Prelude.Types.rangeFromTo [(%con Builtin.MkPair Just 0 [(%con Prelude.Num.Integral at Prelude/Num.idr:L:C--L:C Just 0 [(%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.div [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.mod [!{arg:N}, !{arg:N}])))]), (%con Builtin.MkPair Just 0 [(%con Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [(%con Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd./= [!{arg:N}, !{arg:N}])))]), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.compare [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.> [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.<= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.>= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.max [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.min [!{arg:N}, !{arg:N}])))]), (%con Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [(%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (Prelude.Num.negate [!{arg:N}])), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.- [!{arg:N}, !{arg:N}])))])])]), (cast-Integer-Int [1]), (cast-Integer-Int [100])])])
Prelude.Basics.not = [{arg:N}]: (%case !{arg:N} [(%constcase 0 1), (%constcase 1 0)] Nothing)
Prelude.Basics.intToBool = [{arg:N}]: (%case !{arg:N} [(%constcase 0 1)] Just 0)
Prelude.Basics.flip = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: ((!{arg:N} [!{arg:N}]) [!{arg:N}])
Prelude.Basics.flip = [{arg:N}, {arg:N}, {arg:N}]: ((!{arg:N} [!{arg:N}]) [!{arg:N}])
Prelude.Basics.True = Constructor tag Just 0 arity 0
Prelude.Basics.False = Constructor tag Just 1 arity 0
Prelude.Basics.Bool = Constructor tag Nothing arity 0
Prelude.Basics.&& = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (%force Lazy !{arg:N})), (%constcase 1 1)] Nothing)
Builtin.snd = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing)
Builtin.idris_crash = [{arg:N}, {ext:N}]: (crash [___, !{ext:N}])
Builtin.fst = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing)
Builtin.believe_me = [{arg:N}, {arg:N}, {ext:N}]: (believe_me [___, ___, !{ext:N}])
Builtin.assert_total = [{arg:N}, {arg:N}]: !{arg:N}
Builtin.snd = [{arg:N}]: (%case !{arg:N} [(%concase Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing)
Builtin.idris_crash = [{ext:N}]: (crash [___, !{ext:N}])
Builtin.fst = [{arg:N}]: (%case !{arg:N} [(%concase Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing)
Builtin.believe_me = [{ext:N}]: (believe_me [___, ___, !{ext:N}])
Builtin.assert_total = [{arg:N}]: !{arg:N}
Builtin.MkPair = Constructor tag Just 0 arity 2
Prelude.Types.case block in rangeFromThen = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.countFrom [___, !{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))])), (%constcase 1 (Prelude.Types.countFrom [___, !{arg:N}, (%lam n (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!n]) [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))]))] Nothing)
Prelude.Types.case block in case block in case block in rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (%con Prelude.Types.:: Just 1 [!{arg:N}, (%con Prelude.Types.Nil Just 0 [])])), (%constcase 1 (%con Prelude.Types.Nil Just 0 []))] Nothing)
Prelude.Types.case block in case block in rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeBefore [___, (%lam {arg:N} (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [___, !{arg:N}, (%lam n (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!n]) [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))])])), (%constcase 1 (Prelude.Types.case block in case block in case block in rangeFromThenTo [___, !{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}, (Prelude.Basics.&& [(%case (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing), (%delay Lazy (%case (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing))])]))] Nothing)
Prelude.Types.case block in rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeBefore [___, (%lam {arg:N} (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [___, !{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))])])), (%constcase 1 (Prelude.Types.case block in case block in rangeFromThenTo [___, !{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]) [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing)]))] Nothing)
Prelude.Types.case block in case block in rangeFromTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeUntil [___, (%lam {arg:N} (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [___, !{arg:N}, (%lam x (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!x]) [(%case (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing)]))] Nothing))])])), (%constcase 1 (%con Prelude.Types.:: Just 1 [!{arg:N}, (%con Prelude.Types.Nil Just 0 [])]))] Nothing)
Prelude.Types.case block in rangeFromTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeUntil [___, (%lam {arg:N} (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [___, !{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing)]))] Nothing))])])), (%constcase 1 (Prelude.Types.case block in case block in rangeFromTo [___, !{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing)
Prelude.Types.case block in takeBefore = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (%con Prelude.Types.Nil Just 0 [])), (%constcase 1 (%con Prelude.Types.:: Just 1 [!{arg:N}, (Prelude.Types.takeBefore [___, !{arg:N}, (%force Inf !{arg:N})])]))] Nothing)
Prelude.Types.case block in takeUntil = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (%con Prelude.Types.:: Just 1 [!{arg:N}, (%con Prelude.Types.Nil Just 0 [])])), (%constcase 1 (%con Prelude.Types.:: Just 1 [!{arg:N}, (Prelude.Types.takeUntil [___, !{arg:N}, (%force Inf !{arg:N})])]))] Nothing)
Prelude.Types.case block in prim__integerToNat = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Builtin.believe_me [___, ___, !{arg:N}])), (%constcase 1 0)] Nothing)
Prelude.Types.rangeFrom = [{arg:N}, {arg:N}, {arg:N}]: (Prelude.Types.countFrom [___, !{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [(%case (%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing)]) [!{arg:N}]))] Nothing))])
Prelude.Types.rangeFromTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (Prelude.Types.case block in rangeFromTo [___, !{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)])
Prelude.Types.rangeFromThen = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (Prelude.Types.case block in rangeFromThen [___, !{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)])
Prelude.Types.rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (Prelude.Types.case block in rangeFromThenTo [___, !{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]) [(%case (Builtin.snd [___, ___, (Builtin.snd [___, ___, !{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing)])
Prelude.Types.null = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Nil Just 0 [] (%delay Lazy 0)), (%concase Prelude.Types.:: Just 1 [{e:N}, {e:N}] (%delay Lazy 1))] Nothing)
Prelude.Types.foldr = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Nil Just 0 [] !{arg:N}), (%concase Prelude.Types.:: Just 1 [{e:N}, {e:N}] ((!{arg:N} [!{e:N}]) [(Prelude.Types.foldr [___, ___, !{arg:N}, !{arg:N}, !{e:N}])]))] Nothing)
Prelude.Types.foldl = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Nil Just 0 [] !{arg:N}), (%concase Prelude.Types.:: Just 1 [{e:N}, {e:N}] (Prelude.Types.foldl [___, ___, !{arg:N}, ((!{arg:N} [!{arg:N}]) [!{e:N}]), !{e:N}]))] Nothing)
Prelude.Types.foldlM = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}]: (Prelude.Types.foldl [___, ___, (%lam ma (%lam b (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((((!{e:N} [___]) [___]) [!ma]) [(%lam {eta:N} (Prelude.Basics.flip [___, ___, ___, !{arg:N}, !b, !{eta:N}]))]))] Nothing))), (%case (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{arg:N}]))] Nothing), !{ext:N}])
Prelude.Types.case block in rangeFromThen = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.countFrom [!{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))])), (%constcase 1 (Prelude.Types.countFrom [!{arg:N}, (%lam n (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!n]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))]))] Nothing)
Prelude.Types.case block in case block in case block in rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (%con Prelude.Types.:: Just 1 [!{arg:N}, (%con Prelude.Types.Nil Just 0 [])])), (%constcase 1 (%con Prelude.Types.Nil Just 0 []))] Nothing)
Prelude.Types.case block in case block in rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeBefore [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam n (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!n]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))])])), (%constcase 1 (Prelude.Types.case block in case block in case block in rangeFromThenTo [!{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}, (Prelude.Basics.&& [(%case (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing), (%delay Lazy (%case (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing))])]))] Nothing)
Prelude.Types.case block in rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeBefore [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing))])])), (%constcase 1 (Prelude.Types.case block in case block in rangeFromThenTo [!{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing)]))] Nothing)
Prelude.Types.case block in case block in rangeFromTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeUntil [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam x (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!x]) [(%case (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing)]))] Nothing))])])), (%constcase 1 (%con Prelude.Types.:: Just 1 [!{arg:N}, (%con Prelude.Types.Nil Just 0 [])]))] Nothing)
Prelude.Types.case block in rangeFromTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Prelude.Types.takeUntil [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing)]))] Nothing))])])), (%constcase 1 (Prelude.Types.case block in case block in rangeFromTo [!{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing)
Prelude.Types.case block in takeBefore = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (%con Prelude.Types.Nil Just 0 [])), (%constcase 1 (%con Prelude.Types.:: Just 1 [!{arg:N}, (Prelude.Types.takeBefore [!{arg:N}, (%force Inf !{arg:N})])]))] Nothing)
Prelude.Types.case block in takeUntil = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (%con Prelude.Types.:: Just 1 [!{arg:N}, (%con Prelude.Types.Nil Just 0 [])])), (%constcase 1 (%con Prelude.Types.:: Just 1 [!{arg:N}, (Prelude.Types.takeUntil [!{arg:N}, (%force Inf !{arg:N})])]))] Nothing)
Prelude.Types.case block in prim__integerToNat = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 (Builtin.believe_me [!{arg:N}])), (%constcase 1 0)] Nothing)
Prelude.Types.rangeFrom = [{arg:N}, {arg:N}]: (Prelude.Types.countFrom [!{arg:N}, (%lam {arg:N} (%case (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [(%case (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing)]) [!{arg:N}]))] Nothing))])
Prelude.Types.rangeFromTo = [{arg:N}, {arg:N}, {arg:N}]: (Prelude.Types.case block in rangeFromTo [!{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)])
Prelude.Types.rangeFromThen = [{arg:N}, {arg:N}, {arg:N}]: (Prelude.Types.case block in rangeFromThen [!{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)])
Prelude.Types.rangeFromThenTo = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (Prelude.Types.case block in rangeFromThenTo [!{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}, (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)]))] Nothing)])
Prelude.Types.null = [{arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Nil Just 0 [] (%delay Lazy 0)), (%concase Prelude.Types.:: Just 1 [{e:N}, {e:N}] (%delay Lazy 1))] Nothing)
Prelude.Types.foldr = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Nil Just 0 [] !{arg:N}), (%concase Prelude.Types.:: Just 1 [{e:N}, {e:N}] ((!{arg:N} [!{e:N}]) [(Prelude.Types.foldr [!{arg:N}, !{arg:N}, !{e:N}])]))] Nothing)
Prelude.Types.foldl = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Nil Just 0 [] !{arg:N}), (%concase Prelude.Types.:: Just 1 [{e:N}, {e:N}] (Prelude.Types.foldl [!{arg:N}, ((!{arg:N} [!{arg:N}]) [!{e:N}]), !{e:N}]))] Nothing)
Prelude.Types.foldlM = [{arg:N}, {arg:N}, {arg:N}, {ext:N}]: (Prelude.Types.foldl [(%lam ma (%lam b (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((((!{e:N} [___]) [___]) [!ma]) [(%lam {eta:N} (Prelude.Basics.flip [!{arg:N}, !b, !{eta:N}]))]))] Nothing))), (%case (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing) [(%concase Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{arg:N}]))] Nothing), !{ext:N}])
Prelude.Types.Range at Prelude/Types.idr:L:C--L:C = Constructor tag Just 0 arity 4
Prelude.Types.Range implementation at Prelude/Types.idr:L:C--L:C = [{arg:N}, {arg:N}]: (%con Prelude.Types.Range at Prelude/Types.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.Types.rangeFromTo [___, !{arg:N}, !{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (%lam {arg:N} (Prelude.Types.rangeFromThenTo [___, !{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}])))), (%lam {arg:N} (Prelude.Types.rangeFrom [___, !{arg:N}, !{arg:N}])), (%lam {arg:N} (%lam {arg:N} (Prelude.Types.rangeFromThen [___, !{arg:N}, !{arg:N}, !{arg:N}])))])
Prelude.Types.Foldable implementation at Prelude/Types.idr:L:C--L:C = []: (%con Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [___, ___, !func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [___, ___, !func, !init, !input])))))), (%lam elem (%lam {arg:N} (Prelude.Types.null [___, !{arg:N}]))), (%lam elem (%lam acc (%lam m (%lam {i_con:N} (%lam funcM (%lam init (%lam input (Prelude.Types.foldlM [___, ___, ___, !{i_con:N}, !funcM, !init, !input]))))))))])
Prelude.Types.takeUntil = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Stream.:: Just 0 [{e:N}, {e:N}] (Prelude.Types.case block in takeUntil [___, !{e:N}, !{e:N}, !{arg:N}, (!{arg:N} [!{e:N}])]))] Nothing)
Prelude.Types.takeBefore = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Stream.:: Just 0 [{e:N}, {e:N}] (Prelude.Types.case block in takeBefore [___, !{e:N}, !{e:N}, !{arg:N}, (!{arg:N} [!{e:N}])]))] Nothing)
Prelude.Types.rangeFromTo = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Types.Range at Prelude/Types.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.Types.Range implementation at Prelude/Types.idr:L:C--L:C = [{arg:N}]: (%con Prelude.Types.Range at Prelude/Types.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.Types.rangeFromTo [!{arg:N}, !{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (%lam {arg:N} (Prelude.Types.rangeFromThenTo [!{arg:N}, !{arg:N}, !{arg:N}, !{arg:N}])))), (%lam {arg:N} (Prelude.Types.rangeFrom [!{arg:N}, !{arg:N}])), (%lam {arg:N} (%lam {arg:N} (Prelude.Types.rangeFromThen [!{arg:N}, !{arg:N}, !{arg:N}])))])
Prelude.Types.Foldable implementation at Prelude/Types.idr:L:C--L:C = []: (%con Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [!func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [!func, !init, !input])))))), (%lam elem (%lam {arg:N} (Prelude.Types.null [!{arg:N}]))), (%lam elem (%lam acc (%lam m (%lam {i_con:N} (%lam funcM (%lam init (%lam input (Prelude.Types.foldlM [!{i_con:N}, !funcM, !init, !input]))))))))])
Prelude.Types.takeUntil = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Stream.:: Just 0 [{e:N}, {e:N}] (Prelude.Types.case block in takeUntil [!{e:N}, !{e:N}, !{arg:N}, (!{arg:N} [!{e:N}])]))] Nothing)
Prelude.Types.takeBefore = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Types.Stream.:: Just 0 [{e:N}, {e:N}] (Prelude.Types.case block in takeBefore [!{e:N}, !{e:N}, !{arg:N}, (!{arg:N} [!{e:N}])]))] Nothing)
Prelude.Types.rangeFromTo = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Types.Range at Prelude/Types.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.Types.prim__integerToNat = [{arg:N}]: (Prelude.Types.case block in prim__integerToNat [!{arg:N}, (%case (<=Integer [0, !{arg:N}]) [(%constcase 0 1)] Just 0)])
Prelude.Types.countFrom = [{arg:N}, {arg:N}, {arg:N}]: (%con Prelude.Types.Stream.:: Just 0 [!{arg:N}, (%delay Inf (Prelude.Types.countFrom [___, (!{arg:N} [!{arg:N}]), !{arg:N}]))])
Prelude.Types.countFrom = [{arg:N}, {arg:N}]: (%con Prelude.Types.Stream.:: Just 0 [!{arg:N}, (%delay Inf (Prelude.Types.countFrom [(!{arg:N} [!{arg:N}]), !{arg:N}]))])
Prelude.Types.Z = Constructor tag Just 0 arity 0
Prelude.Types.Stream.Stream = Constructor tag Nothing arity 1
Prelude.Types.Nil = Constructor tag Just 0 arity 0
Prelude.Types.:: = Constructor tag Just 1 arity 2
Prelude.Types.Stream.:: = Constructor tag Just 0 arity 2
Prelude.Num.case block in mod = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 1 (%Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash [___, "Unhandled input for Prelude.Num.case block in mod at Prelude/Num.idr:L:C--L:C"]))
Prelude.Num.case block in div = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 1 (/Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash [___, "Unhandled input for Prelude.Num.case block in div at Prelude/Num.idr:L:C--L:C"]))
Prelude.Num.case block in mod = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 1 (%Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash ["Unhandled input for Prelude.Num.case block in mod at Prelude/Num.idr:L:C--L:C"]))
Prelude.Num.case block in div = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 1 (/Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash ["Unhandled input for Prelude.Num.case block in div at Prelude/Num.idr:L:C--L:C"]))
Prelude.Num.negate = [{arg:N}]: (-Int [0, !{arg:N}])
Prelude.Num.mod = [{arg:N}, {arg:N}]: (Prelude.Num.case block in mod [!{arg:N}, !{arg:N}, (Prelude.EqOrd.== [!{arg:N}, (cast-Integer-Int [0])])])
Prelude.Num.fromInteger = [{ext:N}]: (cast-Integer-Int [!{ext:N}])
@ -69,16 +69,16 @@ Prelude.Num.div = [{arg:N}, {arg:N}]: (Prelude.Num.case block in div [!{arg:N},
Prelude.Num.Num at Prelude/Num.idr:L:C--L:C = Constructor tag Just 0 arity 3
Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C = Constructor tag Just 0 arity 3
Prelude.Num.Integral at Prelude/Num.idr:L:C--L:C = Constructor tag Just 0 arity 3
Prelude.Num.Constraint (Num ty) = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing)
Prelude.Num.Constraint (Num ty) = [{arg:N}]: (%case !{arg:N} [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing)
Prelude.Num.Num implementation at Prelude/Num.idr:L:C--L:C = []: (%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))])
Prelude.Num.Neg implementation at Prelude/Num.idr:L:C--L:C = []: (%con Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [(%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (Prelude.Num.negate [!{arg:N}])), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.- [!{arg:N}, !{arg:N}])))])
Prelude.Num.Integral implementation at Prelude/Num.idr:L:C--L:C = []: (%con Prelude.Num.Integral at Prelude/Num.idr:L:C--L:C Just 0 [(%con Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.div [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.mod [!{arg:N}, !{arg:N}])))])
Prelude.Num.- = [{ext:N}, {ext:N}]: (-Int [!{ext:N}, !{ext:N}])
Prelude.Num.+ = [{ext:N}, {ext:N}]: (+Int [!{ext:N}, !{ext:N}])
Prelude.Num.* = [{ext:N}, {ext:N}]: (*Int [!{ext:N}, !{ext:N}])
Prelude.Num.fromInteger = [{arg:N}, {arg:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [!{ext:N}]))] Nothing)
Prelude.Num.- = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.Num.+ = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.Num.fromInteger = [{arg:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [!{ext:N}]))] Nothing)
Prelude.Num.- = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Num.Neg at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.Num.+ = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.case block in case block in compare = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 1), (%constcase 1 2)] Nothing)
Prelude.EqOrd.case block in compare = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 0), (%constcase 1 (Prelude.EqOrd.case block in case block in compare [!{arg:N}, !{arg:N}, (Prelude.EqOrd.== [!{arg:N}, !{arg:N}])]))] Nothing)
Prelude.EqOrd.case block in max = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 !{arg:N}), (%constcase 1 !{arg:N})] Nothing)
@ -88,7 +88,7 @@ Prelude.EqOrd.max = [{arg:N}, {arg:N}]: (Prelude.EqOrd.case block in max [!{arg:
Prelude.EqOrd.compare = [{arg:N}, {arg:N}]: (Prelude.EqOrd.case block in compare [!{arg:N}, !{arg:N}, (Prelude.EqOrd.< [!{arg:N}, !{arg:N}])])
Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C = Constructor tag Just 0 arity 8
Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C = Constructor tag Just 0 arity 2
Prelude.EqOrd.Constraint (Eq ty) = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] !{e:N})] Nothing)
Prelude.EqOrd.Constraint (Eq ty) = [{arg:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] !{e:N})] Nothing)
Prelude.EqOrd.Ord implementation at Prelude/EqOrd.idr:L:C--L:C = []: (%con Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [(%con Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd./= [!{arg:N}, !{arg:N}])))]), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.compare [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.> [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.<= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.>= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.max [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.min [!{arg:N}, !{arg:N}])))])
Prelude.EqOrd.Eq implementation at Prelude/EqOrd.idr:L:C--L:C = []: (%con Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd./= [!{arg:N}, !{arg:N}])))])
Prelude.EqOrd.> = [{arg:N}, {arg:N}]: (%case (>Int [!{arg:N}, !{arg:N}]) [(%constcase 0 1)] Just 0)
@ -100,33 +100,33 @@ Prelude.EqOrd./= = [{arg:N}, {arg:N}]: (Prelude.Basics.not [(Prelude.EqOrd.== [!
Prelude.EqOrd.LT = Constructor tag Just 0 arity 0
Prelude.EqOrd.GT = Constructor tag Just 2 arity 0
Prelude.EqOrd.EQ = Constructor tag Just 1 arity 0
Prelude.EqOrd.>= = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.> = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.== = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.<= = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.< = [{arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.>= = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.> = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.== = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Eq at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.<= = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.EqOrd.< = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.EqOrd.Ord at Prelude/EqOrd.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{ext:N}]) [!{ext:N}]))] Nothing)
Prelude.Interfaces.Functor at Prelude/Interfaces.idr:L:C--L:C = Constructor tag Just 0 arity 2 (newtype by 1)
Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C = Constructor tag Just 0 arity 4
Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C = Constructor tag Just 0 arity 3
Prelude.Interfaces.Constraint (Applicative m) = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing)
Prelude.Interfaces.sum = [{arg:N}, {arg:N}, {arg:N}, {ext:N}]: (%case (Builtin.fst [___, ___, !{arg:N}]) [(%concase Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] (((((!{e:N} [___]) [___]) [(%lam {eta:N} (%lam {clam:N} (%case (Builtin.snd [___, ___, !{arg:N}]) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{eta:N}]) [!{clam:N}]))] Nothing)))]) [(%case (Builtin.snd [___, ___, !{arg:N}]) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [0]))] Nothing)]) [!{ext:N}]))] Nothing)
Prelude.Interfaces.pure = [{arg:N}, {arg:N}, {arg:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{ext:N}]))] Nothing)
Prelude.Interfaces.foldr = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] (%lam {arg:N} (((((!{e:N} [___]) [___]) [!{ext:N}]) [!{ext:N}]) [!{arg:N}])))] Nothing)
Prelude.Interfaces.foldl = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] (%lam {arg:N} (((((!{e:N} [___]) [___]) [!{ext:N}]) [!{ext:N}]) [!{arg:N}])))] Nothing)
Prelude.Interfaces.>>= = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((((!{e:N} [___]) [___]) [!{ext:N}]) [!{ext:N}]))] Nothing)
PrimIO.case block in unsafePerformIO = [{arg:N}, {arg:N}, {arg:N}, {arg:N}]: (PrimIO.unsafeDestroyWorld [___, ___, !{arg:N}])
PrimIO.case block in case block in io_bind = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (!{arg:N} [!{arg:N}])
PrimIO.case block in io_bind = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}, {arg:N}]: (PrimIO.case block in case block in io_bind [___, ___, ___, ___, ___, !{arg:N}, ___, (!{arg:N} [!{arg:N}])])
PrimIO.unsafePerformIO = [{arg:N}, {arg:N}]: (PrimIO.unsafeCreateWorld [___, (%lam w (PrimIO.case block in unsafePerformIO [___, !{arg:N}, ___, (!{arg:N} [!w])]))])
PrimIO.unsafeDestroyWorld = [{arg:N}, {arg:N}, {arg:N}]: !{arg:N}
PrimIO.unsafeCreateWorld = [{arg:N}, {arg:N}]: (!{arg:N} [%MkWorld])
PrimIO.io_pure = [{arg:N}, {arg:N}, {ext:N}]: !{arg:N}
PrimIO.io_bind = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}]: (PrimIO.case block in io_bind [___, ___, ___, !{arg:N}, ___, (!{arg:N} [!{ext:N}])])
Prelude.Interfaces.Constraint (Applicative m) = [{arg:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] !{e:N})] Nothing)
Prelude.Interfaces.sum = [{arg:N}, {ext:N}]: (%case (Builtin.fst [!{arg:N}]) [(%concase Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] (((((!{e:N} [___]) [___]) [(%lam {eta:N} (%lam {clam:N} (%case (Builtin.snd [!{arg:N}]) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{eta:N}]) [!{clam:N}]))] Nothing)))]) [(%case (Builtin.snd [!{arg:N}]) [(%concase Prelude.Num.Num at Prelude/Num.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [0]))] Nothing)]) [!{ext:N}]))] Nothing)
Prelude.Interfaces.pure = [{arg:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{ext:N}]))] Nothing)
Prelude.Interfaces.foldr = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] (%lam {arg:N} (((((!{e:N} [___]) [___]) [!{ext:N}]) [!{ext:N}]) [!{arg:N}])))] Nothing)
Prelude.Interfaces.foldl = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Foldable at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] (%lam {arg:N} (((((!{e:N} [___]) [___]) [!{ext:N}]) [!{ext:N}]) [!{arg:N}])))] Nothing)
Prelude.Interfaces.>>= = [{arg:N}, {ext:N}, {ext:N}]: (%case !{arg:N} [(%concase Prelude.Interfaces.Monad at Prelude/Interfaces.idr:L:C--L:C Just 0 [{e:N}, {e:N}, {e:N}] ((((!{e:N} [___]) [___]) [!{ext:N}]) [!{ext:N}]))] Nothing)
PrimIO.case block in unsafePerformIO = [{arg:N}, {arg:N}]: (PrimIO.unsafeDestroyWorld [___, !{arg:N}])
PrimIO.case block in case block in io_bind = [{arg:N}, {arg:N}, {arg:N}]: (!{arg:N} [!{arg:N}])
PrimIO.case block in io_bind = [{arg:N}, {arg:N}]: (PrimIO.case block in case block in io_bind [!{arg:N}, ___, (!{arg:N} [!{arg:N}])])
PrimIO.unsafePerformIO = [{arg:N}]: (PrimIO.unsafeCreateWorld [(%lam w (PrimIO.case block in unsafePerformIO [!{arg:N}, (!{arg:N} [!w])]))])
PrimIO.unsafeDestroyWorld = [{arg:N}, {arg:N}]: !{arg:N}
PrimIO.unsafeCreateWorld = [{arg:N}]: (!{arg:N} [%MkWorld])
PrimIO.io_pure = [{arg:N}, {ext:N}]: !{arg:N}
PrimIO.io_bind = [{arg:N}, {arg:N}, {ext:N}]: (PrimIO.case block in io_bind [!{arg:N}, (!{arg:N} [!{ext:N}])])
PrimIO.MkIORes = Constructor tag Just 0 arity 3 (newtype by 1)
PrimIO.MkIO = Constructor tag Just 0 arity 2 (newtype by 1)
Prelude.IO.pure = [{arg:N}, {arg:N}, {ext:N}]: !{arg:N}
Prelude.IO.map = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (!{arg:N} [!{act:N}]))
Prelude.IO.Functor implementation at Prelude/IO.idr:L:C--L:C = [{ext:N}, {ext:N}, {ext:N}, {ext:N}, {ext:N}]: (Prelude.IO.map [___, ___, !{ext:N}, !{ext:N}, !{ext:N}])
Prelude.IO.Applicative implementation at Prelude/IO.idr:L:C--L:C = []: (%con Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam b (%lam a (%lam func (%lam {arg:N} (%lam {eta:N} (Prelude.IO.map [___, ___, !func, !{arg:N}, !{eta:N}])))))), (%lam a (%lam {arg:N} (%lam {eta:N} !{arg:N}))), (%lam b (%lam a (%lam {arg:N} (%lam {arg:N} (%lam {eta:N} (%let {act:N} (!{arg:N} [!{eta:N}]) (%let {act:N} (!{arg:N} [!{eta:N}]) (!{act:N} [!{act:N}]))))))))])
Prelude.IO.<*> = [{arg:N}, {arg:N}, {arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (%let {act:N} (!{arg:N} [!{ext:N}]) (!{act:N} [!{act:N}])))
Prelude.IO.pure = [{arg:N}, {ext:N}]: !{arg:N}
Prelude.IO.map = [{arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (!{arg:N} [!{act:N}]))
Prelude.IO.Functor implementation at Prelude/IO.idr:L:C--L:C = [{ext:N}, {ext:N}, {ext:N}, {ext:N}, {ext:N}]: (Prelude.IO.map [!{ext:N}, !{ext:N}, !{ext:N}])
Prelude.IO.Applicative implementation at Prelude/IO.idr:L:C--L:C = []: (%con Prelude.Interfaces.Applicative at Prelude/Interfaces.idr:L:C--L:C Just 0 [(%lam b (%lam a (%lam func (%lam {arg:N} (%lam {eta:N} (Prelude.IO.map [!func, !{arg:N}, !{eta:N}])))))), (%lam a (%lam {arg:N} (%lam {eta:N} !{arg:N}))), (%lam b (%lam a (%lam {arg:N} (%lam {arg:N} (%lam {eta:N} (%let {act:N} (!{arg:N} [!{eta:N}]) (%let {act:N} (!{arg:N} [!{eta:N}]) (!{act:N} [!{act:N}]))))))))])
Prelude.IO.<*> = [{arg:N}, {arg:N}, {ext:N}]: (%let {act:N} (!{arg:N} [!{ext:N}]) (%let {act:N} (!{arg:N} [!{ext:N}]) (!{act:N} [!{act:N}])))