├── .gitignore ├── CFA ├── AFJ.hs ├── AFJ │ ├── Analysis.hs │ ├── Analysis │ │ └── NonShared.hs │ └── Examples │ │ └── AbstractNonShared.hs ├── CESK.hs ├── CESK │ ├── Analysis.hs │ ├── Analysis │ │ ├── Concrete.hs │ │ └── NonShared.hs │ ├── Examples.hs │ └── Examples │ │ ├── AbstractNonShared.hs │ │ └── Concrete.hs ├── CFAMonads.hs ├── CPS.hs ├── CPS │ ├── Analysis.hs │ ├── Analysis │ │ ├── Concrete.hs │ │ ├── Interpreter.hs │ │ ├── NonShared.hs │ │ ├── Proof.hs │ │ ├── ReallyNonShared.hs │ │ ├── Shape │ │ │ └── Basic.hs │ │ ├── SimpleAnalysis.hs │ │ ├── SingleStore.hs │ │ ├── SingleStore │ │ │ └── SingleStoreGC.hs │ │ └── SingleStoreShape.hs │ ├── Examples.hs │ ├── Examples │ │ ├── AbstractNonShared.hs │ │ ├── AbstractNonSharedCount.hs │ │ ├── AbstractShared.hs │ │ ├── AbstractSharedCount.hs │ │ ├── AbstractSharedShape.hs │ │ └── Concrete.hs │ └── KCFA.hs ├── Lattice.hs ├── Runner.hs └── Store.hs ├── CRAPL-LICENSE.txt ├── Makefile ├── README ├── Util.hs └── obsolete ├── CFA.hs ├── CPS.hs ├── MonadAFJ.hs ├── MonadCESK.hs ├── MonadCFA.hs ├── MonadCFASP.hs ├── MonadCFASPT.hs ├── MonadCFASPTCh.hs ├── MonadKCFA.hs └── XCFA.hs /.gitignore: -------------------------------------------------------------------------------- 1 | *.o 2 | *.hi -------------------------------------------------------------------------------- /CFA/AFJ.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE TypeSynonymInstances #-} 3 | {-# LANGUAGE ParallelListComp #-} 4 | {-# LANGUAGE MultiParamTypeClasses #-} 5 | {-# LANGUAGE ScopedTypeVariables #-} 6 | {-# LANGUAGE ImplicitParams #-} 7 | {-# LANGUAGE TupleSections #-} 8 | 9 | module CFA.AFJ where 10 | 11 | import Data.Set as Set 12 | import Data.Map as Map 13 | import Data.List as L 14 | 15 | import CFA.Lattice 16 | 17 | ---------------------------------------------------------------------- 18 | -- ANF FJ Syntax and gelper functions 19 | ---------------------------------------------------------------------- 20 | 21 | -- `a' stands for the nature of addresses 22 | -- it is also a leaf of the state-space 23 | 24 | type Var = String 25 | type FieldName = Var 26 | type ClassName = String 27 | type MethodName = String 28 | type Lab = String 29 | 30 | type Class = (ClassName, Maybe ClassName, [(ClassName, FieldName)], 31 | Konstr, [Method]) 32 | 33 | type Konstr = (ClassName, [(ClassName, Var)], [Var], [(FieldName, Var)]) 34 | 35 | type Method = (ClassName, MethodName, [(ClassName, Var)], 36 | [(ClassName, Var)], [Stmt]) 37 | 38 | data Stmt = Asgn Var Var Lab 39 | | Lkp Var Var FieldName Lab 40 | | MCall Var Var MethodName [Var] Lab 41 | | New Var ClassName [Var] Lab 42 | | Cast Var ClassName Var Lab 43 | | Ret Var Lab 44 | deriving (Eq, Ord, Show) 45 | 46 | data ClassTable = CTable [Class] 47 | 48 | {---------------- HELPERS ----------------} 49 | 50 | -- retrieve labels 51 | lab :: Stmt -> Lab 52 | lab (Asgn _ _ l) = l 53 | lab (Lkp _ _ _ l) = l 54 | lab (MCall _ _ _ _ l) = l 55 | lab (New _ _ _ l) = l 56 | lab (Cast _ _ _ l) = l 57 | lab (Ret _ l) = l 58 | 59 | -- get class name 60 | name :: Class -> ClassName 61 | name (cn, _, _, _, _) = cn 62 | 63 | super :: Class -> Maybe ClassName 64 | super (_, cn, _, _, _) = cn 65 | 66 | -- findClass 67 | findClass :: ClassTable -> ClassName -> Maybe Class 68 | findClass (CTable classes) cn = L.find (\c -> (name c) == cn) classes 69 | 70 | -- returns ALL fields for the given class name (including super's) 71 | allFields :: ClassTable -> ClassName -> [FieldName] 72 | allFields table cn = maybe [] (\clz@(_, _, pairs, _, _) -> 73 | let properFields = L.map snd pairs 74 | in case (super clz) of 75 | Just superName -> properFields ++ (allFields table superName) 76 | Nothing -> properFields) 77 | (findClass table cn) 78 | 79 | -- transitive method lookup 80 | method :: ClassTable -> ClassName -> MethodName -> Method 81 | method table cn m = 82 | case findClass table cn of 83 | Just clz@(_, _, _, _, methods) -> 84 | case L.find (\(_, mname, _, _, _) -> mname == m) methods of 85 | Just mtd -> mtd 86 | Nothing -> maybe undefined (\sup -> method table sup m) (super clz) 87 | Nothing -> undefined 88 | 89 | -- find class constructors 90 | konstr :: Class -> Konstr 91 | konstr (_, _, _, k, _) = k 92 | 93 | -- get the list of field mappings by a chain of constructors for a given class 94 | classFieldMappings :: ClassTable -> ClassName -> [a] -> [(FieldName, a)] 95 | classFieldMappings table cn values = 96 | maybe [] (\clz -> fieldMappings table (konstr clz) values) 97 | (findClass table cn) 98 | 99 | fieldMappings :: ClassTable -> Konstr -> [a] -> [(FieldName, a)] 100 | fieldMappings table (cn, params, superBindings, thisBindings) values = 101 | let paramBindings = Map.empty // (zip (L.map snd params) values) 102 | thisMappings = L.map (\(f, v) -> (f, paramBindings Map.! v)) thisBindings 103 | superArgs = L.map (\v -> paramBindings Map.! v) superBindings 104 | in case (superArgs, 105 | (findClass table cn) >>= (return . konstr)) of 106 | (h:t, Just superKonstr) -> 107 | thisMappings ++ (fieldMappings table superKonstr superArgs) 108 | _ -> thisMappings 109 | 110 | 111 | -------------------------------------------------------------------------------- /CFA/AFJ/Analysis.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE TypeSynonymInstances #-} 3 | {-# LANGUAGE ParallelListComp #-} 4 | {-# LANGUAGE FlexibleContexts #-} 5 | {-# LANGUAGE MultiParamTypeClasses #-} 6 | {-# LANGUAGE ScopedTypeVariables #-} 7 | {-# LANGUAGE ImplicitParams #-} 8 | {-# LANGUAGE TupleSections #-} 9 | {-# LANGUAGE FunctionalDependencies #-} 10 | {-# LANGUAGE FlexibleInstances #-} 11 | 12 | module CFA.AFJ.Analysis where 13 | 14 | import Data.Map as Map 15 | import Data.Set as Set 16 | import Data.List as L 17 | 18 | import CFA.Lattice 19 | import CFA.Store 20 | import CFA.Runner 21 | 22 | import CFA.AFJ 23 | 24 | ---------------------------------------------------------------------- 25 | -- State-Space 26 | ---------------------------------------------------------------------- 27 | 28 | -- a - for addresses 29 | type BEnv a = Var :-> a 30 | type Kont a = (Var, [Stmt], BEnv a, a) 31 | type PState a = ([Stmt], BEnv a, a) 32 | type Obj a = (ClassName, BEnv a) 33 | 34 | class (Eq a, Ord a) => Address a 35 | 36 | data Storable a = Val (Obj a) 37 | | Cont (Kont a) 38 | deriving (Eq, Ord, Show) 39 | 40 | --------------------------------------------------------------------- 41 | -- k-CFA addresses 42 | ---------------------------------------------------------------------- 43 | 44 | data Addr = AVar Var [Lab] 45 | | ACall MethodName [Lab] 46 | deriving (Eq, Ord, Show) 47 | 48 | type Time = [Lab] 49 | 50 | instance Address Addr 51 | 52 | instance HasInitial Time where 53 | initial = [] 54 | 55 | instance Truncatable Time where 56 | trunc ls = take 1 ls 57 | 58 | --------------------------------------------------------------------- 59 | -- Abstract analysis interface. 60 | ---------------------------------------------------------------------- 61 | 62 | -- Hint: Add new primitives as they appear in the semantics 63 | class Monad m => AFJCESKInterface m a | m -> a where 64 | tick :: PState a -> m () 65 | getObj :: BEnv a -> Var -> m (Obj a) 66 | putObj :: BEnv a -> Var -> Obj a -> m () 67 | getCont :: a -> m (Kont a) 68 | putCont :: MethodName -> (Kont a) -> m a 69 | getConstr :: ClassTable -> ClassName -> m ([Obj a] -> m (BEnv a)) 70 | getMethod :: ClassTable -> Obj a -> MethodName -> m Method 71 | initBEnv :: BEnv a -> [Var] -> [Var] -> m (BEnv a) 72 | 73 | -- stepAnalysis :: ClassTable -> s -> g -> PState a -> (s, [(PState a, g)]) 74 | -- inject :: [Var] -> [Stmt] -> (PState a, s, g) 75 | 76 | mstep :: AFJCESKInterface m a => ClassTable -> PState a -> m (PState a) 77 | mstep table ctx@((Asgn v v' l):succ, β, pk) = do 78 | tick ctx 79 | d <- getObj β v' 80 | putObj β v d 81 | return $! (succ, β, pk) 82 | mstep table ctx@((Ret v l):[], β, pk) = do 83 | tick ctx 84 | (v', s, β', pk') <- getCont pk 85 | d <- getObj β v 86 | putObj β' v' d 87 | return $! (s, β', pk') 88 | mstep table ctx@((Lkp v v' f l):succ, β, pk) = do 89 | tick ctx 90 | (c, β') <- getObj β v' 91 | d <- getObj β' f 92 | putObj β v d 93 | return $! (succ, β, pk) 94 | mstep table ctx@((Cast v cn v' l):succ, β, pk) = do 95 | tick ctx 96 | d <- getObj β v' 97 | putObj β v d 98 | return $! (succ, β, pk) 99 | mstep table ctx@((New v cn vs l):succ, β, pk) = do 100 | tick ctx 101 | ructor <- getConstr table cn 102 | ds <- sequence [getObj β vi | vi <- vs] 103 | β' <- ructor ds 104 | let d' = (cn, β') 105 | putObj β v d' 106 | return $! (succ, β, pk) 107 | mstep table ctx@((MCall v v0 mthd vs l):succ, β, pk) = do 108 | tick ctx 109 | d0 <- getObj β v0 110 | (cn, _ , params, locals, body) 111 | <- getMethod table d0 mthd 112 | ds <- sequence [getObj β vi | vi <- vs] 113 | let κ = (v, succ, β, pk) 114 | pk' <- putCont mthd κ 115 | let vs'' = L.map snd params 116 | let vs''' = L.map snd locals 117 | let β' = Map.empty // ["this" ==> (β ! v0)] 118 | β'' <- initBEnv β' vs'' vs''' 119 | sequence [putObj β'' vi di | vi <- vs'' | di <- ds] 120 | return $! (body, β'', pk') 121 | -- final state 122 | mstep _ c@([], _ , _) = return c 123 | 124 | ---------------------------------------------------------------------- 125 | -- Running the analysis 126 | ---------------------------------------------------------------------- 127 | 128 | runAnalysis :: (Lattice fp , AddStepToFP m (PState a) fp, AFJCESKInterface m a) => 129 | ClassTable -> PState a -> fp 130 | runAnalysis ct p = exploreFP (mstep ct) p 131 | -------------------------------------------------------------------------------- /CFA/AFJ/Analysis/NonShared.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE TypeSynonymInstances #-} 3 | {-# LANGUAGE ParallelListComp #-} 4 | {-# LANGUAGE FlexibleInstances #-} 5 | {-# LANGUAGE FlexibleContexts #-} 6 | {-# LANGUAGE MultiParamTypeClasses #-} 7 | {-# LANGUAGE ScopedTypeVariables #-} 8 | {-# LANGUAGE ImplicitParams #-} 9 | {-# LANGUAGE TupleSections #-} 10 | 11 | module CFA.AFJ.Analysis.NonShared where 12 | 13 | import Data.Map as Map 14 | import Data.Set as Set 15 | import Data.List as L 16 | import Control.Monad.State 17 | import Control.Monad.ListT 18 | import Control.Monad.Identity 19 | 20 | import CFA.Lattice 21 | import CFA.Runner 22 | import CFA.CFAMonads 23 | import CFA.Store 24 | 25 | import CFA.AFJ 26 | import CFA.AFJ.Analysis 27 | 28 | ---------------------------------------------------------------------- 29 | -- ABSTRACT INTERPRETATION 30 | ---------------------------------------------------------------------- 31 | type D a = ℙ (Storable a) 32 | 33 | instance (StoreLike Addr s (D Addr), Truncatable Time) => 34 | AFJCESKInterface (StorePassingSemantics s Time) Addr where 35 | 36 | tick ctx@(stmts, _, _) = modify $ 37 | \t -> trunc ((lab (head stmts)):t) 38 | 39 | getObj β v = lift $ getsNDSet $ 40 | \σ -> Set.map (\(Val d) -> d) $ fetch σ (β!v) 41 | 42 | putObj β v d = lift $ modify $ 43 | \σ -> bind σ (β!v) (Set.singleton (Val d)) 44 | 45 | getCont pk = lift $ getsNDSet $ 46 | \σ -> Set.map (\(Cont κ) -> κ) $ fetch σ pk 47 | 48 | putCont m κ 49 | = do t <- get 50 | let b = alloc_k t m 51 | lift $ modify $ \σ -> bind σ b $ 52 | Set.singleton (Cont κ) 53 | return b 54 | 55 | initBEnv β vs'' vs''' 56 | = do t <- get 57 | let pairs' = L.map (\v -> (v, alloc t v)) vs'' 58 | pairs'' = L.map (\v -> (v, alloc t v)) vs''' 59 | return $ β // pairs' // pairs'' 60 | 61 | getMethod table (cn, _) m = return $ method table cn m 62 | 63 | getConstr table cn = do t <- get 64 | σ <- lift get 65 | -- updates a store and returns an environment of all class fields 66 | return $ createConstr table cn σ t 67 | 68 | createConstr table cn σ t ds = 69 | do σ' <- lift get 70 | let fs = allFields table cn -- compute all fields 71 | as = L.map (alloc t) fs -- appropriate addresses for fields 72 | fBindings = zip fs as -- bindings [field |-> addr] 73 | -- mapping from all class fields to provided arguments 74 | fMappings = Map.empty // classFieldMappings table cn ds 75 | -- heap is updated according to the mappings 76 | pairs = [(ai, Set.singleton (Val $ fMappings ! fi)) 77 | | (fi, ai) <- fBindings] 78 | σ'' = L.foldl (\store (ai, di) -> bind store ai di) σ' pairs 79 | -- new environment is created 80 | β' = Map.empty // fBindings 81 | lift $ modify (\_ -> σ'') 82 | return β' 83 | 84 | 85 | 86 | ---------------------------------------------------------------------------- 87 | -- Store machinery 88 | ---------------------------------------------------------------------------- 89 | 90 | alloc :: (Truncatable Time) => Time -> Var -> Addr 91 | alloc t v = AVar v $ trunc t 92 | 93 | alloc_k :: (Truncatable Time) => Time -> MethodName -> Addr 94 | alloc_k t m = ACall m $ trunc t 95 | 96 | type Store a = a :-> (D a) 97 | 98 | instance StoreLike Addr (Store Addr) (D Addr) where 99 | σ0 = Map.empty 100 | bind σ a d = σ ⨆ [a ==> d] 101 | fetch σ a = σ CFA.Lattice.!! a 102 | replace σ a d = σ ⨆ [a ==> d] 103 | filterStore σ p = Map.filterWithKey (\k -> \v -> p k) σ 104 | 105 | ---------------------------------------------------------------------------- 106 | -- Runner machinery 107 | ---------------------------------------------------------------------------- 108 | 109 | instance (Ord s, StoreLike Addr s (D Addr), Truncatable Time) => 110 | AddStepToFP (StorePassingSemantics s Time) 111 | (PState Addr) 112 | (Set (PState Addr, Time, s)) where 113 | applyStep step = 114 | joinWith 115 | (\(p, t, σ) -> Set.fromList $ L.map (\((a, b), c) -> (a, b, c)) $ 116 | runIdentity $ 117 | collectListT (runStateT (runStateT (step p) t) σ)) 118 | inject p = Set.singleton $ (p, initial, σ0) 119 | 120 | ---------------------------------------------------------------------------- 121 | -- Utility function to inject variables and stamtement to state -- 122 | ---------------------------------------------------------------------------- 123 | injectToState :: [Var] -> [Stmt] -> PState Addr 124 | injectToState vars stmts = let t0 = [] 125 | as = L.map (alloc t0) vars 126 | varBinds = L.zip vars as 127 | a0 = ACall "main" [] 128 | in (stmts, Map.empty // varBinds, a0) 129 | -------------------------------------------------------------------------------- /CFA/AFJ/Examples/AbstractNonShared.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeSynonymInstances #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | 4 | module CFA.AFJ.Examples where 5 | 6 | import Data.Map as Map 7 | import Data.Set as Set 8 | import Data.List as List 9 | 10 | import CFA.Store 11 | import CFA.AFJ 12 | import CFA.AFJ.Analysis 13 | import CFA.Runner 14 | 15 | ---------------------------------------------------------------------- 16 | -- abstract interpreter with a per-state store 17 | ---------------------------------------------------------------------- 18 | 19 | import CFA.AFJ.Analysis.NonShared 20 | 21 | ---------------------------------------------------------------------- 22 | -- example program 23 | ---------------------------------------------------------------------- 24 | 25 | {- 26 | 27 | class B {} 28 | 29 | class A { 30 | private final B myB; 31 | 32 | public A(B givenB) { 33 | this.myB = givenB; 34 | } 35 | 36 | B foo() { 37 | B tmpB; 38 | tmpB = this.myB; 39 | return B; 40 | } 41 | } 42 | 43 | // Main method of the AFJ program 44 | 45 | void main () { 46 | A mainA; 47 | B mainB; 48 | B mainB1; 49 | mainB = new B(); 50 | mainA = new A(mainB); 51 | mainB1 = mainA.foo(); 52 | } 53 | 54 | -} 55 | 56 | bClass :: Class 57 | bClass = ("B", Nothing, [], ("B", [], [], []), []) 58 | 59 | aClass :: Class 60 | aClass = ("A", Nothing, 61 | [("B", "myB")], ("A", [("B", "givenB")], [], [("myB", "givenB")]), 62 | [ 63 | ("B", "foo", [], [("B", "tmpB")], 64 | [Lkp "tmpB" "this" "myB" "l1", 65 | Ret "tmpB" "l2"])]) 66 | 67 | ctable :: ClassTable 68 | ctable = CTable [aClass, bClass] 69 | 70 | mainVars = ["mainA", "mainB", "mainB1"] 71 | mainStmts = [ New "mainB" "B" [] "l3" 72 | , New "mainA" "A" ["mainB"] "l4" 73 | , MCall "mainB1" "mainA" "foo" [] "l5" 74 | ] 75 | 76 | ---------------------------------------------------------------------- 77 | 78 | abstractResult :: ClassTable -> [Var] -> [Stmt] -> Set (PState Addr, Time, Store Addr) 79 | abstractResult ct vars stmts = runAnalysis ct (injectToState vars stmts) 80 | 81 | -- Try: 82 | -- abstractResult ctable mainVars mainStmts -------------------------------------------------------------------------------- /CFA/CESK.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE ParallelListComp #-} 3 | {-# LANGUAGE TypeSynonymInstances #-} 4 | {-# LANGUAGE TupleSections #-} 5 | {-# LANGUAGE MultiParamTypeClasses #-} 6 | {-# LANGUAGE ScopedTypeVariables #-} 7 | {-# LANGUAGE ImplicitParams #-} 8 | 9 | module CFA.CESK where 10 | 11 | -- Imports. 12 | import Data.Map as Map 13 | import Data.Set as Set 14 | 15 | import CFA.Lattice 16 | import CFA.Store 17 | import CFA.Runner 18 | 19 | {---------------- SYNTAX AND STATE-SPACE ----------------} 20 | 21 | -- `a' stands for the nature of addresses 22 | -- it is also a leaf of the state-space 23 | type Var = String 24 | type Lab = String 25 | type Env a = Var :-> a 26 | type Lambda = (Var, Exp) 27 | type PState a = (Exp, Env a, a) 28 | 29 | data Exp = Ref (Var, Lab) 30 | | App (Exp, Exp, Lab) 31 | | Lam (Lambda, Lab) 32 | deriving (Eq, Ord, Show) 33 | 34 | data Clo a = Clo (Lambda, Env a, Lab) 35 | deriving (Eq, Ord, Show) 36 | 37 | data Kont a = Mt 38 | | Ar (Exp, Env a, a) 39 | | Fn (Lambda, Env a, a) 40 | deriving (Eq, Ord, Show) 41 | 42 | data Storable a = Val (Clo a) 43 | | Cont (Kont a) 44 | deriving (Eq, Ord, Show) 45 | 46 | -- retrieve a label 47 | lab :: Exp -> Lab 48 | lab (Ref (_, l)) = l 49 | lab (App (_, _, l)) = l 50 | lab (Lam (_, l)) = l 51 | 52 | ---------------------------------------------------------------------- 53 | -- TIME-STAMPED ADDRESSES (FIXED) 54 | ---------------------------------------------------------------------- 55 | data Time = TLab Lab [Lab] 56 | | TMt [Lab] 57 | deriving (Eq, Ord, Show) 58 | 59 | class (Eq a, Ord a) => Address a 60 | 61 | data Addr = Bind Var [Lab] 62 | | Call Lab [Lab] 63 | deriving (Eq, Ord, Show) 64 | 65 | instance Address Addr 66 | 67 | contour :: Time -> [Lab] 68 | contour (TLab _ c) = c 69 | contour (TMt c) = c 70 | 71 | instance HasInitial Time where 72 | initial = TMt [] 73 | 74 | instance Truncatable Time where 75 | trunc (TMt (lab:ls)) = TMt ls 76 | trunc t = t 77 | 78 | instance HasInitial Addr where 79 | initial = Call "mt" [] 80 | 81 | -------------------------------------------------------------------------------- /CFA/CESK/Analysis.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FlexibleContexts #-} 4 | {-# LANGUAGE FunctionalDependencies #-} 5 | {-# LANGUAGE TypeOperators #-} 6 | {-# LANGUAGE ParallelListComp #-} 7 | {-# LANGUAGE TypeSynonymInstances #-} 8 | 9 | module CFA.CESK.Analysis where 10 | 11 | -- Imports. 12 | import Data.Map as Map 13 | import Data.Set as Set 14 | 15 | import CFA.Lattice 16 | import CFA.CESK 17 | import CFA.Runner 18 | 19 | ---------------------------------------------------------------------- 20 | -- Abstract analysis interface. 21 | ---------------------------------------------------------------------- 22 | class Monad m => LambdaCESKInterface m a | m -> a where 23 | tick :: PState a -> m () 24 | alloc :: () -> m a 25 | getVar :: Env a -> Var -> m (Clo a) 26 | putVar :: Env a -> Var -> a -> Clo a -> m (Env a) 27 | loadCont :: a -> m (Kont a) 28 | storeCont :: a -> Kont a -> m () 29 | 30 | -- A small-step monadic semantics for CESK* machine 31 | -- in Store- and Time- passing style 32 | 33 | ---------------------------------------------------------------------- 34 | -- Generic transition 35 | ---------------------------------------------------------------------- 36 | mstep :: (LambdaCESKInterface m a) => PState a -> m (PState a) 37 | mstep ctx@(Ref (x, l), ρ, a) = do 38 | tick ctx 39 | Clo (v, ρ', l) <- getVar ρ x 40 | return $! (Lam (v, l), ρ', a) 41 | mstep ctx@(App (e0, e1, l), ρ, a) = do 42 | tick ctx 43 | b <- alloc () 44 | storeCont b (Ar (e1, ρ, a)) 45 | return (e0, ρ, b) 46 | mstep ctx@(Lam (v, l), ρ, a) = do 47 | tick ctx 48 | κ <- loadCont a 49 | case κ of 50 | Ar (e, ρ, a) -> do 51 | b <- alloc () 52 | storeCont b κ 53 | return (e, ρ, b) 54 | Fn ((x, e), ρ', c) -> do 55 | b <- alloc () 56 | ρ'' <- putVar ρ' x b (Clo (v, ρ, l)) 57 | return (e, ρ'', c) 58 | 59 | ---------------------------------------------------------------------- 60 | -- Running the analysis 61 | ---------------------------------------------------------------------- 62 | 63 | runAnalysis :: (Lattice fp , AddStepToFP m (PState a) fp, 64 | HasInitial a, LambdaCESKInterface m a) => 65 | Exp -> fp 66 | runAnalysis e = exploreFP mstep (e, Map.empty, initial) 67 | -------------------------------------------------------------------------------- /CFA/CESK/Analysis/Concrete.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE ParallelListComp #-} 3 | {-# LANGUAGE TypeSynonymInstances #-} 4 | {-# LANGUAGE TupleSections #-} 5 | {-# LANGUAGE MultiParamTypeClasses #-} 6 | {-# LANGUAGE ScopedTypeVariables #-} 7 | {-# LANGUAGE ImplicitParams #-} 8 | {-# LANGUAGE FlexibleInstances #-} 9 | 10 | module CFA.CESK.Analysis.Concrete where 11 | 12 | import Data.Map as Map 13 | import Data.Set as Set 14 | import Data.List as List 15 | import Data.Traversable 16 | import Control.Monad.State 17 | import Control.Monad.ListT 18 | import Control.Monad.Identity 19 | import Control.Monad.Reader 20 | 21 | import CFA.Lattice 22 | import CFA.Store 23 | import CFA.CFAMonads 24 | import CFA.Runner 25 | 26 | import CFA.CESK 27 | import CFA.CESK.Analysis 28 | 29 | ---------------------------------------------------------------------- 30 | -- CONCRETE INTERPRETATION 31 | ---------------------------------------------------------------------- 32 | 33 | type Store a = a :-> (Storable a) 34 | 35 | instance LambdaCESKInterface (StorePassingSemantics (Store Addr) (Time, PState Addr)) Addr where 36 | 37 | tick ctx@(Ref (_, _), ρ, a) = modify $ \(t, _) -> (t, ctx) 38 | tick ctx@(App (_, _, l), ρ, a) = modify $ \(t, _) -> (TLab l (contour t), ctx) 39 | tick ctx@(v, ρ, a) = do (t, s) <- get 40 | σ <- lift get 41 | case t of 42 | TLab l ctr -> 43 | let (Cont κ) = σ ! a 44 | in case κ of 45 | Mt -> put (TMt (l:ctr), ctx) 46 | Ar _ -> put (TLab l ctr, ctx) 47 | Fn _ -> put (TMt (l:ctr), ctx) 48 | _ -> return () 49 | 50 | getVar ρ x = lift $ getsNDSet (\σ -> let (Val clo) = σ ! (ρ ! x) 51 | in Set.singleton(clo)) 52 | 53 | putVar ρ x b c = (lift $ modify $ \σ -> σ // [(b, Val c)]) >> 54 | (return $ ρ // [(x, b)]) 55 | 56 | loadCont a = lift $ getsNDSet (\σ -> 57 | let (Cont κ) = σ ! a 58 | in Set.singleton κ) 59 | 60 | storeCont b κ = lift $ modify $ \σ -> σ // [(b, Cont κ)] 61 | 62 | alloc _ = do (t, s) <- get 63 | σ <- lift get 64 | return $ allocC t s σ 65 | 66 | -- Concrete allocator function 67 | -- Turns the last captured time moment into the address 68 | allocC :: Time -> PState Addr -> Store Addr -> Addr 69 | allocC t (App (e0, _, _), _ ,_) σ = Call (lab e0) (contour t) 70 | allocC t (Lam _, _, a) σ = 71 | case σ ! a of 72 | Cont (Ar (e, _, _)) -> Call (lab e) (contour t) 73 | Cont (Fn ((x, _), _, _)) -> Bind x (contour t) 74 | 75 | 76 | instance AddStepToFP (StorePassingSemantics (Store Addr) (Time, PState Addr)) 77 | (PState Addr) 78 | (Set (PState Addr, Time, Store Addr)) where 79 | applyStep step = 80 | joinWith 81 | (\(p, t, σ) -> Set.fromList $ List.map (\((a, (b, c)), d) -> (a, b, d)) $ 82 | runIdentity $ 83 | collectListT (runStateT (runStateT (step p) (t, p)) σ)) 84 | inject p = Set.singleton $ (p, initial, Map.empty) -------------------------------------------------------------------------------- /CFA/CESK/Analysis/NonShared.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE ParallelListComp #-} 3 | {-# LANGUAGE TypeSynonymInstances #-} 4 | {-# LANGUAGE TupleSections #-} 5 | {-# LANGUAGE FlexibleContexts #-} 6 | {-# LANGUAGE FlexibleInstances #-} 7 | {-# LANGUAGE MultiParamTypeClasses #-} 8 | {-# LANGUAGE ScopedTypeVariables #-} 9 | {-# LANGUAGE ImplicitParams #-} 10 | 11 | module CFA.CESK.Analysis.NonShared where 12 | 13 | import Data.Map as Map 14 | import Data.Set as Set 15 | import Data.List as List 16 | 17 | import Control.Monad.State 18 | import Control.Monad.ListT 19 | import Control.Monad.Identity 20 | 21 | import CFA.Lattice 22 | import CFA.CFAMonads 23 | import CFA.Runner 24 | import CFA.Store 25 | 26 | import CFA.CESK 27 | import CFA.CESK.Analysis 28 | 29 | ---------------------------------------------------------------------- 30 | -- ABSTRACT INTERPRETATION 31 | ---------------------------------------------------------------------- 32 | 33 | type D a = ℙ (Storable a) 34 | type NDStore a = a :-> (D a) 35 | 36 | instance (StoreLike Addr s (D Addr), Truncatable Time) => 37 | LambdaCESKInterface (StorePassingSemantics s (Time, PState Addr)) Addr where 38 | 39 | tick ctx@(Ref (_, _), ρ, a) 40 | = modify $ \(t, _) -> (t, ctx) 41 | tick ctx@(App (_, _, l), ρ, a) 42 | = modify $ \(t, _) -> (TLab l (contour t), ctx) 43 | tick ctx@(v, ρ, a) 44 | = do (t, s) <- get 45 | σ <- lift get 46 | case t of 47 | TLab l ctr -> 48 | let ts = [case κ of 49 | Mt -> (TMt (l:ctr), ctx) 50 | Ar _ -> (TLab l ctr, ctx) 51 | Fn _ -> (TMt (l:ctr), ctx) 52 | | Cont κ <- Set.toList $ fetch σ a] 53 | in mapM put ts >> return () 54 | _ -> return () 55 | 56 | getVar ρ x 57 | = lift $ getsNDSet $ (\σ -> Set.map (\(Val c) -> c) $ 58 | fetch σ (ρ ! x)) 59 | 60 | putVar ρ x b c 61 | = (lift $ modify $ \σ -> bind σ b $ Set.singleton $ Val c) >> 62 | (return $ ρ // [(x, b)]) 63 | 64 | loadCont a 65 | = lift $ getsNDSet (\σ -> Set.map (\(Cont κ) -> κ) $ 66 | fetch σ a) 67 | 68 | storeCont b κ 69 | = lift $ modify $ \σ -> bind σ b $ Set.singleton (Cont κ) 70 | 71 | alloc _ = do (t, s) <- get 72 | σ <- lift get 73 | pureND $ allocKCFA t s σ 74 | 75 | -- abstract allocator function 76 | -- nondeterministic because of stored continuations 77 | allocKCFA :: StoreLike Addr s (D Addr) => 78 | Time -> PState Addr -> s -> [Addr] 79 | allocKCFA t (App (e0, _, _), _ ,_) σ 80 | = [Call (lab e0) (contour t)] 81 | allocKCFA t (Lam _, _, a) σ = 82 | [case κ of 83 | Ar (e, _, _) -> Call (lab e) (contour t) 84 | Fn ((x, _), _, _) -> Bind x (contour t) 85 | | Cont κ <- Set.toList $ fetch σ a] 86 | 87 | instance StoreLike Addr (NDStore Addr) (D Addr) where 88 | σ0 = Map.empty 89 | bind σ a d = σ ⨆ [a ==> d] 90 | fetch σ a = σ CFA.Lattice.!! a 91 | replace σ a d = σ ⨆ [a ==> d] 92 | filterStore σ p = Map.filterWithKey (\k -> \v -> p k) σ 93 | 94 | 95 | --------------------------------------------------------------- 96 | -- Runner -- 97 | --------------------------------------------------------------- 98 | instance (Ord s, StoreLike Addr s (D Addr), Truncatable Time) => 99 | AddStepToFP (StorePassingSemantics s (Time, PState Addr)) 100 | (PState Addr) 101 | (Set (PState Addr, Time, s)) where 102 | applyStep step = 103 | joinWith 104 | (\(p, t, σ) -> Set.fromList $ List.map (\((a, (b, c)), d) -> (a, b, d)) $ 105 | runIdentity $ 106 | collectListT (runStateT (runStateT (step p) (t, p)) σ)) 107 | inject p = Set.singleton $ (p, initial, σ0) -------------------------------------------------------------------------------- /CFA/CESK/Examples.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE ParallelListComp #-} 3 | {-# LANGUAGE TypeSynonymInstances #-} 4 | {-# LANGUAGE TupleSections #-} 5 | {-# LANGUAGE MultiParamTypeClasses #-} 6 | {-# LANGUAGE ScopedTypeVariables #-} 7 | {-# LANGUAGE ImplicitParams #-} 8 | 9 | module CFA.CESK.Examples where 10 | 11 | -- Imports. 12 | import Data.Map as Map 13 | import Data.Set as Set 14 | 15 | import CFA.Lattice 16 | import CFA.Store 17 | import CFA.CFAMonads 18 | 19 | import CFA.CESK 20 | import CFA.CESK.Analysis 21 | import CFA.CESK.Analysis.Concrete 22 | 23 | 24 | -- provide the initial state 25 | -- State = (Exp, Env, Addr) 26 | 27 | lam1 = Lam (("x", Ref ("x", "l2")), "l1") 28 | arg1 = Lam (("y", Ref ("y", "l6")), "l5") 29 | app = App (lam1, arg1, "l7") 30 | 31 | state0 = (app, Map.empty, Call "l3" []) 32 | 33 | -- define the denotational semantics of the first step 34 | -- literally, mstep - is a denotational semantics, induced by the 35 | -- abstract machine over time and store 36 | transition :: Concrete () (Time, State Addr, Store Addr) (State Addr) 37 | transition = mstep state0 >>= mstep 38 | 39 | -- run analysis with the suplied time and store 40 | result = cf transition (TMt [], undefined, Map.empty) -------------------------------------------------------------------------------- /CFA/CESK/Examples/AbstractNonShared.hs: -------------------------------------------------------------------------------- 1 | module CFA.CESK.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | 7 | import CFA.Store 8 | import CFA.Runner 9 | 10 | import CFA.CESK 11 | import CFA.CESK.Analysis 12 | 13 | ---------------------------------------------------------------------- 14 | -- abstract interpreter with a per-state store 15 | ---------------------------------------------------------------------- 16 | 17 | import CFA.CESK.Analysis.NonShared 18 | 19 | ---------------------------------------------------------------------- 20 | -- example program 21 | ---------------------------------------------------------------------- 22 | 23 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 24 | idx = Lam (("x", Ref ("x", "l1")), "l2") 25 | idy = Lam (("y", Ref ("y", "l3")), "l4") 26 | 27 | comb = Lam (("f", Lam (("g", App (Ref ("f", "l8"), Ref ("g", "l9"), "l7")), "l6")), "l5") 28 | ex = App (App (comb, idx, "l11"), idy, "l10") 29 | 30 | ---------------------------------------------------------------------- 31 | 32 | abstractResult :: Exp -> Set (PState Addr, Time, NDStore Addr) 33 | abstractResult = runAnalysis 34 | -------------------------------------------------------------------------------- /CFA/CESK/Examples/Concrete.hs: -------------------------------------------------------------------------------- 1 | module CFA.CESK.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | 7 | import CFA.Store 8 | import CFA.Runner 9 | 10 | import CFA.CESK 11 | import CFA.CESK.Analysis 12 | 13 | ---------------------------------------------------------------------- 14 | -- abstract interpreter with a per-state store 15 | ---------------------------------------------------------------------- 16 | 17 | import CFA.CESK.Analysis.Concrete 18 | 19 | ---------------------------------------------------------------------- 20 | -- example program 21 | ---------------------------------------------------------------------- 22 | 23 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 24 | idx = Lam (("x", Ref ("x", "l1")), "l2") 25 | idy = Lam (("y", Ref ("y", "l3")), "l4") 26 | 27 | comb = Lam (("f", Lam (("g", App (Ref ("f", "l8"), Ref ("g", "l9"), "l7")), "l6")), "l5") 28 | ex = App (App (comb, idx, "l11"), idy, "l10") 29 | 30 | ---------------------------------------------------------------------- 31 | 32 | concreteResult :: Exp -> Set (PState Addr, Time, Store Addr) 33 | concreteResult = runAnalysis 34 | -------------------------------------------------------------------------------- /CFA/CFAMonads.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE RankNTypes #-} 3 | {-# LANGUAGE FlexibleInstances #-} 4 | {-# LANGUAGE FlexibleContexts #-} 5 | {-# LANGUAGE FunctionalDependencies #-} 6 | {-# LANGUAGE TypeOperators #-} 7 | {-# LANGUAGE TypeFamilies #-} 8 | {-# LANGUAGE ScopedTypeVariables #-} 9 | {-# LANGUAGE TypeSynonymInstances #-} 10 | {-# LANGUAGE GeneralizedNewtypeDeriving #-} 11 | 12 | module CFA.CFAMonads where 13 | 14 | import Data.Map as Map 15 | import Data.Set as Set 16 | import Data.List as List 17 | import Control.Monad as Monad 18 | import Control.Monad.State 19 | import Control.Monad.ListT as ListT 20 | import Control.Monad.Reader 21 | import Control.Monad.Identity 22 | import Control.Monad.Trans 23 | import Control.Applicative 24 | import Data.Monoid 25 | import Data.Foldable as Foldable hiding (msum) 26 | import Data.Traversable as Trav 27 | import Data.List.Class as ListC 28 | 29 | import Util 30 | 31 | import CFA.Lattice 32 | 33 | mapListItem :: Monad m => (forall v. n v -> m v) -> ListItem (ListT n) v -> ListItem (ListT m) v 34 | mapListItem f Nil = Nil 35 | mapListItem f (Cons v vs) = Cons v (mapListT f vs) 36 | 37 | mapListT :: Monad m => (forall v. n v -> m v) -> ListT n v -> ListT m v 38 | mapListT f m = ListT $ liftM (mapListItem f) $ f $ runListT m 39 | 40 | ---------------------------------------------------------------------- 41 | -- SharedStateListT : a reusable component to implement analyses... 42 | ---------------------------------------------------------------------- 43 | swap :: (a,b) -> (b,a) 44 | swap (a,b) = (b,a) 45 | 46 | -- state merging 47 | mergeState :: Lattice s => [(a, s)] -> ([a], s) 48 | mergeState = mapSnd withJoinMonoid . swap . sequenceA . List.map swap . List.map (mapSnd JoinMonoid) 49 | 50 | duplicateState :: ([a], s) -> [(a,s)] 51 | duplicateState (vs, s) = List.map (flip (,) s) vs 52 | 53 | mergeAndDupState :: Lattice s => [(a, s)] -> [(a, s)] 54 | mergeAndDupState = duplicateState . mergeState 55 | 56 | -- newtype SharedStateT s n m a = SharedStateT { runSharedStateT :: s -> n (s, m a) } 57 | -- newtype StateT s m a = StateT { runStateT :: s -> m (a, s) } 58 | -- Note: we are using "ListT done right" from the Control.Monad.ListT package. 59 | newtype SharedStateListT s n a = 60 | SSListT { runSSListT :: StateT s (ListT n) a } 61 | deriving (Monad, MonadState s, MonadPlus) 62 | 63 | instance MonadTrans (SharedStateListT s) where 64 | lift m = SSListT $ lift $ lift m 65 | 66 | listToListT :: Monad n => n [a] -> ListT n a 67 | listToListT = (>>= ListC.fromList) . lift 68 | 69 | collectListT :: Monad n => ListT n a -> n [a] 70 | collectListT = ListC.toList 71 | 72 | collectSSListT :: Monad n => SharedStateListT s n a -> s -> n [(a,s)] 73 | collectSSListT = (collectListT.) . runStateT . runSSListT 74 | 75 | collectSSListTS :: (Lattice s, Monad n) => SharedStateListT s n a -> s -> n ([a],s) 76 | collectSSListTS = (liftM mergeState.) . collectSSListT 77 | 78 | collectSSListTST :: (Lattice s, Monad n) => SharedStateListT s n a -> StateT s n [a] 79 | collectSSListTST = StateT . collectSSListTS 80 | 81 | makeSSListT :: Monad n => (s -> n [(a,s)]) -> SharedStateListT s n a 82 | makeSSListT f = SSListT $ StateT $ listToListT . f 83 | 84 | makeSSListTS :: (Monad n) => (s -> n ([a],s)) -> SharedStateListT s n a 85 | makeSSListTS f = SSListT $ StateT $ listToListT . liftM duplicateState . f 86 | 87 | -- ask for losing precision 88 | mergeSharedState :: forall s n a . (Monad n, Lattice s) => 89 | SharedStateListT s n a -> SharedStateListT s n a 90 | mergeSharedState (SSListT (StateT f)) = SSListT $ StateT $ listToListT . liftM mergeAndDupState . ListC.toList . f 91 | 92 | mapSharedState :: (Monad m, Monad n, Lattice s) => (forall v. n v -> m v) -> 93 | SharedStateListT s n a -> SharedStateListT s m a 94 | mapSharedState f m = SSListT $ mapStateT (mapListT f) (runSSListT m) 95 | 96 | -- SharedStateListT s n a = s -> n (s, [a]) 97 | -- abstracts 98 | -- StateT s (ListT n) a 99 | -- s -> n [(s, a)] 100 | -- alternative: 101 | -- ListT (StateT s n) a = s -> n ([a], s) 102 | -- bad idea, because each non-deterministic branch gets the modified store from the 103 | -- previous ND branch, instead of them all getting the same store and merging stores 104 | -- afterwards. This causes lost precision, because non-first branches get bigger 105 | -- stores. Also not a good idea because this is then order-dependent, and not actually 106 | -- a Monad. 107 | 108 | toNonShared :: Monad n => SharedStateListT s n a -> StateT s (ListT n) a 109 | toNonShared = runSSListT 110 | 111 | toShared :: (Monad n, Lattice s) => StateT s (ListT n) a -> SharedStateListT s n a 112 | toShared = SSListT 113 | 114 | runSSListT0 :: (Monad n, Lattice s) => SharedStateListT s n a -> n ([a], s) 115 | runSSListT0 m = liftM mergeState $ collectSSListT m bot 116 | 117 | -- type MyAnalysis a = StateT g (SharedStateListT s Identity a) 118 | -- = g -> (SharedStateListT s Identity (g, a) 119 | -- = g -> (s -> (s, [(g, a)])) 120 | 121 | 122 | askSSLT :: (Lattice s, MonadReader env m) => SharedStateListT s m env 123 | askSSLT = do env <- lift ask 124 | return env 125 | 126 | localSSLT :: (Lattice s, MonadReader env m) => (env -> env) -> SharedStateListT s m v -> SharedStateListT s m v 127 | localSSLT f m = mapSharedState (local f) m 128 | 129 | pureND :: MonadPlus m => [a] -> m a 130 | pureND = msum . List.map return 131 | 132 | pureNDSet :: MonadPlus m => Set a -> m a 133 | pureNDSet = pureND . Set.toList 134 | 135 | getsND :: (MonadPlus m, MonadState s m) => (s -> [a]) -> m a 136 | getsND f = gets f >>= pureND 137 | 138 | getsNDSet :: (MonadPlus m, MonadState s m) => 139 | (s -> Set a) -> m a 140 | getsNDSet f = getsND (Set.toList . f) 141 | 142 | getsM :: (MonadState s m, Monad m) => (s -> m a) -> m a 143 | getsM f = get >>= f 144 | 145 | asksM :: (MonadReader s (t m), Monad m, MonadTrans t) => (s -> m a) -> t m a 146 | asksM f = asks f >>= lift 147 | 148 | commaM :: Monad m => m a -> m b -> m (a,b) 149 | commaM = liftM2 (,) 150 | 151 | liftTup :: Monad t => (t a, b) -> t (a,b) 152 | liftTup (ma,b) = liftM (flip (,) b) ma 153 | 154 | joinWith :: (Foldable t, Lattice a) => (b -> a) -> t b -> a 155 | joinWith f = Foldable.foldr ((⊔) . f) bot 156 | 157 | type StorePassingSemantics s g = StateT g (StateT s (ListT Identity)) 158 | 159 | -------------------------------------------------------------------------------- /CFA/CPS.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | 3 | module CFA.CPS where 4 | 5 | import Data.Map as Map 6 | import Data.Set as Set 7 | import Data.List as List 8 | 9 | import CFA.Lattice 10 | 11 | -- Syntax. 12 | type Var = String 13 | 14 | data Lambda = [Var] :=> CExp 15 | deriving (Ord, Eq, Show) 16 | 17 | data AExp = Ref Var 18 | | Lam Lambda 19 | deriving (Ord, Eq, Show) 20 | 21 | data CExp = Call AExp [AExp] 22 | | Exit 23 | deriving (Ord, Eq, Show) 24 | 25 | -- State-space 26 | type PΣ a = (CExp, Env a) 27 | type Env a = Var :-> a 28 | type D a = ℙ (Val a) 29 | data Val a = Clo (Lambda, Env a) 30 | deriving (Eq, Ord, Show) 31 | 32 | ρ0 = Map.empty 33 | 34 | -- Generic store. 35 | type Store a = a :-> (D a) 36 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FlexibleContexts #-} 4 | {-# LANGUAGE FunctionalDependencies #-} 5 | {-# LANGUAGE TypeOperators #-} 6 | {-# LANGUAGE ParallelListComp #-} 7 | {-# LANGUAGE TypeSynonymInstances #-} 8 | {-# LANGUAGE TupleSections #-} 9 | {-# LANGUAGE UndecidableInstances #-} 10 | 11 | module CFA.CPS.Analysis where 12 | 13 | import Data.Map as Map 14 | import Data.Set as Set 15 | import Data.List as List 16 | import Control.Monad 17 | 18 | import CFA.CPS 19 | import CFA.Lattice 20 | import CFA.Store 21 | import CFA.Runner 22 | 23 | ---------------------------------------------------------------------- 24 | -- Abstract analysis interface. 25 | ---------------------------------------------------------------------- 26 | class Monad m => Analysis m a | m -> a where 27 | 28 | fun :: Env a -> AExp -> m (Val a) 29 | arg :: Env a -> AExp -> m (Val a) 30 | 31 | ($=) :: a -> Val a -> m () 32 | 33 | updateEnv :: Env a -> [(Var, a)] -> m (Env a) 34 | -- default implementation 35 | updateEnv ρ bs = return $ ρ // bs 36 | 37 | alloc :: Var -> m a 38 | tick :: Val a -> PΣ a -> m () 39 | 40 | ---------------------------------------------------------------------- 41 | -- Generic transition 42 | ---------------------------------------------------------------------- 43 | mnext :: (Analysis m a) => PΣ a -> m (PΣ a) 44 | 45 | mnext ps@(Call f aes, ρ) = do 46 | proc@(Clo (vs :=> call', ρ')) <- fun ρ f 47 | tick proc ps 48 | as <- mapM alloc vs 49 | ds <- mapM (arg ρ) aes 50 | ρ'' <- updateEnv ρ' [ v ==> a | v <- vs | a <- as ] 51 | sequence [ a $= d | a <- as | d <- ds ] 52 | let sn = (call', ρ'') 53 | return $! sn 54 | mnext ps = return ps 55 | 56 | ---------------------------------------------------------------------- 57 | -- Running the analysis 58 | ---------------------------------------------------------------------- 59 | 60 | runAnalysis :: (Lattice fp , AddStepToFP m (PΣ a) fp, Analysis m a) => 61 | CExp -> fp 62 | runAnalysis e = exploreFP mnext (e, Map.empty) 63 | 64 | 65 | ---------------------------------------------------------------------- 66 | -- Addresses, Stores and Choices 67 | ---------------------------------------------------------------------- 68 | 69 | type ProcCh a = Maybe (Val a) -- Nondeterministic choice. 70 | 71 | -- FunctionalDependencies again: 72 | -- time defines addresses 73 | class (Ord a, Eq a) => Addressable a c | c -> a where 74 | τ0 :: c 75 | valloc :: Var -> c -> a 76 | advance :: Val a -> PΣ a -> c -> c 77 | 78 | ---------------------------------------------------------------------- 79 | -- GC Machinery 80 | ---------------------------------------------------------------------- 81 | 82 | -- Free Variables 83 | free' :: CExp -> Set Var -> Set Var 84 | free' Exit bound = Set.empty 85 | free' (Call f as) bound = List.foldl (\res -> \a -> res ⊔ (free a bound)) 86 | (free f bound) as 87 | 88 | free :: AExp -> Set Var -> Set Var 89 | free (Ref v) bound = if (Set.member v bound) 90 | then Set.empty 91 | else Set.singleton v 92 | free (Lam (vs :=> ce)) bound = free' ce (bound ⊔ (Set.fromList vs)) 93 | 94 | -- `touched' (for expressions and environments) 95 | touched' :: (Ord a) => AExp -> Env a -> Set (Var, a) 96 | touched' f ρ = Set.fromList [(v, ρ!v) | v <- Set.toList(free f Set.empty)] 97 | 98 | -- `touched' for states 99 | touched :: (Ord a) => (PΣ a) -> Set (Var, a) 100 | touched (Call f as, ρ) = (touched' f ρ) ⊔ 101 | Set.fromList [bs | a <- as, 102 | bs <- Set.toList (touched' a ρ)] 103 | touched (Exit, _) = Set.empty 104 | 105 | -- `adjacency' 106 | adjacent :: (Ord a, StoreLike a s (D a)) => (Var, a) -> s -> Set (Var, a) 107 | adjacent (v, addr) σ = Set.fromList [b | Clo (f, ρ) <- Set.toList(fetch σ addr), 108 | b <- Set.toList (touched' (Lam f) ρ)] 109 | 110 | -- `reachability' 111 | reachable :: (Ord a, StoreLike a s (D a)) => (PΣ a) -> s -> Set (Var, a) 112 | reachable state σ = 113 | let collect bs = 114 | -- fixpoint iteration 115 | let newBindings = [b' | b <- Set.toList bs, 116 | b' <- Set.toList (adjacent b σ)] 117 | newResult = bs ⊔ (Set.fromList newBindings) 118 | in if newResult == bs 119 | then bs 120 | else collect newResult 121 | -- reflexive-transitive closure 122 | in collect (touched state) 123 | 124 | ---------------------------------------------------------------------- 125 | -- Anodization 126 | ---------------------------------------------------------------------- 127 | 128 | class StoreLike a s d => AlkaliLike a s d where 129 | addUniqueAddr :: a -> s 130 | deAnodizeStore :: s -> s 131 | deAnodizeEnv :: s -> Env a -> Env a 132 | deAnodizeD :: s -> d -> d 133 | reset :: s -> s 134 | 135 | -- a useful instance 136 | instance (Ord a, StoreLike a s d) => StoreLike a (s, ℙ a) d where 137 | σ0 = (σ0, Set.empty) 138 | bind σ a d = (bind (fst σ) a d, snd σ) 139 | fetch σ a = fetch (fst σ) a 140 | replace σ a d = (replace (fst σ) a d, snd σ) 141 | filterStore σ p = (filterStore (fst σ) p, snd σ) 142 | 143 | 144 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/Concrete.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE ParallelListComp #-} 6 | {-# LANGUAGE TypeSynonymInstances #-} 7 | {-# LANGUAGE GeneralizedNewtypeDeriving #-} 8 | 9 | -- TODO: get rid of this 10 | -- {-# LANGUAGE UndecidableInstances #-} 11 | 12 | module CFA.CPS.Analysis.Concrete where 13 | 14 | import Data.Map as Map 15 | import Data.Set as Set 16 | import Data.List as List 17 | import Data.Foldable as Foldable 18 | import Control.Monad.State 19 | import Control.Monad.Identity 20 | import Control.Applicative 21 | 22 | import Util 23 | 24 | import CFA.CPS 25 | import CFA.Lattice 26 | import CFA.Store 27 | import CFA.CFAMonads 28 | import CFA.CPS.Analysis 29 | import CFA.CPS.Analysis.Runner 30 | 31 | data CAddr = CBind Var Int 32 | deriving (Eq, Ord, Show) 33 | 34 | type DStore a = a :-> (Val a) 35 | 36 | type ΣC = (DStore CAddr, Int) 37 | 38 | 39 | alterStore = mapFst 40 | increaseTime = mapSnd (+1) 41 | 42 | -- is a monad 43 | type Concrete = State ΣC 44 | 45 | readA :: CAddr -> Concrete (Val CAddr) 46 | readA a = gets $ (! a) . fst 47 | 48 | getTime :: Concrete Int 49 | getTime = gets snd 50 | 51 | instance Analysis Concrete CAddr where 52 | fun ρ (Lam l) = return $ Clo (l, ρ) 53 | fun ρ (Ref v) = readA (ρ!v) 54 | 55 | arg ρ (Lam l) = let proc = Clo(l, ρ) in return proc 56 | arg ρ (Ref v) = readA (ρ!v) 57 | 58 | a $= d = modify $ alterStore $ Map.insert a d 59 | 60 | alloc v = CBind v <$> getTime 61 | 62 | tick _ _ go = modify increaseTime >> go 63 | 64 | initialΣC :: ΣC 65 | initialΣC = (Map.empty, 0) 66 | 67 | injectConcrete :: CExp -> (PΣ CAddr, ΣC) 68 | injectConcrete call = ((call, ρ0), initialΣC) 69 | 70 | 71 | -- Add Garbage Collection 72 | instance GarbageCollector Concrete (PΣ CAddr) 73 | 74 | instance AddStepToFP Concrete (PΣ CAddr) (ℙ (PΣ CAddr, ΣC)) where 75 | applyStep step states = Set.map (uncurry $ runState . step) states 76 | inject s = Set.singleton (s, initialΣC) 77 | 78 | 79 | exploreConcrete :: CExp -> ℙ (PΣ CAddr, ΣC) 80 | exploreConcrete = exploreFP 81 | 82 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/Interpreter.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE FlexibleInstances #-} 2 | {-# LANGUAGE FlexibleContexts #-} 3 | {-# LANGUAGE MultiParamTypeClasses #-} 4 | {-# LANGUAGE TypeSynonymInstances #-} 5 | 6 | module CFA.CPS.Analysis.Interpreter where 7 | 8 | import Data.IORef 9 | import Data.Map as Map 10 | import Data.Set as Set 11 | import Control.Applicative 12 | 13 | import CFA.CPS 14 | import CFA.CPS.Analysis 15 | 16 | data IOAddr = MkIOAddr { unIOAddr :: IORef (Val IOAddr) } 17 | deriving (Eq) 18 | instance Show IOAddr where show a = "{IOAddr ...}" 19 | 20 | readIOAddr :: IOAddr -> IO (Val IOAddr) 21 | readIOAddr = readIORef . unIOAddr 22 | 23 | writeIOAddr :: IOAddr -> Val IOAddr -> IO () 24 | writeIOAddr = writeIORef . unIOAddr 25 | 26 | instance Analysis IO IOAddr where 27 | fun ρ (Lam l) = return $ Clo (l, ρ) 28 | fun ρ (Ref v) = readIOAddr (ρ ! v) 29 | 30 | arg ρ (Lam l) = return $ Clo (l, ρ) 31 | arg ρ (Ref v) = readIOAddr (ρ!v) 32 | 33 | addr $= v = writeIOAddr addr v 34 | alloc v = MkIOAddr <$> newIORef undefined 35 | tick _ _ = return () 36 | 37 | interpret :: CExp -> IO (PΣ IOAddr) 38 | interpret e = go (e, ρ0) 39 | where go :: (PΣ IOAddr) -> IO (PΣ IOAddr) 40 | go s = do s' <- mnext s 41 | case s' of x@(Exit, _) -> return x 42 | y -> go y 43 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/NonShared.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FlexibleContexts #-} 4 | {-# LANGUAGE FunctionalDependencies #-} 5 | {-# LANGUAGE TypeOperators #-} 6 | {-# LANGUAGE ParallelListComp #-} 7 | {-# LANGUAGE TypeSynonymInstances #-} 8 | {-# LANGUAGE ScopedTypeVariables #-} 9 | {-# LANGUAGE TypeFamilies #-} 10 | 11 | -- TODO: get rid of this 12 | {-# LANGUAGE UndecidableInstances #-} 13 | 14 | module CFA.CPS.Analysis.NonShared where 15 | 16 | import Data.Map as Map 17 | import Data.Set as Set 18 | import Data.List as List 19 | import Data.Foldable as Foldable hiding (concat) 20 | import Data.Traversable 21 | import Control.Monad.State 22 | import Control.Monad.ListT 23 | import Control.Monad.Identity 24 | import Control.Monad.Reader 25 | 26 | import CFA.CPS 27 | import CFA.Lattice 28 | import CFA.Store 29 | import CFA.CFAMonads 30 | import CFA.CPS.Analysis 31 | import CFA.Runner 32 | 33 | import Util 34 | 35 | type NonSharedAnalysis s g = 36 | StateT g (SharedStateListT s (ListT Identity)) 37 | -- NonSharedAnalysis s g a = 38 | -- StateT g (SharedStateListT s (ListT Identity)) a 39 | -- g -> SharedStateListT s (ListT Identity) (a, g) (more or less) 40 | -- g -> s -> ListT Identity [((a, g),s)] (more or less) 41 | -- g -> s -> Identity [[((a, g),s)]] 42 | -- g -> s -> [[((a, g),s)]] 43 | 44 | instance (Addressable a t, StoreLike a s (D a)) 45 | => Analysis (NonSharedAnalysis s (ProcCh a, t)) a 46 | where 47 | fun ρ (Lam l) = return $ Clo(l, ρ) 48 | fun ρ (Ref v) = lift $ getsM $ lift . pureNDSet . (flip fetch $ ρ!v) 49 | 50 | arg ρ (Lam l) = return $ Clo(l, ρ) 51 | arg ρ (Ref v) = lift $ getsNDSet $ flip fetch (ρ!v) 52 | 53 | a $= d = lift $ modify $ \ σ -> bind σ a (Set.singleton d) 54 | 55 | alloc v = gets $ valloc v . snd 56 | 57 | tick proc ps = modify $ \(_, t) -> (Just proc, advance proc ps t) 58 | 59 | -- Garbage Collection 60 | instance (Lattice s, Eq a, StoreLike a s (D a), Ord a) => 61 | GarbageCollector (NonSharedAnalysis s (ProcCh a, t)) (PΣ a) where 62 | gc ps = do 63 | σ <- lift get 64 | let rs = Set.map (\(v, a) -> a) (reachable ps σ) 65 | lift $ modify $ \ σ -> filterStore σ (\a -> Set.member a rs) 66 | 67 | 68 | initialGuts :: Addressable a t => (ProcCh a, t) 69 | initialGuts = (Nothing, τ0) 70 | 71 | 72 | instance (Ord s, Ord a, Ord t, Addressable a t, Lattice s, StoreLike a s (D a)) => 73 | AddStepToFP (NonSharedAnalysis s (ProcCh a, t)) (PΣ a) 74 | (ℙ ((PΣ a, (ProcCh a, t)), s)) where 75 | applyStep step fp = 76 | joinWith (\((p,g),s) -> Set.fromList $ concat $ runIdentity $ 77 | collectListT (collectSSListT (runStateT (mapStateT mergeSharedState $ do {x <- step p; gc x; return x }) g) s)) 78 | fp 79 | inject p = Set.singleton $ ((p, initialGuts), bot) 80 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/Proof.hs: -------------------------------------------------------------------------------- 1 | module CFA.Analysis.Proof where 2 | 3 | import Data.Set as Set 4 | import Data.Foldable as Foldable 5 | 6 | import CFA.Lattice 7 | import CFA.CPS 8 | import CFA.CPS.Analysis 9 | 10 | import CFA.CPS.Analysis.SingleStore 11 | import CFA.CPS.Analysis.ReallyNonShared 12 | 13 | -- conjecture: 14 | -- if g :: b -> c 15 | -- f :: a -> Set b 16 | -- then 17 | -- joinWith g . joinWith f 18 | -- === 19 | -- joinWith (joinWith g . f) 20 | -- proof: 21 | -- ? 22 | 23 | -- conjecture 2: 24 | -- f ⊑ g then joinWith f ⊑ joinWith g 25 | 26 | -- conjecture 3: 27 | -- joinWith f . map g == joinWith (f . g) 28 | 29 | -- proof: 30 | -- joinWith g . joinWith f 31 | 32 | -- conjecture 4: 33 | -- joinWith (\x -> (f x, g x)) == (\xs -> joinWith f xs, joinWith g xs) 34 | 35 | alpha :: (Lattice s, Ord a, Ord t) => 36 | ℙ ((PΣ a, (ProcCh a, t)), s) -> (ℙ (PΣ a, (ProcCh a, t)), s) 37 | alpha = joinWith (\((p, g), s) -> (Set.singleton (p,g), s)) 38 | 39 | -- lemma 0: 40 | -- alpha a `union` alpha b == alpha (a `union` b) 41 | -- lemma 1: 42 | -- foldr ((⊔) . f) bot 43 | -- ⊑ 44 | -- foldr ((⊔) . g) bot 45 | -- if f ⊑ g 46 | 47 | -- lemma 2: 48 | -- foldr ((⊔) . f) bot . foldr (Set.union . g) Set.empty 49 | -- 50 | 51 | -- to prove: alpha . applyStepToFP mnext ⊑ applyStepToFP mnext . alpha 52 | -- joinWith (\((p, g), s) -> (Set.singleton (p,g), s)) . 53 | -- joinWith (\ ((p,g),s) -> Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s)) 54 | where g == (\((p, g), s) -> (Set.singleton (p,g), s)) 55 | f == (\ ((p,g),s) -> Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s)) 56 | joinWith (joinWith g . f) === 57 | joinWith (joinWith (\((p, g), s) -> (Set.singleton (p,g), s)) . (\ ((p,g),s) -> Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s))) === 58 | joinWith ((\((p,g),s) -> joinWith (\((p, g), s) -> (Set.singleton (p,g), s)) $ Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s))) 59 | joinWith ((\((p,g),s) -> (joinWith (\((p, g), s) -> (Set.singleton (p,g))) $ Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s), 60 | joinWith (\((p, g), s) -> s) $ Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s)))) 61 | 62 | joinWith (\((p,g),s) -> let newStates = Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s) 63 | in (Set.map (\((p, g), s) -> (p,g)) newStates, 64 | joinWith (\((p, g), s) -> s) newStates)) 65 | 66 | joinWith (\((p,g),s) -> let newStates = Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s) 67 | in (joinWith (\((p, g), s) -> Set.singleton (p,g)) newStates, 68 | joinWith (\((p, g), s) -> s) newStates)) 69 | 70 | \states -> joinWith (\((p,g),s) -> let newStates = Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s) 71 | in joinWith (\((p, g), s) -> (Set.singleton (p,g), s)) newStates) states 72 | 73 | \states -> joinWith (\((p,g),s) -> joinWith (\((p', g'),s') -> (Set.singleton (p',g'), s')) (Set.fromList $ runIdentity $ collectListT (runStateT (runStateT (gc $ step p) g) s))) states 74 | 75 | -- step is monotonous... 76 | 77 | 78 | -- \states -> joinWith (\((p,g),_) -> joinWith (\((p', g'),s') -> (Set.singleton (p',g'), s')) $ Set.fromList $ runIdentity $ collectListT $ runStateT (runSSListT $ runStateT (gc $ step p) g) (joinWith (\((p,g),s) -> s) states)) states 79 | 80 | -- \states -> joinWith (\((p,g),_) -> mapFst Set.fromList $ mergeState $ runIdentity $ collectListT $ runStateT (runSSListT $ runStateT (gc $ step p) g) (joinWith (\((p,g),s) -> s) states)) states 81 | 82 | -- \states -> joinWith (\((p,g),_) -> mapFst Set.fromList $ runIdentity $ liftM mergeState $ collectListT $ runStateT (runSSListT $ runStateT (gc $ step p) g) (joinWith (\((p,g),s) -> s) states)) states 83 | 84 | -- \states -> joinWith (\((p,g),_) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) (joinWith (\((p,g),s) -> s) states)) states 85 | 86 | -- (\states -> joinWith (let s' = joinWith (\((p,g),s) -> s) states 87 | -- in \(p,g) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) s' . (\((p, g), s) -> (p,g))) states) 88 | 89 | -- (\states -> let s' = joinWith (\((p,g),s) -> s) states 90 | -- in joinWith (\(p,g) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) s' . (\((p, g), s) -> (p,g))) states) 91 | 92 | -- (\states -> let states' = map (\((p, g), s) -> (p,g)) states 93 | -- s = joinWith (\((p,g),s) -> s) states 94 | -- in joinWith (\(p,g) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) s) states') 95 | -- (\states -> (let (states', s) = (map (\((p, g), s) -> (p,g)) states, joinWith (\((p,g),s) -> s) states)) in joinWith (\(p,g) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) s) states') 96 | -- (\states -> (\(states', s) -> joinWith (\(p,g) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) s) states') $ (map (\((p, g), s) -> (p,g)) states, joinWith (\((p,g),s) -> s) states)) 97 | -- (\(states, s) -> joinWith (\(p,g) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) s) states) . (\states -> (map (\((p, g), s) -> (p,g)) states, joinWith (\((p,g),s) -> s) states)) 98 | -- (\(states, s) -> joinWith (\(p,g) -> mapFst Set.fromList $ runIdentity $ collectSSListTS (runStateT (gc $ step p) g) s) states) . (joinWith (\((p, g), s) -> (Set.singleton (p,g), s))) 99 | 100 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/ReallyNonShared.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FlexibleContexts #-} 4 | {-# LANGUAGE FunctionalDependencies #-} 5 | {-# LANGUAGE TypeOperators #-} 6 | {-# LANGUAGE ParallelListComp #-} 7 | {-# LANGUAGE TypeSynonymInstances #-} 8 | {-# LANGUAGE ScopedTypeVariables #-} 9 | {-# LANGUAGE GeneralizedNewtypeDeriving #-} 10 | 11 | -- TODO: get rid of this 12 | {-# LANGUAGE UndecidableInstances #-} 13 | 14 | module CFA.CPS.Analysis.ReallyNonShared where 15 | 16 | import Data.Maybe 17 | import Data.Map as Map 18 | import Data.Set as Set 19 | import Data.List as List 20 | import Data.Foldable as Foldable 21 | import Data.Traversable 22 | import Control.Monad.State 23 | import Control.Monad.ListT 24 | import Control.Monad.Identity 25 | import Control.Monad.Reader 26 | 27 | import CFA.CPS 28 | import CFA.Lattice 29 | import CFA.Store 30 | import CFA.CFAMonads 31 | import CFA.CPS.Analysis 32 | import CFA.Runner 33 | 34 | import Util 35 | 36 | type ReallyNonSharedAnalysis s g = StateT g (StateT s (ListT Identity)) 37 | 38 | instance (Addressable a t, StoreLike a s (D a)) 39 | => Analysis (ReallyNonSharedAnalysis s (ProcCh a, t)) a 40 | where 41 | fun ρ (Lam l) = return $ Clo(l, ρ) 42 | fun ρ (Ref v) = lift $ getsNDSet $ (flip fetch $ ρ!v) 43 | 44 | arg ρ (Lam l) = return $ Clo(l, ρ) 45 | arg ρ (Ref v) = lift $ getsNDSet $ flip fetch (ρ!v) 46 | 47 | a $= d = lift $ modify $ \σ -> bind σ a (Set.singleton d) 48 | 49 | alloc v = gets (valloc v . snd) 50 | 51 | tick proc ps = modify $ \(_, t) -> (Just proc, advance proc ps t) 52 | 53 | -- Garbage Collection 54 | instance (Lattice s, Eq a, StoreLike a s (D a), Ord a) => 55 | GarbageCollector (ReallyNonSharedAnalysis s g) (PΣ a) where 56 | gc ps = do 57 | σ <- lift get 58 | let rs = Set.map (\(v, a) -> a) (reachable ps σ) 59 | lift $ modify $ \ σ -> filterStore σ (\a -> Set.member a rs) 60 | return () 61 | 62 | initialGuts :: Addressable a t => (ProcCh a, t) 63 | initialGuts = (Nothing, τ0) 64 | 65 | instance Addressable a t => HasInitial (ProcCh a, t) where 66 | initial = initialGuts 67 | 68 | newtype RNSFP a g s = RNSFP { unRNSFP :: ℙ ((PΣ a, g), s) } deriving (Lattice) 69 | 70 | instance (Ord s, Ord a, Ord g, HasInitial g, Lattice s, StoreLike a s (D a)) => 71 | AddStepToFP (ReallyNonSharedAnalysis s g) (PΣ a) 72 | (RNSFP a g s) where 73 | applyStep step (RNSFP fp) = 74 | RNSFP $ joinWith 75 | (\ ((p,g),s) -> Set.fromList $ runIdentity $ 76 | collectListT (runStateT (runStateT (do {x <- step p; gc x; return x}) g) s)) 77 | fp 78 | inject c = RNSFP $ Set.singleton $ ((c, initial), bot) 79 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/Shape/Basic.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE ParallelListComp #-} 6 | {-# LANGUAGE TypeSynonymInstances #-} 7 | 8 | module CFA.CPS.Analysis.Shape.Basic where 9 | 10 | import Data.Map as Map 11 | import Data.Set as Set 12 | import Data.List as List 13 | 14 | import CFA.CPS 15 | import CFA.Lattice 16 | import CFA.Store 17 | 18 | import CFA.CFAMonads 19 | import CFA.CPS.Analysis 20 | 21 | instance (Ord a, StoreLike a s d) => AlkaliLike a (s, ℙ a) d where 22 | addUniqueAddr a = undefined 23 | deAnodizeStore σ = undefined 24 | deAnodizeEnv σ ρ = undefined 25 | deAnodizeD σ d = undefined 26 | reset σ = undefined 27 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/SimpleAnalysis.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FlexibleContexts #-} 4 | {-# LANGUAGE FunctionalDependencies #-} 5 | {-# LANGUAGE TypeOperators #-} 6 | {-# LANGUAGE ParallelListComp #-} 7 | {-# LANGUAGE TypeSynonymInstances #-} 8 | {-# LANGUAGE ScopedTypeVariables #-} 9 | {-# LANGUAGE GeneralizedNewtypeDeriving #-} 10 | 11 | -- TODO: get rid of this 12 | {-# LANGUAGE UndecidableInstances #-} 13 | 14 | module CFA.CPS.Analysis.SimpleAnalysis where 15 | 16 | import Data.Maybe 17 | import Data.Map as Map 18 | import Data.Set as Set 19 | import Data.List as List 20 | import Data.Foldable as Foldable 21 | import Data.Traversable 22 | import Control.Monad.State 23 | import Control.Monad.ListT 24 | import Control.Monad.Identity 25 | import Control.Monad.Reader 26 | 27 | import CFA.CPS 28 | import CFA.Lattice 29 | import CFA.Store 30 | import CFA.CFAMonads 31 | import CFA.CPS.Analysis 32 | import CFA.Runner 33 | 34 | import Util 35 | 36 | ---------------------------------------------------------------------- 37 | -- integer time-stamped concrete interpreter 38 | ---------------------------------------------------------------------- 39 | 40 | 41 | instance Analysis (StorePassingSemantics (Store Integer) Integer) Integer 42 | where 43 | fun ρ (Lam l) = return $ Clo(l, ρ) 44 | fun ρ (Ref v) = lift $ getsNDSet $ \σ -> σ!(ρ!v) 45 | 46 | arg ρ (Lam l) = return $ Clo(l, ρ) 47 | arg ρ (Ref v) = lift $ getsNDSet $ \σ -> σ!(ρ!v) 48 | 49 | a $= d = lift $ modify $ \σ -> Map.insert a (Set.singleton d) σ 50 | alloc v = gets id 51 | tick proc ps = modify $ \t -> t + 1 52 | 53 | instance HasInitial Integer where 54 | initial = 0 55 | 56 | instance (Ord a, StoreLike a s (D a)) 57 | => GarbageCollector (StorePassingSemantics s t) (PΣ a) where 58 | gc ps = do 59 | σ <- lift get 60 | let rs = Set.map (\(v, a) -> a) (reachable ps σ) 61 | lift $ modify $ \ σ -> filterStore σ (\a -> Set.member a rs) 62 | return () 63 | 64 | 65 | instance (Ord s, Ord a, Ord g, HasInitial g, Lattice s) => 66 | AddStepToFP (StorePassingSemantics s g) a 67 | (ℙ ((a, g), s)) where 68 | applyStep step fp = 69 | joinWith 70 | (\((p,t),s) -> Set.fromList $ runIdentity $ 71 | collectListT (runStateT (runStateT (step p) t) s)) 72 | fp 73 | inject p = Set.singleton $ ((p, initial), bot) 74 | 75 | idx = Lam (["x"] :=> Call (Ref "x") []) 76 | idy = Lam (["y"] :=> Exit) 77 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 78 | ex = Call comb [idx, idy] 79 | 80 | result e = runAnalysis e :: ℙ ((PΣ Integer, Integer), Store Integer) -------------------------------------------------------------------------------- /CFA/CPS/Analysis/SingleStore.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FlexibleContexts #-} 4 | {-# LANGUAGE FunctionalDependencies #-} 5 | {-# LANGUAGE TypeOperators #-} 6 | {-# LANGUAGE ParallelListComp #-} 7 | {-# LANGUAGE TypeSynonymInstances #-} 8 | {-# LANGUAGE TypeFamilies #-} 9 | 10 | -- TODO: get rid of this 11 | {-# LANGUAGE UndecidableInstances #-} 12 | 13 | module CFA.CPS.Analysis.SingleStore where 14 | 15 | import Data.Map as Map 16 | import Data.Set as Set 17 | import Data.List as List 18 | import Data.Foldable as Foldable 19 | import Control.Monad.State 20 | import Control.Monad.Reader 21 | import Control.Monad.Identity 22 | 23 | import Debug.Trace 24 | 25 | import CFA.CPS 26 | import CFA.Lattice 27 | import CFA.Store 28 | import CFA.CFAMonads 29 | import CFA.CPS.Analysis 30 | import CFA.Runner 31 | import CFA.CPS.Analysis.ReallyNonShared 32 | 33 | import Util 34 | 35 | alpha :: (Lattice s, Ord a, Ord g) => 36 | ℙ ((a, g), s) -> (ℙ (a, g), s) 37 | alpha = joinWith (\((p, g), s) -> (Set.singleton (p,g), s)) 38 | 39 | gamma :: (Ord a, Ord g, Ord s) => 40 | (ℙ (a, g), s) -> ℙ ((a, g), s) 41 | gamma (states, s) = Set.map (\(p, g) -> ((p,g), s)) states 42 | 43 | instance (Ord g, Ord a, Lattice s, StoreLike a s (D a), Ord s, HasInitial g) => 44 | AddStepToFP (ReallyNonSharedAnalysis s g) (PΣ a) (ℙ (PΣ a, g), s) where 45 | applyStep step = alpha . unRNSFP . applyStep step . RNSFP . gamma 46 | 47 | inject a = (Set.singleton (a, initial), bot) 48 | 49 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/SingleStore/SingleStoreGC.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE ParallelListComp #-} 6 | {-# LANGUAGE TypeSynonymInstances #-} 7 | 8 | -- TODO: get rid of this 9 | {-# LANGUAGE UndecidableInstances #-} 10 | 11 | module CFA.CPS.Analysis.SingleStore.SingleStoreGC where 12 | 13 | import Data.Map as Map 14 | import Data.Set as Set 15 | import Data.List as List 16 | 17 | import CFA.CPS 18 | import CFA.Lattice 19 | import CFA.Store 20 | import CFA.CFAMonads 21 | import CFA.CPS.Analysis 22 | 23 | instance (Ord a, StoreLike a s (D a)) 24 | => GarbageCollector (SharedAnalysis s0 s g) (PΣ a) where 25 | gc ps = SSFA (\σ -> \g -> 26 | let rs = Set.map (\(v, a) -> a) (reachable ps σ) 27 | σ' = filterStore σ (\a -> not (Set.member a rs)) 28 | in (σ', [((), g)])) 29 | -------------------------------------------------------------------------------- /CFA/CPS/Analysis/SingleStoreShape.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE ParallelListComp #-} 6 | {-# LANGUAGE TypeSynonymInstances #-} 7 | 8 | -- TODO: get rid of this 9 | {-# LANGUAGE UndecidableInstances #-} 10 | 11 | module CFA.CPS.Analysis.SingleStoreShape where 12 | 13 | import Data.Map as Map 14 | import Data.Set as Set 15 | import Data.List as List 16 | 17 | import CFA.CPS 18 | import CFA.Lattice 19 | import CFA.Store 20 | import CFA.CFAMonads 21 | 22 | import CFA.CPS.Analysis 23 | import CFA.CPS.Analysis.SingleStore.SingleStoreGC 24 | 25 | -- Shape analysis 26 | instance (Addressable a t, StoreLike a s (D a), AlkaliLike a s (D a)) 27 | => Analysis (SingleStoreAnalysis a) 28 | a 29 | s 30 | (ProcCh a, t) 31 | where 32 | fun ρ (Lam l) = SSFA (\σ -> \(_,t) -> 33 | let proc = Clo(l, ρ) 34 | in (σ, [(proc, (Just proc,t)) ])) 35 | fun ρ (Ref v) = SSFA (\σ -> \(_,t) -> 36 | let procs = fetch σ (ρ!v) 37 | in (σ, [ (proc, (Just proc,t)) | proc <- Set.toList procs ])) 38 | 39 | arg ρ (Lam l) = SSFA (\σ -> \(ch, t) -> 40 | let proc = Clo(l, ρ) 41 | in (σ, [ (Set.singleton proc, (ch, t)) ])) 42 | arg ρ (Ref v) = SSFA (\σ -> \( ch, t) -> 43 | let procs = fetch σ (ρ!v) 44 | in (σ, [ (procs, (ch,t)) ])) 45 | 46 | a $= d = SSFA (\σ -> \(ch, t) -> 47 | let σ' = deAnodizeStore σ 48 | σ'' = bind σ' a (deAnodizeD σ d) 49 | in (σ'', [((), (ch, t))] )) 50 | 51 | alloc v = SSFA (\σ -> \(ch, t) -> 52 | let addr = valloc v t 53 | σ' = addUniqueAddr addr 54 | in (σ', [(addr, (ch, t))])) 55 | 56 | updateEnv ρ bs = SSFA (\σ -> \( ch, t) -> 57 | let ρ' = deAnodizeEnv σ ρ 58 | in (σ, [ (ρ' // bs, (ch,t)) ])) 59 | 60 | tick ps = SSFA (\σ -> \ (Just proc, t) -> 61 | (σ, ([((), (Just proc, advance proc ps t))]))) 62 | 63 | stepAnalysis store config state = runWithStore (mnext state >>= gc) (reset store) config 64 | 65 | inject call = ((call, Map.empty), σ0, (Nothing, τ0)) 66 | -------------------------------------------------------------------------------- /CFA/CPS/Examples.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE ParallelListComp #-} 6 | {-# LANGUAGE TypeSynonymInstances #-} 7 | 8 | module CFA.CPS.Examples where 9 | 10 | import CFA.CPS 11 | import CFA.Lattice 12 | import CFA.CPS.KCFA 13 | 14 | ---------------------------------------------------------------------- 15 | -- example program 16 | ---------------------------------------------------------------------- 17 | 18 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 19 | idx = Lam (["x"] :=> Call (Ref "x") []) 20 | idy = Lam (["y"] :=> Exit) 21 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 22 | ex = Call comb [idx, idy] 23 | 24 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 25 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 26 | omega = Call ucombx [ucomby] 27 | -------------------------------------------------------------------------------- /CFA/CPS/Examples/AbstractNonShared.hs: -------------------------------------------------------------------------------- 1 | module CFA.CPS.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | 7 | import CFA.CPS 8 | import CFA.CFAMonads 9 | import CFA.Lattice 10 | import CFA.CPS.Analysis 11 | import CFA.Runner 12 | import Control.Monad.State 13 | import Control.Monad.List 14 | import Control.Monad.Identity 15 | import Control.Monad.Reader 16 | 17 | ---------------------------------------------------------------------- 18 | -- abstract interpreter with a per-state store 19 | ---------------------------------------------------------------------- 20 | 21 | import CFA.Store 22 | import CFA.CPS.KCFA 23 | import CFA.CPS.Analysis.NonShared 24 | import CFA.CPS.Analysis.ReallyNonShared 25 | 26 | ---------------------------------------------------------------------- 27 | -- example program 28 | ---------------------------------------------------------------------- 29 | 30 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 31 | idx = Lam (["x"] :=> Call (Ref "x") []) 32 | idy = Lam (["y"] :=> Exit) 33 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 34 | ex = Call comb [idx, idy] 35 | 36 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 37 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 38 | omega = Call ucombx [ucomby] 39 | 40 | ---------------------------------------------------------------------- 41 | 42 | instance KCFA KTime where 43 | getK = const 2 44 | 45 | type AbstractGuts = (ProcCh KAddr, KTime) 46 | 47 | reallyNonSharedResultC_ :: CExp -> Set ((PΣ KAddr, AbstractGuts), Store KAddr) 48 | reallyNonSharedResultC_ e = unRNSFP $ exploreFP mnext (e, ρ0) 49 | 50 | nonSharedResultC :: CExp -> Set ((PΣ KAddr, AbstractGuts), Store KAddr) 51 | nonSharedResultC e = exploreFP mnext (e, ρ0) 52 | -------------------------------------------------------------------------------- /CFA/CPS/Examples/AbstractNonSharedCount.hs: -------------------------------------------------------------------------------- 1 | module CFA.CPS.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | 7 | import CFA.CPS 8 | import CFA.Lattice 9 | import CFA.Store 10 | import CFA.CPS.Analysis 11 | import CFA.Runner 12 | 13 | ---------------------------------------------------------------------- 14 | -- abstract interpreter with a per-state store and counting 15 | ---------------------------------------------------------------------- 16 | 17 | import CFA.CPS.KCFA 18 | import CFA.CPS.Analysis.NonShared 19 | import CFA.CPS.Analysis.ReallyNonShared 20 | 21 | ---------------------------------------------------------------------- 22 | -- example program 23 | ---------------------------------------------------------------------- 24 | 25 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 26 | idx = Lam (["x"] :=> Call (Ref "x") []) 27 | idy = Lam (["y"] :=> Exit) 28 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 29 | ex = Call comb [idx, idy] 30 | 31 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 32 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 33 | omega = Call ucombx [ucomby] 34 | 35 | ---------------------------------------------------------------------- 36 | 37 | instance KCFA KTime where 38 | getK = const 1 39 | 40 | type AbstractGuts = (ProcCh KAddr, KTime) 41 | 42 | abstractResultC :: CExp -> Set ((PΣ KAddr, AbstractGuts), StoreWithCount KAddr (D KAddr)) 43 | abstractResultC e = exploreFP mnext (e, ρ0) 44 | -------------------------------------------------------------------------------- /CFA/CPS/Examples/AbstractShared.hs: -------------------------------------------------------------------------------- 1 | module CFA.CPS.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | import Control.Monad.State 7 | import Control.Monad.Reader 8 | import Control.Monad.Identity 9 | 10 | import CFA.CPS 11 | import CFA.CFAMonads 12 | import CFA.Lattice 13 | import CFA.Store 14 | import CFA.CPS.Analysis 15 | import CFA.Runner 16 | import CFA.CPS.Analysis.SingleStore 17 | 18 | ---------------------------------------------------------------------- 19 | -- abstract interpreter with a shared store 20 | ---------------------------------------------------------------------- 21 | 22 | import CFA.CPS.KCFA 23 | import CFA.CPS.Analysis.SingleStore 24 | 25 | ---------------------------------------------------------------------- 26 | -- example program 27 | ---------------------------------------------------------------------- 28 | 29 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 30 | idx = Lam (["x"] :=> Call (Ref "x") []) 31 | idy = Lam (["y"] :=> Exit) 32 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 33 | ex = Call comb [idx, idy] 34 | 35 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 36 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 37 | omega = Call ucombx [ucomby] 38 | 39 | ---------------------------------------------------------------------- 40 | 41 | instance KCFA KTime where 42 | getK = const 1 43 | 44 | 45 | type AbstractGutsSS = (ProcCh KAddr, KTime) 46 | 47 | abstractResultSSC :: CExp -> (Set (PΣ KAddr, AbstractGutsSS), Store KAddr) 48 | abstractResultSSC e = exploreFP mnext (e, ρ0) 49 | -------------------------------------------------------------------------------- /CFA/CPS/Examples/AbstractSharedCount.hs: -------------------------------------------------------------------------------- 1 | module CFA.CPS.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | 7 | import CFA.CPS 8 | import CFA.Lattice 9 | import CFA.Store 10 | 11 | import CFA.CPS.Analysis 12 | import CFA.Runner 13 | 14 | ---------------------------------------------------------------------- 15 | -- abstract interpreter with a shared store and counting 16 | ---------------------------------------------------------------------- 17 | 18 | import CFA.CPS.KCFA 19 | import CFA.CPS.Analysis.SingleStore 20 | 21 | ---------------------------------------------------------------------- 22 | -- example program 23 | ---------------------------------------------------------------------- 24 | 25 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 26 | idx = Lam (["x"] :=> Call (Ref "x") []) 27 | idy = Lam (["y"] :=> Exit) 28 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 29 | ex = Call comb [idx, idy] 30 | 31 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 32 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 33 | omega = Call ucombx [ucomby] 34 | 35 | ---------------------------------------------------------------------- 36 | 37 | instance KCFA KTime where 38 | getK = const 1 39 | 40 | type AbstractGutsSS = (ProcCh KAddr, KTime) 41 | 42 | abstractResultSSC :: CExp -> (Set (PΣ KAddr, AbstractGutsSS), StoreWithCount KAddr (D KAddr)) 43 | abstractResultSSC e = exploreFP mnext (e, ρ0) 44 | -------------------------------------------------------------------------------- /CFA/CPS/Examples/AbstractSharedShape.hs: -------------------------------------------------------------------------------- 1 | module CFA.CPS.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | 7 | import CFA.CPS 8 | import CFA.Lattice 9 | import CFA.Store 10 | 11 | import CFA.CPS.Analysis 12 | import CFA.CPS.Analysis.Runner 13 | 14 | ---------------------------------------------------------------------- 15 | -- abstract interpreter with a shared store 16 | ---------------------------------------------------------------------- 17 | 18 | import CFA.CPS.KCFA 19 | import CFA.CPS.Analysis.SingleStoreShape 20 | import CFA.CPS.Analysis.Shape.Basic 21 | 22 | ---------------------------------------------------------------------- 23 | -- example program 24 | ---------------------------------------------------------------------- 25 | 26 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 27 | idx = Lam (["x"] :=> Call (Ref "x") []) 28 | idy = Lam (["y"] :=> Exit) 29 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 30 | ex = Call comb [idx, idy] 31 | 32 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 33 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 34 | omega = Call ucombx [ucomby] 35 | 36 | ---------------------------------------------------------------------- 37 | 38 | instance KCFA KTime where 39 | getK = const 1 40 | 41 | type AbstractGutsSS = (ProcCh KAddr, KTime) 42 | 43 | abstractResultSh :: CExp -> ((Store KAddr, ℙ KAddr), Set (PΣ KAddr, AbstractGutsSS)) 44 | abstractResultSh = explore -------------------------------------------------------------------------------- /CFA/CPS/Examples/Concrete.hs: -------------------------------------------------------------------------------- 1 | module CFA.CPS.Examples where 2 | 3 | import Data.Map as Map 4 | import Data.Set as Set 5 | import Data.List as List 6 | 7 | import CFA.CPS 8 | import CFA.CPS.Analysis 9 | import CFA.CPS.Analysis.Interpreter 10 | 11 | ---------------------------------------------------------------------- 12 | -- example program 13 | ---------------------------------------------------------------------- 14 | 15 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 16 | idx = Lam (["x"] :=> Call (Ref "x") []) 17 | idy = Lam (["y"] :=> Exit) 18 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 19 | ex = Call comb [idx, idy] 20 | 21 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 22 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 23 | omega = Call ucombx [ucomby] 24 | 25 | ---------------------------------------------------------------------- 26 | 27 | concreteResult :: CExp -> IO (PΣ IOAddr) 28 | concreteResult = interpret 29 | 30 | 31 | 32 | -------------------------------------------------------------------------------- /CFA/CPS/KCFA.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE ParallelListComp #-} 3 | {-# LANGUAGE TypeSynonymInstances #-} 4 | {-# LANGUAGE MultiParamTypeClasses #-} 5 | {-# LANGUAGE FlexibleInstances #-} 6 | {-# LANGUAGE FlexibleContexts #-} 7 | 8 | module CFA.CPS.KCFA where 9 | 10 | import Data.Map as Map 11 | 12 | import CFA.CPS 13 | import CFA.Lattice 14 | import CFA.Store 15 | import CFA.CPS.Analysis 16 | 17 | 18 | class KCFA a where 19 | getK :: a -> Int 20 | 21 | data KTime = KCalls [CExp] 22 | deriving (Eq, Ord, Show) 23 | 24 | data KAddr = KBind Var KTime 25 | deriving (Eq, Ord, Show) 26 | 27 | instance (KCFA KTime) => Addressable KAddr KTime where 28 | τ0 = KCalls [] 29 | valloc v t = KBind v t 30 | advance proc (call, ρ) t@(KCalls calls) = 31 | KCalls $ take (getK t) (call : calls) 32 | 33 | -- Simple store 34 | instance StoreLike KAddr (Store KAddr) (D KAddr) where 35 | σ0 = Map.empty 36 | 37 | bind σ a d = σ ⨆ [a ==> d] 38 | fetch σ a = σ CFA.Lattice.!! a 39 | replace σ a d = σ ⨆ [a ==> d] 40 | filterStore σ p = Map.filterWithKey (\k -> \v -> p k) σ 41 | -------------------------------------------------------------------------------- /CFA/Lattice.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE ParallelListComp #-} 6 | {-# LANGUAGE TypeSynonymInstances #-} 7 | {-# LANGUAGE GeneralizedNewtypeDeriving #-} 8 | 9 | module CFA.Lattice where 10 | 11 | import Data.Map as Map 12 | import Data.Set as Set 13 | import Data.List as List 14 | import Data.Monoid 15 | 16 | -- Abbreviations. 17 | type k :-> v = Map.Map k v 18 | type ℙ a = Set.Set a 19 | 20 | (==>) :: a -> b -> (a,b) 21 | (==>) x y = (x,y) 22 | 23 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 24 | (//) f [] = f 25 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 26 | 27 | (⊎) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 28 | f ⊎ [] = f 29 | f ⊎ ((k,v):tl) = Map.insertWith (⊔) k v (f ⊎ tl) 30 | 31 | 32 | -- Partial order theory. 33 | class Lattice a where 34 | bot :: a 35 | top :: a 36 | (⊑) :: a -> a -> Bool 37 | (⊔) :: a -> a -> a 38 | (⊓) :: a -> a -> a 39 | 40 | -- unit is a trivial lattice 41 | instance Lattice () where 42 | bot = () 43 | top = () 44 | x ⊔ y = top 45 | x ⊓ y = bot 46 | x ⊑ y = True 47 | 48 | instance (Lattice a, Lattice b) => Lattice (a, b) where 49 | bot = (bot, bot) 50 | top = (top, top) 51 | (a1, b1) ⊔ (a2, b2) = (a1 ⊔ a2, b1 ⊔ b2) 52 | (a1, b1) ⊓ (a2, b2) = (a1 ⊓ a2, b1 ⊓ b2) 53 | (a1, b1) ⊑ (a2, b2) = (a1 ⊑ a2) && (b1 ⊑ b2) 54 | 55 | instance (Ord s, Eq s) => Lattice (ℙ s) where 56 | bot = Set.empty 57 | top = error "no representation of universal set" 58 | x ⊔ y = x `Set.union` y 59 | x ⊓ y = x `Set.intersection` y 60 | x ⊑ y = x `Set.isSubsetOf` y 61 | 62 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 63 | bot = Map.empty 64 | top = error "no representation of top map" 65 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 66 | f ⊔ g = Map.unionWith (⊔) f g 67 | f ⊓ g = Map.intersectionWith (⊓) f g 68 | 69 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 70 | f ⨆ [] = f 71 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 72 | 73 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 74 | f !! k = Map.findWithDefault bot k f 75 | 76 | -- a lattice is a monoid in two ways... 77 | newtype JoinMonoid a = JoinMonoid { withJoinMonoid :: a } deriving (Show, Eq, Lattice) 78 | 79 | instance Lattice a => Monoid (JoinMonoid a) where 80 | mempty = bot 81 | mappend = (⊔) 82 | 83 | newtype MeetMonoid a = MeetMonoid { withMeetMonoid :: a } deriving (Show, Eq, Lattice) 84 | 85 | instance Lattice a => Monoid (MeetMonoid a) where 86 | mempty = top 87 | mappend = (⊓) 88 | 89 | ljoin :: Lattice a => [a] -> a 90 | ljoin l = withJoinMonoid $ mconcat $ List.map JoinMonoid l 91 | -------------------------------------------------------------------------------- /CFA/Runner.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FlexibleContexts #-} 4 | {-# LANGUAGE FunctionalDependencies #-} 5 | {-# LANGUAGE TypeOperators #-} 6 | {-# LANGUAGE ScopedTypeVariables #-} 7 | {-# LANGUAGE TypeSynonymInstances #-} 8 | {-# LANGUAGE GeneralizedNewtypeDeriving #-} 9 | {-# LANGUAGE RankNTypes #-} 10 | {-# LANGUAGE GADTs #-} 11 | {-# LANGUAGE TypeFamilies #-} 12 | 13 | module CFA.Runner where 14 | 15 | import Data.Map as Map 16 | import Data.Set as Set 17 | import Data.List as List 18 | import Data.Traversable 19 | import Data.Foldable as Foldable 20 | import Control.Monad.Trans 21 | import Control.Monad.Identity 22 | import Control.Monad.State 23 | import Control.Applicative 24 | 25 | import CFA.Lattice 26 | import CFA.CFAMonads 27 | 28 | class HasInitial a where 29 | initial :: a 30 | 31 | findFP :: forall a m. (Lattice a) => (a -> a) -> a 32 | findFP f = loop bot 33 | where loop :: a -> a 34 | loop c = let c' = f c in if c' ⊑ c then c else loop c' 35 | 36 | class AddStepToFP m a fp | m -> a, fp -> m where 37 | applyStep :: (a -> m a) -> fp -> fp 38 | inject :: a -> fp 39 | 40 | exploreFP :: forall m a fp. 41 | (Lattice fp, AddStepToFP m a fp) => 42 | (a -> m a) -> a -> fp 43 | exploreFP step s = findFP loop 44 | where loop :: fp -> fp 45 | loop acc = inject s ⊔ applyStep step acc 46 | -------------------------------------------------------------------------------- /CFA/Store.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE TypeSynonymInstances #-} 6 | 7 | module CFA.Store where 8 | 9 | import Data.Map as Map 10 | import Data.Set as Set 11 | import Data.List as List 12 | 13 | import CFA.Lattice 14 | 15 | -- Store uniquely defines the type of its addresses 16 | class (Eq a, Lattice s, Lattice d) 17 | => StoreLike a s d | s->a, s->d where 18 | σ0 :: s 19 | bind :: s -> a -> d -> s 20 | replace :: s -> a -> d -> s 21 | fetch :: s -> a -> d 22 | filterStore :: s -> (a -> Bool) -> s 23 | 24 | ---------------------------------------------------------------------- 25 | -- Abstract Garbage Collection 26 | ---------------------------------------------------------------------- 27 | 28 | -- Abstract garbage collector 29 | class Monad m => GarbageCollector m a where 30 | gc :: a -> m () 31 | -- default implementation 32 | gc _ = return () 33 | 34 | ---------------------------------------------------------------------- 35 | -- Abstract Counting 36 | ---------------------------------------------------------------------- 37 | 38 | -- Abstract natural number 39 | data AbsNat = AZero | AOne | AMany 40 | deriving (Ord, Eq, Show) 41 | 42 | instance Lattice AbsNat where 43 | bot = AZero 44 | top = AMany 45 | n ⊑ m = (n == bot) || (m == top) || (n == m) 46 | n ⊔ m = if (n ⊑ m) then m else n 47 | n ⊓ m = if (n ⊑ m) then n else m 48 | 49 | -- Abstract addition 50 | (⊕) :: AbsNat -> AbsNat -> AbsNat 51 | AZero ⊕ n = n 52 | n ⊕ AZero = n 53 | n ⊕ m = AMany 54 | 55 | class ACounter a s where 56 | count :: s -> a -> AbsNat 57 | 58 | -- Counting store 59 | type StoreWithCount a d = a :-> (d, AbsNat) 60 | 61 | instance (Ord a, Lattice d) => ACounter a (StoreWithCount a d) where 62 | -- fetching with default bottom 63 | count σ a = snd $ σ CFA.Lattice.!! a 64 | 65 | -- counter is nullified when filtered 66 | -- and incremented when `bind' is called 67 | instance (Ord a, Lattice d) => StoreLike a (StoreWithCount a d) d where 68 | σ0 = Map.empty 69 | bind σ a d = σ `update_add` [a ==> (d, AOne)] 70 | fetch σ a = fst $ σ CFA.Lattice.!! a 71 | replace σ a d = σ ⨆ [a ==> (d, AZero)] 72 | filterStore σ p = Map.filterWithKey (\k -> \v -> p k) σ 73 | 74 | update_add :: (Ord k, Lattice v) => (k :-> (v, AbsNat)) -> [(k, (v, AbsNat))] -> (k :-> (v, AbsNat)) 75 | update_add f [] = f 76 | update_add f ((k,v):tl) = Map.insertWith (\(x1, y1) -> \(x2, y2) -> (x1 ⊔ x2, y1 ⊕ y2)) k v (update_add f tl) 77 | 78 | ---------------------------------------------------------------------- 79 | -- Utility 80 | ---------------------------------------------------------------------- 81 | 82 | class Truncatable t where 83 | trunc :: t -> t 84 | -------------------------------------------------------------------------------- /CRAPL-LICENSE.txt: -------------------------------------------------------------------------------- 1 | THE CRAPL v0 BETA 0 2 | 3 | 4 | 0. Information about the CRAPL 5 | 6 | If you have questions or concerns about the CRAPL, or you need more 7 | information about this license, please contact: 8 | 9 | Matthew Might 10 | http://matt.might.net/ 11 | 12 | 13 | I. Preamble 14 | 15 | Science thrives on openness. 16 | 17 | In modern science, it is often infeasible to replicate claims without 18 | access to the software underlying those claims. 19 | 20 | Let's all be honest: when scientists write code, aesthetics and 21 | software engineering principles take a back seat to having running, 22 | working code before a deadline. 23 | 24 | So, let's release the ugly. And, let's be proud of that. 25 | 26 | 27 | II. Definitions 28 | 29 | 1. "This License" refers to version 0 beta 0 of the Community 30 | Research and Academic Programming License (the CRAPL). 31 | 32 | 2. "The Program" refers to the medley of source code, shell scripts, 33 | executables, objects, libraries and build files supplied to You, 34 | or these files as modified by You. 35 | 36 | [Any appearance of design in the Program is purely coincidental and 37 | should not in any way be mistaken for evidence of thoughtful 38 | software construction.] 39 | 40 | 3. "You" refers to the person or persons brave and daft enough to use 41 | the Program. 42 | 43 | 4. "The Documentation" refers to the Program. 44 | 45 | 5. "The Author" probably refers to the caffeine-addled graduate 46 | student that got the Program to work moments before a submission 47 | deadline. 48 | 49 | 50 | III. Terms 51 | 52 | 1. By reading this sentence, You have agreed to the terms and 53 | conditions of this License. 54 | 55 | 2. If the Program shows any evidence of having been properly tested 56 | or verfied, You will disregard this evidence. 57 | 58 | 3. You agree to hold the Author free from shame, embarrassment or 59 | ridicule for any hacks, kludges or leaps of faith found within the 60 | Program. 61 | 62 | 4. You recognize that any request for support for the Program will be 63 | discarded with extreme prejudice. 64 | 65 | 5. The Author reserves all rights to the Program, except for any 66 | rights granted under any additional licenses attached to the 67 | Program. 68 | 69 | 70 | IV. Permissions 71 | 72 | 1. You are permitted to use the Program to validate published 73 | scientific claims. 74 | 75 | 2. You are permitted to use the Program to validate scientific claims 76 | submitted for peer review, under the condition that You keep 77 | modifications to the Program confidential until those claims have 78 | been published. 79 | 80 | 2. You are permitted to use and/or modify the Program for the 81 | validation of novel scientific claims if You make a good-faith 82 | attempt to notify the Author of Your work and Your claims prior to 83 | submission for publication. 84 | 85 | 3. If You publicly release any claims or data that were supported or 86 | generated by the Program or a modification thereof, in whole or in 87 | part, You will release any inputs supplied to the Program and any 88 | modifications You made to the Progam. This License will be in 89 | effect for the modified program. 90 | 91 | 92 | V. Disclaimer of Warranty 93 | 94 | THERE IS NO WARRANTY FOR THE PROGRAM, TO THE EXTENT PERMITTED BY 95 | APPLICABLE LAW. EXCEPT WHEN OTHERWISE STATED IN WRITING THE COPYRIGHT 96 | HOLDERS AND/OR OTHER PARTIES PROVIDE THE PROGRAM "AS IS" WITHOUT 97 | WARRANTY OF ANY KIND, EITHER EXPRESSED OR IMPLIED, INCLUDING, BUT NOT 98 | LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR 99 | A PARTICULAR PURPOSE. THE ENTIRE RISK AS TO THE QUALITY AND 100 | PERFORMANCE OF THE PROGRAM IS WITH YOU. SHOULD THE PROGRAM PROVE 101 | DEFECTIVE, YOU ASSUME THE COST OF ALL NECESSARY SERVICING, REPAIR OR 102 | CORRECTION. 103 | 104 | 105 | VI. Limitation of Liability 106 | 107 | IN NO EVENT UNLESS REQUIRED BY APPLICABLE LAW OR AGREED TO IN WRITING 108 | WILL ANY COPYRIGHT HOLDER, OR ANY OTHER PARTY WHO MODIFIES AND/OR 109 | CONVEYS THE PROGRAM AS PERMITTED ABOVE, BE LIABLE TO YOU FOR DAMAGES, 110 | INCLUDING ANY GENERAL, SPECIAL, INCIDENTAL OR CONSEQUENTIAL DAMAGES 111 | ARISING OUT OF THE USE OR INABILITY TO USE THE PROGRAM (INCLUDING BUT 112 | NOT LIMITED TO LOSS OF DATA OR DATA BEING RENDERED INACCURATE OR 113 | LOSSES SUSTAINED BY YOU OR THIRD PARTIES OR A FAILURE OF THE PROGRAM 114 | TO OPERATE WITH ANY OTHER PROGRAMS), EVEN IF SUCH HOLDER OR OTHER 115 | PARTY HAS BEEN ADVISED OF THE POSSIBILITY OF SUCH DAMAGES. 116 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | GHC=ghc -XTypeSynonymInstances -XParallelListComp -XTypeOperators -XMultiParamTypeClasses -XFlexibleInstances 2 | 3 | all: cps cesk afj 4 | 5 | # Abstract machine for Lambda calculus in CPS 6 | cps: CPSSimple CPSConcrete CPSAbstractNonShared CPSAbstractNonSharedCount CPSAbstractShared CPSAbstractSharedCount 7 | 8 | CPSSimple: CFA/CPS/Examples/Concrete.hs 9 | $(GHC) --make CFA/CPS/Analysis/SimpleAnalysis.hs 10 | 11 | CPSConcrete: CFA/CPS/Examples/Concrete.hs 12 | $(GHC) --make CFA/CPS/Examples/Concrete.hs 13 | 14 | CPSAbstractNonShared: CFA/CPS/Examples/AbstractNonShared.hs 15 | $(GHC) --make CFA/CPS/Examples/AbstractNonShared.hs 16 | 17 | CPSAbstractNonSharedCount: CFA/CPS/Examples/AbstractNonSharedCount.hs 18 | $(GHC) --make CFA/CPS/Examples/AbstractNonSharedCount.hs 19 | 20 | CPSAbstractShared: CFA/CPS/Examples/AbstractShared.hs 21 | $(GHC) --make CFA/CPS/Examples/AbstractShared.hs 22 | 23 | CPSAbstractSharedCount: CFA/CPS/Examples/AbstractSharedCount.hs 24 | $(GHC) --make CFA/CPS/Examples/AbstractSharedCount.hs 25 | 26 | # CESK machine for Lambda calculus 27 | cesk: CESKConcrete CESKAbstractNonShared 28 | 29 | CESKConcrete: CFA/CESK/Examples/Concrete.hs 30 | $(GHC) --make CFA/CESK/Examples/Concrete.hs 31 | 32 | CESKAbstractNonShared: CFA/CESK/Examples/AbstractNonShared.hs 33 | $(GHC) --make CFA/CESK/Examples/AbstractNonShared.hs 34 | 35 | # CESK machine for A-Normal Featherweight Java 36 | afj: AFJAbstractNonShared 37 | 38 | AFJAbstractNonShared: CFA/AFJ/Examples/AbstractNonShared.hs 39 | $(GHC) --make CFA/AFJ/Examples/AbstractNonShared.hs 40 | 41 | 42 | clean: 43 | find . -type f -name "*.o" -exec rm -fv {} \; 44 | find . -type f -name "*.hi" -exec rm -fv {} \; 45 | 46 | 47 | -------------------------------------------------------------------------------- /README: -------------------------------------------------------------------------------- 1 | Project structure: 2 | ------------------ 3 | 4 | CFA - General modules, shred between different CFA instances and languages 5 | |- AFJ - A family of analyses based on CESK Abstract Machine 6 | | for A-Normal Featherweight Java 7 | |- CESK - A family of analyses based on CESK Abstract Machine 8 | | for direct-style Lambda-calculus 9 | |- CPS - A family of analyses based on CPS Scheme 10 | 11 | How to build: 12 | ------------- 13 | 14 | Build everything: 15 | 16 | make all 17 | 18 | Build CPS-based analyses: 19 | 20 | make cps 21 | 22 | Build CESK-based analyses: 23 | 24 | make cesk 25 | 26 | Build ANF FJ-based analyses: 27 | 28 | make afj 29 | 30 | Clean: 31 | 32 | make clean -------------------------------------------------------------------------------- /Util.hs: -------------------------------------------------------------------------------- 1 | module Util where 2 | 3 | mapFst :: (a -> b) -> (a,c) -> (b,c) 4 | mapFst f (a,c) = (f a, c) 5 | 6 | mapSnd :: (b -> c) -> (a, b) -> (a,c) 7 | mapSnd f (a,b) = (a, f b) 8 | 9 | -------------------------------------------------------------------------------- /obsolete/CFA.hs: -------------------------------------------------------------------------------- 1 | module Main where 2 | 3 | -- Imports. 4 | import CPS 5 | import Data.Map as Map 6 | import Data.Set as Set 7 | 8 | -- Abbreviations. 9 | type k :-> v = Map.Map k v 10 | type ℙ a = Set.Set a 11 | 12 | (==>) :: a -> b -> (a,b) 13 | (==>) x y = (x,y) 14 | 15 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 16 | (//) f [] = f 17 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 18 | 19 | 20 | -- Partial order theory. 21 | class Lattice a where 22 | bot :: a 23 | top :: a 24 | (⊑) :: a -> a -> Bool 25 | (⊔) :: a -> a -> a 26 | (⊓) :: a -> a -> a 27 | 28 | 29 | instance (Ord s, Eq s) => Lattice (ℙ s) where 30 | bot = Set.empty 31 | top = error "no representation of universal set" 32 | x ⊔ y = x `Set.union` y 33 | x ⊓ y = x `Set.intersection` y 34 | x ⊑ y = x `Set.isSubsetOf` y 35 | 36 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 37 | bot = Map.empty 38 | top = error "no representation of top map" 39 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 40 | f ⊔ g = Map.unionWith (⊔) f g 41 | f ⊓ g = Map.intersectionWith (⊓) f g 42 | 43 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 44 | f ⨆ [] = f 45 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 46 | 47 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 48 | f !! k = Map.findWithDefault bot k f 49 | 50 | 51 | 52 | -- State-space. 53 | type Σ = (CExp, Env, Store, Time) 54 | type Env = Var :-> Addr 55 | type Store = Addr :-> D 56 | type D = ℙ Val 57 | data Val = Clo (Lambda, Env) 58 | deriving (Eq,Ord) 59 | type Addr = (Var,Time) 60 | type Time = [CExp] 61 | 62 | 63 | -- Helpers. 64 | arg :: (AExp, Env, Store) -> D 65 | arg (Ref v, ρ, σ) = σ!(ρ!v) 66 | arg (Lam l, ρ, σ) = Set.singleton (Clo (l, ρ)) 67 | 68 | 69 | -- k-CFA-style allocation. 70 | k = 1 71 | 72 | tick :: (Val,Σ) -> Time 73 | tick (_,(call,_,_,t)) = take k (call:t) 74 | 75 | alloc :: (Var, Time, Val, Σ) -> Addr 76 | alloc (v, t', proc, ς) = (v,t') 77 | 78 | -- Transition. 79 | next :: Σ -> [Σ] 80 | next ς@(Call f aes, ρ, σ, t) = [ (call, ρ'', σ', t') | 81 | proc@(Clo (vs :=> call, ρ')) <- Set.toList (arg (f, ρ, σ)), 82 | let t' = tick (proc, ς), 83 | let as = [ alloc (v, t', proc, ς) | v <- vs], 84 | let ds = [ arg(ae, ρ, σ) | ae <- aes ], 85 | let ρ'' = ρ' // [ v ==> a | v <- vs | a <- as ], 86 | let σ' = σ ⨆ [ a ==> d | a <- as | d <- ds ] ] 87 | 88 | mnext :: Σ -> [Σ] 89 | mnext ς@(Call f aes, ρ, σ, t) = do 90 | proc@(Clo (vs :=> call, ρ')) <- Set.toList (arg (f, ρ, σ)) 91 | let t' = tick (proc, ς) 92 | let as = [ alloc (v, t', proc, ς) | v <- vs] 93 | let ds = [ arg(ae, ρ, σ) | ae <- aes ] 94 | let ρ'' = ρ' // [ v ==> a | v <- vs | a <- as ] 95 | let σ' = σ ⨆ [ a ==> d | a <- as | d <- ds ] 96 | return $ (call, ρ'', σ', t') 97 | 98 | 99 | 100 | main :: IO () 101 | main = do 102 | return () 103 | 104 | -------------------------------------------------------------------------------- /obsolete/CPS.hs: -------------------------------------------------------------------------------- 1 | module CPS where 2 | 3 | -- Syntax. 4 | type Var = String 5 | 6 | data Lambda = [Var] :=> CExp 7 | deriving (Ord, Eq, Show) 8 | 9 | data AExp = Ref Var 10 | | Lam Lambda 11 | deriving (Ord, Eq, Show) 12 | 13 | data CExp = Call AExp [AExp] 14 | | Exit 15 | deriving (Ord, Eq, Show) 16 | 17 | -------------------------------------------------------------------------------- /obsolete/MonadAFJ.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE TypeSynonymInstances #-} 3 | {-# LANGUAGE ParallelListComp #-} 4 | {-# LANGUAGE MultiParamTypeClasses #-} 5 | {-# LANGUAGE ScopedTypeVariables #-} 6 | {-# LANGUAGE ImplicitParams #-} 7 | {-# LANGUAGE TupleSections #-} 8 | 9 | module Main where 10 | 11 | -- Imports. 12 | import Data.Map as Map hiding (map) 13 | import Data.Set as Set hiding (map) 14 | import qualified Data.List as L 15 | 16 | -- Abbreviations. 17 | type k :-> v = Map.Map k v 18 | type ℙ a = Set.Set a 19 | 20 | (==>) :: a -> b -> (a,b) 21 | (==>) x y = (x,y) 22 | 23 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 24 | (//) f [] = f 25 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 26 | 27 | 28 | {--------------------- DOMAIN THEORY ---------------------} 29 | 30 | class Lattice a where 31 | bot :: a 32 | top :: a 33 | (⊑) :: a -> a -> Bool 34 | (⊔) :: a -> a -> a 35 | (⊓) :: a -> a -> a 36 | 37 | 38 | instance (Ord s, Eq s) => Lattice (ℙ s) where 39 | bot = Set.empty 40 | top = error "no representation of universal set" 41 | x ⊔ y = x `Set.union` y 42 | x ⊓ y = x `Set.intersection` y 43 | x ⊑ y = x `Set.isSubsetOf` y 44 | 45 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 46 | bot = Map.empty 47 | top = error "no representation of top map" 48 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 49 | f ⊔ g = Map.unionWith (⊔) f g 50 | f ⊓ g = Map.intersectionWith (⊓) f g 51 | 52 | (⊎) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 53 | f ⊎ [] = f 54 | f ⊎ ((k,v):tl) = Map.insertWith (⊔) k v (f ⊎ tl) 55 | 56 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 57 | f !! k = Map.findWithDefault bot k f 58 | 59 | 60 | {---------------- SYNTAX ----------------} 61 | 62 | -- `a' stands for the nature of addresses 63 | -- it is also a leaf of the state-space 64 | 65 | type Var = String 66 | type FieldName = Var 67 | type ClassName = String 68 | type MethodName = String 69 | type Lab = String 70 | 71 | type Class = (ClassName, Maybe ClassName, [(ClassName, FieldName)], 72 | Konstr, [Method]) 73 | 74 | type Konstr = (ClassName, [(ClassName, Var)], [Var], [(FieldName, Var)]) 75 | 76 | type Method = (ClassName, MethodName, [(ClassName, Var)], 77 | [(ClassName, Var)], [Stmt]) 78 | 79 | data Stmt = Asgn Var Var Lab 80 | | Lkp Var Var FieldName Lab 81 | | MCall Var Var MethodName [Var] Lab 82 | | New Var ClassName [Var] Lab 83 | | Cast Var ClassName Var Lab 84 | | Ret Var Lab 85 | deriving (Eq,Ord) 86 | 87 | data ClassTable = CTable [Class] 88 | 89 | {---------------- HELPERS ----------------} 90 | 91 | -- retrieve labels 92 | lab :: Stmt -> Lab 93 | lab (Asgn _ _ l) = l 94 | lab (Lkp _ _ _ l) = l 95 | lab (MCall _ _ _ _ l) = l 96 | lab (New _ _ _ l) = l 97 | lab (Cast _ _ _ l) = l 98 | lab (Ret _ l) = l 99 | 100 | -- get class name 101 | name :: Class -> ClassName 102 | name (cn, _, _, _, _) = cn 103 | 104 | super :: Class -> Maybe ClassName 105 | super (_, cn, _, _, _) = cn 106 | 107 | -- findClass 108 | findClass :: ClassTable -> ClassName -> Maybe Class 109 | findClass (CTable classes) cn = L.find (\c -> (name c) == cn) classes 110 | 111 | -- returns ALL fields for the given class name (including super's) 112 | allFields :: ClassTable -> ClassName -> [FieldName] 113 | allFields table cn = maybe [] (\clz@(_, _, pairs, _, _) -> 114 | let properFields = L.map snd pairs 115 | in case (super clz) of 116 | Just superName -> properFields ++ (allFields table superName) 117 | Nothing -> properFields) 118 | (findClass table cn) 119 | 120 | -- transitive method lookup 121 | method :: ClassTable -> ClassName -> MethodName -> Method 122 | method table cn m = 123 | case findClass table cn of 124 | Just clz@(_, _, _, _, methods) -> 125 | case L.find (\(_, mname, _, _, _) -> mname == m) methods of 126 | Just mtd -> mtd 127 | Nothing -> maybe undefined (\sup -> method table sup m) (super clz) 128 | Nothing -> undefined 129 | 130 | -- find class constructors 131 | konstr :: Class -> Konstr 132 | konstr (_, _, _, k, _) = k 133 | 134 | -- get the list of field mappings by a chain of constructors for a given class 135 | classFieldMappings :: ClassTable -> ClassName -> [a] -> [(FieldName, a)] 136 | classFieldMappings table cn values = 137 | maybe [] (\clz -> fieldMappings table (konstr clz) values) 138 | (findClass table cn) 139 | 140 | fieldMappings :: ClassTable -> Konstr -> [a] -> [(FieldName, a)] 141 | fieldMappings table (cn, params, superBindings, thisBindings) values = 142 | let paramBindings = Map.empty // (zip (map snd params) values) 143 | thisMappings = map (\(f, v) -> (f, paramBindings ! v)) thisBindings 144 | superArgs = map (\v -> paramBindings ! v) superBindings 145 | in case (superArgs, 146 | (findClass table cn) >>= (return . konstr)) of 147 | (h:t, Just superKonstr) -> 148 | thisMappings ++ (fieldMappings table superKonstr superArgs) 149 | _ -> thisMappings 150 | 151 | 152 | {---------------- STATE-SPACE ----------------} 153 | 154 | -- a - for addresses 155 | 156 | type BEnv a = Var :-> a 157 | type Kont a = (Var, [Stmt], BEnv a, a) 158 | type State a = ([Stmt], BEnv a, a) 159 | type Obj a = (ClassName, BEnv a) 160 | 161 | class (Eq a, Ord a) => Address a 162 | 163 | {----------------- MONADIC SEMANTICS ------------------} 164 | 165 | -- Hint: Add new primitives as they appear in the semantics 166 | class (Address a, Monad m) => JavaSemanticInterface m a where 167 | tick :: State a -> m () 168 | getObj :: BEnv a -> Var -> m (Obj a) 169 | putObj :: BEnv a -> Var -> Obj a -> m () 170 | getCont :: a -> m (Kont a) 171 | putCont :: MethodName -> (Kont a) -> m a 172 | getC :: (?table :: ClassTable) => ClassName -> m ([Obj a] -> m (BEnv a)) 173 | getM :: (?table :: ClassTable) => Obj a -> MethodName -> m Method 174 | initBEnv :: BEnv a -> [Var] -> [Var] -> m (BEnv a) 175 | 176 | 177 | mstep :: (JavaSemanticInterface m a, ?table :: ClassTable) => State a -> m (State a) 178 | mstep ctx@((Asgn v v' l):succ, β, pk) = do 179 | tick ctx 180 | d <- getObj β v' 181 | putObj β v d 182 | return $! (succ, β, pk) 183 | mstep ctx@((Ret v l):[], β, pk) = do 184 | tick ctx 185 | (v', s, β', pk') <- getCont pk 186 | d <- getObj β v 187 | putObj β' v' d 188 | return $! (s, β', pk') 189 | mstep ctx@((Lkp v v' f l):succ, β, pk) = do 190 | tick ctx 191 | (c, β') <- getObj β v' 192 | d <- getObj β' f 193 | putObj β v d 194 | return $! (succ, β, pk) 195 | mstep ctx@((Cast v cn v' l):succ, β, pk) = do 196 | tick ctx 197 | d <- getObj β v' 198 | putObj β v d 199 | return $! (succ, β, pk) 200 | mstep ctx@((New v cn vs l):succ, β, pk) = do 201 | tick ctx 202 | ructor <- getC cn 203 | ds <- sequence [getObj β vi | vi <- vs] 204 | β' <- ructor ds 205 | let d' = (cn, β') 206 | putObj β v d' 207 | return $! (succ, β, pk) 208 | mstep ctx@((MCall v v0 m vs l):succ, β, pk) = do 209 | tick ctx 210 | d0 <- getObj β v0 211 | (cn, _ , params, locals, body) <- getM d0 m 212 | ds <- sequence [getObj β vi | vi <- vs] 213 | let κ = (v, succ, β, pk) 214 | pk' <- putCont m κ 215 | let vs'' = map snd params 216 | let vs''' = map snd locals 217 | let β' = Map.empty // ["this" ==> (β ! v0)] 218 | β'' <- initBEnv β' vs'' vs''' 219 | sequence [putObj β'' vi di | vi <- vs'' | di <- ds] 220 | return $! (body, β'', pk') 221 | 222 | {----------------- TIME-STAMPED ADDRESSES ------------------} 223 | 224 | data Addr = AVar Var [Lab] 225 | | ACall MethodName [Lab] 226 | deriving (Eq, Ord, Show) 227 | 228 | type Time = [Lab] 229 | 230 | instance Address Addr 231 | 232 | 233 | {---------------- CONCRETE INTERPRETATION ----------------} 234 | 235 | data D = Val (Obj Addr) 236 | | Cont (Kont Addr) 237 | deriving (Eq, Ord) 238 | 239 | -- TODO: implement the concrete interpreter 240 | 241 | {---------------- ABSTRACT INTERPRETATION ----------------} 242 | 243 | -- Used EXACTLY the same monad as for KCFA in Lambda 244 | 245 | data KCFA a = KCFA { kf :: !((Time, AStore) -> [(a, Time, AStore)]) } 246 | 247 | --TODO is it possible to abstract over the list structure? 248 | type D1 = (ℙ D) 249 | 250 | type AStore = Addr :-> D1 251 | 252 | k = 1 253 | 254 | instance Monad KCFA where 255 | return a = KCFA (\(t, σ) -> [(a, t, σ)]) 256 | (>>=) (KCFA f) g = KCFA (\ (t, σ) -> 257 | let chs = f(t, σ) 258 | in concatMap (\ (a, t', σ') -> (kf $ g(a))(t', σ')) chs) 259 | 260 | instance JavaSemanticInterface KCFA Addr where 261 | tick ctx@(stmts, _, _) = KCFA (\(t, σ) -> [((), take k ((lab (head stmts)):t), σ)]) 262 | 263 | getObj β v = KCFA (\(t, σ) -> 264 | [(d, t, σ) | Val d <- Set.toList $ σ!(β!v)]) 265 | 266 | putObj β v d = KCFA (\(t, σ) -> 267 | [((), t, σ ⊎ [(β!v) ==> (Set.singleton (Val d))])]) 268 | 269 | getCont pk = KCFA (\(t, σ) -> 270 | [(κ, t, σ) | Cont κ <- Set.toList $ σ ! pk]) 271 | 272 | putCont m κ = KCFA (\(t, σ) -> 273 | let b = alloc_k t m in 274 | [(b, t, σ ⊎ [b ==> (Set.singleton (Cont κ))])]) 275 | 276 | initBEnv β vs'' vs''' = KCFA (\(t, σ) -> 277 | let pairs' = map (\v -> (v, alloc t v)) vs'' 278 | pairs'' = map (\v -> (v, alloc t v)) vs''' in 279 | let β' = β // pairs' // pairs'' in 280 | [(β', t, σ)]) 281 | 282 | getC cn = KCFA (\(t, σ) -> 283 | -- updates a store and returns an environment of all class fields 284 | let ructor = (\ds -> KCFA(\(t', σ') -> 285 | let fs = allFields ?table cn -- compute all fields 286 | as = map (alloc t) fs -- appropriate addresses for fields 287 | fBindings = zip fs as -- bindings [field |-> addr] 288 | -- mapping from all class fields to provided arguments 289 | fMappings = Map.empty // classFieldMappings ?table cn ds 290 | -- heap is updated according to the mappings 291 | σ'' = σ' ⊎ [ai ==> (Set.singleton (Val $ fMappings ! fi)) | (fi, ai) <- fBindings] 292 | -- new environment is create 293 | β' = Map.empty // fBindings 294 | in [(β', t', σ'')])) 295 | in [(ructor, t, σ)]) 296 | 297 | getM (cn, _) m = KCFA (\(t, σ) -> [(method ?table cn m, t, σ)]) 298 | 299 | alloc :: Time -> Var -> Addr 300 | alloc t v = AVar v (take k t) 301 | 302 | alloc_k :: Time -> MethodName -> Addr 303 | alloc_k t m = ACall m (take k t) 304 | -------------------------------------------------------------------------------- /obsolete/MonadCESK.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TypeOperators #-} 2 | {-# LANGUAGE ParallelListComp #-} 3 | {-# LANGUAGE TypeSynonymInstances #-} 4 | {-# LANGUAGE TupleSections #-} 5 | {-# LANGUAGE MultiParamTypeClasses #-} 6 | {-# LANGUAGE ScopedTypeVariables #-} 7 | {-# LANGUAGE ImplicitParams #-} 8 | 9 | module CFA.CESK where 10 | 11 | -- Imports. 12 | import Data.Map as Map 13 | import Data.Set as Set 14 | 15 | import CFA.Lattice 16 | 17 | {---------------- SYNTAX AND STATE-SPACE ----------------} 18 | 19 | -- `a' stands for the nature of addresses 20 | -- it is also a leaf of the state-space 21 | 22 | type Var = String 23 | type Lab = String 24 | type Env a = Var :-> a 25 | type Lambda = (Var, Exp) 26 | type State a = (Exp, Env a, a) 27 | 28 | data Exp = Ref (Var, Lab) 29 | | App (Exp, Exp, Lab) 30 | | Lam (Lambda, Lab) 31 | deriving (Eq, Ord, Show) 32 | 33 | -- retrieve a label 34 | lab :: Exp -> Lab 35 | lab (Ref (_, l)) = l 36 | lab (App (_, _, l)) = l 37 | lab (Lam (_, l)) = l 38 | 39 | data Clo a = Clo (Lambda, Env a, Lab) 40 | deriving (Eq, Ord) 41 | 42 | data Kont a = Mt 43 | | Ar (Exp, Env a, a) 44 | | Fn (Lambda, Env a, a) 45 | deriving (Eq, Ord) 46 | 47 | class (Eq a, Ord a) => Address a 48 | 49 | class (Address a, Monad m) => Analysis m a where 50 | tick :: State a -> m () 51 | getVar :: Env a -> Var -> m (Clo a) 52 | putVar :: Env a -> Var -> Clo a -> m (Env a) 53 | loadCont :: a -> m (Kont a) 54 | storeCont :: Kont a -> m a 55 | 56 | -- A small-step monadic semantics for CESK* machine 57 | -- in Store- and Time- passing style 58 | 59 | -- The underlying monad implemens a `time-state-store-passing style' 60 | -- in order to disguise `alloc' calls, which use both time and state 61 | -- (at least, in the abstract case) 62 | 63 | {------------------ MONADIC SEMANTICS --------------------} 64 | 65 | mstep :: (Analysis m a) => State a -> m (State a) 66 | mstep ctx@(Ref (x, l), ρ, a) = do 67 | tick ctx 68 | Clo (v, ρ', l) <- getVar ρ x 69 | return $! (Lam (v, l), ρ', a) 70 | mstep ctx@(App (e0, e1, l), ρ, a) = do 71 | tick ctx 72 | b <- storeCont (Ar (e1, ρ, a)) 73 | return (e0, ρ, b) 74 | mstep ctx@(Lam (v, l), ρ, a) = do 75 | tick ctx 76 | κ <- loadCont a 77 | case κ of 78 | Ar (e, ρ, a) -> do 79 | b <- storeCont κ 80 | return (e, ρ, b) 81 | Fn ((x, e), ρ', c) -> do 82 | ρ'' <- putVar ρ' x (Clo (v, ρ, l)) 83 | return (e, ρ'', c) 84 | 85 | {----------------- TIME-STAMPED ADDRESSES ------------------} 86 | 87 | data Addr = Bind Var [Lab] 88 | | Call Lab [Lab] 89 | deriving (Eq, Ord, Show) 90 | 91 | data Time = TLab Lab [Lab] 92 | | TMt [Lab] 93 | deriving (Eq, Ord, Show) 94 | 95 | contour :: Time -> [Lab] 96 | contour (TLab _ c) = c 97 | contour (TMt c) = c 98 | 99 | instance Address Addr 100 | 101 | {---------------- CONCRETE INTERPRETATION ----------------} 102 | 103 | data D = Val (Clo Addr) 104 | | Cont (Kont Addr) 105 | deriving (Eq, Ord) 106 | 107 | type Store = Addr :-> D 108 | 109 | data CESKx a = CESKx { 110 | st :: !((Time, State Addr, Store) -> (a, Time, State Addr, Store)) 111 | } 112 | 113 | instance Monad CESKx where 114 | return a = CESKx (\(t, s, σ) -> (a, t, s, σ)) 115 | (>>=) (CESKx f) g = CESKx (\(t, s, σ) -> 116 | let result = f (t, s, σ) 117 | in (\(a, t', s', σ') -> (st $ g(a))(t', s', σ')) result) 118 | 119 | instance Analysis CESKx Addr where 120 | tick ctx@(Ref (_, _), ρ, a) = CESKx (\(t, s, σ) -> ((), t, ctx, σ)) 121 | tick ctx@(App (_, _, l), ρ, a) = CESKx (\(t, s, σ) -> ((), TLab l (contour t), ctx, σ)) 122 | tick ctx@(v, ρ, a) = CESKx (\(TLab l ctr, s, σ) -> 123 | let (Cont κ) = σ ! a 124 | in case κ of 125 | Ar _ -> ((), TLab l ctr, ctx, σ) 126 | Fn _ -> ((), TMt (l:ctr), ctx, σ)) 127 | 128 | getVar ρ x = CESKx (\(t, s, σ) -> 129 | let (Val clo) = σ ! (ρ ! x) 130 | in (clo, t, s, σ)) 131 | 132 | putVar ρ x c = CESKx (\(t, s, σ) -> 133 | let b = alloc t s σ 134 | σ' = σ // [(b, Val c)] 135 | in (ρ // [(x, b)], t, s, σ')) 136 | 137 | loadCont a = CESKx (\(t, s, σ) -> 138 | let (Cont κ) = σ ! a 139 | in (κ, t, s, σ)) 140 | 141 | storeCont κ = CESKx (\(t, s, σ) -> 142 | let b = alloc t s σ 143 | σ' = σ // [(b, Cont κ)] 144 | in (b, t, s, σ')) 145 | 146 | 147 | 148 | -- Concrete allocator function 149 | -- Turns the last captured time moment into the address 150 | alloc :: Time -> State Addr -> Store -> Addr 151 | alloc t (App (e0, _, _), _ ,_) σ = Call (lab e0) (contour t) 152 | alloc t (Lam _, _, a) σ = 153 | case σ ! a of 154 | Cont (Ar (e, _, _)) -> Call (lab e) (contour t) 155 | Cont (Fn ((x, _), _, _)) -> Bind x (contour t) 156 | 157 | {---------------- ABSTRACT INTERPRETATION ----------------} 158 | 159 | data KCFA a = KCFA { kf :: !((Time, State Addr, AStore) -> [(a, Time, State Addr, AStore)]) } 160 | 161 | --TODO is it possible to abstract over the list structure? 162 | type D1 = (ℙ D) 163 | 164 | type AStore = Addr :-> D1 165 | 166 | k = 1 167 | 168 | instance Monad KCFA where 169 | return a = KCFA (\(t, s, σ) -> [(a, t, s, σ)]) 170 | (>>=) (KCFA f) g = KCFA (\ (t, s, σ) -> 171 | let chs = f(t, s, σ) 172 | in concatMap (\ (a, t', s', σ') -> (kf $ g(a))(t', s', σ')) chs) 173 | 174 | instance Analysis KCFA Addr where 175 | tick ctx@(Ref (_, _), ρ, a) = KCFA (\(t, s, σ) -> [((), t, ctx, σ)]) 176 | tick ctx@(App (_, _, l), ρ, a) = KCFA (\(t, s, σ) -> [((), TLab l (contour t), ctx, σ)]) 177 | tick ctx@(v, ρ, a) = KCFA (\(TLab l ctr, s, σ) -> 178 | [case κ of 179 | Ar _ -> ((), TLab l ctr, ctx, σ) 180 | Fn _ -> ((), TMt (take k (l : ctr)), ctx, σ) 181 | | Cont κ <- Set.toList (σ ! a)]) 182 | 183 | getVar ρ x = KCFA (\(t, s, σ) -> 184 | let clos = σ ! (ρ ! x) 185 | in [(clo, t, s, σ) | Val clo <- Set.toList clos]) 186 | 187 | putVar ρ x c = KCFA (\(t, s, σ) -> do 188 | b <- allocKCFA t s σ 189 | let d = (Set.singleton (Val c)) 190 | σ' = σ ⊎ [b ==> d] 191 | ρ' = ρ // [(x, b)] 192 | return (ρ', t, s, σ')) 193 | 194 | loadCont a = KCFA (\(t, s, σ) -> 195 | let ks = σ ! a 196 | in [(κ, t, s, σ) | Cont κ <- Set.toList ks]) 197 | 198 | storeCont κ = KCFA (\(t, s, σ) -> do 199 | b <- allocKCFA t s σ 200 | let d = (Set.singleton (Cont κ)) 201 | σ' = σ ⊎ [b ==> d] 202 | return (b, t, s, σ')) 203 | 204 | -- abstract allocator function 205 | -- similar to the concrete allocator 206 | -- nondeterministic because of stored continuations 207 | allocKCFA :: Time -> State Addr -> AStore -> [Addr] 208 | allocKCFA t (App (e0, _, _), _ ,_) σ = [Call (lab e0) (contour t)] 209 | allocKCFA t (Lam _, _, a) σ = 210 | [case κ of 211 | Ar (e, _, _) -> Call (lab e) (contour t) 212 | Fn ((x, _), _, _) -> Bind x (contour t) 213 | | Cont κ <- Set.toList (σ ! a)] 214 | 215 | 216 | {-------------------------- EXAMPLE -------------------------} 217 | 218 | -- provide the initial state 219 | -- State = (Exp, Env, Addr) 220 | state0 = (Lam (("x", Ref ("x", "l2")), "l1"), Map.empty, Call "l3" []) 221 | 222 | -- define the denotational semantics of the first step 223 | -- literally, mstep - is a denotational semantics, induced by the 224 | -- abstract machine over time and store 225 | transition = mstep state0 :: KCFA (State Addr) 226 | 227 | -- run analysis with the suplied time and store 228 | result = kf transition (TMt [], undefined, Map.empty) -------------------------------------------------------------------------------- /obsolete/MonadCFA.hs: -------------------------------------------------------------------------------- 1 | module Main where 2 | 3 | -- Imports. 4 | import CPS 5 | import Data.Map as Map 6 | import Data.Set as Set 7 | 8 | -- Abbreviations. 9 | type k :-> v = Map.Map k v 10 | type ℙ a = Set.Set a 11 | 12 | (==>) :: a -> b -> (a,b) 13 | (==>) x y = (x,y) 14 | 15 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 16 | (//) f [] = f 17 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 18 | 19 | 20 | -- Partial order theory. 21 | class Lattice a where 22 | bot :: a 23 | top :: a 24 | (⊑) :: a -> a -> Bool 25 | (⊔) :: a -> a -> a 26 | (⊓) :: a -> a -> a 27 | 28 | 29 | instance (Ord s, Eq s) => Lattice (ℙ s) where 30 | bot = Set.empty 31 | top = error "no representation of universal set" 32 | x ⊔ y = x `Set.union` y 33 | x ⊓ y = x `Set.intersection` y 34 | x ⊑ y = x `Set.isSubsetOf` y 35 | 36 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 37 | bot = Map.empty 38 | top = error "no representation of top map" 39 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 40 | f ⊔ g = Map.unionWith (⊔) f g 41 | f ⊓ g = Map.intersectionWith (⊓) f g 42 | 43 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 44 | f ⨆ [] = f 45 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 46 | 47 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 48 | f !! k = Map.findWithDefault bot k f 49 | 50 | 51 | 52 | -- State-space. 53 | type Ctx = (CExp, Env, Time) 54 | type Env = Var :-> Addr 55 | type Store = Addr :-> D 56 | type D = ℙ Val 57 | data Val = Clo (Lambda, Env) 58 | deriving (Eq,Ord) 59 | data Addr = Bind Var Time 60 | deriving (Eq,Ord) 61 | data Time = CallSeq [CExp] 62 | deriving (Eq,Ord) 63 | 64 | 65 | 66 | class Monad m => Analysis m where 67 | funEval :: Env -> AExp -> m Val 68 | argEval :: Env -> AExp -> m D 69 | 70 | ($=) :: Addr -> D -> m () 71 | 72 | alloc :: Var -> m Addr 73 | tick :: (Ctx,Time) -> m Time 74 | 75 | -- class Monad m where 76 | -- (>>=) :: m a -> (a -> m b) -> m b 77 | -- (>>) :: m a -> m b -> m b 78 | -- return :: a -> m a 79 | -- fail :: String -> m a 80 | 81 | 82 | 83 | data Exact a = Exact !(Store -> (a,Store)) 84 | data KCFA a = KCFA { kf :: !((Choices,Store) -> [(a,(Choices,Store))]) } 85 | data Choice = Invoke Val 86 | | Ticked Time 87 | 88 | type Choices = [Choice] 89 | 90 | 91 | timeFrom ((Ticked t):_) = t 92 | 93 | k = 1 94 | 95 | -- return a >>= k = k a 96 | -- m >>= return = m 97 | -- m >>= (\x -> k x >>= h) = (m >>= k) >>= h 98 | 99 | 100 | instance Monad KCFA where 101 | (>>=) (KCFA f) g = KCFA (\ (ch, σ) -> 102 | let chs = f(ch, σ) 103 | in concatMap (\ (a, (ch',σ')) -> (kf $ g(a))(ch', σ')) chs) 104 | return a = KCFA (\ (ch, σ) -> [(a,(ch, σ))]) 105 | 106 | 107 | 108 | instance Analysis KCFA where 109 | funEval ρ (Lam l) = KCFA (\ (ch,σ) -> 110 | let proc = Clo(l, ρ) 111 | in [ (proc,((Invoke proc) : ch, σ)) ]) 112 | funEval ρ (Ref v) = KCFA (\ (ch,σ) -> 113 | let procs = σ!(ρ!v) 114 | in [ (proc,((Invoke proc) : ch, σ)) | proc <- Set.toList procs ]) 115 | 116 | argEval ρ (Lam l) = KCFA (\ (ch,σ) -> 117 | let proc = Clo(l, ρ) 118 | in [ (Set.singleton proc, (ch, σ)) ]) 119 | argEval ρ (Ref v) = KCFA (\ (ch,σ) -> 120 | let procs = σ!(ρ!v) 121 | in [ (procs, (ch, σ)) ]) 122 | 123 | a $= d = KCFA (\ (ch,σ) -> [((),(ch,σ ⨆ [a ==> d]))]) 124 | 125 | alloc v = KCFA (\ (ch, σ) -> 126 | let t' = timeFrom(ch) 127 | in [(Bind v t',(ch, σ))]) 128 | 129 | tick ((call, ρ, _),CallSeq t) = KCFA (\ (ch, σ) -> 130 | [(CallSeq (take k (call:t)), (ch,σ))]) 131 | 132 | 133 | mnext :: (Analysis m) => (CExp,Env,Time) -> m (CExp,Env,Time) 134 | mnext ctx@(Call f aes, ρ, t) = do 135 | clo@(Clo (vs :=> call', ρ')) <- funEval ρ f 136 | t' <- tick(ctx,t) 137 | as <- mapM alloc vs 138 | ds <- mapM (argEval ρ) aes 139 | let ρ'' = ρ' // [ v ==> a | v <- vs | a <- as ] 140 | sequence [ a $= d | a <- as | d <- ds ] 141 | return $! (call', ρ'', t') 142 | 143 | 144 | 145 | main :: IO () 146 | main = do 147 | return () 148 | 149 | 150 | -------------------------------------------------------------------------------- /obsolete/MonadCFASP.hs: -------------------------------------------------------------------------------- 1 | module Main where 2 | 3 | -- Imports. 4 | import CPS 5 | import Data.Map as Map 6 | import Data.Set as Set 7 | 8 | -- Abbreviations. 9 | type k :-> v = Map.Map k v 10 | type ℙ a = Set.Set a 11 | 12 | (==>) :: a -> b -> (a,b) 13 | (==>) x y = (x,y) 14 | 15 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 16 | (//) f [] = f 17 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 18 | 19 | 20 | -- Partial order theory. 21 | class Lattice a where 22 | bot :: a 23 | top :: a 24 | (⊑) :: a -> a -> Bool 25 | (⊔) :: a -> a -> a 26 | (⊓) :: a -> a -> a 27 | 28 | 29 | instance (Ord s, Eq s) => Lattice (ℙ s) where 30 | bot = Set.empty 31 | top = error "no representation of universal set" 32 | x ⊔ y = x `Set.union` y 33 | x ⊓ y = x `Set.intersection` y 34 | x ⊑ y = x `Set.isSubsetOf` y 35 | 36 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 37 | bot = Map.empty 38 | top = error "no representation of top map" 39 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 40 | f ⊔ g = Map.unionWith (⊔) f g 41 | f ⊓ g = Map.intersectionWith (⊓) f g 42 | 43 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 44 | f ⨆ [] = f 45 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 46 | 47 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 48 | f !! k = Map.findWithDefault bot k f 49 | 50 | 51 | 52 | -- State-space. 53 | data Σ = Eval PΣ Store 54 | deriving (Eq,Ord) 55 | type PΣ = (CExp, Env, Time) 56 | type Env = Var :-> Addr 57 | type Store = Addr :-> D 58 | type D = ℙ Val 59 | data Val = Clo (Lambda, Env) 60 | deriving (Eq,Ord) 61 | data Addr = Bind Var Time 62 | deriving (Eq,Ord) 63 | data Time = CallSeq [CExp] 64 | deriving (Eq,Ord) 65 | 66 | 67 | 68 | class Monad m => Analysis m where 69 | fun :: Env -> AExp -> m Val 70 | arg :: Env -> AExp -> m D 71 | 72 | ($=) :: Addr -> D -> m () 73 | 74 | alloc :: Time -> Var -> m Addr 75 | tick :: Val -> PΣ -> m Time 76 | 77 | -- class Monad m where 78 | -- (>>=) :: m a -> (a -> m b) -> m b 79 | -- (>>) :: m a -> m b -> m b 80 | -- return :: a -> m a 81 | -- fail :: String -> m a 82 | 83 | 84 | 85 | data Exact a = Exact !(Store -> (a,Store)) 86 | data KCFA a = KCFA { kf :: !(Store -> [(a,Store)]) } 87 | 88 | 89 | k = 1 90 | 91 | -- return a >>= k = k a 92 | -- m >>= return = m 93 | -- m >>= (\x -> k x >>= h) = (m >>= k) >>= h 94 | 95 | 96 | 97 | instance Monad KCFA where 98 | (>>=) (KCFA f) g = KCFA (\ σ -> 99 | concatMap (\ (a, σ') -> (kf $ g(a))(σ')) (f(σ))) 100 | return a = KCFA (\ σ -> [(a,σ)]) 101 | 102 | 103 | instance Analysis KCFA where 104 | fun ρ (Lam l) = KCFA (\ σ -> 105 | let proc = Clo(l, ρ) 106 | in [ (proc,σ) ]) 107 | fun ρ (Ref v) = KCFA (\ σ -> 108 | let procs = σ!(ρ!v) 109 | in [ (proc,σ) | proc <- Set.toList procs ]) 110 | 111 | arg ρ (Lam l) = KCFA (\ σ -> 112 | let proc = Clo(l, ρ) 113 | in [ (Set.singleton proc, σ) ]) 114 | arg ρ (Ref v) = KCFA (\ σ -> 115 | let procs = σ!(ρ!v) 116 | in [ (procs, σ) ]) 117 | 118 | a $= d = KCFA (\ σ -> [((),σ ⨆ [a ==> d])]) 119 | 120 | alloc t' v = KCFA (\ σ -> [(Bind v t', σ)]) 121 | 122 | tick clo (call, ρ, CallSeq t) = KCFA (\ σ -> 123 | [(CallSeq (take k (call:t)), σ)]) 124 | 125 | 126 | mnext :: (Analysis m) => PΣ -> m PΣ 127 | mnext ps@(Call f aes, ρ, t) = do 128 | proc@(Clo (vs :=> call', ρ')) <- fun ρ f 129 | t' <- tick proc ps 130 | as <- mapM (alloc t') vs 131 | ds <- mapM (arg ρ) aes 132 | let ρ'' = ρ' // [ v ==> a | v <- vs | a <- as ] 133 | sequence [ a $= d | a <- as | d <- ds ] 134 | return $! (call', ρ'', t') 135 | 136 | 137 | 138 | main :: IO () 139 | main = do 140 | return () 141 | 142 | 143 | -------------------------------------------------------------------------------- /obsolete/MonadCFASPT.hs: -------------------------------------------------------------------------------- 1 | module Main where 2 | 3 | -- Imports. 4 | import CPS 5 | import Data.Map as Map 6 | import Data.Set as Set 7 | 8 | -- Abbreviations. 9 | type k :-> v = Map.Map k v 10 | type ℙ a = Set.Set a 11 | 12 | (==>) :: a -> b -> (a,b) 13 | (==>) x y = (x,y) 14 | 15 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 16 | (//) f [] = f 17 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 18 | 19 | 20 | -- Partial order theory. 21 | class Lattice a where 22 | bot :: a 23 | top :: a 24 | (⊑) :: a -> a -> Bool 25 | (⊔) :: a -> a -> a 26 | (⊓) :: a -> a -> a 27 | 28 | 29 | instance (Ord s, Eq s) => Lattice (ℙ s) where 30 | bot = Set.empty 31 | top = error "no representation of universal set" 32 | x ⊔ y = x `Set.union` y 33 | x ⊓ y = x `Set.intersection` y 34 | x ⊑ y = x `Set.isSubsetOf` y 35 | 36 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 37 | bot = Map.empty 38 | top = error "no representation of top map" 39 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 40 | f ⊔ g = Map.unionWith (⊔) f g 41 | f ⊓ g = Map.intersectionWith (⊓) f g 42 | 43 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 44 | f ⨆ [] = f 45 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 46 | 47 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 48 | f !! k = Map.findWithDefault bot k f 49 | 50 | 51 | 52 | -- State-space. 53 | data Σ = Eval PΣ Store Time 54 | deriving (Eq,Ord) 55 | type PΣ = (CExp, Env) 56 | type Env = Var :-> Addr 57 | type Store = Addr :-> D 58 | type D = ℙ Val 59 | data Val = Clo (Lambda, Env) 60 | deriving (Eq,Ord) 61 | data Addr = Bind Var Time 62 | deriving (Eq,Ord) 63 | data Time = CallSeq [CExp] 64 | deriving (Eq,Ord) 65 | 66 | 67 | 68 | class Monad m => Analysis m where 69 | fun :: Env -> AExp -> m Val 70 | arg :: Env -> AExp -> m D 71 | 72 | ($=) :: Addr -> D -> m () 73 | 74 | alloc :: Var -> m Addr 75 | tick :: Val -> PΣ -> m () 76 | 77 | -- class Monad m where 78 | -- (>>=) :: m a -> (a -> m b) -> m b 79 | -- (>>) :: m a -> m b -> m b 80 | -- return :: a -> m a 81 | -- fail :: String -> m a 82 | 83 | 84 | 85 | data KCFA a = KCFA { kf :: !((Store,Time) -> [(a,Store,Time)]) } 86 | 87 | k = 1 88 | 89 | -- return a >>= k = k a 90 | -- m >>= return = m 91 | -- m >>= (\x -> k x >>= h) = (m >>= k) >>= h 92 | 93 | 94 | 95 | instance Monad KCFA where 96 | (>>=) (KCFA f) g = KCFA (\ (σ,t) -> 97 | let chs = f(σ,t) 98 | in concatMap (\ (a,σ',t') -> (kf $ g(a))(σ',t')) chs) 99 | return a = KCFA (\ (σ,t) -> [(a,σ,t)]) 100 | 101 | 102 | instance Analysis KCFA where 103 | fun ρ (Lam l) = KCFA (\ (σ,t) -> 104 | let proc = Clo(l, ρ) 105 | in [ (proc,σ,t) ]) 106 | fun ρ (Ref v) = KCFA (\ (σ,t) -> 107 | let procs = σ!(ρ!v) 108 | in [ (proc,σ,t) | proc <- Set.toList procs ]) 109 | 110 | arg ρ (Lam l) = KCFA (\ (σ,t) -> 111 | let proc = Clo(l, ρ) 112 | in [ (Set.singleton proc, σ, t) ]) 113 | arg ρ (Ref v) = KCFA (\ (σ,t) -> 114 | let procs = σ!(ρ!v) 115 | in [ (procs, σ, t) ]) 116 | 117 | a $= d = KCFA (\ (σ,t) -> [((),σ ⨆ [a ==> d],t)] ) 118 | 119 | alloc v = KCFA (\ (σ,t) -> [(Bind v t, σ, t)]) 120 | 121 | tick clo (call, ρ) = KCFA (\ (σ,CallSeq t) -> 122 | [((), σ, CallSeq (take k (call:t)))]) 123 | 124 | 125 | mnext :: (Analysis m) => PΣ -> m PΣ 126 | mnext ps@(Call f aes, ρ) = do 127 | proc@(Clo (vs :=> call', ρ')) <- fun ρ f 128 | tick proc ps 129 | as <- mapM alloc vs 130 | ds <- mapM (arg ρ) aes 131 | let ρ'' = ρ' // [ v ==> a | v <- vs | a <- as ] 132 | sequence [ a $= d | a <- as | d <- ds ] 133 | return $! (call', ρ'') 134 | 135 | 136 | 137 | main :: IO () 138 | main = do 139 | return () 140 | 141 | 142 | -------------------------------------------------------------------------------- /obsolete/MonadCFASPTCh.hs: -------------------------------------------------------------------------------- 1 | module Main where 2 | 3 | -- Imports. 4 | import CPS 5 | import Data.Map as Map 6 | import Data.Set as Set 7 | 8 | -- Abbreviations. 9 | type k :-> v = Map.Map k v 10 | type ℙ a = Set.Set a 11 | 12 | (==>) :: a -> b -> (a,b) 13 | (==>) x y = (x,y) 14 | 15 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 16 | (//) f [] = f 17 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 18 | 19 | 20 | -- Partial order theory. 21 | class Lattice a where 22 | bot :: a 23 | top :: a 24 | (⊑) :: a -> a -> Bool 25 | (⊔) :: a -> a -> a 26 | (⊓) :: a -> a -> a 27 | 28 | 29 | instance (Ord s, Eq s) => Lattice (ℙ s) where 30 | bot = Set.empty 31 | top = error "no representation of universal set" 32 | x ⊔ y = x `Set.union` y 33 | x ⊓ y = x `Set.intersection` y 34 | x ⊑ y = x `Set.isSubsetOf` y 35 | 36 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 37 | bot = Map.empty 38 | top = error "no representation of top map" 39 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 40 | f ⊔ g = Map.unionWith (⊔) f g 41 | f ⊓ g = Map.intersectionWith (⊓) f g 42 | 43 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 44 | f ⨆ [] = f 45 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 46 | 47 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 48 | f !! k = Map.findWithDefault bot k f 49 | 50 | 51 | 52 | -- State-space. 53 | data Σ = Eval PΣ Store Time 54 | deriving (Eq,Ord) 55 | type PΣ = (CExp, Env) 56 | type Env = Var :-> Addr 57 | type Store = Addr :-> D 58 | type D = ℙ Val 59 | data Val = Clo (Lambda, Env) 60 | deriving (Eq,Ord) 61 | data Addr = Bind Var Time 62 | deriving (Eq,Ord) 63 | data Time = CallSeq [CExp] 64 | deriving (Eq,Ord) 65 | 66 | 67 | 68 | class Monad m => Analysis m where 69 | fun :: Env -> AExp -> m Val 70 | arg :: Env -> AExp -> m D 71 | 72 | ($=) :: Addr -> D -> m () 73 | 74 | alloc :: Var -> m Addr 75 | tick :: PΣ -> m () 76 | 77 | -- class Monad m where 78 | -- (>>=) :: m a -> (a -> m b) -> m b 79 | -- (>>) :: m a -> m b -> m b 80 | -- return :: a -> m a 81 | -- fail :: String -> m a 82 | 83 | 84 | type ProcCh = Maybe Val 85 | 86 | data KCFA a = KCFA { kf :: !((ProcCh,Store,Time) -> [(a,ProcCh,Store,Time)]) } 87 | 88 | k = 1 89 | 90 | -- return a >>= k = k a 91 | -- m >>= return = m 92 | -- m >>= (\x -> k x >>= h) = (m >>= k) >>= h 93 | 94 | 95 | 96 | instance Monad KCFA where 97 | (>>=) (KCFA f) g = KCFA (\ (ch,σ,t) -> 98 | concatMap (\ (a,ch,σ',t') -> (kf $ g(a))(ch,σ',t')) (f(ch,σ,t))) 99 | return a = KCFA (\ (ch,σ,t) -> [(a,ch,σ,t)]) 100 | 101 | 102 | instance Analysis KCFA where 103 | fun ρ (Lam l) = KCFA (\ (_,σ,t) -> 104 | let proc = Clo(l, ρ) 105 | in [ (proc,Just proc,σ,t) ]) 106 | fun ρ (Ref v) = KCFA (\ (_,σ,t) -> 107 | let procs = σ!(ρ!v) 108 | in [ (proc,Just proc,σ,t) | proc <- Set.toList procs ]) 109 | 110 | arg ρ (Lam l) = KCFA (\ (ch,σ,t) -> 111 | let proc = Clo(l, ρ) 112 | in [ (Set.singleton proc, ch, σ, t) ]) 113 | arg ρ (Ref v) = KCFA (\ (ch,σ,t) -> 114 | let procs = σ!(ρ!v) 115 | in [ (procs, ch, σ, t) ]) 116 | 117 | a $= d = KCFA (\ (ch,σ,t) -> [((),ch,σ ⨆ [a ==> d],t)] ) 118 | 119 | alloc v = KCFA (\ (ch,σ,t) -> [(Bind v t, ch, σ, t)]) 120 | 121 | tick (call, ρ) = KCFA (\ (ch,σ,CallSeq t) -> 122 | [((), ch, σ, CallSeq (take k (call:t)))]) 123 | 124 | 125 | mnext :: (Analysis m) => PΣ -> m PΣ 126 | mnext ps@(Call f aes, ρ) = do 127 | proc@(Clo (vs :=> call', ρ')) <- fun ρ f 128 | tick ps 129 | as <- mapM alloc vs 130 | ds <- mapM (arg ρ) aes 131 | let ρ'' = ρ' // [ v ==> a | v <- vs | a <- as ] 132 | sequence [ a $= d | a <- as | d <- ds ] 133 | return $! (call', ρ'') 134 | 135 | 136 | 137 | main :: IO () 138 | main = do 139 | return () 140 | 141 | 142 | -------------------------------------------------------------------------------- /obsolete/MonadKCFA.hs: -------------------------------------------------------------------------------- 1 | module Main where 2 | 3 | -- Imports. 4 | import CPS 5 | import Data.Map as Map 6 | import Data.Set as Set 7 | 8 | -- Abbreviations. 9 | type k :-> v = Map.Map k v 10 | type ℙ a = Set.Set a 11 | 12 | (==>) :: a -> b -> (a,b) 13 | (==>) x y = (x,y) 14 | 15 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 16 | (//) f [] = f 17 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 18 | 19 | 20 | -- Partial order theory. 21 | class Lattice a where 22 | bot :: a 23 | top :: a 24 | (⊑) :: a -> a -> Bool 25 | (⊔) :: a -> a -> a 26 | (⊓) :: a -> a -> a 27 | 28 | 29 | instance (Ord s, Eq s) => Lattice (ℙ s) where 30 | bot = Set.empty 31 | top = error "no representation of universal set" 32 | x ⊔ y = x `Set.union` y 33 | x ⊓ y = x `Set.intersection` y 34 | x ⊑ y = x `Set.isSubsetOf` y 35 | 36 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 37 | bot = Map.empty 38 | top = error "no representation of top map" 39 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 40 | f ⊔ g = Map.unionWith (⊔) f g 41 | f ⊓ g = Map.intersectionWith (⊓) f g 42 | 43 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 44 | f ⨆ [] = f 45 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 46 | 47 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 48 | f !! k = Map.findWithDefault bot k f 49 | 50 | 51 | 52 | -- State-space. 53 | data Σ = Eval Ctx Store 54 | deriving (Eq,Ord) 55 | type Ctx = (CExp, Env, Time) 56 | type Env = Var :-> Addr 57 | type Store = Addr :-> D 58 | type D = ℙ Val 59 | data Val = Clo (Lambda, Env) 60 | deriving (Eq,Ord) 61 | data Addr = Bind Var Time 62 | deriving (Eq,Ord) 63 | data Time = CallSeq [CExp] 64 | deriving (Eq,Ord) 65 | 66 | 67 | 68 | class Monad m => Analysis m where 69 | fun :: Env -> AExp -> m Val 70 | arg :: Env -> AExp -> m D 71 | 72 | ($=) :: Addr -> D -> m () 73 | 74 | alloc :: Time -> Var -> m Addr 75 | tick :: (Ctx,Time) -> m Time 76 | 77 | -- class Monad m where 78 | -- (>>=) :: m a -> (a -> m b) -> m b 79 | -- (>>) :: m a -> m b -> m b 80 | -- return :: a -> m a 81 | -- fail :: String -> m a 82 | 83 | 84 | 85 | data Exact a = Exact !(Store -> (a,Store)) 86 | data KCFA a = KCFA { kf :: !(Store -> [(a,Store)]) } 87 | 88 | 89 | k = 1 90 | 91 | -- return a >>= k = k a 92 | -- m >>= return = m 93 | -- m >>= (\x -> k x >>= h) = (m >>= k) >>= h 94 | 95 | 96 | 97 | instance Monad KCFA where 98 | (>>=) (KCFA f) g = KCFA (\ σ -> 99 | let chs = f(σ) 100 | in concatMap (\ (a, σ') -> (kf $ g(a))(σ')) chs) 101 | return a = KCFA (\ σ -> [(a,σ)]) 102 | 103 | 104 | instance Analysis KCFA where 105 | fun ρ (Lam l) = KCFA (\ σ -> 106 | let proc = Clo(l, ρ) 107 | in [ (proc,σ) ]) 108 | fun ρ (Ref v) = KCFA (\ σ -> 109 | let procs = σ!(ρ!v) 110 | in [ (proc,σ) | proc <- Set.toList procs ]) 111 | 112 | arg ρ (Lam l) = KCFA (\ σ -> 113 | let proc = Clo(l, ρ) 114 | in [ (Set.singleton proc, σ) ]) 115 | arg ρ (Ref v) = KCFA (\ σ -> 116 | let procs = σ!(ρ!v) 117 | in [ (procs, σ) ]) 118 | 119 | a $= d = KCFA (\ σ -> [((),σ ⨆ [a ==> d])]) 120 | 121 | alloc t' v = KCFA (\ σ -> [(Bind v t', σ)]) 122 | 123 | tick ((call, ρ, _),CallSeq t) = KCFA (\ σ -> 124 | [(CallSeq (take k (call:t)), σ)]) 125 | 126 | 127 | mnext :: (Analysis m) => (CExp,Env,Time) -> m (CExp,Env,Time) 128 | mnext ctx@(Call f aes, ρ, t) = do 129 | clo@(Clo (vs :=> call', ρ')) <- fun ρ f 130 | t' <- tick(ctx,t) 131 | as <- mapM (alloc t') vs 132 | ds <- mapM (arg ρ) aes 133 | let ρ'' = ρ' // [ v ==> a | v <- vs | a <- as ] 134 | sequence [ a $= d | a <- as | d <- ds ] 135 | return $! (call', ρ'', t') 136 | 137 | 138 | 139 | main :: IO () 140 | main = do 141 | return () 142 | 143 | 144 | -------------------------------------------------------------------------------- /obsolete/XCFA.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE MultiParamTypeClasses #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE FunctionalDependencies #-} 4 | {-# LANGUAGE TypeOperators #-} 5 | {-# LANGUAGE ParallelListComp #-} 6 | {-# LANGUAGE TypeSynonymInstances #-} 7 | 8 | -- TODO: get rid of this 9 | {-# LANGUAGE UndecidableInstances #-} 10 | 11 | module Main where 12 | 13 | -- Imports. 14 | import CPS 15 | import Data.Map as Map 16 | import Data.Set as Set 17 | import Data.List as List 18 | 19 | -- debug 20 | import Debug.Trace 21 | 22 | -- Abbreviations. 23 | type k :-> v = Map.Map k v 24 | type ℙ a = Set.Set a 25 | 26 | (==>) :: a -> b -> (a,b) 27 | (==>) x y = (x,y) 28 | 29 | (//) :: Ord k => (k :-> v) -> [(k,v)] -> (k :-> v) 30 | (//) f [] = f 31 | (//) f ((x,y):tl) = Map.insert x y (f // tl) 32 | 33 | 34 | -- Partial order theory. 35 | class Lattice a where 36 | bot :: a 37 | top :: a 38 | (⊑) :: a -> a -> Bool 39 | (⊔) :: a -> a -> a 40 | (⊓) :: a -> a -> a 41 | 42 | -- unit is a trivial lattice 43 | instance Lattice () where 44 | bot = () 45 | top = () 46 | x ⊔ y = top 47 | x ⊓ y = bot 48 | x ⊑ y = True 49 | 50 | instance (Lattice a, Lattice b) => Lattice (a, b) where 51 | bot = (bot, bot) 52 | top = (top, top) 53 | (a1, b1) ⊔ (a2, b2) = (a1 ⊔ a2, b1 ⊔ b2) 54 | (a1, b1) ⊓ (a2, b2) = (a1 ⊓ a2, b1 ⊓ b2) 55 | (a1, b1) ⊑ (a2, b2) = (a1 ⊑ a2) && (b1 ⊑ b2) 56 | 57 | instance (Ord s, Eq s) => Lattice (ℙ s) where 58 | bot = Set.empty 59 | top = error "no representation of universal set" 60 | x ⊔ y = x `Set.union` y 61 | x ⊓ y = x `Set.intersection` y 62 | x ⊑ y = x `Set.isSubsetOf` y 63 | 64 | instance (Ord k, Lattice v) => Lattice (k :-> v) where 65 | bot = Map.empty 66 | top = error "no representation of top map" 67 | f ⊑ g = Map.isSubmapOfBy (⊑) f g 68 | f ⊔ g = Map.unionWith (⊔) f g 69 | f ⊓ g = Map.intersectionWith (⊓) f g 70 | 71 | (⨆) :: (Ord k, Lattice v) => (k :-> v) -> [(k,v)] -> (k :-> v) 72 | f ⨆ [] = f 73 | f ⨆ ((k,v):tl) = Map.insertWith (⊔) k v (f ⨆ tl) 74 | 75 | (!!) :: (Ord k, Lattice v) => (k :-> v) -> k -> v 76 | f !! k = Map.findWithDefault bot k f 77 | 78 | -- State-space. 79 | type PΣ a = (CExp, Env a) 80 | type Env a = Var :-> a 81 | type D a = ℙ (Val a) 82 | data Val a = Clo (Lambda, Env a) 83 | deriving (Eq, Ord, Show) 84 | 85 | ρ0 = Map.empty 86 | 87 | -- Generic store. 88 | type Store a = a :-> (D a) 89 | 90 | -- Abstract analysis interface. 91 | -- Type parameter "g" is for guts and is passed along 92 | -- Address is a part of the Semantic interface, but not of the analysis!!! 93 | -- So `a' parameter should not be in the analysis, only `g' for guts 94 | 95 | -- Do we really need functional dependencies between variables? 96 | -- 1. Indeed, guts define the underlying monad 97 | -- 2. Guts define the shared component 98 | -- !! No more dependencies is needed 99 | 100 | class Monad (m s g) => Analysis m a s g | g -> m, m -> s, g -> a where 101 | fun :: (Env a) -> AExp -> m s g (Val a) 102 | arg :: (Env a) -> AExp -> m s g (D a) 103 | 104 | ($=) :: a -> (D a) -> m s g () 105 | 106 | updateEnv :: Env a -> [(Var, a)] -> m s g (Env a) 107 | -- default implementation 108 | updateEnv ρ bs = return $ ρ // bs 109 | 110 | alloc :: Var -> m s g a 111 | tick :: (PΣ a) -> m s g () 112 | 113 | stepAnalysis :: s -> g -> PΣ a -> (s, [(PΣ a, g)]) 114 | inject :: CExp -> (PΣ a, s, g) 115 | 116 | -- Abstract garbage collector 117 | class Monad m => GarbageCollector m a where 118 | gc :: (PΣ a) -> m (PΣ a) 119 | -- default implementation 120 | gc = return 121 | 122 | -- Generic transition 123 | mnext :: Analysis m a s g => (PΣ a) -> m s g (PΣ a) 124 | mnext ps@(Call f aes, ρ) = do 125 | proc@(Clo (vs :=> call', ρ')) <- fun ρ f 126 | tick ps 127 | as <- mapM alloc vs 128 | ds <- mapM (arg ρ) aes 129 | ρ'' <- updateEnv ρ' [ v ==> a | v <- vs | a <- as ] 130 | sequence [ a $= d | a <- as | d <- ds ] 131 | return $! (call', ρ'') 132 | mnext ps@(Exit, ρ) = return $! ps 133 | 134 | ---------------------------------------------------------------------- 135 | -- Example: Concrete Semantics 136 | ---------------------------------------------------------------------- 137 | 138 | data Concrete s g b = Concrete { cf :: g -> (b, g) } 139 | 140 | instance Monad (Concrete s g) where 141 | (>>=) (Concrete f) g = Concrete (\guts -> 142 | let (b, guts') = f guts 143 | in (cf $ g(b)) guts') 144 | return b = Concrete (\guts -> (b, guts)) 145 | 146 | data CAddr = CBind Var Int 147 | deriving (Eq, Ord, Show) 148 | 149 | instance Analysis (Concrete) 150 | CAddr 151 | () 152 | (Store CAddr, Int) 153 | where 154 | fun ρ (Lam l) = Concrete (\ (σ,t) -> 155 | let proc = Clo(l, ρ) 156 | in (proc,(σ,t))) 157 | fun ρ (Ref v) = Concrete (\ (σ,t) -> 158 | let [proc] = Set.toList $ σ!(ρ!v) 159 | in (proc,(σ,t))) 160 | 161 | arg ρ (Lam l) = Concrete (\ (σ,t) -> 162 | let proc = Clo(l, ρ) 163 | in (Set.singleton proc, (σ, t))) 164 | arg ρ (Ref v) = Concrete (\ (σ,t) -> 165 | let procs = σ!(ρ!v) 166 | in (procs, (σ, t))) 167 | 168 | a $= d = Concrete (\ (σ,t) -> ((), (σ ⨆ [a ==> d],t)) ) 169 | 170 | alloc v = Concrete (\ (σ,t) -> (CBind v t, (σ, t))) 171 | 172 | tick (call, ρ) = Concrete (\ (σ,n) -> ((), (σ, n+1))) 173 | 174 | stepAnalysis _ config state = ((), [cf (mnext state) config]) 175 | 176 | inject call = ((call, Map.empty), (), (bot, 0)) 177 | 178 | -- Add Garbage Collection 179 | instance GarbageCollector (Concrete () (Store CAddr, Int)) CAddr 180 | 181 | ---------------------------------------------------------------------- 182 | -- Addresses, Stores and Choices 183 | ---------------------------------------------------------------------- 184 | 185 | type ProcCh a = Maybe (Val a) -- Nondeterministic choice. 186 | 187 | -- FunctionalDependencies again: 188 | -- time defines addresses 189 | class (Ord a, Eq a) => Addressable a c | c -> a where 190 | τ0 :: c 191 | valloc :: Var -> c -> a 192 | advance :: Val a -> PΣ a -> c -> c 193 | 194 | -- and again: 195 | -- Store uniquely defines the type of its addresses 196 | class (Eq a, Lattice s) => StoreLike a s | s -> a where 197 | σ0 :: s 198 | bind :: s -> a -> (D a)-> s 199 | replace :: s -> a -> (D a) -> s 200 | fetch :: s -> a -> (D a) 201 | filterStore :: s -> (a -> Bool) -> s 202 | 203 | ---------------------------------------------------------------------- 204 | -- Generic analysis with no shared component 205 | ---------------------------------------------------------------------- 206 | 207 | -- GenericAnalysis :: * -> * -> * 208 | -- parametrized by guts and passed result 209 | data GenericAnalysis s g b = GCFA { 210 | gf :: g -> [(b, g)] 211 | } 212 | 213 | -- Curry GenericAnalysis for the fixed guts 214 | instance Monad (GenericAnalysis s g) where 215 | (>>=) (GCFA f) g = GCFA (\ guts -> 216 | concatMap (\ (a, guts') -> (gf $ g(a)) guts') (f guts)) 217 | return a = GCFA (\ guts -> [(a,guts)]) 218 | 219 | instance (Addressable a t, StoreLike a s) 220 | => Analysis (GenericAnalysis) 221 | a -- address type 222 | () -- no shared result 223 | (ProcCh a, s, t) -- Generic Analysis' guts 224 | where 225 | fun ρ (Lam l) = GCFA (\ (_,σ,t) -> 226 | let proc = Clo(l, ρ) 227 | in [ (proc, (Just proc,σ,t)) ]) 228 | fun ρ (Ref v) = GCFA (\ (_,σ,t) -> 229 | let procs = fetch σ (ρ!v) 230 | in [ (proc, (Just proc,σ,t)) | proc <- Set.toList procs ]) 231 | 232 | arg ρ (Lam l) = GCFA (\ (ch,σ,t) -> 233 | let proc = Clo(l, ρ) 234 | in [ (Set.singleton proc, (ch, σ, t)) ]) 235 | arg ρ (Ref v) = GCFA (\ (ch,σ,t) -> 236 | let procs = fetch σ (ρ!v) 237 | in [ (procs, (ch, σ, t)) ]) 238 | 239 | a $= d = GCFA (\ (ch,σ,t) -> [((),(ch,bind σ a d,t))] ) 240 | 241 | alloc v = GCFA (\ (ch,σ,t) -> [(valloc v t, (ch, σ, t))]) 242 | 243 | tick ps = GCFA (\ (Just proc, σ, t) -> 244 | [((), (Just proc, σ, advance proc ps t))]) 245 | 246 | stepAnalysis _ config state = ((), gf (mnext state) config) 247 | 248 | inject call = ((call, Map.empty), (), (Nothing, σ0, τ0)) 249 | 250 | -- Garbage Collection 251 | instance GarbageCollector (GenericAnalysis () (ProcCh a, s, t)) a 252 | 253 | ---------------------------------------------------------------------- 254 | -- Single store-threading analysis. 255 | ---------------------------------------------------------------------- 256 | 257 | data SingleStoreAnalysis a s g b = SSFA { runWithStore :: s -> g -> (s, [(b, g)]) } 258 | 259 | -- TODO redefine store-like logic 260 | instance Lattice s => Monad (SingleStoreAnalysis a s g) where 261 | (>>=) (SSFA f) g = SSFA (\st -> \guts -> 262 | let (st', pairs) = f st guts -- make an f-step 263 | -- get new results via g :: [(st, [(b, g)])] 264 | newResults = List.map (\(a, guts') -> (runWithStore $ g(a)) st' guts') pairs 265 | -- merge stores and concatenate the results :: (st, [(b, g)]) 266 | -- requires a lattice structure of a store 267 | in foldl (\(s, bg) -> \(s', bg') -> (s ⊔ s', bg ++ bg')) 268 | (st', []) newResults) 269 | 270 | return a = SSFA (\s -> \guts -> (s, [(a, guts)])) 271 | 272 | instance (Addressable a t, StoreLike a s) 273 | => Analysis (SingleStoreAnalysis a) 274 | a -- address type 275 | s -- shared store 276 | (ProcCh a, t) -- SingleStore Analysis' guts 277 | where 278 | fun ρ (Lam l) = SSFA (\σ -> \(_,t) -> 279 | let proc = Clo(l, ρ) 280 | in (σ, [(proc, (Just proc,t)) ])) 281 | fun ρ (Ref v) = SSFA (\σ -> \(_,t) -> 282 | let procs = fetch σ (ρ!v) 283 | in (σ, [ (proc, (Just proc,t)) | proc <- Set.toList procs ])) 284 | 285 | arg ρ (Lam l) = SSFA (\σ -> \(ch, t) -> 286 | let proc = Clo(l, ρ) 287 | in (σ, [ (Set.singleton proc, (ch, t)) ])) 288 | arg ρ (Ref v) = SSFA (\σ -> \(ch, t) -> 289 | let procs = fetch σ (ρ!v) 290 | in (σ, [ (procs, (ch,t)) ])) 291 | 292 | a $= d = SSFA (\σ -> \(ch, t) -> 293 | let σ' = bind σ a d 294 | in (σ', [((), (ch, t))] )) 295 | 296 | alloc v = SSFA (\σ -> \(ch, t) -> (σ, [(valloc v t, (ch, t))])) 297 | 298 | tick ps = SSFA (\σ -> \ (Just proc, t) -> 299 | (σ, ([((), (Just proc, advance proc ps t))]))) 300 | 301 | stepAnalysis store config state = runWithStore (mnext state >>= gc) store config 302 | 303 | inject call = ((call, Map.empty), σ0, (Nothing, τ0)) 304 | 305 | 306 | -- Enhance GC for single-store analysis 307 | 308 | instance (Ord a, StoreLike a s) 309 | => GarbageCollector ((SingleStoreAnalysis a) s g) a where 310 | gc ps = SSFA (\σ -> \g -> 311 | let rs = Set.map (\(v, a) -> a) (reachable ps σ) 312 | σ' = filterStore σ (\a -> not (Set.member a rs)) 313 | in (σ', [(ps, g)])) 314 | 315 | ---------------------------------------------------------------------- 316 | -- Abstract Garbage Collection 317 | ---------------------------------------------------------------------- 318 | 319 | -- Free Variables 320 | free' :: CExp -> Set Var -> Set Var 321 | free' Exit bound = Set.empty 322 | free' (Call f as) bound = foldl (\res -> \a -> res ⊔ (free a bound)) 323 | (free f bound) as 324 | 325 | free :: AExp -> Set Var -> Set Var 326 | free (Ref v) bound = if (Set.member v bound) 327 | then Set.empty 328 | else Set.singleton v 329 | free (Lam (vs :=> ce)) bound = free' ce (bound ⊔ (Set.fromList vs)) 330 | 331 | -- `touched' (for expressions and environments) 332 | touched' :: (Ord a) => AExp -> Env a -> Set (Var, a) 333 | touched' f ρ = Set.fromList [(v, ρ!v) | v <- Set.toList(free f Set.empty)] 334 | 335 | -- `touched' for states 336 | touched :: (Ord a) => (PΣ a) -> Set (Var, a) 337 | touched (Call f as, ρ) = (touched' f ρ) ⊔ 338 | Set.fromList [bs | a <- as, 339 | bs <- Set.toList (touched' a ρ)] 340 | touched (Exit, _) = Set.empty 341 | 342 | -- `adjacency' 343 | adjacent :: (Ord a, StoreLike a s) => (Var, a) -> s -> Set (Var, a) 344 | adjacent (v, addr) σ = Set.fromList [b | Clo (f, ρ) <- Set.toList(fetch σ addr), 345 | b <- Set.toList (touched' (Lam f) ρ)] 346 | 347 | -- `reachability' 348 | reachable :: (Ord a, StoreLike a s) => (PΣ a) -> s -> Set (Var, a) 349 | reachable state σ = 350 | let collect bs = 351 | -- fixpoint iteration 352 | let newBindings = [b' | b <- Set.toList bs, 353 | b' <- Set.toList (adjacent b σ)] 354 | newResult = bs ⊔ (Set.fromList newBindings) 355 | in if newResult == bs 356 | then bs 357 | else collect newResult 358 | -- reflexive-transitive closure 359 | in collect (touched state) 360 | 361 | 362 | ---------------------------------------------------------------------- 363 | -- Abstract Counting 364 | ---------------------------------------------------------------------- 365 | 366 | -- Abstract natural number 367 | data AbsNat = AZero | AOne | AMany 368 | deriving (Ord, Eq, Show) 369 | 370 | instance Lattice AbsNat where 371 | bot = AZero 372 | top = AMany 373 | n ⊑ m = (n == bot) || (m == top) || (n == m) 374 | n ⊔ m = if (n ⊑ m) then m else n 375 | n ⊓ m = if (n ⊑ m) then n else m 376 | 377 | -- Abstract addition 378 | (⊕) :: AbsNat -> AbsNat -> AbsNat 379 | AZero ⊕ n = n 380 | n ⊕ AZero = n 381 | n ⊕ m = AMany 382 | 383 | class StoreLike a s => ACounter a s where 384 | count :: s -> a -> AbsNat 385 | 386 | -- Counting store 387 | type StoreWithCount a = a :-> ((D a), AbsNat) 388 | 389 | instance (Ord a) => ACounter a (StoreWithCount a) where 390 | -- fetching with default bottom 391 | count σ a = snd $ σ Main.!! a 392 | 393 | -- counter is nullified when filtered 394 | -- and incremented when `bind' is called 395 | instance (Ord a) => StoreLike a (StoreWithCount a) where 396 | σ0 = Map.empty 397 | bind σ a d = σ `update_add` [a ==> (d, AOne)] 398 | fetch σ a = fst $ σ Main.!! a 399 | replace σ a d = σ ⨆ [a ==> (d, AZero)] 400 | filterStore σ p = Map.filterWithKey (\k -> \v -> p k) σ 401 | 402 | update_add :: (Ord k, Lattice v) => (k :-> (v, AbsNat)) -> [(k, (v, AbsNat))] -> (k :-> (v, AbsNat)) 403 | update_add f [] = f 404 | update_add f ((k,v):tl) = Map.insertWith (\(x1, y1) -> \(x2, y2) -> (x1 ⊔ x2, y1 ⊕ y2)) k v (update_add f tl) 405 | 406 | ---------------------------------------------------------------------- 407 | -- Anodization 408 | ---------------------------------------------------------------------- 409 | 410 | class StoreLike a s => AlkaliLike a s where 411 | addUniqueAddr :: a -> s 412 | deAnodizeStore :: s -> s 413 | deAnodizeEnv :: s -> Env a -> Env a 414 | deAnodizeD :: s -> D a -> D a 415 | reset :: s -> s 416 | 417 | -- a usefule instance 418 | instance (Ord a, StoreLike a s) => StoreLike a (s, ℙ a) where 419 | σ0 = (σ0, Set.empty) 420 | bind σ a d = (bind (fst σ) a d, snd σ) 421 | fetch σ a = fetch (fst σ) a 422 | replace σ a d = (replace (fst σ) a d, snd σ) 423 | filterStore σ p = (filterStore (fst σ) p, snd σ) 424 | 425 | -- Shape analysis 426 | -- instance (Addressable a t, StoreLike a s, AlkaliLike a s) 427 | -- => Analysis (SingleStoreAnalysis a) 428 | -- a 429 | -- s 430 | -- (ProcCh a, t) 431 | -- where 432 | -- fun ρ (Lam l) = SSFA (\σ -> \(_,t) -> 433 | -- let proc = Clo(l, ρ) 434 | -- in (σ, [(proc, (Just proc,t)) ])) 435 | -- fun ρ (Ref v) = SSFA (\σ -> \(_,t) -> 436 | -- let procs = fetch σ (ρ!v) 437 | -- in (σ, [ (proc, (Just proc,t)) | proc <- Set.toList procs ])) 438 | 439 | -- arg ρ (Lam l) = SSFA (\σ -> \( ch, t) -> 440 | -- let proc = Clo(l, ρ) 441 | -- in (σ, [ (Set.singleton proc, (ch, t)) ])) 442 | -- arg ρ (Ref v) = SSFA (\σ -> \( ch, t) -> 443 | -- let procs = fetch σ (ρ!v) 444 | -- in (σ, [ (procs, (ch,t)) ])) 445 | 446 | -- a $= d = SSFA (\σ -> \(ch, t) -> 447 | -- let σ' = deAnodizeStore σ 448 | -- σ'' = bind σ' a (deAnodizeD σ d) 449 | -- in (σ'', [((), (ch, t))] )) 450 | 451 | -- alloc v = SSFA (\σ -> \(ch, t) -> 452 | -- let addr = valloc v t 453 | -- σ' = addUniqueAddr addr 454 | -- in (σ', [(addr, (ch, t))])) 455 | 456 | -- updateEnv ρ bs = SSFA (\σ -> \( ch, t) -> 457 | -- let ρ' = deAnodizeEnv σ ρ 458 | -- in (σ, [ (ρ' // bs, (ch,t)) ])) 459 | 460 | -- tick ps = SSFA (\σ -> \ (Just proc, t) -> 461 | -- (σ, ([((), (Just proc, advance proc ps t))]))) 462 | 463 | -- stepAnalysis store config state = runWithStore (mnext state >>= gc) (reset store) config 464 | 465 | -- inject call = ((call, Map.empty), σ0, (Nothing, τ0)) 466 | 467 | 468 | -- Particular implementation of anodization 469 | ---------------------------------------------------------------------- 470 | 471 | 472 | -- instance (Ord a, StoreLike a s) => AlkaliLike a (s, ℙ a) where 473 | 474 | 475 | ---------------------------------------------------------------------- 476 | -- KCFA addresses and stores 477 | ---------------------------------------------------------------------- 478 | 479 | k = 1 480 | 481 | data KTime = KCalls [CExp] 482 | deriving (Eq, Ord, Show) 483 | 484 | data KAddr = KBind Var KTime 485 | deriving (Eq, Ord, Show) 486 | 487 | instance Addressable KAddr KTime where 488 | τ0 = KCalls [] 489 | valloc v t = KBind v t 490 | advance proc (call, ρ) (KCalls calls) = 491 | KCalls $ take k (call : calls) 492 | 493 | -- Simple store 494 | instance StoreLike KAddr (Store KAddr) where 495 | σ0 = Map.empty 496 | 497 | bind σ a d = σ ⨆ [a ==> d] 498 | fetch σ a = σ Main.!! a 499 | replace σ a d = σ ⨆ [a ==> d] 500 | filterStore σ p = Map.filterWithKey (\k -> \v -> p k) σ 501 | 502 | 503 | ---------------------------------------------------------------------- 504 | -- Running the analysis 505 | ---------------------------------------------------------------------- 506 | 507 | -- Abstract state-space exploration algorithm 508 | -- TODO: remove step counting and trace output 509 | loop :: (Analysis m a s g, Ord a, Ord g, Show a, Show g, Lattice s) => 510 | [(PΣ a, g)] -> (s, Set (PΣ a, g)) -> Int -> (s, Set (PΣ a, g)) 511 | 512 | loop worklist v@(shared, oldStates) step = 513 | -- trace output 514 | trace ("Worklist [step " ++ show step ++ "]:\n" ++ show worklist ++ "\n") $ 515 | let newStoreStates = List.map (\(state, config) -> 516 | stepAnalysis shared config state) 517 | worklist 518 | -- compute a new shared component and new states 519 | (shared', states') = foldl (\(s, bg) -> \(s', bg') -> (s ⊔ s', bg ++ bg')) 520 | (shared, []) newStoreStates 521 | newWorkList = List.filter (\elem -> not (Set.member elem oldStates)) states' 522 | in if List.null newWorkList 523 | then v 524 | else let newVisited = (shared', oldStates ⊔ (Set.fromList newWorkList)) 525 | in loop newWorkList newVisited (step + 1) 526 | 527 | -- compute an approximation 528 | explore :: (Analysis m a s g, Ord a, Ord g, Show a, Show g, Lattice s) => 529 | CExp -> (s, Set (PΣ a, g)) 530 | 531 | explore program = 532 | let (state0, σ0, g0) = inject program 533 | in loop [(state0, g0)] (σ0, Set.empty) 0 534 | 535 | ---------------------------------------------------------------------- 536 | -- example program 537 | ---------------------------------------------------------------------- 538 | 539 | -- ((λ (f g) (f g)) (λ (x) x) (λ (y) Exit)) 540 | idx = Lam (["x"] :=> Call (Ref "x") []) 541 | idy = Lam (["y"] :=> Exit) 542 | comb = Lam (["f", "g"] :=> Call (Ref "f") [Ref "g"]) 543 | ex = Call comb [idx, idy] 544 | 545 | ucombx = Lam (["x"] :=> Call (Ref "x") [Ref "x"]) 546 | ucomby = Lam (["y"] :=> Call (Ref "y") [Ref "y"]) 547 | omega = Call ucombx [ucomby] 548 | 549 | -- a particular analysis is chosen by the type specification 550 | 551 | -- concrete interpreter 552 | ---------------------------------------------------------------------- 553 | type ConcreteGuts = (Store CAddr, Int) 554 | 555 | concreteResult :: CExp -> ((), Set (PΣ CAddr, ConcreteGuts)) 556 | concreteResult = explore 557 | 558 | -- abstract interpreter with a per-state store 559 | ---------------------------------------------------------------------- 560 | type AbstractGuts = (ProcCh KAddr, Store KAddr, KTime) 561 | 562 | abstractResult :: CExp -> ((), Set (PΣ KAddr, AbstractGuts)) 563 | abstractResult = explore 564 | 565 | -- abstract interpreter with a per-state store and counting 566 | ---------------------------------------------------------------------- 567 | type AbstractGutsWithCounting = (ProcCh KAddr, StoreWithCount KAddr, KTime) 568 | 569 | abstractResultC :: CExp -> ((), Set (PΣ KAddr, AbstractGutsWithCounting)) 570 | abstractResultC = explore 571 | 572 | type AbstractGutsSS = (ProcCh KAddr, KTime) 573 | 574 | ---------------------------------------------------------------------- 575 | -- Functional dependencies are in the conflict for same guts, 576 | -- so only one analysis should be uncommented 577 | ---------------------------------------------------------------------- 578 | 579 | -- abstract interpreter with a single-threaded store 580 | ---------------------------------------------------------------------- 581 | -- abstractResultSS :: CExp -> (Store KAddr, Set (PΣ KAddr, AbstractGutsSS)) 582 | -- abstractResultSS = explore 583 | 584 | -- abstract interpreter with a single-threaded store and counting 585 | ---------------------------------------------------------------------- 586 | abstractResultSSC :: CExp -> (StoreWithCount KAddr, Set (PΣ KAddr, AbstractGutsSS)) 587 | abstractResultSSC = explore 588 | 589 | 590 | -- abstract interpreter with a single-threaded store and shape analysis 591 | ---------------------------------------------------------------------- 592 | -- abstractResultSh :: CExp -> ((Store KAddr, ℙ KAddr), Set (PΣ KAddr, AbstractGutsSS)) 593 | -- abstractResultSh = explore 594 | --------------------------------------------------------------------------------