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