src/HOL/Tools/Quickcheck/Narrowing_Engine.hs
author haftmann
Sun, 23 Feb 2014 10:33:43 +0100
changeset 55684 ee49b4f7edc8
parent 55676 fb46f1c379b5
permissions -rw-r--r--
keep only identifiers public which are explicitly requested or demanded by dependencies
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41933
10f254a4e5b9 adapting Main file generation for Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
     1
module Narrowing_Engine where {
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     2
46335
0fd9ab902b5a using fully qualified module names in Haskell source, which seems to be required by GHC 7.0.4
bulwahn
parents: 45760
diff changeset
     3
import Control.Monad;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     4
import Control.Exception;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     5
import System.IO;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     6
import System.Exit;
55676
fb46f1c379b5 avoid ad-hoc patching of generated code
haftmann
parents: 46758
diff changeset
     7
import qualified Typerep;
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
     8
import qualified Generated_Code;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     9
41908
3bd9a21366d2 changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
bulwahn
parents: 41905
diff changeset
    10
type Pos = [Int];
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    11
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    12
-- Term refinement
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    13
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
    14
new :: Pos -> [[Generated_Code.Narrowing_type]] -> [Generated_Code.Narrowing_term];
46758
4106258260b3 choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
bulwahn
parents: 46335
diff changeset
    15
new p ps = [ Generated_Code.Narrowing_constructor c (zipWith (\i t -> Generated_Code.Narrowing_variable (p++[i]) t) [0..] ts)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    16
           | (c, ts) <- zip [0..] ps ];
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    17
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
    18
refine :: Generated_Code.Narrowing_term -> Pos -> [Generated_Code.Narrowing_term];
46758
4106258260b3 choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
bulwahn
parents: 46335
diff changeset
    19
refine (Generated_Code.Narrowing_variable p (Generated_Code.Narrowing_sum_of_products ss)) [] = new p ss;
4106258260b3 choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
bulwahn
parents: 46335
diff changeset
    20
refine (Generated_Code.Narrowing_constructor c xs) p = map (Generated_Code.Narrowing_constructor c) (refineList xs p);
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    21
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
    22
refineList :: [Generated_Code.Narrowing_term] -> Pos -> [[Generated_Code.Narrowing_term]];
41908
3bd9a21366d2 changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
bulwahn
parents: 41905
diff changeset
    23
refineList xs (i:is) = let (ls, x:rs) = splitAt i xs in [ls ++ y:rs | y <- refine x is];
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    24
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    25
-- Find total instantiations of a partial value
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    26
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
    27
total :: Generated_Code.Narrowing_term -> [Generated_Code.Narrowing_term];
46758
4106258260b3 choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
bulwahn
parents: 46335
diff changeset
    28
total (Generated_Code.Narrowing_constructor c xs) = [Generated_Code.Narrowing_constructor c ys | ys <- mapM total xs];
4106258260b3 choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
bulwahn
parents: 46335
diff changeset
    29
total (Generated_Code.Narrowing_variable p (Generated_Code.Narrowing_sum_of_products ss)) = [y | x <- new p ss, y <- total x];
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    30
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    31
-- Answers
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    32
45725
2987b29518aa the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents: 45685
diff changeset
    33
answeri :: a -> (Bool -> a -> IO b) -> (Pos -> IO b) -> IO b;
45003
7591039fb6b4 catch PatternMatchFail exceptions in narrowing-based quickcheck
bulwahn
parents: 44751
diff changeset
    34
answeri a known unknown =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    35
  try (evaluate a) >>= (\res ->
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    36
     case res of
45725
2987b29518aa the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents: 45685
diff changeset
    37
       Right b -> known True b
41908
3bd9a21366d2 changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
bulwahn
parents: 41905
diff changeset
    38
       Left (ErrorCall ('\0':p)) -> unknown (map fromEnum p)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    39
       Left e -> throw e);
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    40
45725
2987b29518aa the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents: 45685
diff changeset
    41
answer :: Bool -> 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: 45756
diff changeset
    42
answer genuine_only a known unknown =
45003
7591039fb6b4 catch PatternMatchFail exceptions in narrowing-based quickcheck
bulwahn
parents: 44751
diff changeset
    43
  Control.Exception.catch (answeri a known unknown) 
45760
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
    44
    (\ (PatternMatchFail _) -> known False genuine_only);
45003
7591039fb6b4 catch PatternMatchFail exceptions in narrowing-based quickcheck
bulwahn
parents: 44751
diff changeset
    45
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    46
-- Refute
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    47
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    48
str_of_list [] = "[]";
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    49
str_of_list (x:xs) = "(" ++ x ++ " :: " ++ str_of_list xs ++ ")";
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    50
45725
2987b29518aa the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents: 45685
diff changeset
    51
report :: Bool -> Result -> [Generated_Code.Narrowing_term] -> IO Int;
2987b29518aa the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents: 45685
diff changeset
    52
report genuine r xs = putStrLn ("SOME (" ++ (if genuine then "true" else "false") ++
2987b29518aa the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents: 45685
diff changeset
    53
  ", " ++ (str_of_list $ zipWith ($) (showArgs r) xs) ++ ")") >> hFlush stdout >> exitWith ExitSuccess;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    54
45725
2987b29518aa the narrowing also indicates if counterexample is potentially spurious
bulwahn
parents: 45685
diff changeset
    55
eval :: Bool -> Bool -> (Bool -> Bool -> IO a) -> (Pos -> IO a) -> IO a;
45760
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
    56
eval genuine_only p k u = answer genuine_only p k u;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    57
45685
e2e928af750b quickcheck narrowing also shows potential counterexamples
bulwahn
parents: 45081
diff changeset
    58
ref :: Bool -> Result -> [Generated_Code.Narrowing_term] -> IO Int;
45760
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
    59
ref genuine_only r xs = eval genuine_only (apply_fun r xs) (\genuine res -> if res then return 1 else report genuine r xs)
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
    60
  (\p -> sumMapM (ref genuine_only r) 1 (refineList xs p));
45756
295658b28d3b quickcheck narrowing continues searching after found a potentially spurious counterexample
bulwahn
parents: 45725
diff changeset
    61
45685
e2e928af750b quickcheck narrowing also shows potential counterexamples
bulwahn
parents: 45081
diff changeset
    62
refute :: Bool -> Result -> IO Int;
45760
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
    63
refute genuine_only r = ref genuine_only r (args r);
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    64
41908
3bd9a21366d2 changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
bulwahn
parents: 41905
diff changeset
    65
sumMapM :: (a -> IO Int) -> Int -> [a] -> IO Int;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    66
sumMapM f n [] = return n;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    67
sumMapM f n (a:as) = seq n (do m <- f a ; sumMapM f (n+m) as);
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    68
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    69
-- Testable
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    70
55676
fb46f1c379b5 avoid ad-hoc patching of generated code
haftmann
parents: 46758
diff changeset
    71
instance Show Typerep.Typerep where {
fb46f1c379b5 avoid ad-hoc patching of generated code
haftmann
parents: 46758
diff changeset
    72
  show (Typerep.Typerep c ts) = "Type (\"" ++ c ++ "\", " ++ show ts ++ ")";
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    73
};
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    74
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
    75
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
    76
  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
    77
  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
    78
  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
    79
  show (Generated_Code.Free s ty) = "Free (\"" ++ s ++  "\", " ++ show ty ++ ")";
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    80
};
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    81
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    82
data Result =
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
    83
  Result { args     :: [Generated_Code.Narrowing_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
    84
         , showArgs :: [Generated_Code.Narrowing_term -> String]
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
    85
         , apply_fun    :: [Generated_Code.Narrowing_term] -> Bool
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    86
         };
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    87
41908
3bd9a21366d2 changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
bulwahn
parents: 41905
diff changeset
    88
data P = P (Int -> Int -> Result);
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    89
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
    90
run :: Testable a => ([Generated_Code.Narrowing_term] -> a) -> Int -> Int -> Result;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    91
run a = let P f = property a in f;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    92
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    93
class Testable a where {
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
    94
  property :: ([Generated_Code.Narrowing_term] -> a) -> P;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    95
};
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    96
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    97
instance Testable Bool where {
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    98
  property app = P $ \n d -> Result [] [] (app . reverse);
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    99
};
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   100
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
   101
instance (Generated_Code.Partial_term_of a, Generated_Code.Narrowing a, Testable b) => Testable (a -> b) where {
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   102
  property f = P $ \n d ->
46758
4106258260b3 choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
bulwahn
parents: 46335
diff changeset
   103
    let Generated_Code.Narrowing_cons t c = Generated_Code.narrowing d
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
   104
        c' = Generated_Code.conv c
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   105
        r = run (\(x:xs) -> f xs (c' x)) (n+1) d
46758
4106258260b3 choosing longer constant names in Quickcheck_Narrowing to reduce the chances of name clashes in Quickcheck-Narrowing
bulwahn
parents: 46335
diff changeset
   106
    in  r { args = Generated_Code.Narrowing_variable [n] t : args r,
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
   107
      showArgs = (show . Generated_Code.partial_term_of (Generated_Code.Type :: Generated_Code.Itself a)) : showArgs r };
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   108
};
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   109
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   110
-- Top-level interface
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   111
45685
e2e928af750b quickcheck narrowing also shows potential counterexamples
bulwahn
parents: 45081
diff changeset
   112
depthCheck :: Testable a => Bool -> Int -> a -> IO ();
45760
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
   113
depthCheck genuine_only d p =
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
   114
  (refute genuine_only $ run (const p) 0 d) >> putStrLn ("NONE") >> hFlush stdout;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   115
45685
e2e928af750b quickcheck narrowing also shows potential counterexamples
bulwahn
parents: 45081
diff changeset
   116
smallCheck :: Testable a => Bool -> Int -> a -> IO ();
45760
3b5a735897c3 inverted flag potential to genuine_only in the quickcheck narrowing Haskell code
bulwahn
parents: 45756
diff changeset
   117
smallCheck genuine_only d p = mapM_ (\d -> depthCheck genuine_only d p) [0..d] >> putStrLn ("NONE") >> hFlush stdout;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   118
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   119
}