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