src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy
changeset 46585 f462e49eaf11
parent 45790 3355c27c93a4
child 46879 a8b1236e0837
equal deleted inserted replaced
46584:a935175fe6b6 46585:f462e49eaf11
       
     1 (*  Title:      HOL/ex/Quickcheck_Narrowing_Examples.thy
       
     2     Author:     Lukas Bulwahn
       
     3     Copyright   2011 TU Muenchen
       
     4 *)
       
     5 
       
     6 header {* Examples for narrowing-based testing  *}
       
     7 
       
     8 theory Quickcheck_Narrowing_Examples
       
     9 imports Main
       
    10 begin
       
    11 
       
    12 declare [[quickcheck_timeout = 3600]]
       
    13 
       
    14 subsection {* Minimalistic examples *}
       
    15 
       
    16 lemma
       
    17   "x = y"
       
    18 quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample]
       
    19 oops
       
    20 
       
    21 lemma
       
    22   "x = y"
       
    23 quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
       
    24 oops
       
    25 
       
    26 subsection {* Simple examples with existentials *}
       
    27 
       
    28 lemma
       
    29   "\<exists> y :: nat. \<forall> x. x = y"
       
    30 quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
       
    31 oops
       
    32 (*
       
    33 lemma
       
    34   "\<exists> y :: int. \<forall> x. x = y"
       
    35 quickcheck[tester = narrowing, size = 2]
       
    36 oops
       
    37 *)
       
    38 lemma
       
    39   "x > 1 --> (\<exists>y :: nat. x < y & y <= 1)"
       
    40 quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
       
    41 oops
       
    42 
       
    43 lemma
       
    44   "x > 2 --> (\<exists>y :: nat. x < y & y <= 2)"
       
    45 quickcheck[tester = narrowing, finite_types = false, size = 10]
       
    46 oops
       
    47 
       
    48 lemma
       
    49   "\<forall> x. \<exists> y :: nat. x > 3 --> (y < x & y > 3)"
       
    50 quickcheck[tester = narrowing, finite_types = false, size = 7]
       
    51 oops
       
    52 
       
    53 text {* A false conjecture derived from an partial copy-n-paste of @{thm not_distinct_decomp} *}
       
    54 lemma
       
    55   "~ distinct ws ==> EX xs ys zs y. ws = xs @ [y] @ ys @ [y]"
       
    56 quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
       
    57 oops
       
    58 
       
    59 text {* A false conjecture derived from theorems @{thm split_list_first} and @{thm split_list_last} *}  
       
    60 lemma
       
    61   "x : set xs ==> EX ys zs. xs = ys @ x # zs & x ~: set zs & x ~: set ys"
       
    62 quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
       
    63 oops
       
    64 
       
    65 text {* A false conjecture derived from @{thm map_eq_Cons_conv} with a typo *}
       
    66 lemma
       
    67   "(map f xs = y # ys) = (EX z zs. xs = z' # zs & f z = y & map f zs = ys)"
       
    68 quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
       
    69 oops
       
    70 
       
    71 lemma "a # xs = ys @ [a] ==> EX zs. xs = zs @ [a] & ys = a#zs"
       
    72 quickcheck[narrowing, expect = counterexample]
       
    73 quickcheck[exhaustive, random, expect = no_counterexample]
       
    74 oops
       
    75 
       
    76 subsection {* Simple list examples *}
       
    77 
       
    78 lemma "rev xs = xs"
       
    79 quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
       
    80 oops
       
    81 
       
    82 (* FIXME: integer has strange representation! *)
       
    83 lemma "rev xs = xs"
       
    84 quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample]
       
    85 oops
       
    86 (*
       
    87 lemma "rev xs = xs"
       
    88   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
       
    89 oops
       
    90 *)
       
    91 subsection {* Simple examples with functions *}
       
    92 
       
    93 lemma "map f xs = map g xs"
       
    94   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
       
    95 oops
       
    96 
       
    97 lemma "map f xs = map f ys ==> xs = ys"
       
    98   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
       
    99 oops
       
   100 
       
   101 lemma
       
   102   "list_all2 P (rev xs) (rev ys) = list_all2 P xs (rev ys)"
       
   103   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
       
   104 oops
       
   105 
       
   106 lemma "map f xs = F f xs"
       
   107   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
       
   108 oops
       
   109 
       
   110 lemma "map f xs = F f xs"
       
   111   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
       
   112 oops
       
   113 
       
   114 (* requires some work...*)
       
   115 (*
       
   116 lemma "R O S = S O R"
       
   117   quickcheck[tester = narrowing, size = 2]
       
   118 oops
       
   119 *)
       
   120 
       
   121 subsection {* Simple examples with inductive predicates *}
       
   122 
       
   123 inductive even where
       
   124   "even 0" |
       
   125   "even n ==> even (Suc (Suc n))"
       
   126 
       
   127 code_pred even .
       
   128 
       
   129 lemma "even (n - 2) ==> even n"
       
   130 quickcheck[narrowing, expect = counterexample]
       
   131 oops
       
   132 
       
   133 
       
   134 subsection {* AVL Trees *}
       
   135 
       
   136 datatype 'a tree = ET |  MKT 'a "'a tree" "'a tree" nat
       
   137 
       
   138 primrec set_of :: "'a tree \<Rightarrow> 'a set"
       
   139 where
       
   140 "set_of ET = {}" |
       
   141 "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
       
   142 
       
   143 primrec height :: "'a tree \<Rightarrow> nat"
       
   144 where
       
   145 "height ET = 0" |
       
   146 "height (MKT x l r h) = max (height l) (height r) + 1"
       
   147 
       
   148 primrec avl :: "'a tree \<Rightarrow> bool"
       
   149 where
       
   150 "avl ET = True" |
       
   151 "avl (MKT x l r h) =
       
   152  ((height l = height r \<or> height l = 1+height r \<or> height r = 1+height l) \<and> 
       
   153   h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
       
   154 
       
   155 primrec is_ord :: "('a::order) tree \<Rightarrow> bool"
       
   156 where
       
   157 "is_ord ET = True" |
       
   158 "is_ord (MKT n l r h) =
       
   159  ((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)"
       
   160 
       
   161 primrec is_in :: "('a::order) \<Rightarrow> 'a tree \<Rightarrow> bool"
       
   162 where
       
   163  "is_in k ET = False" |
       
   164  "is_in k (MKT n l r h) = (if k = n then True else
       
   165                            if k < n then (is_in k l)
       
   166                            else (is_in k r))"
       
   167 
       
   168 primrec ht :: "'a tree \<Rightarrow> nat"
       
   169 where
       
   170 "ht ET = 0" |
       
   171 "ht (MKT x l r h) = h"
       
   172 
       
   173 definition
       
   174  mkt :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
       
   175 "mkt x l r = MKT x l r (max (ht l) (ht r) + 1)"
       
   176 
       
   177 (* replaced MKT lrn lrl lrr by MKT lrr lrl *)
       
   178 fun l_bal where
       
   179 "l_bal(n, MKT ln ll lr h, r) =
       
   180  (if ht ll < ht lr
       
   181   then case lr of ET \<Rightarrow> ET (* impossible *)
       
   182                 | MKT lrn lrr lrl lrh \<Rightarrow>
       
   183                   mkt lrn (mkt ln ll lrl) (mkt n lrr r)
       
   184   else mkt ln ll (mkt n lr r))"
       
   185 
       
   186 fun r_bal where
       
   187 "r_bal(n, l, MKT rn rl rr h) =
       
   188  (if ht rl > ht rr
       
   189   then case rl of ET \<Rightarrow> ET (* impossible *)
       
   190                 | MKT rln rll rlr h \<Rightarrow> mkt rln (mkt n l rll) (mkt rn rlr rr)
       
   191   else mkt rn (mkt n l rl) rr)"
       
   192 
       
   193 primrec insrt :: "'a::order \<Rightarrow> 'a tree \<Rightarrow> 'a tree"
       
   194 where
       
   195 "insrt x ET = MKT x ET ET 1" |
       
   196 "insrt x (MKT n l r h) = 
       
   197    (if x=n
       
   198     then MKT n l r h
       
   199     else if x<n
       
   200          then let l' = insrt x l; hl' = ht l'; hr = ht r
       
   201               in if hl' = 2+hr then l_bal(n,l',r)
       
   202                  else MKT n l' r (1 + max hl' hr)
       
   203          else let r' = insrt x r; hl = ht l; hr' = ht r'
       
   204               in if hr' = 2+hl then r_bal(n,l,r')
       
   205                  else MKT n l r' (1 + max hl hr'))"
       
   206 
       
   207 
       
   208 subsubsection {* Necessary setup for code generation *}
       
   209 
       
   210 primrec set_of'
       
   211 where 
       
   212   "set_of' ET = []"
       
   213 | "set_of' (MKT n l r h) = n # (set_of' l @ set_of' r)"
       
   214 
       
   215 lemma set_of':
       
   216   "set (set_of' t) = set_of t"
       
   217 by (induct t) auto
       
   218 
       
   219 lemma is_ord_mkt:
       
   220   "is_ord (MKT n l r h) = ((ALL n': set (set_of' l). n' < n) & (ALL n': set (set_of' r). n < n') & is_ord l & is_ord r)"
       
   221 by (simp add: set_of')
       
   222 
       
   223 declare is_ord.simps(1)[code] is_ord_mkt[code]
       
   224 
       
   225 subsubsection {* Invalid Lemma due to typo in lbal *}
       
   226 
       
   227 lemma is_ord_l_bal:
       
   228  "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))"
       
   229 quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, expect = counterexample]
       
   230 oops
       
   231 
       
   232 subsection {* Examples with hd and last *}
       
   233 
       
   234 lemma
       
   235   "hd (xs @ ys) = hd ys"
       
   236 quickcheck[narrowing, expect = counterexample]
       
   237 oops
       
   238 
       
   239 lemma
       
   240   "last(xs @ ys) = last xs"
       
   241 quickcheck[narrowing, expect = counterexample]
       
   242 oops
       
   243 
       
   244 subsection {* Examples with underspecified/partial functions *}
       
   245 
       
   246 lemma
       
   247   "xs = [] ==> hd xs \<noteq> x"
       
   248 quickcheck[narrowing, expect = no_counterexample]
       
   249 oops
       
   250 
       
   251 lemma
       
   252   "xs = [] ==> hd xs = x"
       
   253 quickcheck[narrowing, expect = no_counterexample]
       
   254 oops
       
   255 
       
   256 lemma "xs = [] ==> hd xs = x ==> x = y"
       
   257 quickcheck[narrowing, expect = no_counterexample]
       
   258 oops
       
   259 
       
   260 lemma
       
   261   "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
       
   262 quickcheck[narrowing, expect = no_counterexample]
       
   263 oops
       
   264 
       
   265 lemma
       
   266   "hd (map f xs) = f (hd xs)"
       
   267 quickcheck[narrowing, expect = no_counterexample]
       
   268 oops
       
   269 
       
   270 end