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