| author | wenzelm | 
| Tue, 20 Feb 2018 22:04:04 +0100 | |
| changeset 67681 | b5058ba95e32 | 
| parent 55676 | fb46f1c379b5 | 
| permissions | -rw-r--r-- | 
| 43313 | 1 | {-
 | 
| 2 | A narrowing-based Evaluator for Formulas in Prefix Normal Form based on the compilation technique of LazySmallCheck | |
| 3 | -} | |
| 4 | module Narrowing_Engine where | |
| 5 | ||
| 46408 
2520cd337056
using fully qualified module names in Haskell source, which seems to be required by GHC 7.0.4 (also cf. 0fd9ab902b5a)
 bulwahn parents: 
45760diff
changeset | 6 | import Control.Monad | 
| 43313 | 7 | import Control.Exception | 
| 46408 
2520cd337056
using fully qualified module names in Haskell source, which seems to be required by GHC 7.0.4 (also cf. 0fd9ab902b5a)
 bulwahn parents: 
45760diff
changeset | 8 | import System.IO | 
| 43313 | 9 | import System.Exit | 
| 46408 
2520cd337056
using fully qualified module names in Haskell source, which seems to be required by GHC 7.0.4 (also cf. 0fd9ab902b5a)
 bulwahn parents: 
45760diff
changeset | 10 | import Data.Maybe | 
| 
2520cd337056
using fully qualified module names in Haskell source, which seems to be required by GHC 7.0.4 (also cf. 0fd9ab902b5a)
 bulwahn parents: 
45760diff
changeset | 11 | import Data.List (partition, findIndex) | 
| 45081 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 12 | import qualified Generated_Code | 
| 55676 | 13 | import qualified Typerep | 
| 43313 | 14 | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 15 | type Pos = [Int] | 
| 43313 | 16 | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 17 | -- Refinement Tree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 18 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 19 | data Quantifier = Existential | Universal | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 20 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 21 | data Truth = Eval Bool | Unevaluated | Unknown deriving Eq | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 22 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 23 | conj :: Truth -> Truth -> Truth | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 24 | conj (Eval True) b = b | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 25 | conj (Eval False) _ = Eval False | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 26 | conj b (Eval True) = b | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 27 | conj _ (Eval False) = Eval False | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 28 | conj Unevaluated _ = Unevaluated | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 29 | conj _ Unevaluated = Unevaluated | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 30 | conj Unknown Unknown = Unknown | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 31 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 32 | disj :: Truth -> Truth -> Truth | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 33 | disj (Eval True) _ = Eval True | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 34 | disj (Eval False) b = b | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 35 | disj _ (Eval True) = Eval True | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 36 | disj b (Eval False) = b | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 37 | disj Unknown _ = Unknown | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 38 | disj _ Unknown = Unknown | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 39 | disj Unevaluated Unevaluated = Unevaluated | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 40 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 41 | ball ts = foldl (\s t -> conj s (value_of t)) (Eval True) ts | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 42 | bexists ts = foldl (\s t -> disj s (value_of t)) (Eval False) ts | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 43 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 44 | data Tree = Leaf Truth | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 45 | | Variable Quantifier Truth Pos Generated_Code.Narrowing_type Tree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 46 | | Constructor Quantifier Truth Pos [Tree] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 47 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 48 | value_of :: Tree -> Truth | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 49 | value_of (Leaf r) = r | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 50 | value_of (Variable _ r _ _ _) = r | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 51 | value_of (Constructor _ r _ _) = r | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 52 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 53 | data Edge = V Pos Generated_Code.Narrowing_type | C Pos Int | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 54 | type Path = [Edge] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 55 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 56 | position_of :: Edge -> Pos | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 57 | position_of (V pos _) = pos | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 58 | position_of (C pos _) = pos | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 59 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 60 | -- Operation find: finds first relevant unevaluated node and returns its path | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 61 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 62 | find :: Tree -> Path | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 63 | find (Leaf Unevaluated) = [] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 64 | find (Variable _ _ pos ty t) = V pos ty : (find t) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 65 | find (Constructor _ _ pos ts) = C pos i : find (ts !! i) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 66 | where | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 67 | Just i = findIndex (\t -> value_of t == Unevaluated) ts | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 68 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 69 | -- Operation update: updates the leaf and the cached truth values results along the path | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 70 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 71 | update :: Path -> Truth -> Tree -> Tree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 72 | update [] v (Leaf _) = Leaf v | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 73 | update (V _ _ : es) v (Variable q r p ty t) = Variable q (value_of t') p ty t' | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 74 | where | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 75 | t' = update es v t | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 76 | update (C _ i : es) v (Constructor q r pos ts) = Constructor q r' pos ts' | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 77 | where | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 78 | (xs, y : ys) = splitAt i ts | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 79 | y' = update es v y | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 80 | ts' = xs ++ (y' : ys) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 81 | r' = valueOf ts' | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 82 |     valueOf = case q of { Universal -> ball; Existential -> bexists}
 | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 83 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 84 | -- Operation: refineTree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 85 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 86 | replace :: (Tree -> Tree) -> Path -> Tree -> Tree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 87 | replace f [] t = (f t) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 88 | replace f (V _ _ : es) (Variable q r pos ty t) = Variable q r pos ty (replace f es t) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 89 | replace f (C _ i : es) (Constructor q r pos ts) = Constructor q r pos (xs ++ (replace f es y : ys)) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 90 | where | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 91 | (xs, y : ys) = splitAt i ts | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 92 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 93 | refine_tree :: [Edge] -> Pos -> Tree -> Tree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 94 | refine_tree es p t = replace refine (path_of_position p es) t | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 95 | where | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 96 | path_of_position p es = takeWhile (\e -> position_of e /= p) es | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 97 | refine (Variable q r p (Generated_Code.Narrowing_sum_of_products ps) t) = | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 98 | Constructor q r p [ foldr (\(i,ty) t -> Variable q r (p++[i]) ty t) t (zip [0..] ts) | ts <- ps ] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 99 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 100 | -- refute | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 101 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 102 | refute :: ([Generated_Code.Narrowing_term] -> Bool) -> Bool -> Int -> Tree -> IO Tree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 103 | refute exec genuine_only d t = ref t | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 104 | where | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 105 | ref t = | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 106 | let path = find t in | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 107 | do | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 108 | t' <- answer genuine_only (exec (terms_of [] path)) (\b -> return (update path (Eval b) t)) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 109 | (\p -> return (if length p < d then refine_tree path p t else update path Unknown t)); | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 110 | case value_of t' of | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 111 | Unevaluated -> ref t' | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 112 | _ -> return t' | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 113 | |
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 114 | depthCheck :: Bool -> Int -> Generated_Code.Property -> IO () | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 115 | depthCheck genuine_only d p = refute (checkOf p) genuine_only d (treeOf 0 p) >>= (\t -> | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 116 | case value_of t of | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 117 |    Eval False -> putStrLn ("SOME (" ++ show (counterexampleOf (reifysOf p) (exampleOf 0 t)) ++ ")")  
 | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 118 |    _ -> putStrLn ("NONE"))
 | 
| 43313 | 119 | |
| 120 | -- Term refinement | |
| 121 | ||
| 122 | -- Operation: termOf | |
| 123 | ||
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 124 | term_of :: Pos -> Path -> Generated_Code.Narrowing_term | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 125 | term_of p (C [] i : es) = Generated_Code.Narrowing_constructor i (terms_of p es) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 126 | term_of p [V [] ty] = Generated_Code.Narrowing_variable p ty | 
| 43313 | 127 | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 128 | terms_of :: Pos -> Path -> [Generated_Code.Narrowing_term] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 129 | terms_of p es = terms_of' 0 es | 
| 43313 | 130 | where | 
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 131 | terms_of' i [] = [] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 132 | terms_of' i (e : es) = (t : terms_of' (i + 1) rs) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 133 | where | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 134 | (ts, rs) = Data.List.partition (\e -> head (position_of e) == i) (e : es) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 135 | t = term_of (p ++ [i]) (map (map_pos tail) ts) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 136 | map_pos f (V p ty) = V (f p) ty | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 137 | map_pos f (C p ts) = C (f p) ts | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 138 | |
| 43313 | 139 | -- Answers | 
| 140 | ||
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 141 | data Answer = Known Bool | Refine Pos; | 
| 43313 | 142 | |
| 45003 
7591039fb6b4
catch PatternMatchFail exceptions in narrowing-based quickcheck
 bulwahn parents: 
44751diff
changeset | 143 | answeri :: a -> (a -> IO b) -> (Pos -> IO b) -> IO b; | 
| 
7591039fb6b4
catch PatternMatchFail exceptions in narrowing-based quickcheck
 bulwahn parents: 
44751diff
changeset | 144 | answeri a known unknown = | 
| 43313 | 145 | do res <- try (evaluate a) | 
| 146 | case res of | |
| 147 | Right b -> known b | |
| 148 |        Left (ErrorCall ('\0':p)) -> unknown (map fromEnum p)
 | |
| 149 | Left e -> throw e | |
| 150 | ||
| 45685 
e2e928af750b
quickcheck narrowing also shows potential counterexamples
 bulwahn parents: 
45081diff
changeset | 151 | answer :: Bool -> Bool -> (Bool -> IO b) -> (Pos -> IO b) -> IO b; | 
| 45760 
3b5a735897c3
inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
 bulwahn parents: 
45685diff
changeset | 152 | answer genuine_only a known unknown = | 
| 45003 
7591039fb6b4
catch PatternMatchFail exceptions in narrowing-based quickcheck
 bulwahn parents: 
44751diff
changeset | 153 | Control.Exception.catch (answeri a known unknown) | 
| 45760 
3b5a735897c3
inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
 bulwahn parents: 
45685diff
changeset | 154 | (\ (PatternMatchFail _) -> known genuine_only) | 
| 45003 
7591039fb6b4
catch PatternMatchFail exceptions in narrowing-based quickcheck
 bulwahn parents: 
44751diff
changeset | 155 | |
| 43313 | 156 | |
| 157 | -- presentation of counterexample | |
| 158 | ||
| 159 | ||
| 55676 | 160 | instance Show Typerep.Typerep where {
 | 
| 161 | show (Typerep.Typerep c ts) = "Type (\"" ++ c ++ "\", " ++ show ts ++ ")"; | |
| 43313 | 162 | }; | 
| 163 | ||
| 45081 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 164 | instance Show Generated_Code.Term where {
 | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 165 | show (Generated_Code.Const c t) = "Const (\"" ++ c ++ "\", " ++ show t ++ ")"; | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 166 |   show (Generated_Code.App s t) = "(" ++ show s ++ ") $ (" ++ show t ++ ")";
 | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 167 | show (Generated_Code.Abs s ty t) = "Abs (\"" ++ s ++ "\", " ++ show ty ++ ", " ++ show t ++ ")"; | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 168 | show (Generated_Code.Free s ty) = "Free (\"" ++ s ++ "\", " ++ show ty ++ ")"; | 
| 43313 | 169 | }; | 
| 170 | {-
 | |
| 171 | posOf :: Edge -> Pos | |
| 172 | posOf (VN pos _) = pos | |
| 173 | posOf (CtrB pos _) = pos | |
| 174 | ||
| 175 | tailPosEdge :: Edge -> Edge | |
| 176 | tailPosEdge (VN pos ty) = VN (tail pos) ty | |
| 177 | tailPosEdge (CtrB pos ts) = CtrB (tail pos) ts | |
| 178 | ||
| 179 | termOf :: Pos -> Tree -> (Narrowing_term, Tree) | |
| 180 | termOf pos = if Ctr i (termListOf (pos ++ [i]) ) | |
| 181 | termOf pos [VN [] ty] = Var pos ty | |
| 182 | ||
| 183 | termListOf :: Pos -> [Narrowing_term] | |
| 184 | termListOf pos es = termListOf' 0 es | |
| 185 | where | |
| 186 | termListOf' i [] = [] | |
| 187 | termListOf' i (e : es) = | |
| 188 | let | |
| 189 | (ts, rs) = List.partition (\e -> head (posOf e) == i) (e : es) | |
| 190 | t = termOf (pos ++ [i]) (map tailPosEdge ts) | |
| 191 | in | |
| 192 | (t : termListOf' (i + 1) rs) | |
| 193 | ||
| 194 | termlist_of :: Pos -> QuantTree -> ([Term], QuantTree) | |
| 195 | ||
| 196 | termlist_of p' (Node r) | |
| 197 | ||
| 198 | term_of p' (VarNode _ _ p ty t) = if p == p' then | |
| 199 | (Some (Var ty), t) | |
| 200 | else | |
| 201 | (None, t) | |
| 202 | term_of p' (CtrBranch q _ p ts) = | |
| 203 | if p == p' then | |
| 204 | let | |
| 205 | i = findindex (\t -> evalOf t == Eval False) | |
| 206 | in | |
| 207 | Ctr i (termlist_of (p ++ [i]) (ts ! i) []) | |
| 208 | else | |
| 209 | error "" | |
| 210 | -} | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 211 | termlist_of :: Pos -> ([Generated_Code.Narrowing_term], Tree) -> ([Generated_Code.Narrowing_term], Tree) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 212 | termlist_of p' (terms, Leaf b) = (terms, Leaf b) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 213 | termlist_of p' (terms, Variable q r p ty t) = if p' == take (length p') p then | 
| 46758 
4106258260b3
choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
 bulwahn parents: 
46408diff
changeset | 214 | termlist_of p' (terms ++ [Generated_Code.Narrowing_variable p ty], t) | 
| 43313 | 215 | else | 
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 216 | (terms, Variable q r p ty t) | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 217 | termlist_of p' (terms, Constructor q r p ts) = if p' == take (length p') p then | 
| 43313 | 218 | let | 
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 219 | Just i = findIndex (\t -> value_of t == Eval False) ts | 
| 43313 | 220 | (subterms, t') = fixp (\j -> termlist_of (p ++ [j])) 0 ([], ts !! i) | 
| 221 | in | |
| 46758 
4106258260b3
choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
 bulwahn parents: 
46408diff
changeset | 222 | (terms ++ [Generated_Code.Narrowing_constructor i subterms], t') | 
| 43313 | 223 | else | 
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 224 | (terms, Constructor q r p ts) | 
| 43313 | 225 | where | 
| 226 | fixp f j s = if length (fst (f j s)) == length (fst s) then s else fixp f (j + 1) (f j s) | |
| 227 | ||
| 228 | ||
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 229 | alltermlist_of :: Pos -> ([Generated_Code.Narrowing_term], Tree) -> [([Generated_Code.Narrowing_term], Tree)] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 230 | alltermlist_of p' (terms, Leaf b) = [(terms, Leaf b)] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 231 | alltermlist_of p' (terms, Variable q r p ty t) = if p' == take (length p') p then | 
| 46758 
4106258260b3
choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
 bulwahn parents: 
46408diff
changeset | 232 | alltermlist_of p' (terms ++ [Generated_Code.Narrowing_variable p ty], t) | 
| 43313 | 233 | else | 
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 234 | [(terms, Variable q r p ty t)] | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 235 | alltermlist_of p' (terms, Constructor q r p ts) = | 
| 43313 | 236 | if p' == take (length p') p then | 
| 237 | let | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 238 | its = filter (\(i, t) -> value_of t == Eval False) (zip [0..] ts) | 
| 43313 | 239 | in | 
| 240 | concatMap | |
| 46758 
4106258260b3
choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
 bulwahn parents: 
46408diff
changeset | 241 | (\(i, t) -> map (\(subterms, t') -> (terms ++ [Generated_Code.Narrowing_constructor i subterms], t')) | 
| 43313 | 242 | (fixp (\j -> alltermlist_of (p ++ [j])) 0 ([], t))) its | 
| 243 | else | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 244 | [(terms, Constructor q r p ts)] | 
| 43313 | 245 | where | 
| 246 | fixp f j s = case (f j s) of | |
| 247 | [s'] -> if length (fst s') == length (fst s) then [s'] else concatMap (fixp f (j + 1)) (f j s) | |
| 248 | _ -> concatMap (fixp f (j + 1)) (f j s) | |
| 249 | ||
| 45081 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 250 | data Example = UnivExample Generated_Code.Narrowing_term Example | ExExample [(Generated_Code.Narrowing_term, Example)] | EmptyExample | 
| 43313 | 251 | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 252 | quantifierOf (Variable q _ _ _ _) = q | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 253 | quantifierOf (Constructor q _ _ _) = q | 
| 43313 | 254 | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 255 | exampleOf :: Int -> Tree -> Example | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 256 | exampleOf _ (Leaf _) = EmptyExample | 
| 43313 | 257 | exampleOf p t = | 
| 258 | case quantifierOf t of | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 259 | Universal -> | 
| 43313 | 260 | let | 
| 261 | ([term], rt) = termlist_of [p] ([], t) | |
| 262 | in UnivExample term (exampleOf (p + 1) rt) | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 263 | Existential -> | 
| 43313 | 264 | ExExample (map (\([term], rt) -> (term, exampleOf (p + 1) rt)) (alltermlist_of [p] ([], t))) | 
| 265 | ||
| 45081 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 266 | data Counterexample = Universal_Counterexample (Generated_Code.Term, Counterexample) | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 267 | | Existential_Counterexample [(Generated_Code.Term, Counterexample)] | Empty_Assignment | 
| 43313 | 268 | |
| 269 | instance Show Counterexample where {
 | |
| 270 | show Empty_Assignment = "Narrowing_Generators.Empty_Assignment"; | |
| 271 | show (Universal_Counterexample x) = "Narrowing_Generators.Universal_Counterexample" ++ show x; | |
| 272 | show (Existential_Counterexample x) = "Narrowing_Generators.Existential_Counterexample" ++ show x; | |
| 273 | }; | |
| 274 | ||
| 45081 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 275 | counterexampleOf :: [Generated_Code.Narrowing_term -> Generated_Code.Term] -> Example -> Counterexample | 
| 43313 | 276 | counterexampleOf [] EmptyExample = Empty_Assignment | 
| 277 | counterexampleOf (reify : reifys) (UnivExample t ex) = Universal_Counterexample (reify t, counterexampleOf reifys ex) | |
| 278 | counterexampleOf (reify : reifys) (ExExample exs) = Existential_Counterexample (map (\(t, ex) -> (reify t, counterexampleOf reifys ex)) exs) | |
| 279 | ||
| 45081 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 280 | checkOf :: Generated_Code.Property -> [Generated_Code.Narrowing_term] -> Bool | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 281 | checkOf (Generated_Code.Property b) = (\[] -> b) | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 282 | checkOf (Generated_Code.Universal _ f _) = (\(t : ts) -> checkOf (f t) ts) | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 283 | checkOf (Generated_Code.Existential _ f _) = (\(t : ts) -> checkOf (f t) ts) | 
| 43313 | 284 | |
| 49193 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 285 | treeOf :: Int -> Generated_Code.Property -> Tree | 
| 
0067d83414c8
clearer names for functions in Quickcheck's narrowing engine
 bulwahn parents: 
47090diff
changeset | 286 | treeOf n (Generated_Code.Property _) = Leaf Unevaluated | 
| 49253 
4b11240d80bf
replacing own dummy value by Haskell's Prelude.undefined
 bulwahn parents: 
49193diff
changeset | 287 | treeOf n (Generated_Code.Universal ty f _) = Variable Universal Unevaluated [n] ty (treeOf (n + 1) (f undefined)) | 
| 
4b11240d80bf
replacing own dummy value by Haskell's Prelude.undefined
 bulwahn parents: 
49193diff
changeset | 288 | treeOf n (Generated_Code.Existential ty f _) = Variable Existential Unevaluated [n] ty (treeOf (n + 1) (f undefined)) | 
| 43313 | 289 | |
| 45081 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 290 | reifysOf :: Generated_Code.Property -> [Generated_Code.Narrowing_term -> Generated_Code.Term] | 
| 
f00e52acbd42
importing the Generated_Code module qualified to reduce the probability of name clashes between the static code and the generated code in the narrowing-based Quickcheck
 bulwahn parents: 
45003diff
changeset | 291 | reifysOf (Generated_Code.Property _) = [] | 
| 49253 
4b11240d80bf
replacing own dummy value by Haskell's Prelude.undefined
 bulwahn parents: 
49193diff
changeset | 292 | reifysOf (Generated_Code.Universal _ f r) = (r : (reifysOf (f undefined))) | 
| 
4b11240d80bf
replacing own dummy value by Haskell's Prelude.undefined
 bulwahn parents: 
49193diff
changeset | 293 | reifysOf (Generated_Code.Existential _ f r) = (r : (reifysOf (f undefined))) | 
| 43313 | 294 |