| author | berghofe | 
| Thu, 03 Mar 2011 11:36:10 +0100 | |
| changeset 41880 | 01f3e2df796c | 
| parent 41278 | 8e1cde88aae6 | 
| child 42208 | 02513eb26eb7 | 
| permissions | -rw-r--r-- | 
| 33197 | 1 | (* Title: HOL/Nitpick_Examples/Refute_Nits.thy | 
| 2 | Author: Jasmin Blanchette, TU Muenchen | |
| 35076 
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
 blanchet parents: 
34083diff
changeset | 3 | Copyright 2009, 2010 | 
| 33197 | 4 | |
| 5 | Refute examples adapted to Nitpick. | |
| 6 | *) | |
| 7 | ||
| 8 | header {* Refute Examples Adapted to Nitpick *}
 | |
| 9 | ||
| 10 | theory Refute_Nits | |
| 11 | imports Main | |
| 12 | begin | |
| 13 | ||
| 41278 
8e1cde88aae6
added a timestamp to Nitpick in verbose mode for debugging purposes;
 blanchet parents: 
40341diff
changeset | 14 | nitpick_params [verbose, card = 1\<midarrow>6, max_potential = 0, | 
| 40341 
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
 blanchet parents: 
38185diff
changeset | 15 | sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60] | 
| 34083 
652719832159
make Nitpick tests more robust by specifying SAT solver, singlethreading (in Kodkod, not in Isabelle), and higher time limits
 blanchet parents: 
33735diff
changeset | 16 | |
| 33197 | 17 | lemma "P \<and> Q" | 
| 18 | apply (rule conjI) | |
| 19 | nitpick [expect = genuine] 1 | |
| 20 | nitpick [expect = genuine] 2 | |
| 21 | nitpick [expect = genuine] | |
| 22 | nitpick [card = 5, expect = genuine] | |
| 33735 
0c0e7b2ecf2e
use SAT solver that's available everywhere for this example
 blanchet parents: 
33197diff
changeset | 23 | nitpick [sat_solver = SAT4J, expect = genuine] 2 | 
| 33197 | 24 | oops | 
| 25 | ||
| 26 | subsection {* Examples and Test Cases *}
 | |
| 27 | ||
| 28 | subsubsection {* Propositional logic *}
 | |
| 29 | ||
| 30 | lemma "True" | |
| 31 | nitpick [expect = none] | |
| 32 | apply auto | |
| 33 | done | |
| 34 | ||
| 35 | lemma "False" | |
| 36 | nitpick [expect = genuine] | |
| 37 | oops | |
| 38 | ||
| 39 | lemma "P" | |
| 40 | nitpick [expect = genuine] | |
| 41 | oops | |
| 42 | ||
| 43 | lemma "\<not> P" | |
| 44 | nitpick [expect = genuine] | |
| 45 | oops | |
| 46 | ||
| 47 | lemma "P \<and> Q" | |
| 48 | nitpick [expect = genuine] | |
| 49 | oops | |
| 50 | ||
| 51 | lemma "P \<or> Q" | |
| 52 | nitpick [expect = genuine] | |
| 53 | oops | |
| 54 | ||
| 55 | lemma "P \<longrightarrow> Q" | |
| 56 | nitpick [expect = genuine] | |
| 57 | oops | |
| 58 | ||
| 59 | lemma "(P\<Colon>bool) = Q" | |
| 60 | nitpick [expect = genuine] | |
| 61 | oops | |
| 62 | ||
| 63 | lemma "(P \<or> Q) \<longrightarrow> (P \<and> Q)" | |
| 64 | nitpick [expect = genuine] | |
| 65 | oops | |
| 66 | ||
| 67 | subsubsection {* Predicate logic *}
 | |
| 68 | ||
| 69 | lemma "P x y z" | |
| 70 | nitpick [expect = genuine] | |
| 71 | oops | |
| 72 | ||
| 73 | lemma "P x y \<longrightarrow> P y x" | |
| 74 | nitpick [expect = genuine] | |
| 75 | oops | |
| 76 | ||
| 77 | lemma "P (f (f x)) \<longrightarrow> P x \<longrightarrow> P (f x)" | |
| 78 | nitpick [expect = genuine] | |
| 79 | oops | |
| 80 | ||
| 81 | subsubsection {* Equality *}
 | |
| 82 | ||
| 83 | lemma "P = True" | |
| 84 | nitpick [expect = genuine] | |
| 85 | oops | |
| 86 | ||
| 87 | lemma "P = False" | |
| 88 | nitpick [expect = genuine] | |
| 89 | oops | |
| 90 | ||
| 91 | lemma "x = y" | |
| 92 | nitpick [expect = genuine] | |
| 93 | oops | |
| 94 | ||
| 95 | lemma "f x = g x" | |
| 96 | nitpick [expect = genuine] | |
| 97 | oops | |
| 98 | ||
| 99 | lemma "(f\<Colon>'a\<Rightarrow>'b) = g" | |
| 100 | nitpick [expect = genuine] | |
| 101 | oops | |
| 102 | ||
| 103 | lemma "(f\<Colon>('d\<Rightarrow>'d)\<Rightarrow>('c\<Rightarrow>'d)) = g"
 | |
| 104 | nitpick [expect = genuine] | |
| 105 | oops | |
| 106 | ||
| 107 | lemma "distinct [a, b]" | |
| 108 | nitpick [expect = genuine] | |
| 109 | apply simp | |
| 110 | nitpick [expect = genuine] | |
| 111 | oops | |
| 112 | ||
| 113 | subsubsection {* First-Order Logic *}
 | |
| 114 | ||
| 115 | lemma "\<exists>x. P x" | |
| 116 | nitpick [expect = genuine] | |
| 117 | oops | |
| 118 | ||
| 119 | lemma "\<forall>x. P x" | |
| 120 | nitpick [expect = genuine] | |
| 121 | oops | |
| 122 | ||
| 123 | lemma "\<exists>!x. P x" | |
| 124 | nitpick [expect = genuine] | |
| 125 | oops | |
| 126 | ||
| 127 | lemma "Ex P" | |
| 128 | nitpick [expect = genuine] | |
| 129 | oops | |
| 130 | ||
| 131 | lemma "All P" | |
| 132 | nitpick [expect = genuine] | |
| 133 | oops | |
| 134 | ||
| 135 | lemma "Ex1 P" | |
| 136 | nitpick [expect = genuine] | |
| 137 | oops | |
| 138 | ||
| 139 | lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)" | |
| 140 | nitpick [expect = genuine] | |
| 141 | oops | |
| 142 | ||
| 143 | lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>y. \<forall>x. P x y)" | |
| 144 | nitpick [expect = genuine] | |
| 145 | oops | |
| 146 | ||
| 147 | lemma "(\<exists>x. P x) \<longrightarrow> (\<exists>!x. P x)" | |
| 148 | nitpick [expect = genuine] | |
| 149 | oops | |
| 150 | ||
| 151 | text {* A true statement (also testing names of free and bound variables being identical) *}
 | |
| 152 | ||
| 153 | lemma "(\<forall>x y. P x y \<longrightarrow> P y x) \<longrightarrow> (\<forall>x. P x y) \<longrightarrow> P y x" | |
| 154 | nitpick [expect = none] | |
| 155 | apply fast | |
| 156 | done | |
| 157 | ||
| 158 | text {* "A type has at most 4 elements." *}
 | |
| 159 | ||
| 160 | lemma "\<not> distinct [a, b, c, d, e]" | |
| 161 | nitpick [expect = genuine] | |
| 162 | apply simp | |
| 163 | nitpick [expect = genuine] | |
| 164 | oops | |
| 165 | ||
| 166 | lemma "distinct [a, b, c, d]" | |
| 167 | nitpick [expect = genuine] | |
| 168 | apply simp | |
| 169 | nitpick [expect = genuine] | |
| 170 | oops | |
| 171 | ||
| 172 | text {* "Every reflexive and symmetric relation is transitive." *}
 | |
| 173 | ||
| 174 | lemma "\<lbrakk>\<forall>x. P x x; \<forall>x y. P x y \<longrightarrow> P y x\<rbrakk> \<Longrightarrow> P x y \<longrightarrow> P y z \<longrightarrow> P x z" | |
| 175 | nitpick [expect = genuine] | |
| 176 | oops | |
| 177 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 178 | text {* The ``Drinker's theorem'' *}
 | 
| 33197 | 179 | |
| 180 | lemma "\<exists>x. f x = g x \<longrightarrow> f = g" | |
| 181 | nitpick [expect = none] | |
| 182 | apply (auto simp add: ext) | |
| 183 | done | |
| 184 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 185 | text {* And an incorrect version of it *}
 | 
| 33197 | 186 | |
| 187 | lemma "(\<exists>x. f x = g x) \<longrightarrow> f = g" | |
| 188 | nitpick [expect = genuine] | |
| 189 | oops | |
| 190 | ||
| 191 | text {* "Every function has a fixed point." *}
 | |
| 192 | ||
| 193 | lemma "\<exists>x. f x = x" | |
| 194 | nitpick [expect = genuine] | |
| 195 | oops | |
| 196 | ||
| 197 | text {* "Function composition is commutative." *}
 | |
| 198 | ||
| 199 | lemma "f (g x) = g (f x)" | |
| 200 | nitpick [expect = genuine] | |
| 201 | oops | |
| 202 | ||
| 203 | text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *}
 | |
| 204 | ||
| 205 | lemma "((P\<Colon>('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
 | |
| 206 | nitpick [expect = genuine] | |
| 207 | oops | |
| 208 | ||
| 209 | subsubsection {* Higher-Order Logic *}
 | |
| 210 | ||
| 211 | lemma "\<exists>P. P" | |
| 212 | nitpick [expect = none] | |
| 213 | apply auto | |
| 214 | done | |
| 215 | ||
| 216 | lemma "\<forall>P. P" | |
| 217 | nitpick [expect = genuine] | |
| 218 | oops | |
| 219 | ||
| 220 | lemma "\<exists>!P. P" | |
| 221 | nitpick [expect = none] | |
| 222 | apply auto | |
| 223 | done | |
| 224 | ||
| 225 | lemma "\<exists>!P. P x" | |
| 226 | nitpick [expect = genuine] | |
| 227 | oops | |
| 228 | ||
| 229 | lemma "P Q \<or> Q x" | |
| 230 | nitpick [expect = genuine] | |
| 231 | oops | |
| 232 | ||
| 233 | lemma "x \<noteq> All" | |
| 234 | nitpick [expect = genuine] | |
| 235 | oops | |
| 236 | ||
| 237 | lemma "x \<noteq> Ex" | |
| 238 | nitpick [expect = genuine] | |
| 239 | oops | |
| 240 | ||
| 241 | lemma "x \<noteq> Ex1" | |
| 242 | nitpick [expect = genuine] | |
| 243 | oops | |
| 244 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 245 | text {* ``The transitive closure of an arbitrary relation is non-empty.'' *}
 | 
| 33197 | 246 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35342diff
changeset | 247 | definition "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 33197 | 248 | "trans P \<equiv> (ALL x y z. P x y \<longrightarrow> P y z \<longrightarrow> P x z)" | 
| 35421 | 249 | |
| 250 | definition | |
| 251 | "subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
 | |
| 33197 | 252 | "subset P Q \<equiv> (ALL x y. P x y \<longrightarrow> Q x y)" | 
| 35421 | 253 | |
| 254 | definition | |
| 255 | "trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
 | |
| 33197 | 256 | "trans_closure P Q \<equiv> (subset Q P) \<and> (trans P) \<and> (ALL R. subset Q R \<longrightarrow> trans R \<longrightarrow> subset P R)" | 
| 257 | ||
| 258 | lemma "trans_closure T P \<longrightarrow> (\<exists>x y. T x y)" | |
| 259 | nitpick [expect = genuine] | |
| 260 | oops | |
| 261 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 262 | text {* ``The union of transitive closures is equal to the transitive closure of unions.'' *}
 | 
| 33197 | 263 | |
| 264 | lemma "(\<forall>x y. (P x y \<or> R x y) \<longrightarrow> T x y) \<longrightarrow> trans T \<longrightarrow> (\<forall>Q. (\<forall>x y. (P x y \<or> R x y) \<longrightarrow> Q x y) \<longrightarrow> trans Q \<longrightarrow> subset T Q) | |
| 265 | \<longrightarrow> trans_closure TP P | |
| 266 | \<longrightarrow> trans_closure TR R | |
| 267 | \<longrightarrow> (T x y = (TP x y \<or> TR x y))" | |
| 268 | nitpick [expect = genuine] | |
| 269 | oops | |
| 270 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 271 | text {* ``Every surjective function is invertible.'' *}
 | 
| 33197 | 272 | |
| 273 | lemma "(\<forall>y. \<exists>x. y = f x) \<longrightarrow> (\<exists>g. \<forall>x. g (f x) = x)" | |
| 274 | nitpick [expect = genuine] | |
| 275 | oops | |
| 276 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 277 | text {* ``Every invertible function is surjective.'' *}
 | 
| 33197 | 278 | |
| 279 | lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)" | |
| 280 | nitpick [expect = genuine] | |
| 281 | oops | |
| 282 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 283 | text {* ``Every point is a fixed point of some function.'' *}
 | 
| 33197 | 284 | |
| 285 | lemma "\<exists>f. f x = x" | |
| 286 | nitpick [card = 1\<midarrow>7, expect = none] | |
| 287 | apply (rule_tac x = "\<lambda>x. x" in exI) | |
| 288 | apply simp | |
| 289 | done | |
| 290 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 291 | text {* Axiom of Choice: first an incorrect version *}
 | 
| 33197 | 292 | |
| 293 | lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>!f. \<forall>x. P x (f x))" | |
| 294 | nitpick [expect = genuine] | |
| 295 | oops | |
| 296 | ||
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 297 | text {* And now two correct ones *}
 | 
| 33197 | 298 | |
| 299 | lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>f. \<forall>x. P x (f x))" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 300 | nitpick [card = 1-4, expect = none] | 
| 33197 | 301 | apply (simp add: choice) | 
| 302 | done | |
| 303 | ||
| 304 | lemma "(\<forall>x. \<exists>!y. P x y) \<longrightarrow> (\<exists>!f. \<forall>x. P x (f x))" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 305 | nitpick [card = 1-3, expect = none] | 
| 33197 | 306 | apply auto | 
| 307 | apply (simp add: ex1_implies_ex choice) | |
| 308 | apply (fast intro: ext) | |
| 309 | done | |
| 310 | ||
| 311 | subsubsection {* Metalogic *}
 | |
| 312 | ||
| 313 | lemma "\<And>x. P x" | |
| 314 | nitpick [expect = genuine] | |
| 315 | oops | |
| 316 | ||
| 317 | lemma "f x \<equiv> g x" | |
| 318 | nitpick [expect = genuine] | |
| 319 | oops | |
| 320 | ||
| 321 | lemma "P \<Longrightarrow> Q" | |
| 322 | nitpick [expect = genuine] | |
| 323 | oops | |
| 324 | ||
| 325 | lemma "\<lbrakk>P; Q; R\<rbrakk> \<Longrightarrow> S" | |
| 326 | nitpick [expect = genuine] | |
| 327 | oops | |
| 328 | ||
| 329 | lemma "(x \<equiv> all) \<Longrightarrow> False" | |
| 330 | nitpick [expect = genuine] | |
| 331 | oops | |
| 332 | ||
| 333 | lemma "(x \<equiv> (op \<equiv>)) \<Longrightarrow> False" | |
| 334 | nitpick [expect = genuine] | |
| 335 | oops | |
| 336 | ||
| 337 | lemma "(x \<equiv> (op \<Longrightarrow>)) \<Longrightarrow> False" | |
| 338 | nitpick [expect = genuine] | |
| 339 | oops | |
| 340 | ||
| 341 | subsubsection {* Schematic Variables *}
 | |
| 342 | ||
| 36319 | 343 | schematic_lemma "?P" | 
| 33197 | 344 | nitpick [expect = none] | 
| 345 | apply auto | |
| 346 | done | |
| 347 | ||
| 36319 | 348 | schematic_lemma "x = ?y" | 
| 33197 | 349 | nitpick [expect = none] | 
| 350 | apply auto | |
| 351 | done | |
| 352 | ||
| 353 | subsubsection {* Abstractions *}
 | |
| 354 | ||
| 355 | lemma "(\<lambda>x. x) = (\<lambda>x. y)" | |
| 356 | nitpick [expect = genuine] | |
| 357 | oops | |
| 358 | ||
| 359 | lemma "(\<lambda>f. f x) = (\<lambda>f. True)" | |
| 360 | nitpick [expect = genuine] | |
| 361 | oops | |
| 362 | ||
| 363 | lemma "(\<lambda>x. x) = (\<lambda>y. y)" | |
| 364 | nitpick [expect = none] | |
| 365 | apply simp | |
| 366 | done | |
| 367 | ||
| 368 | subsubsection {* Sets *}
 | |
| 369 | ||
| 370 | lemma "P (A\<Colon>'a set)" | |
| 371 | nitpick [expect = genuine] | |
| 372 | oops | |
| 373 | ||
| 374 | lemma "P (A\<Colon>'a set set)" | |
| 375 | nitpick [expect = genuine] | |
| 376 | oops | |
| 377 | ||
| 378 | lemma "{x. P x} = {y. P y}"
 | |
| 379 | nitpick [expect = none] | |
| 380 | apply simp | |
| 381 | done | |
| 382 | ||
| 383 | lemma "x \<in> {x. P x}"
 | |
| 384 | nitpick [expect = genuine] | |
| 385 | oops | |
| 386 | ||
| 387 | lemma "P (op \<in>)" | |
| 388 | nitpick [expect = genuine] | |
| 389 | oops | |
| 390 | ||
| 391 | lemma "P (op \<in> x)" | |
| 392 | nitpick [expect = genuine] | |
| 393 | oops | |
| 394 | ||
| 395 | lemma "P Collect" | |
| 396 | nitpick [expect = genuine] | |
| 397 | oops | |
| 398 | ||
| 399 | lemma "A Un B = A Int B" | |
| 400 | nitpick [expect = genuine] | |
| 401 | oops | |
| 402 | ||
| 403 | lemma "(A Int B) Un C = (A Un C) Int B" | |
| 404 | nitpick [expect = genuine] | |
| 405 | oops | |
| 406 | ||
| 407 | lemma "Ball A P \<longrightarrow> Bex A P" | |
| 408 | nitpick [expect = genuine] | |
| 409 | oops | |
| 410 | ||
| 411 | subsubsection {* @{const undefined} *}
 | |
| 412 | ||
| 413 | lemma "undefined" | |
| 414 | nitpick [expect = genuine] | |
| 415 | oops | |
| 416 | ||
| 417 | lemma "P undefined" | |
| 418 | nitpick [expect = genuine] | |
| 419 | oops | |
| 420 | ||
| 421 | lemma "undefined x" | |
| 422 | nitpick [expect = genuine] | |
| 423 | oops | |
| 424 | ||
| 425 | lemma "undefined undefined" | |
| 426 | nitpick [expect = genuine] | |
| 427 | oops | |
| 428 | ||
| 429 | subsubsection {* @{const The} *}
 | |
| 430 | ||
| 431 | lemma "The P" | |
| 432 | nitpick [expect = genuine] | |
| 433 | oops | |
| 434 | ||
| 435 | lemma "P The" | |
| 436 | nitpick [expect = genuine] | |
| 437 | oops | |
| 438 | ||
| 439 | lemma "P (The P)" | |
| 440 | nitpick [expect = genuine] | |
| 441 | oops | |
| 442 | ||
| 443 | lemma "(THE x. x=y) = z" | |
| 444 | nitpick [expect = genuine] | |
| 445 | oops | |
| 446 | ||
| 447 | lemma "Ex P \<longrightarrow> P (The P)" | |
| 448 | nitpick [expect = genuine] | |
| 449 | oops | |
| 450 | ||
| 451 | subsubsection {* @{const Eps} *}
 | |
| 452 | ||
| 453 | lemma "Eps P" | |
| 454 | nitpick [expect = genuine] | |
| 455 | oops | |
| 456 | ||
| 457 | lemma "P Eps" | |
| 458 | nitpick [expect = genuine] | |
| 459 | oops | |
| 460 | ||
| 461 | lemma "P (Eps P)" | |
| 462 | nitpick [expect = genuine] | |
| 463 | oops | |
| 464 | ||
| 465 | lemma "(SOME x. x=y) = z" | |
| 466 | nitpick [expect = genuine] | |
| 467 | oops | |
| 468 | ||
| 469 | lemma "Ex P \<longrightarrow> P (Eps P)" | |
| 470 | nitpick [expect = none] | |
| 471 | apply (auto simp add: someI) | |
| 472 | done | |
| 473 | ||
| 474 | subsubsection {* Operations on Natural Numbers *}
 | |
| 475 | ||
| 476 | lemma "(x\<Colon>nat) + y = 0" | |
| 477 | nitpick [expect = genuine] | |
| 478 | oops | |
| 479 | ||
| 480 | lemma "(x\<Colon>nat) = x + x" | |
| 481 | nitpick [expect = genuine] | |
| 482 | oops | |
| 483 | ||
| 484 | lemma "(x\<Colon>nat) - y + y = x" | |
| 485 | nitpick [expect = genuine] | |
| 486 | oops | |
| 487 | ||
| 488 | lemma "(x\<Colon>nat) = x * x" | |
| 489 | nitpick [expect = genuine] | |
| 490 | oops | |
| 491 | ||
| 492 | lemma "(x\<Colon>nat) < x + y" | |
| 493 | nitpick [card = 1, expect = genuine] | |
| 494 | nitpick [card = 2-5, expect = genuine] | |
| 495 | oops | |
| 496 | ||
| 497 | text {* \<times> *}
 | |
| 498 | ||
| 499 | lemma "P (x\<Colon>'a\<times>'b)" | |
| 500 | nitpick [expect = genuine] | |
| 501 | oops | |
| 502 | ||
| 503 | lemma "\<forall>x\<Colon>'a\<times>'b. P x" | |
| 504 | nitpick [expect = genuine] | |
| 505 | oops | |
| 506 | ||
| 507 | lemma "P (x, y)" | |
| 508 | nitpick [expect = genuine] | |
| 509 | oops | |
| 510 | ||
| 511 | lemma "P (fst x)" | |
| 512 | nitpick [expect = genuine] | |
| 513 | oops | |
| 514 | ||
| 515 | lemma "P (snd x)" | |
| 516 | nitpick [expect = genuine] | |
| 517 | oops | |
| 518 | ||
| 519 | lemma "P Pair" | |
| 520 | nitpick [expect = genuine] | |
| 521 | oops | |
| 522 | ||
| 523 | lemma "prod_rec f p = f (fst p) (snd p)" | |
| 524 | nitpick [expect = none] | |
| 525 | by (case_tac p) auto | |
| 526 | ||
| 527 | lemma "prod_rec f (a, b) = f a b" | |
| 528 | nitpick [expect = none] | |
| 529 | by auto | |
| 530 | ||
| 531 | lemma "P (prod_rec f x)" | |
| 532 | nitpick [expect = genuine] | |
| 533 | oops | |
| 534 | ||
| 535 | lemma "P (case x of Pair a b \<Rightarrow> f a b)" | |
| 536 | nitpick [expect = genuine] | |
| 537 | oops | |
| 538 | ||
| 539 | subsubsection {* Subtypes (typedef), typedecl *}
 | |
| 540 | ||
| 541 | text {* A completely unspecified non-empty subset of @{typ "'a"}: *}
 | |
| 542 | ||
| 543 | typedef 'a myTdef = "insert (undefined\<Colon>'a) (undefined\<Colon>'a set)" | |
| 544 | by auto | |
| 545 | ||
| 546 | lemma "(x\<Colon>'a myTdef) = y" | |
| 547 | nitpick [expect = genuine] | |
| 548 | oops | |
| 549 | ||
| 550 | typedecl myTdecl | |
| 551 | ||
| 552 | typedef 'a T_bij = "{(f\<Colon>'a\<Rightarrow>'a). \<forall>y. \<exists>!x. f x = y}"
 | |
| 553 | by auto | |
| 554 | ||
| 555 | lemma "P (f\<Colon>(myTdecl myTdef) T_bij)" | |
| 556 | nitpick [expect = genuine] | |
| 557 | oops | |
| 558 | ||
| 559 | subsubsection {* Inductive Datatypes *}
 | |
| 560 | ||
| 561 | text {* unit *}
 | |
| 562 | ||
| 563 | lemma "P (x\<Colon>unit)" | |
| 564 | nitpick [expect = genuine] | |
| 565 | oops | |
| 566 | ||
| 567 | lemma "\<forall>x\<Colon>unit. P x" | |
| 568 | nitpick [expect = genuine] | |
| 569 | oops | |
| 570 | ||
| 571 | lemma "P ()" | |
| 572 | nitpick [expect = genuine] | |
| 573 | oops | |
| 574 | ||
| 575 | lemma "unit_rec u x = u" | |
| 576 | nitpick [expect = none] | |
| 577 | apply simp | |
| 578 | done | |
| 579 | ||
| 580 | lemma "P (unit_rec u x)" | |
| 581 | nitpick [expect = genuine] | |
| 582 | oops | |
| 583 | ||
| 584 | lemma "P (case x of () \<Rightarrow> u)" | |
| 585 | nitpick [expect = genuine] | |
| 586 | oops | |
| 587 | ||
| 588 | text {* option *}
 | |
| 589 | ||
| 590 | lemma "P (x\<Colon>'a option)" | |
| 591 | nitpick [expect = genuine] | |
| 592 | oops | |
| 593 | ||
| 594 | lemma "\<forall>x\<Colon>'a option. P x" | |
| 595 | nitpick [expect = genuine] | |
| 596 | oops | |
| 597 | ||
| 598 | lemma "P None" | |
| 599 | nitpick [expect = genuine] | |
| 600 | oops | |
| 601 | ||
| 602 | lemma "P (Some x)" | |
| 603 | nitpick [expect = genuine] | |
| 604 | oops | |
| 605 | ||
| 606 | lemma "option_rec n s None = n" | |
| 607 | nitpick [expect = none] | |
| 608 | apply simp | |
| 609 | done | |
| 610 | ||
| 611 | lemma "option_rec n s (Some x) = s x" | |
| 612 | nitpick [expect = none] | |
| 613 | apply simp | |
| 614 | done | |
| 615 | ||
| 616 | lemma "P (option_rec n s x)" | |
| 617 | nitpick [expect = genuine] | |
| 618 | oops | |
| 619 | ||
| 620 | lemma "P (case x of None \<Rightarrow> n | Some u \<Rightarrow> s u)" | |
| 621 | nitpick [expect = genuine] | |
| 622 | oops | |
| 623 | ||
| 624 | text {* + *}
 | |
| 625 | ||
| 626 | lemma "P (x\<Colon>'a+'b)" | |
| 627 | nitpick [expect = genuine] | |
| 628 | oops | |
| 629 | ||
| 630 | lemma "\<forall>x\<Colon>'a+'b. P x" | |
| 631 | nitpick [expect = genuine] | |
| 632 | oops | |
| 633 | ||
| 634 | lemma "P (Inl x)" | |
| 635 | nitpick [expect = genuine] | |
| 636 | oops | |
| 637 | ||
| 638 | lemma "P (Inr x)" | |
| 639 | nitpick [expect = genuine] | |
| 640 | oops | |
| 641 | ||
| 642 | lemma "P Inl" | |
| 643 | nitpick [expect = genuine] | |
| 644 | oops | |
| 645 | ||
| 646 | lemma "sum_rec l r (Inl x) = l x" | |
| 647 | nitpick [expect = none] | |
| 648 | apply simp | |
| 649 | done | |
| 650 | ||
| 651 | lemma "sum_rec l r (Inr x) = r x" | |
| 652 | nitpick [expect = none] | |
| 653 | apply simp | |
| 654 | done | |
| 655 | ||
| 656 | lemma "P (sum_rec l r x)" | |
| 657 | nitpick [expect = genuine] | |
| 658 | oops | |
| 659 | ||
| 660 | lemma "P (case x of Inl a \<Rightarrow> l a | Inr b \<Rightarrow> r b)" | |
| 661 | nitpick [expect = genuine] | |
| 662 | oops | |
| 663 | ||
| 664 | text {* Non-recursive datatypes *}
 | |
| 665 | ||
| 666 | datatype T1 = A | B | |
| 667 | ||
| 668 | lemma "P (x\<Colon>T1)" | |
| 669 | nitpick [expect = genuine] | |
| 670 | oops | |
| 671 | ||
| 672 | lemma "\<forall>x\<Colon>T1. P x" | |
| 673 | nitpick [expect = genuine] | |
| 674 | oops | |
| 675 | ||
| 676 | lemma "P A" | |
| 677 | nitpick [expect = genuine] | |
| 678 | oops | |
| 679 | ||
| 680 | lemma "P B" | |
| 681 | nitpick [expect = genuine] | |
| 682 | oops | |
| 683 | ||
| 684 | lemma "T1_rec a b A = a" | |
| 685 | nitpick [expect = none] | |
| 686 | apply simp | |
| 687 | done | |
| 688 | ||
| 689 | lemma "T1_rec a b B = b" | |
| 690 | nitpick [expect = none] | |
| 691 | apply simp | |
| 692 | done | |
| 693 | ||
| 694 | lemma "P (T1_rec a b x)" | |
| 695 | nitpick [expect = genuine] | |
| 696 | oops | |
| 697 | ||
| 698 | lemma "P (case x of A \<Rightarrow> a | B \<Rightarrow> b)" | |
| 699 | nitpick [expect = genuine] | |
| 700 | oops | |
| 701 | ||
| 702 | datatype 'a T2 = C T1 | D 'a | |
| 703 | ||
| 704 | lemma "P (x\<Colon>'a T2)" | |
| 705 | nitpick [expect = genuine] | |
| 706 | oops | |
| 707 | ||
| 708 | lemma "\<forall>x\<Colon>'a T2. P x" | |
| 709 | nitpick [expect = genuine] | |
| 710 | oops | |
| 711 | ||
| 712 | lemma "P D" | |
| 713 | nitpick [expect = genuine] | |
| 714 | oops | |
| 715 | ||
| 716 | lemma "T2_rec c d (C x) = c x" | |
| 717 | nitpick [expect = none] | |
| 718 | apply simp | |
| 719 | done | |
| 720 | ||
| 721 | lemma "T2_rec c d (D x) = d x" | |
| 722 | nitpick [expect = none] | |
| 723 | apply simp | |
| 724 | done | |
| 725 | ||
| 726 | lemma "P (T2_rec c d x)" | |
| 727 | nitpick [expect = genuine] | |
| 728 | oops | |
| 729 | ||
| 730 | lemma "P (case x of C u \<Rightarrow> c u | D v \<Rightarrow> d v)" | |
| 731 | nitpick [expect = genuine] | |
| 732 | oops | |
| 733 | ||
| 734 | datatype ('a, 'b) T3 = E "'a \<Rightarrow> 'b"
 | |
| 735 | ||
| 736 | lemma "P (x\<Colon>('a, 'b) T3)"
 | |
| 737 | nitpick [expect = genuine] | |
| 738 | oops | |
| 739 | ||
| 740 | lemma "\<forall>x\<Colon>('a, 'b) T3. P x"
 | |
| 741 | nitpick [expect = genuine] | |
| 742 | oops | |
| 743 | ||
| 744 | lemma "P E" | |
| 745 | nitpick [expect = genuine] | |
| 746 | oops | |
| 747 | ||
| 748 | lemma "T3_rec e (E x) = e x" | |
| 749 | nitpick [card = 1\<midarrow>4, expect = none] | |
| 750 | apply simp | |
| 751 | done | |
| 752 | ||
| 753 | lemma "P (T3_rec e x)" | |
| 754 | nitpick [expect = genuine] | |
| 755 | oops | |
| 756 | ||
| 757 | lemma "P (case x of E f \<Rightarrow> e f)" | |
| 758 | nitpick [expect = genuine] | |
| 759 | oops | |
| 760 | ||
| 761 | text {* Recursive datatypes *}
 | |
| 762 | ||
| 763 | text {* nat *}
 | |
| 764 | ||
| 765 | lemma "P (x\<Colon>nat)" | |
| 766 | nitpick [expect = genuine] | |
| 767 | oops | |
| 768 | ||
| 769 | lemma "\<forall>x\<Colon>nat. P x" | |
| 770 | nitpick [expect = genuine] | |
| 771 | oops | |
| 772 | ||
| 773 | lemma "P (Suc 0)" | |
| 774 | nitpick [expect = genuine] | |
| 775 | oops | |
| 776 | ||
| 777 | lemma "P Suc" | |
| 778 | nitpick [card = 1\<midarrow>7, expect = none] | |
| 779 | oops | |
| 780 | ||
| 781 | lemma "nat_rec zero suc 0 = zero" | |
| 782 | nitpick [expect = none] | |
| 783 | apply simp | |
| 784 | done | |
| 785 | ||
| 786 | lemma "nat_rec zero suc (Suc x) = suc x (nat_rec zero suc x)" | |
| 787 | nitpick [expect = none] | |
| 788 | apply simp | |
| 789 | done | |
| 790 | ||
| 791 | lemma "P (nat_rec zero suc x)" | |
| 792 | nitpick [expect = genuine] | |
| 793 | oops | |
| 794 | ||
| 795 | lemma "P (case x of 0 \<Rightarrow> zero | Suc n \<Rightarrow> suc n)" | |
| 796 | nitpick [expect = genuine] | |
| 797 | oops | |
| 798 | ||
| 799 | text {* 'a list *}
 | |
| 800 | ||
| 801 | lemma "P (xs\<Colon>'a list)" | |
| 802 | nitpick [expect = genuine] | |
| 803 | oops | |
| 804 | ||
| 805 | lemma "\<forall>xs\<Colon>'a list. P xs" | |
| 806 | nitpick [expect = genuine] | |
| 807 | oops | |
| 808 | ||
| 809 | lemma "P [x, y]" | |
| 810 | nitpick [expect = genuine] | |
| 811 | oops | |
| 812 | ||
| 813 | lemma "list_rec nil cons [] = nil" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 814 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 815 | apply simp | 
| 816 | done | |
| 817 | ||
| 818 | lemma "list_rec nil cons (x#xs) = cons x xs (list_rec nil cons xs)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 819 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 820 | apply simp | 
| 821 | done | |
| 822 | ||
| 823 | lemma "P (list_rec nil cons xs)" | |
| 824 | nitpick [expect = genuine] | |
| 825 | oops | |
| 826 | ||
| 827 | lemma "P (case x of Nil \<Rightarrow> nil | Cons a b \<Rightarrow> cons a b)" | |
| 828 | nitpick [expect = genuine] | |
| 829 | oops | |
| 830 | ||
| 831 | lemma "(xs\<Colon>'a list) = ys" | |
| 832 | nitpick [expect = genuine] | |
| 833 | oops | |
| 834 | ||
| 835 | lemma "a # xs = b # xs" | |
| 836 | nitpick [expect = genuine] | |
| 837 | oops | |
| 838 | ||
| 839 | datatype BitList = BitListNil | Bit0 BitList | Bit1 BitList | |
| 840 | ||
| 841 | lemma "P (x\<Colon>BitList)" | |
| 842 | nitpick [expect = genuine] | |
| 843 | oops | |
| 844 | ||
| 845 | lemma "\<forall>x\<Colon>BitList. P x" | |
| 846 | nitpick [expect = genuine] | |
| 847 | oops | |
| 848 | ||
| 849 | lemma "P (Bit0 (Bit1 BitListNil))" | |
| 850 | nitpick [expect = genuine] | |
| 851 | oops | |
| 852 | ||
| 853 | lemma "BitList_rec nil bit0 bit1 BitListNil = nil" | |
| 854 | nitpick [expect = none] | |
| 855 | apply simp | |
| 856 | done | |
| 857 | ||
| 858 | lemma "BitList_rec nil bit0 bit1 (Bit0 xs) = bit0 xs (BitList_rec nil bit0 bit1 xs)" | |
| 859 | nitpick [expect = none] | |
| 860 | apply simp | |
| 861 | done | |
| 862 | ||
| 863 | lemma "BitList_rec nil bit0 bit1 (Bit1 xs) = bit1 xs (BitList_rec nil bit0 bit1 xs)" | |
| 864 | nitpick [expect = none] | |
| 865 | apply simp | |
| 866 | done | |
| 867 | ||
| 868 | lemma "P (BitList_rec nil bit0 bit1 x)" | |
| 869 | nitpick [expect = genuine] | |
| 870 | oops | |
| 871 | ||
| 872 | datatype 'a BinTree = Leaf 'a | Node "'a BinTree" "'a BinTree" | |
| 873 | ||
| 874 | lemma "P (x\<Colon>'a BinTree)" | |
| 875 | nitpick [expect = genuine] | |
| 876 | oops | |
| 877 | ||
| 878 | lemma "\<forall>x\<Colon>'a BinTree. P x" | |
| 879 | nitpick [expect = genuine] | |
| 880 | oops | |
| 881 | ||
| 882 | lemma "P (Node (Leaf x) (Leaf y))" | |
| 883 | nitpick [expect = genuine] | |
| 884 | oops | |
| 885 | ||
| 886 | lemma "BinTree_rec l n (Leaf x) = l x" | |
| 887 | nitpick [expect = none] | |
| 888 | apply simp | |
| 889 | done | |
| 890 | ||
| 891 | lemma "BinTree_rec l n (Node x y) = n x y (BinTree_rec l n x) (BinTree_rec l n y)" | |
| 35087 | 892 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 893 | apply simp | 
| 894 | done | |
| 895 | ||
| 896 | lemma "P (BinTree_rec l n x)" | |
| 897 | nitpick [expect = genuine] | |
| 898 | oops | |
| 899 | ||
| 900 | lemma "P (case x of Leaf a \<Rightarrow> l a | Node a b \<Rightarrow> n a b)" | |
| 901 | nitpick [expect = genuine] | |
| 902 | oops | |
| 903 | ||
| 904 | text {* Mutually recursive datatypes *}
 | |
| 905 | ||
| 906 | datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp" | |
| 907 | and 'a bexp = Equal "'a aexp" "'a aexp" | |
| 908 | ||
| 909 | lemma "P (x\<Colon>'a aexp)" | |
| 910 | nitpick [expect = genuine] | |
| 911 | oops | |
| 912 | ||
| 913 | lemma "\<forall>x\<Colon>'a aexp. P x" | |
| 914 | nitpick [expect = genuine] | |
| 915 | oops | |
| 916 | ||
| 917 | lemma "P (ITE (Equal (Number x) (Number y)) (Number x) (Number y))" | |
| 918 | nitpick [expect = genuine] | |
| 919 | oops | |
| 920 | ||
| 921 | lemma "P (x\<Colon>'a bexp)" | |
| 922 | nitpick [expect = genuine] | |
| 923 | oops | |
| 924 | ||
| 925 | lemma "\<forall>x\<Colon>'a bexp. P x" | |
| 926 | nitpick [expect = genuine] | |
| 927 | oops | |
| 928 | ||
| 929 | lemma "aexp_bexp_rec_1 number ite equal (Number x) = number x" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 930 | nitpick [card = 1\<midarrow>3, expect = none] | 
| 33197 | 931 | apply simp | 
| 932 | done | |
| 933 | ||
| 934 | lemma "aexp_bexp_rec_1 number ite equal (ITE x y z) = ite x y z (aexp_bexp_rec_2 number ite equal x) (aexp_bexp_rec_1 number ite equal y) (aexp_bexp_rec_1 number ite equal z)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 935 | nitpick [card = 1\<midarrow>3, expect = none] | 
| 33197 | 936 | apply simp | 
| 937 | done | |
| 938 | ||
| 939 | lemma "P (aexp_bexp_rec_1 number ite equal x)" | |
| 940 | nitpick [expect = genuine] | |
| 941 | oops | |
| 942 | ||
| 943 | lemma "P (case x of Number a \<Rightarrow> number a | ITE b a1 a2 \<Rightarrow> ite b a1 a2)" | |
| 944 | nitpick [expect = genuine] | |
| 945 | oops | |
| 946 | ||
| 947 | lemma "aexp_bexp_rec_2 number ite equal (Equal x y) = equal x y (aexp_bexp_rec_1 number ite equal x) (aexp_bexp_rec_1 number ite equal y)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 948 | nitpick [card = 1\<midarrow>3, expect = none] | 
| 33197 | 949 | apply simp | 
| 950 | done | |
| 951 | ||
| 952 | lemma "P (aexp_bexp_rec_2 number ite equal x)" | |
| 953 | nitpick [expect = genuine] | |
| 954 | oops | |
| 955 | ||
| 956 | lemma "P (case x of Equal a1 a2 \<Rightarrow> equal a1 a2)" | |
| 957 | nitpick [expect = genuine] | |
| 958 | oops | |
| 959 | ||
| 960 | datatype X = A | B X | C Y | |
| 961 | and Y = D X | E Y | F | |
| 962 | ||
| 963 | lemma "P (x\<Colon>X)" | |
| 964 | nitpick [expect = genuine] | |
| 965 | oops | |
| 966 | ||
| 967 | lemma "P (y\<Colon>Y)" | |
| 968 | nitpick [expect = genuine] | |
| 969 | oops | |
| 970 | ||
| 971 | lemma "P (B (B A))" | |
| 972 | nitpick [expect = genuine] | |
| 973 | oops | |
| 974 | ||
| 975 | lemma "P (B (C F))" | |
| 976 | nitpick [expect = genuine] | |
| 977 | oops | |
| 978 | ||
| 979 | lemma "P (C (D A))" | |
| 980 | nitpick [expect = genuine] | |
| 981 | oops | |
| 982 | ||
| 983 | lemma "P (C (E F))" | |
| 984 | nitpick [expect = genuine] | |
| 985 | oops | |
| 986 | ||
| 987 | lemma "P (D (B A))" | |
| 988 | nitpick [expect = genuine] | |
| 989 | oops | |
| 990 | ||
| 991 | lemma "P (D (C F))" | |
| 992 | nitpick [expect = genuine] | |
| 993 | oops | |
| 994 | ||
| 995 | lemma "P (E (D A))" | |
| 996 | nitpick [expect = genuine] | |
| 997 | oops | |
| 998 | ||
| 999 | lemma "P (E (E F))" | |
| 1000 | nitpick [expect = genuine] | |
| 1001 | oops | |
| 1002 | ||
| 1003 | lemma "P (C (D (C F)))" | |
| 1004 | nitpick [expect = genuine] | |
| 1005 | oops | |
| 1006 | ||
| 1007 | lemma "X_Y_rec_1 a b c d e f A = a" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1008 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 1009 | apply simp | 
| 1010 | done | |
| 1011 | ||
| 1012 | lemma "X_Y_rec_1 a b c d e f (B x) = b x (X_Y_rec_1 a b c d e f x)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1013 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 1014 | apply simp | 
| 1015 | done | |
| 1016 | ||
| 1017 | lemma "X_Y_rec_1 a b c d e f (C y) = c y (X_Y_rec_2 a b c d e f y)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1018 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 1019 | apply simp | 
| 1020 | done | |
| 1021 | ||
| 1022 | lemma "X_Y_rec_2 a b c d e f (D x) = d x (X_Y_rec_1 a b c d e f x)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1023 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 1024 | apply simp | 
| 1025 | done | |
| 1026 | ||
| 1027 | lemma "X_Y_rec_2 a b c d e f (E y) = e y (X_Y_rec_2 a b c d e f y)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1028 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 1029 | apply simp | 
| 1030 | done | |
| 1031 | ||
| 1032 | lemma "X_Y_rec_2 a b c d e f F = f" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1033 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 1034 | apply simp | 
| 1035 | done | |
| 1036 | ||
| 1037 | lemma "P (X_Y_rec_1 a b c d e f x)" | |
| 1038 | nitpick [expect = genuine] | |
| 1039 | oops | |
| 1040 | ||
| 1041 | lemma "P (X_Y_rec_2 a b c d e f y)" | |
| 1042 | nitpick [expect = genuine] | |
| 1043 | oops | |
| 1044 | ||
| 1045 | text {* Other datatype examples *}
 | |
| 1046 | ||
| 1047 | text {* Indirect recursion is implemented via mutual recursion. *}
 | |
| 1048 | ||
| 1049 | datatype XOpt = CX "XOpt option" | DX "bool \<Rightarrow> XOpt option" | |
| 1050 | ||
| 1051 | lemma "P (x\<Colon>XOpt)" | |
| 1052 | nitpick [expect = genuine] | |
| 1053 | oops | |
| 1054 | ||
| 1055 | lemma "P (CX None)" | |
| 1056 | nitpick [expect = genuine] | |
| 1057 | oops | |
| 1058 | ||
| 1059 | lemma "P (CX (Some (CX None)))" | |
| 1060 | nitpick [expect = genuine] | |
| 1061 | oops | |
| 1062 | ||
| 1063 | lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (CX x) = cx x (XOpt_rec_2 cx dx n1 s1 n2 s2 x)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1064 | nitpick [card = 1\<midarrow>5, expect = none] | 
| 33197 | 1065 | apply simp | 
| 1066 | done | |
| 1067 | ||
| 1068 | lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\<lambda>b. XOpt_rec_3 cx dx n1 s1 n2 s2 (x b))" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1069 | nitpick [card = 1\<midarrow>3, expect = none] | 
| 33197 | 1070 | apply simp | 
| 1071 | done | |
| 1072 | ||
| 1073 | lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 None = n1" | |
| 1074 | nitpick [card = 1\<midarrow>4, expect = none] | |
| 1075 | apply simp | |
| 1076 | done | |
| 1077 | ||
| 1078 | lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 (Some x) = s1 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)" | |
| 1079 | nitpick [card = 1\<midarrow>4, expect = none] | |
| 1080 | apply simp | |
| 1081 | done | |
| 1082 | ||
| 1083 | lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 None = n2" | |
| 1084 | nitpick [card = 1\<midarrow>4, expect = none] | |
| 1085 | apply simp | |
| 1086 | done | |
| 1087 | ||
| 1088 | lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 (Some x) = s2 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)" | |
| 1089 | nitpick [card = 1\<midarrow>4, expect = none] | |
| 1090 | apply simp | |
| 1091 | done | |
| 1092 | ||
| 1093 | lemma "P (XOpt_rec_1 cx dx n1 s1 n2 s2 x)" | |
| 1094 | nitpick [expect = genuine] | |
| 1095 | oops | |
| 1096 | ||
| 1097 | lemma "P (XOpt_rec_2 cx dx n1 s1 n2 s2 x)" | |
| 1098 | nitpick [expect = genuine] | |
| 1099 | oops | |
| 1100 | ||
| 1101 | lemma "P (XOpt_rec_3 cx dx n1 s1 n2 s2 x)" | |
| 1102 | nitpick [expect = genuine] | |
| 1103 | oops | |
| 1104 | ||
| 1105 | datatype 'a YOpt = CY "('a \<Rightarrow> 'a YOpt) option"
 | |
| 1106 | ||
| 1107 | lemma "P (x\<Colon>'a YOpt)" | |
| 1108 | nitpick [expect = genuine] | |
| 1109 | oops | |
| 1110 | ||
| 1111 | lemma "P (CY None)" | |
| 1112 | nitpick [expect = genuine] | |
| 1113 | oops | |
| 1114 | ||
| 1115 | lemma "P (CY (Some (\<lambda>a. CY None)))" | |
| 1116 | nitpick [expect = genuine] | |
| 1117 | oops | |
| 1118 | ||
| 1119 | lemma "YOpt_rec_1 cy n s (CY x) = cy x (YOpt_rec_2 cy n s x)" | |
| 1120 | nitpick [card = 1\<midarrow>2, expect = none] | |
| 1121 | apply simp | |
| 1122 | done | |
| 1123 | ||
| 1124 | lemma "YOpt_rec_2 cy n s None = n" | |
| 1125 | nitpick [card = 1\<midarrow>2, expect = none] | |
| 1126 | apply simp | |
| 1127 | done | |
| 1128 | ||
| 1129 | lemma "YOpt_rec_2 cy n s (Some x) = s x (\<lambda>a. YOpt_rec_1 cy n s (x a))" | |
| 1130 | nitpick [card = 1\<midarrow>2, expect = none] | |
| 1131 | apply simp | |
| 1132 | done | |
| 1133 | ||
| 1134 | lemma "P (YOpt_rec_1 cy n s x)" | |
| 1135 | nitpick [expect = genuine] | |
| 1136 | oops | |
| 1137 | ||
| 1138 | lemma "P (YOpt_rec_2 cy n s x)" | |
| 1139 | nitpick [expect = genuine] | |
| 1140 | oops | |
| 1141 | ||
| 1142 | datatype Trie = TR "Trie list" | |
| 1143 | ||
| 1144 | lemma "P (x\<Colon>Trie)" | |
| 1145 | nitpick [expect = genuine] | |
| 1146 | oops | |
| 1147 | ||
| 1148 | lemma "\<forall>x\<Colon>Trie. P x" | |
| 1149 | nitpick [expect = genuine] | |
| 1150 | oops | |
| 1151 | ||
| 1152 | lemma "P (TR [TR []])" | |
| 1153 | nitpick [expect = genuine] | |
| 1154 | oops | |
| 1155 | ||
| 1156 | lemma "Trie_rec_1 tr nil cons (TR x) = tr x (Trie_rec_2 tr nil cons x)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1157 | nitpick [card = 1\<midarrow>4, expect = none] | 
| 33197 | 1158 | apply simp | 
| 1159 | done | |
| 1160 | ||
| 1161 | lemma "Trie_rec_2 tr nil cons [] = nil" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1162 | nitpick [card = 1\<midarrow>4, expect = none] | 
| 33197 | 1163 | apply simp | 
| 1164 | done | |
| 1165 | ||
| 1166 | lemma "Trie_rec_2 tr nil cons (x#xs) = cons x xs (Trie_rec_1 tr nil cons x) (Trie_rec_2 tr nil cons xs)" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1167 | nitpick [card = 1\<midarrow>4, expect = none] | 
| 33197 | 1168 | apply simp | 
| 1169 | done | |
| 1170 | ||
| 1171 | lemma "P (Trie_rec_1 tr nil cons x)" | |
| 1172 | nitpick [card = 1, expect = genuine] | |
| 1173 | oops | |
| 1174 | ||
| 1175 | lemma "P (Trie_rec_2 tr nil cons x)" | |
| 1176 | nitpick [card = 1, expect = genuine] | |
| 1177 | oops | |
| 1178 | ||
| 1179 | datatype InfTree = Leaf | Node "nat \<Rightarrow> InfTree" | |
| 1180 | ||
| 1181 | lemma "P (x\<Colon>InfTree)" | |
| 1182 | nitpick [expect = genuine] | |
| 1183 | oops | |
| 1184 | ||
| 1185 | lemma "\<forall>x\<Colon>InfTree. P x" | |
| 1186 | nitpick [expect = genuine] | |
| 1187 | oops | |
| 1188 | ||
| 1189 | lemma "P (Node (\<lambda>n. Leaf))" | |
| 1190 | nitpick [expect = genuine] | |
| 1191 | oops | |
| 1192 | ||
| 1193 | lemma "InfTree_rec leaf node Leaf = leaf" | |
| 1194 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1195 | apply simp | |
| 1196 | done | |
| 1197 | ||
| 1198 | lemma "InfTree_rec leaf node (Node x) = node x (\<lambda>n. InfTree_rec leaf node (x n))" | |
| 1199 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1200 | apply simp | |
| 1201 | done | |
| 1202 | ||
| 1203 | lemma "P (InfTree_rec leaf node x)" | |
| 1204 | nitpick [expect = genuine] | |
| 1205 | oops | |
| 1206 | ||
| 1207 | datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a \<Rightarrow> 'a lambda" | |
| 1208 | ||
| 1209 | lemma "P (x\<Colon>'a lambda)" | |
| 1210 | nitpick [expect = genuine] | |
| 1211 | oops | |
| 1212 | ||
| 1213 | lemma "\<forall>x\<Colon>'a lambda. P x" | |
| 1214 | nitpick [expect = genuine] | |
| 1215 | oops | |
| 1216 | ||
| 1217 | lemma "P (Lam (\<lambda>a. Var a))" | |
| 1218 | nitpick [card = 1\<midarrow>5, expect = none] | |
| 1219 | nitpick [card 'a = 4, card "'a lambda" = 5, expect = genuine] | |
| 1220 | oops | |
| 1221 | ||
| 1222 | lemma "lambda_rec var app lam (Var x) = var x" | |
| 1223 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1224 | apply simp | |
| 1225 | done | |
| 1226 | ||
| 1227 | lemma "lambda_rec var app lam (App x y) = app x y (lambda_rec var app lam x) (lambda_rec var app lam y)" | |
| 1228 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1229 | apply simp | |
| 1230 | done | |
| 1231 | ||
| 1232 | lemma "lambda_rec var app lam (Lam x) = lam x (\<lambda>a. lambda_rec var app lam (x a))" | |
| 1233 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1234 | apply simp | |
| 1235 | done | |
| 1236 | ||
| 1237 | lemma "P (lambda_rec v a l x)" | |
| 1238 | nitpick [expect = genuine] | |
| 1239 | oops | |
| 1240 | ||
| 1241 | text {* Taken from "Inductive datatypes in HOL", p. 8: *}
 | |
| 1242 | ||
| 1243 | datatype ('a, 'b) T = C "'a \<Rightarrow> bool" | D "'b list"
 | |
| 1244 | datatype 'c U = E "('c, 'c U) T"
 | |
| 1245 | ||
| 1246 | lemma "P (x\<Colon>'c U)" | |
| 1247 | nitpick [expect = genuine] | |
| 1248 | oops | |
| 1249 | ||
| 1250 | lemma "\<forall>x\<Colon>'c U. P x" | |
| 1251 | nitpick [expect = genuine] | |
| 1252 | oops | |
| 1253 | ||
| 1254 | lemma "P (E (C (\<lambda>a. True)))" | |
| 1255 | nitpick [expect = genuine] | |
| 1256 | oops | |
| 1257 | ||
| 1258 | lemma "U_rec_1 e c d nil cons (E x) = e x (U_rec_2 e c d nil cons x)" | |
| 1259 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1260 | apply simp | |
| 1261 | done | |
| 1262 | ||
| 1263 | lemma "U_rec_2 e c d nil cons (C x) = c x" | |
| 1264 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1265 | apply simp | |
| 1266 | done | |
| 1267 | ||
| 1268 | lemma "U_rec_2 e c d nil cons (D x) = d x (U_rec_3 e c d nil cons x)" | |
| 1269 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1270 | apply simp | |
| 1271 | done | |
| 1272 | ||
| 1273 | lemma "U_rec_3 e c d nil cons [] = nil" | |
| 1274 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1275 | apply simp | |
| 1276 | done | |
| 1277 | ||
| 1278 | lemma "U_rec_3 e c d nil cons (x#xs) = cons x xs (U_rec_1 e c d nil cons x) (U_rec_3 e c d nil cons xs)" | |
| 1279 | nitpick [card = 1\<midarrow>3, expect = none] | |
| 1280 | apply simp | |
| 1281 | done | |
| 1282 | ||
| 1283 | lemma "P (U_rec_1 e c d nil cons x)" | |
| 1284 | nitpick [expect = genuine] | |
| 1285 | oops | |
| 1286 | ||
| 1287 | lemma "P (U_rec_2 e c d nil cons x)" | |
| 1288 | nitpick [card = 1, expect = genuine] | |
| 1289 | oops | |
| 1290 | ||
| 1291 | lemma "P (U_rec_3 e c d nil cons x)" | |
| 1292 | nitpick [card = 1, expect = genuine] | |
| 1293 | oops | |
| 1294 | ||
| 1295 | subsubsection {* Records *}
 | |
| 1296 | ||
| 1297 | record ('a, 'b) point =
 | |
| 1298 | xpos :: 'a | |
| 1299 | ypos :: 'b | |
| 1300 | ||
| 1301 | lemma "(x\<Colon>('a, 'b) point) = y"
 | |
| 1302 | nitpick [expect = genuine] | |
| 1303 | oops | |
| 1304 | ||
| 1305 | record ('a, 'b, 'c) extpoint = "('a, 'b) point" +
 | |
| 1306 | ext :: 'c | |
| 1307 | ||
| 1308 | lemma "(x\<Colon>('a, 'b, 'c) extpoint) = y"
 | |
| 1309 | nitpick [expect = genuine] | |
| 1310 | oops | |
| 1311 | ||
| 1312 | subsubsection {* Inductively Defined Sets *}
 | |
| 1313 | ||
| 1314 | inductive_set undefinedSet :: "'a set" where | |
| 1315 | "undefined \<in> undefinedSet" | |
| 1316 | ||
| 1317 | lemma "x \<in> undefinedSet" | |
| 1318 | nitpick [expect = genuine] | |
| 1319 | oops | |
| 1320 | ||
| 1321 | inductive_set evenCard :: "'a set set" | |
| 1322 | where | |
| 1323 | "{} \<in> evenCard" |
 | |
| 1324 | "\<lbrakk>S \<in> evenCard; x \<notin> S; y \<notin> S; x \<noteq> y\<rbrakk> \<Longrightarrow> S \<union> {x, y} \<in> evenCard"
 | |
| 1325 | ||
| 1326 | lemma "S \<in> evenCard" | |
| 1327 | nitpick [expect = genuine] | |
| 1328 | oops | |
| 1329 | ||
| 1330 | inductive_set | |
| 1331 | even :: "nat set" | |
| 1332 | and odd :: "nat set" | |
| 1333 | where | |
| 1334 | "0 \<in> even" | | |
| 1335 | "n \<in> even \<Longrightarrow> Suc n \<in> odd" | | |
| 1336 | "n \<in> odd \<Longrightarrow> Suc n \<in> even" | |
| 1337 | ||
| 1338 | lemma "n \<in> odd" | |
| 1339 | nitpick [expect = genuine] | |
| 1340 | oops | |
| 1341 | ||
| 1342 | consts f :: "'a \<Rightarrow> 'a" | |
| 1343 | ||
| 1344 | inductive_set a_even :: "'a set" and a_odd :: "'a set" where | |
| 1345 | "undefined \<in> a_even" | | |
| 1346 | "x \<in> a_even \<Longrightarrow> f x \<in> a_odd" | | |
| 1347 | "x \<in> a_odd \<Longrightarrow> f x \<in> a_even" | |
| 1348 | ||
| 1349 | lemma "x \<in> a_odd" | |
| 1350 | nitpick [expect = genuine] | |
| 1351 | oops | |
| 1352 | ||
| 1353 | subsubsection {* Examples Involving Special Functions *}
 | |
| 1354 | ||
| 1355 | lemma "card x = 0" | |
| 1356 | nitpick [expect = genuine] | |
| 1357 | oops | |
| 1358 | ||
| 1359 | lemma "finite x" | |
| 1360 | nitpick [expect = none] | |
| 1361 | oops | |
| 1362 | ||
| 1363 | lemma "xs @ [] = ys @ []" | |
| 1364 | nitpick [expect = genuine] | |
| 1365 | oops | |
| 1366 | ||
| 1367 | lemma "xs @ ys = ys @ xs" | |
| 1368 | nitpick [expect = genuine] | |
| 1369 | oops | |
| 1370 | ||
| 1371 | lemma "f (lfp f) = lfp f" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1372 | nitpick [card = 2, expect = genuine] | 
| 33197 | 1373 | oops | 
| 1374 | ||
| 1375 | lemma "f (gfp f) = gfp f" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1376 | nitpick [card = 2, expect = genuine] | 
| 33197 | 1377 | oops | 
| 1378 | ||
| 1379 | lemma "lfp f = gfp f" | |
| 35284 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35087diff
changeset | 1380 | nitpick [card = 2, expect = genuine] | 
| 33197 | 1381 | oops | 
| 1382 | ||
| 1383 | subsubsection {* Axiomatic Type Classes and Overloading *}
 | |
| 1384 | ||
| 1385 | text {* A type class without axioms: *}
 | |
| 1386 | ||
| 35338 | 1387 | class classA | 
| 33197 | 1388 | |
| 1389 | lemma "P (x\<Colon>'a\<Colon>classA)" | |
| 1390 | nitpick [expect = genuine] | |
| 1391 | oops | |
| 1392 | ||
| 1393 | text {* An axiom with a type variable (denoting types which have at least two elements): *}
 | |
| 1394 | ||
| 35338 | 1395 | class classC = | 
| 1396 | assumes classC_ax: "\<exists>x y. x \<noteq> y" | |
| 33197 | 1397 | |
| 1398 | lemma "P (x\<Colon>'a\<Colon>classC)" | |
| 1399 | nitpick [expect = genuine] | |
| 1400 | oops | |
| 1401 | ||
| 1402 | lemma "\<exists>x y. (x\<Colon>'a\<Colon>classC) \<noteq> y" | |
| 1403 | nitpick [expect = none] | |
| 1404 | sorry | |
| 1405 | ||
| 1406 | text {* A type class for which a constant is defined: *}
 | |
| 1407 | ||
| 35338 | 1408 | class classD = | 
| 1409 | fixes classD_const :: "'a \<Rightarrow> 'a" | |
| 1410 | assumes classD_ax: "classD_const (classD_const x) = classD_const x" | |
| 33197 | 1411 | |
| 1412 | lemma "P (x\<Colon>'a\<Colon>classD)" | |
| 1413 | nitpick [expect = genuine] | |
| 1414 | oops | |
| 1415 | ||
| 1416 | text {* A type class with multiple superclasses: *}
 | |
| 1417 | ||
| 35338 | 1418 | class classE = classC + classD | 
| 33197 | 1419 | |
| 1420 | lemma "P (x\<Colon>'a\<Colon>classE)" | |
| 1421 | nitpick [expect = genuine] | |
| 1422 | oops | |
| 1423 | ||
| 1424 | text {* OFCLASS: *}
 | |
| 1425 | ||
| 1426 | lemma "OFCLASS('a\<Colon>type, type_class)"
 | |
| 1427 | nitpick [expect = none] | |
| 1428 | apply intro_classes | |
| 1429 | done | |
| 1430 | ||
| 1431 | lemma "OFCLASS('a\<Colon>classC, type_class)"
 | |
| 1432 | nitpick [expect = none] | |
| 1433 | apply intro_classes | |
| 1434 | done | |
| 1435 | ||
| 1436 | lemma "OFCLASS('a\<Colon>type, classC_class)"
 | |
| 1437 | nitpick [expect = genuine] | |
| 1438 | oops | |
| 1439 | ||
| 1440 | text {* Overloading: *}
 | |
| 1441 | ||
| 1442 | consts inverse :: "'a \<Rightarrow> 'a" | |
| 1443 | ||
| 1444 | defs (overloaded) | |
| 1445 | inverse_bool: "inverse (b\<Colon>bool) \<equiv> \<not> b" | |
| 1446 | inverse_set: "inverse (S\<Colon>'a set) \<equiv> -S" | |
| 1447 | inverse_pair: "inverse p \<equiv> (inverse (fst p), inverse (snd p))" | |
| 1448 | ||
| 1449 | lemma "inverse b" | |
| 1450 | nitpick [expect = genuine] | |
| 1451 | oops | |
| 1452 | ||
| 1453 | lemma "P (inverse (S\<Colon>'a set))" | |
| 1454 | nitpick [expect = genuine] | |
| 1455 | oops | |
| 1456 | ||
| 1457 | lemma "P (inverse (p\<Colon>'a\<times>'b))" | |
| 1458 | nitpick [expect = genuine] | |
| 1459 | oops | |
| 1460 | ||
| 1461 | end |