| author | wenzelm | 
| Wed, 23 Mar 2022 12:21:13 +0100 | |
| changeset 75311 | 5960bae73afe | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 46879 | 1 | (* Title: HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy | 
| 41907 | 2 | Author: Lukas Bulwahn | 
| 3 | Copyright 2011 TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 63167 | 6 | section \<open>Examples for narrowing-based testing\<close> | 
| 41907 | 7 | |
| 41932 
e8f113ce8a94
adapting Quickcheck_Narrowing and example file to new names
 bulwahn parents: 
41931diff
changeset | 8 | theory Quickcheck_Narrowing_Examples | 
| 43318 | 9 | imports Main | 
| 41907 | 10 | begin | 
| 11 | ||
| 45773 
7f2366dc8c0f
increasing quickcheck's timeout in the example theory to avoid failures on the testing infrastructure
 bulwahn parents: 
45658diff
changeset | 12 | declare [[quickcheck_timeout = 3600]] | 
| 
7f2366dc8c0f
increasing quickcheck's timeout in the example theory to avoid failures on the testing infrastructure
 bulwahn parents: 
45658diff
changeset | 13 | |
| 63167 | 14 | subsection \<open>Minimalistic examples\<close> | 
| 43356 | 15 | |
| 42020 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 16 | lemma | 
| 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 17 | "x = y" | 
| 42021 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 bulwahn parents: 
42020diff
changeset | 18 | quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample] | 
| 41907 | 19 | oops | 
| 43356 | 20 | |
| 42020 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 21 | lemma | 
| 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 22 | "x = y" | 
| 43356 | 23 | quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] | 
| 41907 | 24 | oops | 
| 25 | ||
| 63167 | 26 | subsection \<open>Simple examples with existentials\<close> | 
| 43239 | 27 | |
| 28 | lemma | |
| 29 | "\<exists> y :: nat. \<forall> x. x = y" | |
| 30 | quickcheck[tester = narrowing, finite_types = false, expect = counterexample] | |
| 31 | oops | |
| 43356 | 32 | (* | 
| 33 | lemma | |
| 34 | "\<exists> y :: int. \<forall> x. x = y" | |
| 35 | quickcheck[tester = narrowing, size = 2] | |
| 36 | oops | |
| 37 | *) | |
| 43239 | 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 | |
| 67613 | 44 | "x > 2 \<longrightarrow> (\<exists>y :: nat. x < y \<and> y \<le> 2)" | 
| 43239 | 45 | quickcheck[tester = narrowing, finite_types = false, size = 10] | 
| 46 | oops | |
| 47 | ||
| 48 | lemma | |
| 67613 | 49 | "\<forall>x. \<exists>y :: nat. x > 3 \<longrightarrow> (y < x \<and> y > 3)" | 
| 43239 | 50 | quickcheck[tester = narrowing, finite_types = false, size = 7] | 
| 51 | oops | |
| 52 | ||
| 63167 | 53 | text \<open>A false conjecture derived from an partial copy-n-paste of @{thm not_distinct_decomp}\<close>
 | 
| 43239 | 54 | lemma | 
| 67613 | 55 | "\<not> distinct ws \<Longrightarrow> \<exists>xs ys zs y. ws = xs @ [y] @ ys @ [y]" | 
| 43239 | 56 | quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] | 
| 57 | oops | |
| 58 | ||
| 63167 | 59 | text \<open>A false conjecture derived from theorems @{thm split_list_first} and @{thm split_list_last}\<close>  
 | 
| 43239 | 60 | lemma | 
| 67613 | 61 | "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs \<and> x \<notin> set ys" | 
| 43239 | 62 | quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] | 
| 63 | oops | |
| 64 | ||
| 63167 | 65 | text \<open>A false conjecture derived from @{thm map_eq_Cons_conv} with a typo\<close>
 | 
| 43239 | 66 | lemma | 
| 67613 | 67 | "(map f xs = y # ys) = (\<exists>z zs. xs = z' # zs & f z = y \<and> map f zs = ys)" | 
| 43239 | 68 | quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] | 
| 69 | oops | |
| 70 | ||
| 67613 | 71 | lemma "a # xs = ys @ [a] \<Longrightarrow> \<exists>zs. xs = zs @ [a] \<and> ys = a#zs" | 
| 45441 
fb4ac1dd4fde
adding some test cases for preprocessing and narrowing
 bulwahn parents: 
45082diff
changeset | 72 | quickcheck[narrowing, expect = counterexample] | 
| 
fb4ac1dd4fde
adding some test cases for preprocessing and narrowing
 bulwahn parents: 
45082diff
changeset | 73 | quickcheck[exhaustive, random, expect = no_counterexample] | 
| 
fb4ac1dd4fde
adding some test cases for preprocessing and narrowing
 bulwahn parents: 
45082diff
changeset | 74 | oops | 
| 
fb4ac1dd4fde
adding some test cases for preprocessing and narrowing
 bulwahn parents: 
45082diff
changeset | 75 | |
| 63167 | 76 | subsection \<open>Simple list examples\<close> | 
| 42020 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 77 | |
| 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 78 | lemma "rev xs = xs" | 
| 43239 | 79 | quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] | 
| 42020 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 80 | oops | 
| 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 81 | |
| 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 82 | lemma "rev xs = xs" | 
| 43239 | 83 | quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample] | 
| 42021 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 bulwahn parents: 
42020diff
changeset | 84 | oops | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46879diff
changeset | 85 | |
| 42021 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 bulwahn parents: 
42020diff
changeset | 86 | lemma "rev xs = xs" | 
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42021diff
changeset | 87 | quickcheck[tester = narrowing, finite_types = true, expect = counterexample] | 
| 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42021diff
changeset | 88 | oops | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46879diff
changeset | 89 | |
| 63167 | 90 | subsection \<open>Simple examples with functions\<close> | 
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42021diff
changeset | 91 | |
| 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42021diff
changeset | 92 | lemma "map f xs = map g xs" | 
| 43239 | 93 | quickcheck[tester = narrowing, finite_types = false, expect = counterexample] | 
| 42020 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 94 | oops | 
| 
2da02764d523
translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
 bulwahn parents: 
41963diff
changeset | 95 | |
| 42023 
1bd4b6d1c482
adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
 bulwahn parents: 
42021diff
changeset | 96 | lemma "map f xs = map f ys ==> xs = ys" | 
| 43239 | 97 | quickcheck[tester = narrowing, finite_types = false, expect = counterexample] | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 98 | oops | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 99 | |
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 100 | lemma | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 101 | "list_all2 P (rev xs) (rev ys) = list_all2 P xs (rev ys)" | 
| 45773 
7f2366dc8c0f
increasing quickcheck's timeout in the example theory to avoid failures on the testing infrastructure
 bulwahn parents: 
45658diff
changeset | 102 | quickcheck[tester = narrowing, finite_types = false, expect = counterexample] | 
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 103 | oops | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 104 | |
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 105 | lemma "map f xs = F f xs" | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 106 | quickcheck[tester = narrowing, finite_types = false, expect = counterexample] | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 107 | oops | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 108 | |
| 43239 | 109 | lemma "map f xs = F f xs" | 
| 110 | quickcheck[tester = narrowing, finite_types = false, expect = counterexample] | |
| 111 | oops | |
| 112 | ||
| 113 | (* requires some work...*) | |
| 114 | (* | |
| 42024 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 115 | lemma "R O S = S O R" | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 116 | quickcheck[tester = narrowing, size = 2] | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 117 | oops | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 118 | *) | 
| 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 bulwahn parents: 
42023diff
changeset | 119 | |
| 63167 | 120 | subsection \<open>Simple examples with inductive predicates\<close> | 
| 45082 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 121 | |
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 122 | inductive even where | 
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 123 | "even 0" | | 
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 124 | "even n ==> even (Suc (Suc n))" | 
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 125 | |
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 126 | code_pred even . | 
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 127 | |
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 128 | lemma "even (n - 2) ==> even n" | 
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 129 | quickcheck[narrowing, expect = counterexample] | 
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 130 | oops | 
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 131 | |
| 
54c4e12bb5e0
adding an example with inductive predicates to quickcheck narrowing examples
 bulwahn parents: 
43356diff
changeset | 132 | |
| 63167 | 133 | subsection \<open>AVL Trees\<close> | 
| 41907 | 134 | |
| 58310 | 135 | datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat | 
| 41907 | 136 | |
| 137 | primrec set_of :: "'a tree \<Rightarrow> 'a set" | |
| 138 | where | |
| 139 | "set_of ET = {}" |
 | |
| 140 | "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)" | |
| 141 | ||
| 142 | primrec height :: "'a tree \<Rightarrow> nat" | |
| 143 | where | |
| 144 | "height ET = 0" | | |
| 145 | "height (MKT x l r h) = max (height l) (height r) + 1" | |
| 146 | ||
| 147 | primrec avl :: "'a tree \<Rightarrow> bool" | |
| 148 | where | |
| 149 | "avl ET = True" | | |
| 150 | "avl (MKT x l r h) = | |
| 151 | ((height l = height r \<or> height l = 1+height r \<or> height r = 1+height l) \<and> | |
| 152 | h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)" | |
| 153 | ||
| 154 | primrec is_ord :: "('a::order) tree \<Rightarrow> bool"
 | |
| 155 | where | |
| 156 | "is_ord ET = True" | | |
| 157 | "is_ord (MKT n l r h) = | |
| 158 | ((\<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)" | |
| 159 | ||
| 160 | primrec is_in :: "('a::order) \<Rightarrow> 'a tree \<Rightarrow> bool"
 | |
| 161 | where | |
| 162 | "is_in k ET = False" | | |
| 163 | "is_in k (MKT n l r h) = (if k = n then True else | |
| 164 | if k < n then (is_in k l) | |
| 165 | else (is_in k r))" | |
| 166 | ||
| 167 | primrec ht :: "'a tree \<Rightarrow> nat" | |
| 168 | where | |
| 169 | "ht ET = 0" | | |
| 170 | "ht (MKT x l r h) = h" | |
| 171 | ||
| 172 | definition | |
| 173 | mkt :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where | |
| 174 | "mkt x l r = MKT x l r (max (ht l) (ht r) + 1)" | |
| 175 | ||
| 176 | (* replaced MKT lrn lrl lrr by MKT lrr lrl *) | |
| 177 | fun l_bal where | |
| 178 | "l_bal(n, MKT ln ll lr h, r) = | |
| 179 | (if ht ll < ht lr | |
| 67414 | 180 | then case lr of ET \<Rightarrow> ET \<comment> \<open>impossible\<close> | 
| 41907 | 181 | | MKT lrn lrr lrl lrh \<Rightarrow> | 
| 182 | mkt lrn (mkt ln ll lrl) (mkt n lrr r) | |
| 183 | else mkt ln ll (mkt n lr r))" | |
| 184 | ||
| 185 | fun r_bal where | |
| 186 | "r_bal(n, l, MKT rn rl rr h) = | |
| 187 | (if ht rl > ht rr | |
| 67414 | 188 | then case rl of ET \<Rightarrow> ET \<comment> \<open>impossible\<close> | 
| 41907 | 189 | | MKT rln rll rlr h \<Rightarrow> mkt rln (mkt n l rll) (mkt rn rlr rr) | 
| 190 | else mkt rn (mkt n l rl) rr)" | |
| 191 | ||
| 192 | primrec insrt :: "'a::order \<Rightarrow> 'a tree \<Rightarrow> 'a tree" | |
| 193 | where | |
| 194 | "insrt x ET = MKT x ET ET 1" | | |
| 195 | "insrt x (MKT n l r h) = | |
| 196 | (if x=n | |
| 197 | then MKT n l r h | |
| 198 | else if x<n | |
| 199 | then let l' = insrt x l; hl' = ht l'; hr = ht r | |
| 200 | in if hl' = 2+hr then l_bal(n,l',r) | |
| 201 | else MKT n l' r (1 + max hl' hr) | |
| 202 | else let r' = insrt x r; hl = ht l; hr' = ht r' | |
| 203 | in if hr' = 2+hl then r_bal(n,l,r') | |
| 204 | else MKT n l r' (1 + max hl hr'))" | |
| 205 | ||
| 206 | ||
| 63167 | 207 | subsubsection \<open>Necessary setup for code generation\<close> | 
| 41907 | 208 | |
| 209 | primrec set_of' | |
| 210 | where | |
| 211 | "set_of' ET = []" | |
| 212 | | "set_of' (MKT n l r h) = n # (set_of' l @ set_of' r)" | |
| 213 | ||
| 214 | lemma set_of': | |
| 215 | "set (set_of' t) = set_of t" | |
| 216 | by (induct t) auto | |
| 217 | ||
| 218 | lemma is_ord_mkt: | |
| 67613 | 219 | "is_ord (MKT n l r h) = ((\<forall>n' \<in> set (set_of' l). n' < n) \<and> (\<forall>n' \<in> set (set_of' r). n < n') \<and> is_ord l \<and> is_ord r)" | 
| 41907 | 220 | by (simp add: set_of') | 
| 221 | ||
| 222 | declare is_ord.simps(1)[code] is_ord_mkt[code] | |
| 223 | ||
| 69597 | 224 | subsubsection \<open>Invalid Lemma due to typo in \<^const>\<open>l_bal\<close>\<close> | 
| 41907 | 225 | |
| 226 | lemma is_ord_l_bal: | |
| 227 | "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))" | |
| 45773 
7f2366dc8c0f
increasing quickcheck's timeout in the example theory to avoid failures on the testing infrastructure
 bulwahn parents: 
45658diff
changeset | 228 | quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, expect = counterexample] | 
| 41907 | 229 | oops | 
| 230 | ||
| 63167 | 231 | subsection \<open>Examples with hd and last\<close> | 
| 45790 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 232 | |
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 233 | lemma | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 234 | "hd (xs @ ys) = hd ys" | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 235 | quickcheck[narrowing, expect = counterexample] | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 236 | oops | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 237 | |
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 238 | lemma | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 239 | "last(xs @ ys) = last xs" | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 240 | quickcheck[narrowing, expect = counterexample] | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 241 | oops | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 242 | |
| 63167 | 243 | subsection \<open>Examples with underspecified/partial functions\<close> | 
| 45790 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 244 | |
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 245 | lemma | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 246 | "xs = [] ==> hd xs \<noteq> x" | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 247 | quickcheck[narrowing, expect = no_counterexample] | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 248 | oops | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 249 | |
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 250 | lemma | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 251 | "xs = [] ==> hd xs = x" | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 252 | quickcheck[narrowing, expect = no_counterexample] | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 253 | oops | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 254 | |
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 255 | lemma "xs = [] ==> hd xs = x ==> x = y" | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 256 | quickcheck[narrowing, expect = no_counterexample] | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 257 | oops | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 258 | |
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 259 | lemma | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 260 | "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)" | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 261 | quickcheck[narrowing, expect = no_counterexample] | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 262 | oops | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 263 | |
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 264 | lemma | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 265 | "hd (map f xs) = f (hd xs)" | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 266 | quickcheck[narrowing, expect = no_counterexample] | 
| 
3355c27c93a4
adding examples for quickcheck narrowing about partial functions
 bulwahn parents: 
45773diff
changeset | 267 | oops | 
| 41907 | 268 | |
| 269 | end |