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