| author | paulson | 
| Thu, 05 Oct 2006 10:46:26 +0200 | |
| changeset 20864 | bb75b876b260 | 
| parent 18789 | 8a5c6fc3ad27 | 
| child 21502 | 7f3ea2b3bab6 | 
| permissions | -rw-r--r-- | 
| 14350 | 1 | (* Title: HOL/ex/Refute_Examples.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tjark Weber | |
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 4 | Copyright 2003-2005 | 
| 14350 | 5 | *) | 
| 6 | ||
| 7 | (* See 'HOL/Refute.thy' for help. *) | |
| 8 | ||
| 9 | header {* Examples for the 'refute' command *}
 | |
| 10 | ||
| 15297 | 11 | theory Refute_Examples imports Main | 
| 12 | ||
| 13 | begin | |
| 14350 | 14 | |
| 18774 | 15 | refute_params [satsolver="dpll"] | 
| 16 | ||
| 14350 | 17 | lemma "P \<and> Q" | 
| 18 | apply (rule conjI) | |
| 19 |   refute 1  -- {* refutes @{term "P"} *}
 | |
| 20 |   refute 2  -- {* refutes @{term "Q"} *}
 | |
| 21 |   refute    -- {* equivalent to 'refute 1' *}
 | |
| 14455 | 22 |     -- {* here 'refute 3' would cause an exception, since we only have 2 subgoals *}
 | 
| 14465 | 23 |   refute [maxsize=5]           -- {* we can override parameters ... *}
 | 
| 24 |   refute [satsolver="dpll"] 2  -- {* ... and specify a subgoal at the same time *}
 | |
| 14350 | 25 | oops | 
| 26 | ||
| 14455 | 27 | section {* Examples and Test Cases *}
 | 
| 14350 | 28 | |
| 29 | subsection {* Propositional logic *}
 | |
| 30 | ||
| 31 | lemma "True" | |
| 32 | refute | |
| 33 | apply auto | |
| 34 | done | |
| 35 | ||
| 36 | lemma "False" | |
| 37 | refute | |
| 38 | oops | |
| 39 | ||
| 40 | lemma "P" | |
| 41 | refute | |
| 42 | oops | |
| 43 | ||
| 44 | lemma "~ P" | |
| 45 | refute | |
| 46 | oops | |
| 47 | ||
| 48 | lemma "P & Q" | |
| 49 | refute | |
| 50 | oops | |
| 51 | ||
| 52 | lemma "P | Q" | |
| 53 | refute | |
| 54 | oops | |
| 55 | ||
| 56 | lemma "P \<longrightarrow> Q" | |
| 57 | refute | |
| 58 | oops | |
| 59 | ||
| 60 | lemma "(P::bool) = Q" | |
| 61 | refute | |
| 62 | oops | |
| 63 | ||
| 64 | lemma "(P | Q) \<longrightarrow> (P & Q)" | |
| 65 | refute | |
| 66 | oops | |
| 67 | ||
| 68 | subsection {* Predicate logic *}
 | |
| 69 | ||
| 14455 | 70 | lemma "P x y z" | 
| 14350 | 71 | refute | 
| 72 | oops | |
| 73 | ||
| 74 | lemma "P x y \<longrightarrow> P y x" | |
| 75 | refute | |
| 76 | oops | |
| 77 | ||
| 14455 | 78 | lemma "P (f (f x)) \<longrightarrow> P x \<longrightarrow> P (f x)" | 
| 79 | refute | |
| 80 | oops | |
| 81 | ||
| 14350 | 82 | subsection {* Equality *}
 | 
| 83 | ||
| 84 | lemma "P = True" | |
| 85 | refute | |
| 86 | oops | |
| 87 | ||
| 88 | lemma "P = False" | |
| 89 | refute | |
| 90 | oops | |
| 91 | ||
| 92 | lemma "x = y" | |
| 93 | refute | |
| 94 | oops | |
| 95 | ||
| 96 | lemma "f x = g x" | |
| 97 | refute | |
| 98 | oops | |
| 99 | ||
| 100 | lemma "(f::'a\<Rightarrow>'b) = g" | |
| 101 | refute | |
| 102 | oops | |
| 103 | ||
| 104 | lemma "(f::('d\<Rightarrow>'d)\<Rightarrow>('c\<Rightarrow>'d)) = g"
 | |
| 105 | refute | |
| 106 | oops | |
| 107 | ||
| 108 | lemma "distinct [a,b]" | |
| 14809 | 109 | refute | 
| 14350 | 110 | apply simp | 
| 111 | refute | |
| 112 | oops | |
| 113 | ||
| 114 | subsection {* First-Order Logic *}
 | |
| 115 | ||
| 116 | lemma "\<exists>x. P x" | |
| 117 | refute | |
| 118 | oops | |
| 119 | ||
| 120 | lemma "\<forall>x. P x" | |
| 121 | refute | |
| 122 | oops | |
| 123 | ||
| 124 | lemma "EX! x. P x" | |
| 125 | refute | |
| 126 | oops | |
| 127 | ||
| 128 | lemma "Ex P" | |
| 129 | refute | |
| 130 | oops | |
| 131 | ||
| 132 | lemma "All P" | |
| 133 | refute | |
| 134 | oops | |
| 135 | ||
| 136 | lemma "Ex1 P" | |
| 137 | refute | |
| 138 | oops | |
| 139 | ||
| 140 | lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)" | |
| 141 | refute | |
| 142 | oops | |
| 143 | ||
| 144 | lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>y. \<forall>x. P x y)" | |
| 145 | refute | |
| 146 | oops | |
| 147 | ||
| 148 | lemma "(\<exists>x. P x) \<longrightarrow> (EX! x. P x)" | |
| 149 | refute | |
| 150 | oops | |
| 151 | ||
| 152 | text {* A true statement (also testing names of free and bound variables being identical) *}
 | |
| 153 | ||
| 154 | lemma "(\<forall>x y. P x y \<longrightarrow> P y x) \<longrightarrow> (\<forall>x. P x y) \<longrightarrow> P y x" | |
| 18774 | 155 | refute [maxsize=4] | 
| 14350 | 156 | apply fast | 
| 157 | done | |
| 158 | ||
| 18789 | 159 | text {* "A type has at most 4 elements." *}
 | 
| 14350 | 160 | |
| 18789 | 161 | lemma "a=b | a=c | a=d | a=e | b=c | b=d | b=e | c=d | c=e | d=e" | 
| 14455 | 162 | refute | 
| 163 | oops | |
| 164 | ||
| 18789 | 165 | lemma "\<forall>a b c d e. a=b | a=c | a=d | a=e | b=c | b=d | b=e | c=d | c=e | d=e" | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 166 |   refute  -- {* quantification causes an expansion of the formula; the
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 167 | previous version with free variables is refuted much faster *} | 
| 14350 | 168 | oops | 
| 169 | ||
| 170 | text {* "Every reflexive and symmetric relation is transitive." *}
 | |
| 171 | ||
| 172 | 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" | |
| 14489 | 173 | refute | 
| 14350 | 174 | oops | 
| 175 | ||
| 14465 | 176 | text {* The "Drinker's theorem" ... *}
 | 
| 14350 | 177 | |
| 178 | lemma "\<exists>x. f x = g x \<longrightarrow> f = g" | |
| 14809 | 179 | refute [maxsize=4] | 
| 14350 | 180 | apply (auto simp add: ext) | 
| 181 | done | |
| 182 | ||
| 14465 | 183 | text {* ... and an incorrect version of it *}
 | 
| 14350 | 184 | |
| 185 | lemma "(\<exists>x. f x = g x) \<longrightarrow> f = g" | |
| 186 | refute | |
| 187 | oops | |
| 188 | ||
| 189 | text {* "Every function has a fixed point." *}
 | |
| 190 | ||
| 191 | lemma "\<exists>x. f x = x" | |
| 192 | refute | |
| 193 | oops | |
| 194 | ||
| 195 | text {* "Function composition is commutative." *}
 | |
| 196 | ||
| 197 | lemma "f (g x) = g (f x)" | |
| 198 | refute | |
| 199 | oops | |
| 200 | ||
| 201 | text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *}
 | |
| 202 | ||
| 203 | lemma "((P::('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
 | |
| 204 | refute | |
| 205 | oops | |
| 206 | ||
| 207 | subsection {* Higher-Order Logic *}
 | |
| 208 | ||
| 209 | lemma "\<exists>P. P" | |
| 210 | refute | |
| 211 | apply auto | |
| 212 | done | |
| 213 | ||
| 214 | lemma "\<forall>P. P" | |
| 215 | refute | |
| 216 | oops | |
| 217 | ||
| 218 | lemma "EX! P. P" | |
| 219 | refute | |
| 220 | apply auto | |
| 221 | done | |
| 222 | ||
| 223 | lemma "EX! P. P x" | |
| 224 | refute | |
| 225 | oops | |
| 226 | ||
| 227 | lemma "P Q | Q x" | |
| 228 | refute | |
| 229 | oops | |
| 230 | ||
| 14455 | 231 | lemma "P All" | 
| 232 | refute | |
| 233 | oops | |
| 234 | ||
| 235 | lemma "P Ex" | |
| 236 | refute | |
| 237 | oops | |
| 238 | ||
| 239 | lemma "P Ex1" | |
| 240 | refute | |
| 241 | oops | |
| 242 | ||
| 14350 | 243 | text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *}
 | 
| 244 | ||
| 245 | constdefs | |
| 246 |   "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
 | |
| 247 | "trans P == (ALL x y z. P x y \<longrightarrow> P y z \<longrightarrow> P x z)" | |
| 248 |   "subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
 | |
| 249 | "subset P Q == (ALL x y. P x y \<longrightarrow> Q x y)" | |
| 250 |   "trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
 | |
| 251 | "trans_closure P Q == (subset Q P) & (trans P) & (ALL R. subset Q R \<longrightarrow> trans R \<longrightarrow> subset P R)" | |
| 252 | ||
| 253 | lemma "trans_closure T P \<longrightarrow> (\<exists>x y. T x y)" | |
| 254 | refute | |
| 255 | oops | |
| 256 | ||
| 257 | text {* "The union of transitive closures is equal to the transitive closure of unions." *}
 | |
| 258 | ||
| 259 | lemma "(\<forall>x y. (P x y | R x y) \<longrightarrow> T x y) \<longrightarrow> trans T \<longrightarrow> (\<forall>Q. (\<forall>x y. (P x y | R x y) \<longrightarrow> Q x y) \<longrightarrow> trans Q \<longrightarrow> subset T Q) | |
| 260 | \<longrightarrow> trans_closure TP P | |
| 261 | \<longrightarrow> trans_closure TR R | |
| 262 | \<longrightarrow> (T x y = (TP x y | TR x y))" | |
| 16910 | 263 | refute | 
| 14350 | 264 | oops | 
| 265 | ||
| 266 | text {* "Every surjective function is invertible." *}
 | |
| 267 | ||
| 268 | lemma "(\<forall>y. \<exists>x. y = f x) \<longrightarrow> (\<exists>g. \<forall>x. g (f x) = x)" | |
| 269 | refute | |
| 270 | oops | |
| 271 | ||
| 272 | text {* "Every invertible function is surjective." *}
 | |
| 273 | ||
| 274 | lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)" | |
| 275 | refute | |
| 276 | oops | |
| 277 | ||
| 278 | text {* Every point is a fixed point of some function. *}
 | |
| 279 | ||
| 280 | lemma "\<exists>f. f x = x" | |
| 14809 | 281 | refute [maxsize=4] | 
| 14350 | 282 | apply (rule_tac x="\<lambda>x. x" in exI) | 
| 283 | apply simp | |
| 284 | done | |
| 285 | ||
| 14465 | 286 | text {* Axiom of Choice: first an incorrect version ... *}
 | 
| 14350 | 287 | |
| 288 | lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (EX!f. \<forall>x. P x (f x))" | |
| 289 | refute | |
| 290 | oops | |
| 291 | ||
| 14465 | 292 | text {* ... and now two correct ones *}
 | 
| 14350 | 293 | |
| 294 | lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>f. \<forall>x. P x (f x))" | |
| 14809 | 295 | refute [maxsize=4] | 
| 14350 | 296 | apply (simp add: choice) | 
| 297 | done | |
| 298 | ||
| 299 | lemma "(\<forall>x. EX!y. P x y) \<longrightarrow> (EX!f. \<forall>x. P x (f x))" | |
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 300 | refute [maxsize=2] | 
| 14350 | 301 | apply auto | 
| 302 | apply (simp add: ex1_implies_ex choice) | |
| 303 | apply (fast intro: ext) | |
| 304 | done | |
| 305 | ||
| 306 | subsection {* Meta-logic *}
 | |
| 307 | ||
| 308 | lemma "!!x. P x" | |
| 309 | refute | |
| 310 | oops | |
| 311 | ||
| 312 | lemma "f x == g x" | |
| 313 | refute | |
| 314 | oops | |
| 315 | ||
| 316 | lemma "P \<Longrightarrow> Q" | |
| 317 | refute | |
| 318 | oops | |
| 319 | ||
| 320 | lemma "\<lbrakk> P; Q; R \<rbrakk> \<Longrightarrow> S" | |
| 321 | refute | |
| 322 | oops | |
| 323 | ||
| 324 | subsection {* Schematic variables *}
 | |
| 325 | ||
| 326 | lemma "?P" | |
| 327 | refute | |
| 328 | apply auto | |
| 329 | done | |
| 330 | ||
| 331 | lemma "x = ?y" | |
| 332 | refute | |
| 333 | apply auto | |
| 334 | done | |
| 335 | ||
| 336 | subsection {* Abstractions *}
 | |
| 337 | ||
| 338 | lemma "(\<lambda>x. x) = (\<lambda>x. y)" | |
| 339 | refute | |
| 340 | oops | |
| 341 | ||
| 342 | lemma "(\<lambda>f. f x) = (\<lambda>f. True)" | |
| 343 | refute | |
| 344 | oops | |
| 345 | ||
| 346 | lemma "(\<lambda>x. x) = (\<lambda>y. y)" | |
| 347 | refute | |
| 348 | apply simp | |
| 349 | done | |
| 350 | ||
| 351 | subsection {* Sets *}
 | |
| 352 | ||
| 353 | lemma "P (A::'a set)" | |
| 354 | refute | |
| 355 | oops | |
| 356 | ||
| 357 | lemma "P (A::'a set set)" | |
| 358 | refute | |
| 359 | oops | |
| 360 | ||
| 361 | lemma "{x. P x} = {y. P y}"
 | |
| 14489 | 362 | refute | 
| 14350 | 363 | apply simp | 
| 364 | done | |
| 365 | ||
| 366 | lemma "x : {x. P x}"
 | |
| 367 | refute | |
| 368 | oops | |
| 369 | ||
| 14455 | 370 | lemma "P op:" | 
| 371 | refute | |
| 372 | oops | |
| 373 | ||
| 374 | lemma "P (op: x)" | |
| 375 | refute | |
| 376 | oops | |
| 377 | ||
| 378 | lemma "P Collect" | |
| 379 | refute | |
| 380 | oops | |
| 381 | ||
| 14350 | 382 | lemma "A Un B = A Int B" | 
| 383 | refute | |
| 384 | oops | |
| 385 | ||
| 386 | lemma "(A Int B) Un C = (A Un C) Int B" | |
| 387 | refute | |
| 388 | oops | |
| 389 | ||
| 390 | lemma "Ball A P \<longrightarrow> Bex A P" | |
| 14455 | 391 | refute | 
| 392 | oops | |
| 393 | ||
| 394 | subsection {* arbitrary *}
 | |
| 395 | ||
| 396 | lemma "arbitrary" | |
| 397 | refute | |
| 398 | oops | |
| 399 | ||
| 400 | lemma "P arbitrary" | |
| 401 | refute | |
| 402 | oops | |
| 403 | ||
| 404 | lemma "arbitrary x" | |
| 405 | refute | |
| 406 | oops | |
| 407 | ||
| 408 | lemma "arbitrary arbitrary" | |
| 409 | refute | |
| 410 | oops | |
| 411 | ||
| 412 | subsection {* The *}
 | |
| 413 | ||
| 414 | lemma "The P" | |
| 415 | refute | |
| 416 | oops | |
| 417 | ||
| 418 | lemma "P The" | |
| 14350 | 419 | refute | 
| 420 | oops | |
| 421 | ||
| 14455 | 422 | lemma "P (The P)" | 
| 423 | refute | |
| 424 | oops | |
| 425 | ||
| 426 | lemma "(THE x. x=y) = z" | |
| 427 | refute | |
| 428 | oops | |
| 429 | ||
| 430 | lemma "Ex P \<longrightarrow> P (The P)" | |
| 14489 | 431 | refute | 
| 14455 | 432 | oops | 
| 433 | ||
| 434 | subsection {* Eps *}
 | |
| 435 | ||
| 436 | lemma "Eps P" | |
| 437 | refute | |
| 438 | oops | |
| 439 | ||
| 440 | lemma "P Eps" | |
| 441 | refute | |
| 442 | oops | |
| 443 | ||
| 444 | lemma "P (Eps P)" | |
| 445 | refute | |
| 446 | oops | |
| 447 | ||
| 448 | lemma "(SOME x. x=y) = z" | |
| 449 | refute | |
| 450 | oops | |
| 451 | ||
| 452 | lemma "Ex P \<longrightarrow> P (Eps P)" | |
| 14489 | 453 | refute [maxsize=3] | 
| 14455 | 454 | apply (auto simp add: someI) | 
| 455 | done | |
| 456 | ||
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 457 | (******************************************************************************) | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 458 | |
| 14809 | 459 | subsection {* Subtypes (typedef), typedecl *}
 | 
| 460 | ||
| 15161 | 461 | text {* A completely unspecified non-empty subset of @{typ "'a"}: *}
 | 
| 462 | ||
| 14809 | 463 | typedef 'a myTdef = "insert (arbitrary::'a) (arbitrary::'a set)" | 
| 464 | by auto | |
| 465 | ||
| 466 | lemma "(x::'a myTdef) = y" | |
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 467 | refute | 
| 14809 | 468 | oops | 
| 469 | ||
| 470 | typedecl myTdecl | |
| 471 | ||
| 472 | typedef 'a T_bij = "{(f::'a\<Rightarrow>'a). \<forall>y. \<exists>!x. f x = y}"
 | |
| 473 | by auto | |
| 474 | ||
| 475 | lemma "P (f::(myTdecl myTdef) T_bij)" | |
| 476 | refute | |
| 477 | oops | |
| 478 | ||
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 479 | (******************************************************************************) | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 480 | |
| 14455 | 481 | subsection {* Inductive datatypes *}
 | 
| 14350 | 482 | |
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 483 | text {* With quick\_and\_dirty set, the datatype package does not generate
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 484 | certain axioms for recursion operators. Without these axioms, refute may | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 485 | find spurious countermodels. *} | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 486 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 487 | ML {* reset quick_and_dirty; *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 488 | |
| 14350 | 489 | subsubsection {* unit *}
 | 
| 490 | ||
| 491 | lemma "P (x::unit)" | |
| 492 | refute | |
| 493 | oops | |
| 494 | ||
| 495 | lemma "\<forall>x::unit. P x" | |
| 496 | refute | |
| 497 | oops | |
| 498 | ||
| 499 | lemma "P ()" | |
| 500 | refute | |
| 501 | oops | |
| 502 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 503 | lemma "P (unit_rec u x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 504 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 505 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 506 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 507 | lemma "P (case x of () \<Rightarrow> u)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 508 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 509 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 510 | |
| 14455 | 511 | subsubsection {* option *}
 | 
| 512 | ||
| 513 | lemma "P (x::'a option)" | |
| 514 | refute | |
| 515 | oops | |
| 516 | ||
| 517 | lemma "\<forall>x::'a option. P x" | |
| 518 | refute | |
| 519 | oops | |
| 520 | ||
| 14809 | 521 | lemma "P None" | 
| 522 | refute | |
| 523 | oops | |
| 524 | ||
| 14455 | 525 | lemma "P (Some x)" | 
| 526 | refute | |
| 527 | oops | |
| 528 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 529 | lemma "P (option_rec n s x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 530 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 531 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 532 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 533 | lemma "P (case x of None \<Rightarrow> n | Some u \<Rightarrow> s u)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 534 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 535 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 536 | |
| 14350 | 537 | subsubsection {* * *}
 | 
| 538 | ||
| 539 | lemma "P (x::'a*'b)" | |
| 14455 | 540 | refute | 
| 14350 | 541 | oops | 
| 542 | ||
| 543 | lemma "\<forall>x::'a*'b. P x" | |
| 14455 | 544 | refute | 
| 14350 | 545 | oops | 
| 546 | ||
| 547 | lemma "P (x,y)" | |
| 14455 | 548 | refute | 
| 14350 | 549 | oops | 
| 550 | ||
| 551 | lemma "P (fst x)" | |
| 14455 | 552 | refute | 
| 14350 | 553 | oops | 
| 554 | ||
| 555 | lemma "P (snd x)" | |
| 14455 | 556 | refute | 
| 557 | oops | |
| 558 | ||
| 559 | lemma "P Pair" | |
| 560 | refute | |
| 14350 | 561 | oops | 
| 562 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 563 | lemma "P (prod_rec p x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 564 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 565 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 566 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 567 | lemma "P (case x of Pair a b \<Rightarrow> p a b)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 568 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 569 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 570 | |
| 14350 | 571 | subsubsection {* + *}
 | 
| 572 | ||
| 573 | lemma "P (x::'a+'b)" | |
| 14455 | 574 | refute | 
| 14350 | 575 | oops | 
| 576 | ||
| 577 | lemma "\<forall>x::'a+'b. P x" | |
| 14455 | 578 | refute | 
| 14350 | 579 | oops | 
| 580 | ||
| 581 | lemma "P (Inl x)" | |
| 14455 | 582 | refute | 
| 14350 | 583 | oops | 
| 584 | ||
| 585 | lemma "P (Inr x)" | |
| 14455 | 586 | refute | 
| 587 | oops | |
| 588 | ||
| 589 | lemma "P Inl" | |
| 590 | refute | |
| 14350 | 591 | oops | 
| 592 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 593 | lemma "P (sum_rec l r x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 594 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 595 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 596 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 597 | lemma "P (case x of Inl a \<Rightarrow> l a | Inr b \<Rightarrow> r b)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 598 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 599 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 600 | |
| 14350 | 601 | subsubsection {* Non-recursive datatypes *}
 | 
| 602 | ||
| 14455 | 603 | datatype T1 = A | B | 
| 14350 | 604 | |
| 605 | lemma "P (x::T1)" | |
| 606 | refute | |
| 607 | oops | |
| 608 | ||
| 609 | lemma "\<forall>x::T1. P x" | |
| 610 | refute | |
| 611 | oops | |
| 612 | ||
| 14455 | 613 | lemma "P A" | 
| 14350 | 614 | refute | 
| 615 | oops | |
| 616 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 617 | lemma "P (T1_rec a b x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 618 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 619 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 620 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 621 | lemma "P (case x of A \<Rightarrow> a | B \<Rightarrow> b)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 622 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 623 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 624 | |
| 14455 | 625 | datatype 'a T2 = C T1 | D 'a | 
| 626 | ||
| 627 | lemma "P (x::'a T2)" | |
| 14350 | 628 | refute | 
| 629 | oops | |
| 630 | ||
| 14455 | 631 | lemma "\<forall>x::'a T2. P x" | 
| 14350 | 632 | refute | 
| 633 | oops | |
| 634 | ||
| 14455 | 635 | lemma "P D" | 
| 14350 | 636 | refute | 
| 637 | oops | |
| 638 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 639 | lemma "P (T2_rec c d x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 640 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 641 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 642 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 643 | lemma "P (case x of C u \<Rightarrow> c u | D v \<Rightarrow> d v)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 644 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 645 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 646 | |
| 14455 | 647 | datatype ('a,'b) T3 = E "'a \<Rightarrow> 'b"
 | 
| 648 | ||
| 14809 | 649 | lemma "P (x::('a,'b) T3)"
 | 
| 650 | refute | |
| 651 | oops | |
| 652 | ||
| 653 | lemma "\<forall>x::('a,'b) T3. P x"
 | |
| 654 | refute | |
| 655 | oops | |
| 656 | ||
| 14455 | 657 | lemma "P E" | 
| 658 | refute | |
| 14350 | 659 | oops | 
| 660 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 661 | lemma "P (T3_rec e x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 662 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 663 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 664 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 665 | lemma "P (case x of E f \<Rightarrow> e f)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 666 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 667 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 668 | |
| 14350 | 669 | subsubsection {* Recursive datatypes *}
 | 
| 670 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 671 | text {* nat *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 672 | |
| 14809 | 673 | lemma "P (x::nat)" | 
| 674 | refute | |
| 675 | oops | |
| 14350 | 676 | |
| 14809 | 677 | lemma "\<forall>x::nat. P x" | 
| 678 | refute | |
| 14350 | 679 | oops | 
| 680 | ||
| 14809 | 681 | lemma "P (Suc 0)" | 
| 682 | refute | |
| 14350 | 683 | oops | 
| 684 | ||
| 14809 | 685 | lemma "P Suc" | 
| 686 |   refute  -- {* @{term "Suc"} is a partial function (regardless of the size
 | |
| 687 |                 of the model), hence @{term "P Suc"} is undefined, hence no
 | |
| 688 | model will be found *} | |
| 14350 | 689 | oops | 
| 690 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 691 | lemma "P (nat_rec zero suc x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 692 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 693 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 694 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 695 | lemma "P (case x of 0 \<Rightarrow> zero | Suc n \<Rightarrow> suc n)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 696 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 697 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 698 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 699 | text {* 'a list *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 700 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 701 | lemma "P (xs::'a list)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 702 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 703 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 704 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 705 | lemma "\<forall>xs::'a list. P xs" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 706 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 707 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 708 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 709 | lemma "P [x, y]" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 710 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 711 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 712 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 713 | lemma "P (list_rec nil cons xs)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 714 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 715 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 716 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 717 | lemma "P (case x of Nil \<Rightarrow> nil | Cons a b \<Rightarrow> cons a b)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 718 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 719 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 720 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 721 | lemma "(xs::'a list) = ys" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 722 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 723 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 724 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 725 | lemma "a # xs = b # xs" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 726 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 727 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 728 | |
| 14350 | 729 | datatype 'a BinTree = Leaf 'a | Node "'a BinTree" "'a BinTree" | 
| 730 | ||
| 731 | lemma "P (x::'a BinTree)" | |
| 14809 | 732 | refute | 
| 14350 | 733 | oops | 
| 734 | ||
| 735 | lemma "\<forall>x::'a BinTree. P x" | |
| 14809 | 736 | refute | 
| 737 | oops | |
| 738 | ||
| 739 | lemma "P (Node (Leaf x) (Leaf y))" | |
| 740 | refute | |
| 14350 | 741 | oops | 
| 742 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 743 | lemma "P (BinTree_rec l n x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 744 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 745 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 746 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 747 | lemma "P (case x of Leaf a \<Rightarrow> l a | Node a b \<Rightarrow> n a b)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 748 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 749 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 750 | |
| 14350 | 751 | subsubsection {* Mutually recursive datatypes *}
 | 
| 752 | ||
| 753 | datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp" | |
| 754 | and 'a bexp = Equal "'a aexp" "'a aexp" | |
| 755 | ||
| 756 | lemma "P (x::'a aexp)" | |
| 14809 | 757 | refute | 
| 14350 | 758 | oops | 
| 759 | ||
| 760 | lemma "\<forall>x::'a aexp. P x" | |
| 14809 | 761 | refute | 
| 14350 | 762 | oops | 
| 763 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 764 | lemma "P (ITE (Equal (Number x) (Number y)) (Number x) (Number y))" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 765 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 766 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 767 | |
| 14350 | 768 | lemma "P (x::'a bexp)" | 
| 14809 | 769 | refute | 
| 14350 | 770 | oops | 
| 771 | ||
| 772 | lemma "\<forall>x::'a bexp. P x" | |
| 14809 | 773 | refute | 
| 14350 | 774 | oops | 
| 775 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 776 | lemma "P (aexp_bexp_rec_1 number ite equal x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 777 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 778 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 779 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 780 | lemma "P (case x of Number a \<Rightarrow> number a | ITE b a1 a2 \<Rightarrow> ite b a1 a2)" | 
| 14809 | 781 | refute | 
| 14350 | 782 | oops | 
| 783 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 784 | lemma "P (aexp_bexp_rec_2 number ite equal x)" | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 785 | refute | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 786 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 787 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 788 | lemma "P (case x of Equal a1 a2 \<Rightarrow> equal a1 a2)" | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 789 | refute | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 790 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 791 | |
| 14350 | 792 | subsubsection {* Other datatype examples *}
 | 
| 793 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 794 | datatype Trie = TR "Trie list" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 795 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 796 | lemma "P (x::Trie)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 797 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 798 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 799 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 800 | lemma "\<forall>x::Trie. P x" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 801 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 802 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 803 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 804 | lemma "P (TR [TR []])" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 805 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 806 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 807 | |
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 808 | lemma "P (Trie_rec_1 a b c x)" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 809 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 810 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 811 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 812 | lemma "P (Trie_rec_2 a b c x)" | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 813 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 814 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 815 | |
| 14809 | 816 | datatype InfTree = Leaf | Node "nat \<Rightarrow> InfTree" | 
| 14350 | 817 | |
| 818 | lemma "P (x::InfTree)" | |
| 14809 | 819 | refute | 
| 14350 | 820 | oops | 
| 821 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 822 | lemma "\<forall>x::InfTree. P x" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 823 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 824 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 825 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 826 | lemma "P (Node (\<lambda>n. Leaf))" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 827 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 828 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 829 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 830 | lemma "P (InfTree_rec leaf node x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 831 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 832 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 833 | |
| 14350 | 834 | datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a \<Rightarrow> 'a lambda" | 
| 835 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 836 | lemma "P (x::'a lambda)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 837 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 838 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 839 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 840 | lemma "\<forall>x::'a lambda. P x" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 841 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 842 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 843 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 844 | lemma "P (Lam (\<lambda>a. Var a))" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 845 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 846 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 847 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 848 | lemma "P (lambda_rec v a l x)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 849 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 850 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 851 | |
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 852 | text {* Taken from "Inductive datatypes in HOL", p.8: *}
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 853 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 854 | datatype ('a, 'b) T = C "'a \<Rightarrow> bool" | D "'b list"
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 855 | datatype 'c U = E "('c, 'c U) T"
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 856 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 857 | lemma "P (x::'c U)" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 858 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 859 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 860 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 861 | lemma "\<forall>x::'c U. P x" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 862 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 863 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 864 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 865 | lemma "P (E (C (\<lambda>a. True)))" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 866 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 867 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 868 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 869 | lemma "P (U_rec_1 e f g h i x)" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 870 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 871 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 872 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 873 | lemma "P (U_rec_2 e f g h i x)" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 874 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 875 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 876 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 877 | lemma "P (U_rec_3 e f g h i x)" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 878 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 879 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 880 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 881 | (******************************************************************************) | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 882 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 883 | subsection {* Records *}
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 884 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 885 | (*TODO: make use of pair types, rather than typedef, for record types*) | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 886 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 887 | record ('a, 'b) point =
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 888 | xpos :: 'a | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 889 | ypos :: 'b | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 890 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 891 | lemma "(x::('a, 'b) point) = y"
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 892 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 893 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 894 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 895 | record ('a, 'b, 'c) extpoint = "('a, 'b) point" +
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 896 | ext :: 'c | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 897 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 898 | lemma "(x::('a, 'b, 'c) extpoint) = y"
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 899 | refute | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 900 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 901 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 902 | (******************************************************************************) | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 903 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 904 | subsection {* Inductively defined sets *}
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 905 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 906 | consts | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 907 | arbitrarySet :: "'a set" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 908 | inductive arbitrarySet | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 909 | intros | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 910 | "arbitrary : arbitrarySet" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 911 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 912 | lemma "x : arbitrarySet" | 
| 16050 | 913 | refute | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 914 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 915 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 916 | consts | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 917 | evenCard :: "'a set set" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 918 | inductive evenCard | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 919 | intros | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 920 |   "{} : evenCard"
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 921 |   "\<lbrakk> S : evenCard; x \<notin> S; y \<notin> S; x \<noteq> y \<rbrakk> \<Longrightarrow> S \<union> {x, y} : evenCard"
 | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 922 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 923 | lemma "S : evenCard" | 
| 16050 | 924 | refute | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 925 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 926 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 927 | consts | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 928 | even :: "nat set" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 929 | odd :: "nat set" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 930 | inductive even odd | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 931 | intros | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 932 | "0 : even" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 933 | "n : even \<Longrightarrow> Suc n : odd" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 934 | "n : odd \<Longrightarrow> Suc n : even" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 935 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 936 | lemma "n : odd" | 
| 16050 | 937 |   (*refute*)  -- {* unfortunately, this little example already takes too long *}
 | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 938 | oops | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 939 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 940 | (******************************************************************************) | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 941 | |
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 942 | subsection {* Examples involving special functions *}
 | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 943 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 944 | lemma "card x = 0" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 945 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 946 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 947 | |
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 948 | lemma "finite x" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 949 |   refute  -- {* no finite countermodel exists *}
 | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 950 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 951 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 952 | lemma "(x::nat) + y = 0" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 953 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 954 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 955 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 956 | lemma "(x::nat) = x + x" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 957 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 958 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 959 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 960 | lemma "(x::nat) - y + y = x" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 961 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 962 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 963 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 964 | lemma "(x::nat) = x * x" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 965 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 966 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 967 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 968 | lemma "(x::nat) < x + y" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 969 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 970 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 971 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 972 | lemma "a @ [] = b @ []" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 973 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 974 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 975 | |
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 976 | lemma "a @ b = b @ a" | 
| 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 977 | refute | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 978 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 979 | |
| 16050 | 980 | lemma "f (lfp f) = lfp f" | 
| 981 | refute | |
| 982 | oops | |
| 983 | ||
| 984 | lemma "f (gfp f) = gfp f" | |
| 985 | refute | |
| 986 | oops | |
| 987 | ||
| 988 | lemma "lfp f = gfp f" | |
| 989 | refute | |
| 990 | oops | |
| 991 | ||
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15547diff
changeset | 992 | (******************************************************************************) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 993 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 994 | subsection {* Axiomatic type classes and overloading *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 995 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 996 | text {* A type class without axioms: *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 997 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 998 | axclass classA | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 999 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1000 | lemma "P (x::'a::classA)" | 
| 14809 | 1001 | refute | 
| 1002 | oops | |
| 1003 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1004 | text {* The axiom of this type class does not contain any type variables, but is internally converted into one that does: *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1005 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1006 | axclass classB | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1007 | classB_ax: "P | ~ P" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1008 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1009 | lemma "P (x::'a::classB)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1010 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1011 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1012 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1013 | text {* An axiom with a type variable (denoting types which have at least two elements): *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1014 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1015 | axclass classC < type | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1016 | classC_ax: "\<exists>x y. x \<noteq> y" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1017 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1018 | lemma "P (x::'a::classC)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1019 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1020 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1021 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1022 | lemma "\<exists>x y. (x::'a::classC) \<noteq> y" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1023 |   refute  -- {* no countermodel exists *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1024 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1025 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1026 | text {* A type class for which a constant is defined: *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1027 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1028 | consts | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1029 | classD_const :: "'a \<Rightarrow> 'a" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1030 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1031 | axclass classD < type | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1032 | classD_ax: "classD_const (classD_const x) = classD_const x" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1033 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1034 | lemma "P (x::'a::classD)" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1035 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1036 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1037 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1038 | text {* A type class with multiple superclasses: *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1039 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1040 | axclass classE < classC, classD | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1041 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1042 | lemma "P (x::'a::classE)" | 
| 14809 | 1043 | refute | 
| 1044 | oops | |
| 1045 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1046 | lemma "P (x::'a::{classB, classE})"
 | 
| 14809 | 1047 | refute | 
| 1048 | oops | |
| 1049 | ||
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1050 | text {* OFCLASS: *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1051 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1052 | lemma "OFCLASS('a::type, type_class)"
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1053 |   refute  -- {* no countermodel exists *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1054 | apply intro_classes | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1055 | done | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1056 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1057 | lemma "OFCLASS('a::classC, type_class)"
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1058 |   refute  -- {* no countermodel exists *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1059 | apply intro_classes | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1060 | done | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1061 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1062 | lemma "OFCLASS('a, classB_class)"
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1063 |   refute  -- {* no countermodel exists *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1064 | apply intro_classes | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1065 | apply simp | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1066 | done | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1067 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1068 | lemma "OFCLASS('a::type, classC_class)"
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1069 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1070 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1071 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1072 | text {* Overloading: *}
 | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1073 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1074 | consts inverse :: "'a \<Rightarrow> 'a" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1075 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1076 | defs (overloaded) | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1077 | inverse_bool: "inverse (b::bool) == ~ b" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1078 | inverse_set : "inverse (S::'a set) == -S" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1079 | inverse_pair: "inverse p == (inverse (fst p), inverse (snd p))" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1080 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1081 | lemma "inverse b" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1082 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1083 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1084 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1085 | lemma "P (inverse (S::'a set))" | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1086 | refute | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1087 | oops | 
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1088 | |
| 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15297diff
changeset | 1089 | lemma "P (inverse (p::'a\<times>'b))" | 
| 14809 | 1090 | refute | 
| 14350 | 1091 | oops | 
| 1092 | ||
| 18774 | 1093 | refute_params [satsolver="auto"] | 
| 1094 | ||
| 14350 | 1095 | end |