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