mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-15 06:13:26 +03:00
[cleanup] Cleanup after rebasing
This commit is contained in:
parent
26b94f09f7
commit
1e6f9dad71
@ -1001,9 +1001,9 @@ footer = do
|
|||||||
// main function
|
// main function
|
||||||
int main(int argc, char *argv[])
|
int main(int argc, char *argv[])
|
||||||
{
|
{
|
||||||
\{
|
\{ ifThenElse (contains "idris_support.h" !(get HeaderFiles))
|
||||||
(if contains "idris_support.h" !(get HeaderFiles)
|
"idris2_setArgs(argc, argv);"
|
||||||
then "idris2_setArgs(argc, argv);" else "")
|
""
|
||||||
}
|
}
|
||||||
Value *mainExprVal = __mainExpression_0();
|
Value *mainExprVal = __mainExpression_0();
|
||||||
trampoline(mainExprVal);
|
trampoline(mainExprVal);
|
||||||
|
@ -99,11 +99,10 @@ schHeader chez libs whole
|
|||||||
|
|
||||||
\{ showSep "\n" (map (\x => "(load-shared-object \"" ++ escapeStringChez x ++ "\")") libs) }
|
\{ showSep "\n" (map (\x => "(load-shared-object \"" ++ escapeStringChez x ++ "\")") libs) }
|
||||||
|
|
||||||
\{
|
\{ ifThenElse whole
|
||||||
(if whole
|
"(let ()"
|
||||||
then "(let ()"
|
"(source-directories (cons (getenv \"IDRIS2_INC_SRC\") (source-directories)))"
|
||||||
else "(source-directories (cons (getenv \"IDRIS2_INC_SRC\") (source-directories)))")
|
}
|
||||||
}
|
|
||||||
|
|
||||||
"""
|
"""
|
||||||
|
|
||||||
@ -112,8 +111,8 @@ schFooter prof whole = """
|
|||||||
|
|
||||||
(collect 4)
|
(collect 4)
|
||||||
(blodwen-run-finalisers)
|
(blodwen-run-finalisers)
|
||||||
\{ (if prof then "(profile-dump-html)" else "") }
|
\{ ifThenElse prof "(profile-dump-html)" "" }
|
||||||
\{ (if whole then ")" else "") }
|
\{ ifThenElse whole ")" "" }
|
||||||
"""
|
"""
|
||||||
|
|
||||||
showChezChar : Char -> String -> String
|
showChezChar : Char -> String -> String
|
||||||
@ -391,9 +390,9 @@ getFgnCall version (n, fc, d) = schFgnDef fc n d version
|
|||||||
|
|
||||||
export
|
export
|
||||||
startChezPreamble : String
|
startChezPreamble : String
|
||||||
startChezPreamble = #"""
|
startChezPreamble = """
|
||||||
#!/bin/sh
|
#!/bin/sh
|
||||||
# \#{ (generatedString "Chez") }
|
# \{ (generatedString "Chez") }
|
||||||
|
|
||||||
set -e # exit on any error
|
set -e # exit on any error
|
||||||
|
|
||||||
@ -403,44 +402,44 @@ startChezPreamble = #"""
|
|||||||
DIR=$(dirname "$(readlink -f -- "$0")")
|
DIR=$(dirname "$(readlink -f -- "$0")")
|
||||||
fi
|
fi
|
||||||
|
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
startChez : String -> String -> String
|
startChez : String -> String -> String
|
||||||
startChez appdir target = startChezPreamble ++ #"""
|
startChez appdir target = startChezPreamble ++ """
|
||||||
export LD_LIBRARY_PATH="$DIR/\#{ appdir }":$LD_LIBRARY_PATH
|
export LD_LIBRARY_PATH="$DIR/\{ appdir }":$LD_LIBRARY_PATH
|
||||||
export DYLD_LIBRARY_PATH="$DIR/\#{ appdir }":$DYLD_LIBRARY_PATH
|
export DYLD_LIBRARY_PATH="$DIR/\{ appdir }":$DYLD_LIBRARY_PATH
|
||||||
export IDRIS2_INC_SRC="$DIR/\#{ appdir }"
|
export IDRIS2_INC_SRC="$DIR/\{ appdir }"
|
||||||
|
|
||||||
"$DIR/\#{ target }" "$@"
|
"$DIR/\{ target }" "$@"
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
startChezCmd : String -> String -> String -> String
|
startChezCmd : String -> String -> String -> String -> String
|
||||||
startChezCmd chez appdir target progType = #"""
|
startChezCmd chez appdir target progType = """
|
||||||
@echo off
|
@echo off
|
||||||
|
|
||||||
rem \#{ (generatedString "Chez") }
|
rem \{ (generatedString "Chez") }
|
||||||
|
|
||||||
set APPDIR=%~dp0
|
set APPDIR=%~dp0
|
||||||
set PATH=%APPDIR%\#{ appdir };%PATH%
|
set PATH=%APPDIR%\{ appdir };%PATH%
|
||||||
set IDRIS2_INC_SRC=%APPDIR%\#{ appdir }
|
set IDRIS2_INC_SRC=%APPDIR%\{ appdir }
|
||||||
|
|
||||||
"\#{ chez }" \#{ progType } "%APPDIR%\#{ target }" %*
|
"\{ chez }" \{ progType } "%APPDIR%\{ target }" %*
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
startChezWinSh : String -> String -> String -> String
|
startChezWinSh : String -> String -> String -> String -> String
|
||||||
startChezWinSh chez appdir target progType = #"""
|
startChezWinSh chez appdir target progType = """
|
||||||
#!/bin/sh
|
#!/bin/sh
|
||||||
# \#{ (generatedString "Chez") }
|
# \{ (generatedString "Chez") }
|
||||||
|
|
||||||
set -e # exit on any error
|
set -e # exit on any error
|
||||||
|
|
||||||
DIR=$(dirname "$(readlink -f -- "$0" || cygpath -a -- "$0")")
|
DIR=$(dirname "$(readlink -f -- "$0" || cygpath -a -- "$0")")
|
||||||
PATH="$DIR/\#{ appdir }":$PATH
|
PATH="$DIR/\{ appdir }":$PATH
|
||||||
|
|
||||||
export IDRIS2_INC_SRC="$DIR/\#{ appdir }"
|
export IDRIS2_INC_SRC="$DIR/\{ appdir }"
|
||||||
|
|
||||||
"\#{ chez }" \#{ progType } "$DIR/\#{ target }" "$@"
|
"\{ chez }" \{ progType } "$DIR/\{ target }" "$@"
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
||| Compile a TT expression to Chez Scheme
|
||| Compile a TT expression to Chez Scheme
|
||||||
compileToSS : Ref Ctxt Defs ->
|
compileToSS : Ref Ctxt Defs ->
|
||||||
|
@ -47,9 +47,8 @@ schHeader libs compilationUnits = """
|
|||||||
[(i3osx ti3osx a6osx ta6osx tarm64osx) (load-shared-object "libc.dylib")]
|
[(i3osx ti3osx a6osx ta6osx tarm64osx) (load-shared-object "libc.dylib")]
|
||||||
[(i3nt ti3nt a6nt ta6nt) (load-shared-object "msvcrt.dll")]
|
[(i3nt ti3nt a6nt ta6nt) (load-shared-object "msvcrt.dll")]
|
||||||
[else (load-shared-object "libc.so")]
|
[else (load-shared-object "libc.so")]
|
||||||
\{
|
\{ unlines [" (load-shared-object \"" ++ escapeStringChez lib ++ "\")" | lib <- libs] })
|
||||||
unlines [" (load-shared-object \"" ++ escapeStringChez lib ++ "\")" | lib <- libs]
|
|
||||||
})
|
|
||||||
"""
|
"""
|
||||||
|
|
||||||
schFooter : String
|
schFooter : String
|
||||||
@ -60,47 +59,47 @@ schFooter = """
|
|||||||
"""
|
"""
|
||||||
|
|
||||||
startChez : String -> String -> String -> String
|
startChez : String -> String -> String -> String
|
||||||
startChez chez appDirSh targetSh = Chez.startChezPreamble ++ #"""
|
startChez chez appDirSh targetSh = Chez.startChezPreamble ++ """
|
||||||
export LD_LIBRARY_PATH="$DIR/\#{ appDirSh }":$LD_LIBRARY_PATH
|
export LD_LIBRARY_PATH="$DIR/\{ appDirSh }":$LD_LIBRARY_PATH
|
||||||
export DYLD_LIBRARY_PATH="$DIR/\#{ appDirSh }":$DYLD_LIBRARY_PATH
|
export DYLD_LIBRARY_PATH="$DIR/\{ appDirSh }":$DYLD_LIBRARY_PATH
|
||||||
|
|
||||||
"\#{ chez }" -q \
|
"\{ chez }" -q \
|
||||||
--libdirs "$DIR/\#{ appDirSh }" \
|
--libdirs "$DIR/\{ appDirSh }" \
|
||||||
--program "$DIR/\#{ targetSh }" \
|
--program "$DIR/\{ targetSh }" \
|
||||||
"$@"
|
"$@"
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
startChezCmd : String -> String -> String -> String
|
startChezCmd : String -> String -> String -> String
|
||||||
startChezCmd chez appDirSh targetSh = #"""
|
startChezCmd chez appDirSh targetSh = """
|
||||||
@echo off
|
@echo off
|
||||||
|
|
||||||
rem \#{ (generatedString "ChezSep") }
|
rem \{ (generatedString "ChezSep") }
|
||||||
|
|
||||||
set APPDIR=%~dp0
|
set APPDIR=%~dp0
|
||||||
set PATH=%APPDIR%\#{ appDirSh };%PATH%
|
set PATH=%APPDIR%\{ appDirSh };%PATH%
|
||||||
|
|
||||||
"\#{ chez }" -q \
|
"\{ chez }" -q \
|
||||||
--libdirs "%APPDIR%\#{ appDirSh }" \
|
--libdirs "%APPDIR%\{ appDirSh }" \
|
||||||
--program "%APPDIR%\#{ targetSh }" \
|
--program "%APPDIR%\{ targetSh }" \
|
||||||
%*
|
%*
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
startChezWinSh : String -> String -> String -> String
|
startChezWinSh : String -> String -> String -> String
|
||||||
startChezWinSh chez appDirSh targetSh = #"""
|
startChezWinSh chez appDirSh targetSh = """
|
||||||
#!/bin/sh
|
#!/bin/sh
|
||||||
# \#{ (generatedString "ChezSep") }
|
# \{ (generatedString "ChezSep") }
|
||||||
|
|
||||||
set -e # exit on any error
|
set -e # exit on any error
|
||||||
|
|
||||||
DIR=$(dirname "$(readlink -f -- "$0" || cygpath -a -- "$0")")
|
DIR=$(dirname "$(readlink -f -- "$0" || cygpath -a -- "$0")")
|
||||||
PATH="$DIR/\#{ appDirSh }":$PATH
|
PATH="$DIR/\{ appDirSh }":$PATH
|
||||||
|
|
||||||
"\#{ chez }" --program "$DIR/\#{ targetSh }" "$@"
|
"\{ chez }" --program "$DIR/\{ targetSh }" "$@"
|
||||||
"\#{ chez }" -q \
|
"\{ chez }" -q \
|
||||||
--libdirs "$DIR/\#{ appDirSh }" \
|
--libdirs "$DIR/\{ appDirSh }" \
|
||||||
--program "$DIR/\#{ targetSh }" \
|
--program "$DIR/\{ targetSh }" \
|
||||||
"$@"
|
"$@"
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
-- TODO: parallelise this
|
-- TODO: parallelise this
|
||||||
compileChezLibraries : (chez : String) -> (libDir : String) -> (ssFiles : List String) -> Core ()
|
compileChezLibraries : (chez : String) -> (libDir : String) -> (ssFiles : List String) -> Core ()
|
||||||
|
@ -53,7 +53,7 @@ schHeader prof libs = """
|
|||||||
(require rnrs/io/ports-6) ; for files
|
(require rnrs/io/ports-6) ; for files
|
||||||
(require srfi/19) ; for file handling and data
|
(require srfi/19) ; for file handling and data
|
||||||
(require ffi/unsafe ffi/unsafe/define) ; for calling C
|
(require ffi/unsafe ffi/unsafe/define) ; for calling C
|
||||||
\{ (if prof then "(require profile)" else "") }
|
\{ ifThenElse prof "(require profile)" "" }
|
||||||
(require racket/flonum) ; for float-typed transcendental functions
|
(require racket/flonum) ; for float-typed transcendental functions
|
||||||
\{ libs }
|
\{ libs }
|
||||||
(let ()
|
(let ()
|
||||||
@ -335,9 +335,9 @@ getFgnCall : {auto f : Ref Done (List String) } ->
|
|||||||
getFgnCall appdir (n, fc, d) = schFgnDef appdir fc n d
|
getFgnCall appdir (n, fc, d) = schFgnDef appdir fc n d
|
||||||
|
|
||||||
startRacket : String -> String -> String -> String
|
startRacket : String -> String -> String -> String
|
||||||
startRacket racket appdir target = #"""
|
startRacket racket appdir target = """
|
||||||
#!/bin/sh
|
#!/bin/sh
|
||||||
# \#{ (generatedString "Racket") }
|
# \{ (generatedString "Racket") }
|
||||||
|
|
||||||
set -e # exit on any error
|
set -e # exit on any error
|
||||||
|
|
||||||
@ -347,36 +347,36 @@ startRacket racket appdir target = #"""
|
|||||||
DIR=$(dirname "$(readlink -f -- "$0")")
|
DIR=$(dirname "$(readlink -f -- "$0")")
|
||||||
fi
|
fi
|
||||||
|
|
||||||
export LD_LIBRARY_PATH="$DIR/\#{ appdir }":$LD_LIBRARY_PATH
|
export LD_LIBRARY_PATH="$DIR/\{ appdir }":$LD_LIBRARY_PATH
|
||||||
export DYLD_LIBRARY_PATH="$DIR/\#{ appdir }":$DYLD_LIBRARY_PATH
|
export DYLD_LIBRARY_PATH="$DIR/\{ appdir }":$DYLD_LIBRARY_PATH
|
||||||
|
|
||||||
\#{ racket } "$DIR/\#{ target }" "$@"
|
\{ racket } "$DIR/\{ target }" "$@"
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
startRacketCmd : String -> String -> String -> String
|
startRacketCmd : String -> String -> String -> String
|
||||||
startRacketCmd racket appdir target = #"""
|
startRacketCmd racket appdir target = """
|
||||||
@echo off
|
@echo off
|
||||||
|
|
||||||
rem \#{ (generatedString "Racket") }
|
rem \{ (generatedString "Racket") }
|
||||||
|
|
||||||
set APPDIR=%~dp0
|
set APPDIR=%~dp0
|
||||||
set PATH=%APPDIR%\#{ appdir };%PATH%
|
set PATH=%APPDIR%\{ appdir };%PATH%
|
||||||
|
|
||||||
\#{ racket } "%APPDIR%\#{ target }" %*
|
\{ racket } "%APPDIR%\{ target }" %*
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
startRacketWinSh : String -> String -> String -> String
|
startRacketWinSh : String -> String -> String -> String
|
||||||
startRacketWinSh racket appdir target = #"""
|
startRacketWinSh racket appdir target = """
|
||||||
#!/bin/sh
|
#!/bin/sh
|
||||||
# \#{ (generatedString "Racket") }
|
# \{ (generatedString "Racket") }
|
||||||
|
|
||||||
set -e # exit on any error
|
set -e # exit on any error
|
||||||
|
|
||||||
DIR=$(dirname "$(readlink -f -- "$0" || cygpath -a -- "$0")")
|
DIR=$(dirname "$(readlink -f -- "$0" || cygpath -a -- "$0")")
|
||||||
PATH="$DIR/\#{ appdir }":$PATH
|
PATH="$DIR/\{ appdir }":$PATH
|
||||||
|
|
||||||
\#{ racket } "$DIR/\#{ target }" "$@"
|
\{ racket } "$DIR/\{ target }" "$@"
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
compileToRKT : Ref Ctxt Defs ->
|
compileToRKT : Ref Ctxt Defs ->
|
||||||
String -> ClosedTerm -> (outfile : String) -> Core ()
|
String -> ClosedTerm -> (outfile : String) -> Core ()
|
||||||
|
@ -41,12 +41,12 @@ outputDirWithDefault d = fromMaybe (build_dir d </> "exec") (output_dir d)
|
|||||||
public export
|
public export
|
||||||
toString : Dirs -> String
|
toString : Dirs -> String
|
||||||
toString d@(MkDirs wdir sdir bdir ldir odir dfix edirs pdirs ldirs ddirs) = """
|
toString d@(MkDirs wdir sdir bdir ldir odir dfix edirs pdirs ldirs ddirs) = """
|
||||||
Working Directory: \{ show wdir }
|
Working Directory: \{ wdir }
|
||||||
Source Directory: \{ show sdir }
|
Source Directory: \{ show sdir }
|
||||||
Build Directory: \{ show bdir }
|
Build Directory: \{ bdir }
|
||||||
Local Depend Directory: \{ show ldir }
|
Local Depend Directory: \{ ldir }
|
||||||
Output Directory: \{ show (outputDirWithDefault d) }
|
Output Directory: \{ (outputDirWithDefault d) }
|
||||||
Installation Prefix: \{ show dfix }
|
Installation Prefix: \{ dfix }
|
||||||
Extra Directories: \{ show edirs }
|
Extra Directories: \{ show edirs }
|
||||||
Package Directories: \{ show pdirs }
|
Package Directories: \{ show pdirs }
|
||||||
CG Library Directories: \{ show ldirs }
|
CG Library Directories: \{ show ldirs }
|
||||||
|
@ -248,15 +248,15 @@ opts x _ = pure $ if (x `elem` optionFlags)
|
|||||||
|
|
||||||
-- bash autocompletion script using the given function name
|
-- bash autocompletion script using the given function name
|
||||||
completionScript : (fun : String) -> String
|
completionScript : (fun : String) -> String
|
||||||
completionScript fun = let fun' = "_" ++ fun in #"""
|
completionScript fun = let fun' = "_" ++ fun in """
|
||||||
\#{ fun' }()
|
\{ fun' }()
|
||||||
{
|
{
|
||||||
ED=$([ -z $2 ] && echo "--" || echo $2)
|
ED=$([ -z $2 ] && echo "--" || echo $2)
|
||||||
COMPREPLY=($(idris2 --bash-completion $ED $3))
|
COMPREPLY=($(idris2 --bash-completion $ED $3))
|
||||||
}
|
}
|
||||||
|
|
||||||
complete -F \#{ fun' } -o default idris2
|
complete -F \{ fun' } -o default idris2
|
||||||
"""#
|
"""
|
||||||
|
|
||||||
--------------------------------------------------------------------------------
|
--------------------------------------------------------------------------------
|
||||||
-- Processing Options
|
-- Processing Options
|
||||||
|
Loading…
Reference in New Issue
Block a user