src/HOL/ex/SML_Quickcheck_Examples.thy
changeset 44829 bc5e767f0f46
parent 44552 bc7d63c7fd6f
equal deleted inserted replaced
44828:64f88ef1835e 44829:bc5e767f0f46
     1 (*  Title:      HOL/ex/SML_Quickcheck_Examples.thy
       
     2     Author:     Stefan Berghofer, Lukas Bulwahn
       
     3     Copyright   2011 TU Muenchen
       
     4 *)
       
     5 
       
     6 theory SML_Quickcheck_Examples
       
     7 imports "~~/src/HOL/Library/SML_Quickcheck"
       
     8 begin
       
     9 
       
    10 text {*
       
    11 This is a regression test for the 'quickcheck' counterexample generator based on the
       
    12 SML code generator.
       
    13 
       
    14 The counterexample generator has been superseded by counterexample generators based on
       
    15 the generic code generator.
       
    16 *}
       
    17 
       
    18 declare [[quickcheck_finite_types = false]]
       
    19 declare [[quickcheck_timeout = 600.0]]
       
    20 
       
    21 subsection {* Lists *}
       
    22 
       
    23 theorem "map g (map f xs) = map (g o f) xs"
       
    24   quickcheck[SML, expect = no_counterexample]
       
    25   oops
       
    26 
       
    27 theorem "map g (map f xs) = map (f o g) xs"
       
    28   quickcheck[SML, expect = counterexample]
       
    29   oops
       
    30 
       
    31 theorem "rev (xs @ ys) = rev ys @ rev xs"
       
    32   quickcheck[SML, expect = no_counterexample]
       
    33   oops
       
    34 
       
    35 theorem "rev (xs @ ys) = rev xs @ rev ys"
       
    36   quickcheck[SML, expect = counterexample]
       
    37   oops
       
    38 
       
    39 theorem "rev (rev xs) = xs"
       
    40   quickcheck[SML, expect = no_counterexample]
       
    41   oops
       
    42 
       
    43 theorem "rev xs = xs"
       
    44   quickcheck[tester = SML, expect = counterexample]
       
    45 oops
       
    46 
       
    47 
       
    48 text {* An example involving functions inside other data structures *}
       
    49 
       
    50 primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
       
    51   "app [] x = x"
       
    52   | "app (f # fs) x = app fs (f x)"
       
    53 
       
    54 lemma "app (fs @ gs) x = app gs (app fs x)"
       
    55   quickcheck[SML, expect = no_counterexample]
       
    56   by (induct fs arbitrary: x) simp_all
       
    57 
       
    58 lemma "app (fs @ gs) x = app fs (app gs x)"
       
    59   quickcheck[SML, expect = counterexample]
       
    60   oops
       
    61 
       
    62 primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
       
    63   "occurs a [] = 0"
       
    64   | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
       
    65 
       
    66 primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
       
    67   "del1 a [] = []"
       
    68   | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
       
    69 
       
    70 text {* A lemma, you'd think to be true from our experience with delAll *}
       
    71 lemma "Suc (occurs a (del1 a xs)) = occurs a xs"
       
    72   -- {* Wrong. Precondition needed.*}
       
    73   quickcheck[SML, expect = counterexample]
       
    74   oops
       
    75 
       
    76 lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
       
    77   quickcheck[SML, expect = counterexample]
       
    78     -- {* Also wrong.*}
       
    79   oops
       
    80 
       
    81 lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
       
    82   quickcheck[SML, expect = no_counterexample]
       
    83   by (induct xs) auto
       
    84 
       
    85 primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
       
    86   "replace a b [] = []"
       
    87   | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) 
       
    88                             else (x#(replace a b xs)))"
       
    89 
       
    90 lemma "occurs a xs = occurs b (replace a b xs)"
       
    91   quickcheck[SML, expect = counterexample]
       
    92   -- {* Wrong. Precondition needed.*}
       
    93   oops
       
    94 
       
    95 lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)"
       
    96   quickcheck[SML, expect = no_counterexample]
       
    97   by (induct xs) simp_all
       
    98 
       
    99 
       
   100 subsection {* Trees *}
       
   101 
       
   102 datatype 'a tree = Twig |  Leaf 'a | Branch "'a tree" "'a tree"
       
   103 
       
   104 primrec leaves :: "'a tree \<Rightarrow> 'a list" where
       
   105   "leaves Twig = []"
       
   106   | "leaves (Leaf a) = [a]"
       
   107   | "leaves (Branch l r) = (leaves l) @ (leaves r)"
       
   108 
       
   109 primrec plant :: "'a list \<Rightarrow> 'a tree" where
       
   110   "plant [] = Twig "
       
   111   | "plant (x#xs) = Branch (Leaf x) (plant xs)"
       
   112 
       
   113 primrec mirror :: "'a tree \<Rightarrow> 'a tree" where
       
   114   "mirror (Twig) = Twig "
       
   115   | "mirror (Leaf a) = Leaf a "
       
   116   | "mirror (Branch l r) = Branch (mirror r) (mirror l)"
       
   117 
       
   118 theorem "plant (rev (leaves xt)) = mirror xt"
       
   119   quickcheck[SML, expect = counterexample]
       
   120     --{* Wrong! *} 
       
   121   oops
       
   122 
       
   123 theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt"
       
   124   quickcheck[SML, expect = counterexample]
       
   125     --{* Wrong! *} 
       
   126   oops
       
   127 
       
   128 datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree"
       
   129 
       
   130 primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where
       
   131   "inOrder (Tip a)= [a]"
       
   132   | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
       
   133 
       
   134 primrec root :: "'a ntree \<Rightarrow> 'a" where
       
   135   "root (Tip a) = a"
       
   136   | "root (Node f x y) = f"
       
   137 
       
   138 theorem "hd (inOrder xt) = root xt"
       
   139   quickcheck[SML, expect = counterexample]
       
   140   --{* Wrong! *} 
       
   141   oops
       
   142 
       
   143 end