bulwahn@41933: module Narrowing_Engine where { bulwahn@41905: bulwahn@41905: import Monad; bulwahn@41905: import Control.Exception; bulwahn@41905: import System.IO; bulwahn@41905: import System.Exit; bulwahn@45081: import qualified Generated_Code; bulwahn@41905: bulwahn@41908: type Pos = [Int]; bulwahn@41905: bulwahn@41905: -- Term refinement bulwahn@41905: bulwahn@45081: new :: Pos -> [[Generated_Code.Narrowing_type]] -> [Generated_Code.Narrowing_term]; bulwahn@45081: new p ps = [ Generated_Code.Ctr c (zipWith (\i t -> Generated_Code.Var (p++[i]) t) [0..] ts) bulwahn@41905: | (c, ts) <- zip [0..] ps ]; bulwahn@41905: bulwahn@45081: refine :: Generated_Code.Narrowing_term -> Pos -> [Generated_Code.Narrowing_term]; bulwahn@45081: refine (Generated_Code.Var p (Generated_Code.SumOfProd ss)) [] = new p ss; bulwahn@45081: refine (Generated_Code.Ctr c xs) p = map (Generated_Code.Ctr c) (refineList xs p); bulwahn@41905: bulwahn@45081: refineList :: [Generated_Code.Narrowing_term] -> Pos -> [[Generated_Code.Narrowing_term]]; bulwahn@41908: refineList xs (i:is) = let (ls, x:rs) = splitAt i xs in [ls ++ y:rs | y <- refine x is]; bulwahn@41905: bulwahn@41905: -- Find total instantiations of a partial value bulwahn@41905: bulwahn@45081: total :: Generated_Code.Narrowing_term -> [Generated_Code.Narrowing_term]; bulwahn@45081: total (Generated_Code.Ctr c xs) = [Generated_Code.Ctr c ys | ys <- mapM total xs]; bulwahn@45081: total (Generated_Code.Var p (Generated_Code.SumOfProd ss)) = [y | x <- new p ss, y <- total x]; bulwahn@41905: bulwahn@41905: -- Answers bulwahn@41905: bulwahn@45003: answeri :: a -> (a -> IO b) -> (Pos -> IO b) -> IO b; bulwahn@45003: answeri a known unknown = bulwahn@41905: try (evaluate a) >>= (\res -> bulwahn@41905: case res of bulwahn@41905: Right b -> known b bulwahn@41908: Left (ErrorCall ('\0':p)) -> unknown (map fromEnum p) bulwahn@41905: Left e -> throw e); bulwahn@41905: bulwahn@45685: answer :: Bool -> Bool -> (Bool -> IO b) -> (Pos -> IO b) -> IO b; bulwahn@45685: answer potential a known unknown = bulwahn@45003: Control.Exception.catch (answeri a known unknown) bulwahn@45685: (\ (PatternMatchFail _) -> known (not potential)); bulwahn@45003: bulwahn@41905: -- Refute bulwahn@41905: bulwahn@41905: str_of_list [] = "[]"; bulwahn@41905: str_of_list (x:xs) = "(" ++ x ++ " :: " ++ str_of_list xs ++ ")"; bulwahn@41905: bulwahn@45081: report :: Result -> [Generated_Code.Narrowing_term] -> IO Int; bulwahn@43079: report r xs = putStrLn ("SOME (" ++ (str_of_list $ zipWith ($) (showArgs r) xs) ++ ")") >> hFlush stdout >> exitWith ExitSuccess; bulwahn@41905: bulwahn@45685: eval :: Bool -> Bool -> (Bool -> IO a) -> (Pos -> IO a) -> IO a; bulwahn@45685: eval potential p k u = answer potential p (\p -> answer potential p k u) u; bulwahn@41905: bulwahn@45685: ref :: Bool -> Result -> [Generated_Code.Narrowing_term] -> IO Int; bulwahn@45685: ref potential r xs = eval potential (apply_fun r xs) (\res -> if res then return 1 else report r xs) bulwahn@45685: (\p -> sumMapM (ref potential r) 1 (refineList xs p)); bulwahn@41905: bulwahn@45685: refute :: Bool -> Result -> IO Int; bulwahn@45685: refute potential r = ref potential r (args r); bulwahn@41905: bulwahn@41908: sumMapM :: (a -> IO Int) -> Int -> [a] -> IO Int; bulwahn@41905: sumMapM f n [] = return n; bulwahn@41905: sumMapM f n (a:as) = seq n (do m <- f a ; sumMapM f (n+m) as); bulwahn@41905: bulwahn@41905: -- Testable bulwahn@41905: bulwahn@45081: instance Show Generated_Code.Typerep where { bulwahn@45081: show (Generated_Code.Typerep c ts) = "Type (\"" ++ c ++ "\", " ++ show ts ++ ")"; bulwahn@41905: }; bulwahn@41905: bulwahn@45081: instance Show Generated_Code.Term where { bulwahn@45081: show (Generated_Code.Const c t) = "Const (\"" ++ c ++ "\", " ++ show t ++ ")"; bulwahn@45081: show (Generated_Code.App s t) = "(" ++ show s ++ ") $ (" ++ show t ++ ")"; bulwahn@45081: show (Generated_Code.Abs s ty t) = "Abs (\"" ++ s ++ "\", " ++ show ty ++ ", " ++ show t ++ ")"; bulwahn@45081: show (Generated_Code.Free s ty) = "Free (\"" ++ s ++ "\", " ++ show ty ++ ")"; bulwahn@41905: }; bulwahn@41905: bulwahn@41905: data Result = bulwahn@45081: Result { args :: [Generated_Code.Narrowing_term] bulwahn@45081: , showArgs :: [Generated_Code.Narrowing_term -> String] bulwahn@45081: , apply_fun :: [Generated_Code.Narrowing_term] -> Bool bulwahn@41905: }; bulwahn@41905: bulwahn@41908: data P = P (Int -> Int -> Result); bulwahn@41905: bulwahn@45081: run :: Testable a => ([Generated_Code.Narrowing_term] -> a) -> Int -> Int -> Result; bulwahn@41905: run a = let P f = property a in f; bulwahn@41905: bulwahn@41905: class Testable a where { bulwahn@45081: property :: ([Generated_Code.Narrowing_term] -> a) -> P; bulwahn@41905: }; bulwahn@41905: bulwahn@41905: instance Testable Bool where { bulwahn@41905: property app = P $ \n d -> Result [] [] (app . reverse); bulwahn@41905: }; bulwahn@41905: bulwahn@45081: instance (Generated_Code.Partial_term_of a, Generated_Code.Narrowing a, Testable b) => Testable (a -> b) where { bulwahn@41905: property f = P $ \n d -> bulwahn@45081: let Generated_Code.C t c = Generated_Code.narrowing d bulwahn@45081: c' = Generated_Code.conv c bulwahn@41905: r = run (\(x:xs) -> f xs (c' x)) (n+1) d bulwahn@45081: in r { args = Generated_Code.Var [n] t : args r, bulwahn@45081: showArgs = (show . Generated_Code.partial_term_of (Generated_Code.Type :: Generated_Code.Itself a)) : showArgs r }; bulwahn@41905: }; bulwahn@41905: bulwahn@41905: -- Top-level interface bulwahn@41905: bulwahn@45685: depthCheck :: Testable a => Bool -> Int -> a -> IO (); bulwahn@45685: depthCheck potential d p = bulwahn@45685: (refute potential $ run (const p) 0 d) >> putStrLn ("NONE") >> hFlush stdout; bulwahn@41905: bulwahn@45685: smallCheck :: Testable a => Bool -> Int -> a -> IO (); bulwahn@45685: smallCheck potential d p = mapM_ (\d -> depthCheck potential d p) [0..d] >> putStrLn ("NONE") >> hFlush stdout; bulwahn@41905: bulwahn@41905: } bulwahn@41905: