mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 15:07:37 +03:00
Merge pull request #1393 from melted/select_foreign
Choose foreign string based on priority
This commit is contained in:
commit
1db440d3cf
@ -342,23 +342,20 @@ exists f
|
|||||||
| Left err => pure False
|
| Left err => pure False
|
||||||
closeFile ok
|
closeFile ok
|
||||||
pure True
|
pure True
|
||||||
|
-- Select the most preferred target from an ordered list of choices and
|
||||||
-- Parse a calling convention into a backend/target for the call, and
|
-- parse the calling convention into a backend/target for the call, and
|
||||||
-- a comma separated list of any other location data.
|
-- a comma separated list of any other location data. For example
|
||||||
|
-- the chez backend would supply ["scheme,chez", "scheme", "C"]. For a function with
|
||||||
|
-- more than one string, a string with "scheme" would be preferred over one
|
||||||
|
-- with "C" and "scheme,chez" would be preferred to both.
|
||||||
-- e.g. "scheme:display" - call the scheme function 'display'
|
-- e.g. "scheme:display" - call the scheme function 'display'
|
||||||
-- "C:puts,libc,stdio.h" - call the C function 'puts' which is in
|
-- "C:puts,libc,stdio.h" - call the C function 'puts' which is in
|
||||||
-- the library libc and the header stdio.h
|
-- the library libc and the header stdio.h
|
||||||
-- Returns Nothing if the string is empty (which a backend can interpret
|
-- Returns Nothing if there is no match.
|
||||||
-- however it likes)
|
|
||||||
export
|
export
|
||||||
parseCC : String -> Maybe (String, List String)
|
parseCC : List String -> List String -> Maybe (String, List String)
|
||||||
parseCC "" = Nothing
|
parseCC [] _ = Nothing
|
||||||
parseCC str
|
parseCC (target::ts) strs = findTarget target strs <|> parseCC ts strs
|
||||||
= case span (/= ':') str of
|
|
||||||
(target, "") => Just (trim target, [])
|
|
||||||
(target, opts) => Just (trim target,
|
|
||||||
map trim (getOpts
|
|
||||||
(assert_total (strTail opts))))
|
|
||||||
where
|
where
|
||||||
getOpts : String -> List String
|
getOpts : String -> List String
|
||||||
getOpts "" = []
|
getOpts "" = []
|
||||||
@ -366,6 +363,17 @@ parseCC str
|
|||||||
= case span (/= ',') str of
|
= case span (/= ',') str of
|
||||||
(opt, "") => [opt]
|
(opt, "") => [opt]
|
||||||
(opt, rest) => opt :: getOpts (assert_total (strTail rest))
|
(opt, rest) => opt :: getOpts (assert_total (strTail rest))
|
||||||
|
hasTarget : String -> String -> Bool
|
||||||
|
hasTarget target str = case span (/= ':') str of
|
||||||
|
(targetSpec, _) => targetSpec == target
|
||||||
|
findTarget : String -> List String -> Maybe (String, List String)
|
||||||
|
findTarget target [] = Nothing
|
||||||
|
findTarget target (s::xs) = if hasTarget target s
|
||||||
|
then case span (/= ':') s of
|
||||||
|
(t, "") => Just (trim t, [])
|
||||||
|
(t, opts) => Just (trim t, map trim (getOpts
|
||||||
|
(assert_total (strTail opts))))
|
||||||
|
else findTarget target xs
|
||||||
|
|
||||||
export
|
export
|
||||||
dylib_suffix : String
|
dylib_suffix : String
|
||||||
|
@ -272,10 +272,9 @@ schemeCall fc sfn argns ret
|
|||||||
useCC : {auto c : Ref Ctxt Defs} ->
|
useCC : {auto c : Ref Ctxt Defs} ->
|
||||||
{auto l : Ref Loaded (List String)} ->
|
{auto l : Ref Loaded (List String)} ->
|
||||||
String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String)
|
String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String)
|
||||||
useCC appdir fc [] args ret = throw (NoForeignCC fc)
|
useCC appdir fc ccs args ret
|
||||||
useCC appdir fc (cc :: ccs) args ret
|
= case parseCC ["scheme,chez", "scheme", "C"] ccs of
|
||||||
= case parseCC cc of
|
Nothing => throw (NoForeignCC fc)
|
||||||
Nothing => useCC appdir fc ccs args ret
|
|
||||||
Just ("scheme,chez", [sfn]) =>
|
Just ("scheme,chez", [sfn]) =>
|
||||||
do body <- schemeCall fc sfn (map fst args) ret
|
do body <- schemeCall fc sfn (map fst args) ret
|
||||||
pure ("", body)
|
pure ("", body)
|
||||||
@ -284,7 +283,7 @@ useCC appdir fc (cc :: ccs) args ret
|
|||||||
pure ("", body)
|
pure ("", body)
|
||||||
Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret
|
Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret
|
||||||
Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret
|
Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret
|
||||||
_ => useCC appdir fc ccs args ret
|
_ => throw (NoForeignCC fc)
|
||||||
|
|
||||||
-- For every foreign arg type, return a name, and whether to pass it to the
|
-- For every foreign arg type, return a name, and whether to pass it to the
|
||||||
-- foreign call (we don't pass '%World')
|
-- foreign call (we don't pass '%World')
|
||||||
|
@ -301,15 +301,14 @@ schemeCall fc sfn argns ret
|
|||||||
useCC : {auto c : Ref Ctxt Defs} ->
|
useCC : {auto c : Ref Ctxt Defs} ->
|
||||||
{auto l : Ref Loaded (List String)} ->
|
{auto l : Ref Loaded (List String)} ->
|
||||||
FC -> List String -> List (Name, CFType) -> CFType -> Core (Maybe String, (String, String))
|
FC -> List String -> List (Name, CFType) -> CFType -> Core (Maybe String, (String, String))
|
||||||
useCC fc [] args ret = throw (NoForeignCC fc)
|
useCC fc ccs args ret
|
||||||
useCC fc (cc :: ccs) args ret
|
= case parseCC ["scheme,gambit", "scheme", "C"] ccs of
|
||||||
= case parseCC cc of
|
Nothing => throw (NoForeignCC fc)
|
||||||
Nothing => useCC fc ccs args ret
|
|
||||||
Just ("scheme,gambit", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), ""))
|
Just ("scheme,gambit", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), ""))
|
||||||
Just ("scheme", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), ""))
|
Just ("scheme", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), ""))
|
||||||
Just ("C", [cfn, clib]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret))
|
Just ("C", [cfn, clib]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret))
|
||||||
Just ("C", [cfn, clib, chdr]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret))
|
Just ("C", [cfn, clib, chdr]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret))
|
||||||
_ => useCC fc ccs args ret
|
_ => throw (NoForeignCC fc)
|
||||||
where
|
where
|
||||||
fnWrapName : String -> String -> String
|
fnWrapName : String -> String -> String
|
||||||
fnWrapName cfn schemeArgName = schemeArgName ++ "-" ++ cfn ++ "-cFunWrap"
|
fnWrapName cfn schemeArgName = schemeArgName ++ "-" ++ cfn ++ "-cFunWrap"
|
||||||
|
@ -260,10 +260,9 @@ useCC : {auto f : Ref Done (List String) } ->
|
|||||||
{auto c : Ref Ctxt Defs} ->
|
{auto c : Ref Ctxt Defs} ->
|
||||||
{auto l : Ref Loaded (List String)} ->
|
{auto l : Ref Loaded (List String)} ->
|
||||||
String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String)
|
String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String)
|
||||||
useCC appdir fc [] args ret = throw (NoForeignCC fc)
|
useCC appdir fc ccs args ret
|
||||||
useCC appdir fc (cc :: ccs) args ret
|
= case parseCC ["scheme,racket", "scheme", "C"] ccs of
|
||||||
= case parseCC cc of
|
Nothing => throw (NoForeignCC fc)
|
||||||
Nothing => useCC appdir fc ccs args ret
|
|
||||||
Just ("scheme,racket", [sfn]) =>
|
Just ("scheme,racket", [sfn]) =>
|
||||||
do body <- schemeCall fc sfn (map fst args) ret
|
do body <- schemeCall fc sfn (map fst args) ret
|
||||||
pure ("", body)
|
pure ("", body)
|
||||||
@ -272,7 +271,7 @@ useCC appdir fc (cc :: ccs) args ret
|
|||||||
pure ("", body)
|
pure ("", body)
|
||||||
Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret
|
Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret
|
||||||
Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret
|
Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret
|
||||||
_ => useCC appdir fc ccs args ret
|
_ => throw (NoForeignCC fc)
|
||||||
|
|
||||||
-- For every foreign arg type, return a name, and whether to pass it to the
|
-- For every foreign arg type, return a name, and whether to pass it to the
|
||||||
-- foreign call (we don't pass '%World')
|
-- foreign call (we don't pass '%World')
|
||||||
|
Loading…
Reference in New Issue
Block a user