src/HOL/ex/Refute_Examples.thy
author haftmann
Tue, 23 Feb 2010 10:11:12 +0100
changeset 35315 fbdc860d87a3
parent 34120 f9920a3ddf50
child 35416 d8d7d1b785af
permissions -rw-r--r--
dropped axclass
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
    Author:     Tjark Weber
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
     3
    Copyright   2003-2007
32968
wenzelm
parents: 28524
diff changeset
     4
wenzelm
parents: 28524
diff changeset
     5
See HOL/Refute.thy for help.
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     6
*)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     7
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     8
header {* Examples for the 'refute' command *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     9
15297
0aff5d912422 imports (new syntax for theory headers)
webertj
parents: 15161
diff changeset
    10
theory Refute_Examples imports Main
0aff5d912422 imports (new syntax for theory headers)
webertj
parents: 15161
diff changeset
    11
begin
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    12
18774
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
    13
refute_params [satsolver="dpll"]
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
    14
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    15
lemma "P \<and> Q"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    16
  apply (rule conjI)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    17
  refute 1  -- {* refutes @{term "P"} *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    18
  refute 2  -- {* refutes @{term "Q"} *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    19
  refute    -- {* equivalent to 'refute 1' *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    20
    -- {* here 'refute 3' would cause an exception, since we only have 2 subgoals *}
14465
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
    21
  refute [maxsize=5]           -- {* we can override parameters ... *}
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
    22
  refute [satsolver="dpll"] 2  -- {* ... and specify a subgoal at the same time *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    23
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    24
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
    25
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
    26
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
    27
subsection {* Examples and Test Cases *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    28
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
    29
subsubsection {* Propositional logic *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    30
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    31
lemma "True"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    32
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    33
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    34
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    35
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    36
lemma "False"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    37
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    38
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    39
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    40
lemma "P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    41
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    42
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    43
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    44
lemma "~ P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    45
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    46
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    47
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    48
lemma "P & Q"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    49
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    50
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    51
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    52
lemma "P | Q"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    53
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    54
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    55
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    56
lemma "P \<longrightarrow> Q"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    57
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    58
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    59
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    60
lemma "(P::bool) = Q"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    61
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    62
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    63
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    64
lemma "(P | Q) \<longrightarrow> (P & Q)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    65
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    66
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    67
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
    68
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
    69
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
    70
subsubsection {* Predicate logic *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    71
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    72
lemma "P x y z"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    73
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    74
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    75
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    76
lemma "P x y \<longrightarrow> P y x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    77
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    78
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    79
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    80
lemma "P (f (f x)) \<longrightarrow> P x \<longrightarrow> P (f x)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    81
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    82
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    83
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
    84
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
    85
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
    86
subsubsection {* Equality *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    87
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    88
lemma "P = True"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    89
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    90
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    91
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    92
lemma "P = False"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    93
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    94
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    95
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    96
lemma "x = y"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    97
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    98
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    99
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   100
lemma "f x = g x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   101
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   102
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   103
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   104
lemma "(f::'a\<Rightarrow>'b) = g"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   105
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   106
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   107
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   108
lemma "(f::('d\<Rightarrow>'d)\<Rightarrow>('c\<Rightarrow>'d)) = g"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   109
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   110
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   111
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   112
lemma "distinct [a,b]"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   113
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   114
  apply simp
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   115
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   116
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   117
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   118
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   119
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   120
subsubsection {* First-Order Logic *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   121
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   122
lemma "\<exists>x. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   123
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   124
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   125
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   126
lemma "\<forall>x. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   127
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   128
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   129
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   130
lemma "EX! x. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   131
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   132
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   133
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   134
lemma "Ex P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   135
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   136
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   137
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   138
lemma "All P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   139
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   140
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   141
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   142
lemma "Ex1 P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   143
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   144
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   145
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   146
lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   147
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   148
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   149
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   150
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
   151
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   152
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   153
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   154
lemma "(\<exists>x. P x) \<longrightarrow> (EX! x. P x)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   155
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   156
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   157
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   158
text {* A true statement (also testing names of free and bound variables being identical) *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   159
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   160
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
   161
  refute [maxsize=4]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   162
  apply fast
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   163
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   164
18789
8a5c6fc3ad27 smaller example to prevent timeout
webertj
parents: 18774
diff changeset
   165
text {* "A type has at most 4 elements." *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   166
18789
8a5c6fc3ad27 smaller example to prevent timeout
webertj
parents: 18774
diff changeset
   167
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
   168
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   169
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   170
18789
8a5c6fc3ad27 smaller example to prevent timeout
webertj
parents: 18774
diff changeset
   171
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
   172
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   173
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   174
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   175
text {* "Every reflexive and symmetric relation is transitive." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   176
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   177
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
   178
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   179
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   180
14465
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
   181
text {* The "Drinker's theorem" ... *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   182
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   183
lemma "\<exists>x. f x = g x \<longrightarrow> f = g"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   184
  refute [maxsize=4]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   185
  apply (auto simp add: ext)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   186
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   187
14465
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
   188
text {* ... and an incorrect version of it *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   189
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   190
lemma "(\<exists>x. f x = g x) \<longrightarrow> f = g"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   191
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   192
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   193
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   194
text {* "Every function has a fixed point." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   195
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   196
lemma "\<exists>x. f x = x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   197
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   198
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   199
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   200
text {* "Function composition is commutative." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   201
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   202
lemma "f (g x) = g (f x)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   203
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   204
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   205
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   206
text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   207
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   208
lemma "((P::('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   209
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   210
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   211
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   212
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   213
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   214
subsubsection {* Higher-Order Logic *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   215
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   216
lemma "\<exists>P. P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   217
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   218
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   219
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   220
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   221
lemma "\<forall>P. P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   222
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   223
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   224
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   225
lemma "EX! P. P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   226
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   227
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   228
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   229
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   230
lemma "EX! P. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   231
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   232
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   233
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   234
lemma "P Q | Q x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   235
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   236
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   237
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   238
lemma "x \<noteq> All"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   239
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   240
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   241
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   242
lemma "x \<noteq> Ex"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   243
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   244
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   245
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   246
lemma "x \<noteq> Ex1"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   247
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   248
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   249
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   250
text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   251
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   252
constdefs
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   253
  "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   254
  "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
   255
  "subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   256
  "subset P Q == (ALL x y. P x y \<longrightarrow> Q x y)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   257
  "trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   258
  "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
   259
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   260
lemma "trans_closure T P \<longrightarrow> (\<exists>x y. T x y)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   261
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   262
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   263
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   264
text {* "The union of transitive closures is equal to the transitive closure of unions." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   265
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   266
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
   267
        \<longrightarrow> trans_closure TP P
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   268
        \<longrightarrow> trans_closure TR R
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   269
        \<longrightarrow> (T x y = (TP x y | TR x y))"
16910
19b4bf469fb2 minor parameter changes
webertj
parents: 16050
diff changeset
   270
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   271
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   272
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   273
text {* "Every surjective function is invertible." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   274
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   275
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
   276
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   277
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   278
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   279
text {* "Every invertible function is surjective." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   280
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   281
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
   282
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   283
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   284
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   285
text {* Every point is a fixed point of some function. *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   286
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   287
lemma "\<exists>f. f x = x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   288
  refute [maxsize=4]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   289
  apply (rule_tac x="\<lambda>x. x" in exI)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   290
  apply simp
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   291
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   292
14465
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
   293
text {* Axiom of Choice: first an incorrect version ... *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   294
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   295
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
   296
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   297
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   298
14465
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
   299
text {* ... and now two correct ones *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   300
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   301
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
   302
  refute [maxsize=4]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   303
  apply (simp add: choice)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   304
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   305
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   306
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
   307
  refute [maxsize=2]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   308
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   309
    apply (simp add: ex1_implies_ex choice)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   310
  apply (fast intro: ext)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   311
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   312
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   313
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   314
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   315
subsubsection {* Meta-logic *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   316
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   317
lemma "!!x. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   318
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   319
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   320
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   321
lemma "f x == g x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   322
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   323
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   324
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   325
lemma "P \<Longrightarrow> Q"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   326
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   327
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   328
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   329
lemma "\<lbrakk> P; Q; R \<rbrakk> \<Longrightarrow> S"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   330
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   331
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   332
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   333
lemma "(x == all) \<Longrightarrow> False"
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   334
  refute
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   335
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   336
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   337
lemma "(x == (op ==)) \<Longrightarrow> False"
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   338
  refute
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   339
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   340
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   341
lemma "(x == (op \<Longrightarrow>)) \<Longrightarrow> False"
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   342
  refute
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   343
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   344
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   345
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   346
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   347
subsubsection {* Schematic variables *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   348
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   349
lemma "?P"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   350
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   351
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   352
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   353
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   354
lemma "x = ?y"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   355
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   356
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   357
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   358
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   359
(******************************************************************************)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   360
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   361
subsubsection {* Abstractions *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   362
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   363
lemma "(\<lambda>x. x) = (\<lambda>x. y)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   364
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   365
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   366
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   367
lemma "(\<lambda>f. f x) = (\<lambda>f. True)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   368
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   369
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   370
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   371
lemma "(\<lambda>x. x) = (\<lambda>y. y)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   372
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   373
  apply simp
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   374
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   375
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   376
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   377
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   378
subsubsection {* Sets *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   379
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   380
lemma "P (A::'a set)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   381
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   382
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   383
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   384
lemma "P (A::'a set set)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   385
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   386
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   387
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   388
lemma "{x. P x} = {y. P y}"
14489
3676def6b8b9 satsolver=dpll
webertj
parents: 14465
diff changeset
   389
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   390
  apply simp
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   391
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   392
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   393
lemma "x : {x. P x}"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   394
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   395
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   396
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   397
lemma "P op:"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   398
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   399
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   400
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   401
lemma "P (op: x)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   402
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   403
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   404
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   405
lemma "P Collect"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   406
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   407
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   408
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   409
lemma "A Un B = A Int B"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   410
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   411
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   412
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   413
lemma "(A Int B) Un C = (A Un C) Int B"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   414
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   415
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   416
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   417
lemma "Ball A P \<longrightarrow> Bex A P"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   418
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   419
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   420
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   421
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   422
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   423
subsubsection {* undefined *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   424
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   425
lemma "undefined"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   426
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   427
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   428
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   429
lemma "P undefined"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   430
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   431
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   432
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   433
lemma "undefined x"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   434
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   435
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   436
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   437
lemma "undefined undefined"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   438
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   439
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   440
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   441
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   442
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   443
subsubsection {* The *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   444
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   445
lemma "The P"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   446
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   447
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   448
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   449
lemma "P The"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   450
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   451
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   452
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   453
lemma "P (The P)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   454
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   455
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   456
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   457
lemma "(THE x. x=y) = z"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   458
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   459
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   460
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   461
lemma "Ex P \<longrightarrow> P (The P)"
14489
3676def6b8b9 satsolver=dpll
webertj
parents: 14465
diff changeset
   462
  refute
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   463
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   464
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   465
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   466
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   467
subsubsection {* Eps *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   468
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   469
lemma "Eps P"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   470
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   471
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   472
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   473
lemma "P Eps"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   474
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   475
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   476
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   477
lemma "P (Eps P)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   478
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   479
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   480
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   481
lemma "(SOME x. x=y) = z"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   482
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   483
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   484
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   485
lemma "Ex P \<longrightarrow> P (Eps P)"
14489
3676def6b8b9 satsolver=dpll
webertj
parents: 14465
diff changeset
   486
  refute [maxsize=3]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   487
  apply (auto simp add: someI)
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   488
done
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   489
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   490
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   491
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   492
subsubsection {* Subtypes (typedef), typedecl *}
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   493
15161
065ce5385a06 comment modified
webertj
parents: 14809
diff changeset
   494
text {* A completely unspecified non-empty subset of @{typ "'a"}: *}
065ce5385a06 comment modified
webertj
parents: 14809
diff changeset
   495
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   496
typedef 'a myTdef = "insert (undefined::'a) (undefined::'a set)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   497
  by auto
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   498
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   499
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
   500
  refute
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   501
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   502
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   503
typedecl myTdecl
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   504
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   505
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
   506
  by auto
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   507
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   508
lemma "P (f::(myTdecl myTdef) T_bij)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   509
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   510
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   511
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   512
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   513
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   514
subsubsection {* Inductive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   515
21502
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 18789
diff changeset
   516
text {* With @{text quick_and_dirty} set, the datatype package does
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 18789
diff changeset
   517
  not generate certain axioms for recursion operators.  Without these
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 18789
diff changeset
   518
  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
   519
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   520
text {* unit *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   521
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   522
lemma "P (x::unit)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   523
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   524
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   525
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   526
lemma "\<forall>x::unit. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   527
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   528
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   529
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   530
lemma "P ()"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   531
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   532
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   533
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   534
lemma "unit_rec u x = u"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   535
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   536
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   537
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   538
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
   539
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
   540
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   544
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   545
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   546
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   547
text {* option *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   548
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   549
lemma "P (x::'a option)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   550
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   551
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   552
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   553
lemma "\<forall>x::'a option. P x"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   554
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   555
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   556
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   557
lemma "P None"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   558
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   559
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   560
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   561
lemma "P (Some x)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   562
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   563
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   564
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   565
lemma "option_rec n s None = n"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   566
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   567
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   568
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   569
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   570
lemma "option_rec n s (Some x) = s x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   571
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   572
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   573
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   574
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
   575
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
   576
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   577
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   578
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   579
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
   580
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   581
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   582
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   583
text {* * *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   584
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   585
lemma "P (x::'a*'b)"
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 "\<forall>x::'a*'b. P x"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   590
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   591
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   592
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   593
lemma "P (x, y)"
14455
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
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   597
lemma "P (fst x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   598
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   599
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   600
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   601
lemma "P (snd x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   602
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   603
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   604
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   605
lemma "P Pair"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   606
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   607
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   608
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   609
lemma "prod_rec p (a, b) = p a b"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   610
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   611
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   612
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   613
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
   614
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
   615
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   616
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   617
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   618
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
   619
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   620
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   621
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   622
text {* + *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   623
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   624
lemma "P (x::'a+'b)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   625
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   626
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   627
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   628
lemma "\<forall>x::'a+'b. P x"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   629
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   630
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   631
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   632
lemma "P (Inl x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   633
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   634
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   635
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   636
lemma "P (Inr x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   637
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   638
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   639
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   640
lemma "P Inl"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   641
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   642
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   643
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   644
lemma "sum_rec l r (Inl x) = l x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   645
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   646
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   647
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   648
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   649
lemma "sum_rec l r (Inr x) = r x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   650
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   651
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   652
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   653
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
   654
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
   655
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   659
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   660
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   661
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   662
text {* Non-recursive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   663
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   664
datatype T1 = A | B
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   665
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   666
lemma "P (x::T1)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   667
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   668
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   669
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   670
lemma "\<forall>x::T1. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   671
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   672
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   673
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   674
lemma "P A"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   675
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   676
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   677
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   678
lemma "P B"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   679
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   680
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   681
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   682
lemma "T1_rec a b A = a"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   683
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   684
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   685
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   686
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   687
lemma "T1_rec a b B = b"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   688
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   689
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   690
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   691
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
   692
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
   693
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   694
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   695
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   696
lemma "P (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
   697
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   698
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   699
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   700
datatype 'a T2 = C T1 | D 'a
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   701
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   702
lemma "P (x::'a T2)"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   703
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   704
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   705
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   706
lemma "\<forall>x::'a T2. P x"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   707
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   708
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   709
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   710
lemma "P D"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   711
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   712
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   713
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   714
lemma "T2_rec c d (C x) = c x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   715
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   716
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   717
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   718
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   719
lemma "T2_rec c d (D x) = d x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   720
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   721
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   722
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   723
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
   724
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
   725
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   729
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   732
datatype ('a,'b) T3 = E "'a \<Rightarrow> 'b"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   733
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   734
lemma "P (x::('a,'b) T3)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   735
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   736
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   737
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   738
lemma "\<forall>x::('a,'b) T3. P x"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   739
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   740
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   741
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   742
lemma "P E"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   743
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   744
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   745
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   746
lemma "T3_rec e (E x) = e x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   747
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   748
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   749
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   750
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
   751
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
   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 "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
   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
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   759
text {* Recursive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   760
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
   761
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
   762
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   763
lemma "P (x::nat)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   764
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   765
oops
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   766
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   767
lemma "\<forall>x::nat. P x"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   768
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   769
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   770
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   771
lemma "P (Suc 0)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   772
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   773
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   774
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   775
lemma "P Suc"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   776
  refute  -- {* @{term "Suc"} is a partial function (regardless of the size
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   777
                of the model), hence @{term "P Suc"} is undefined, hence no
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   778
                model will be found *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   779
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   780
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   781
lemma "nat_rec zero suc 0 = zero"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   782
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   783
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   784
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   785
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   786
lemma "nat_rec zero suc (Suc x) = suc x (nat_rec zero suc x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   787
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   788
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   789
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   790
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
   791
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
   792
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   793
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   794
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   795
lemma "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
   796
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   797
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   798
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   799
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
   800
f08e2d83681e major code change: 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
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
   802
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   803
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   804
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   805
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
   806
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   807
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   808
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   809
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
   810
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   813
lemma "list_rec nil cons [] = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   814
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   815
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   816
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   817
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   818
lemma "list_rec nil cons (x#xs) = cons x xs (list_rec nil cons xs)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   819
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   820
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   821
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   822
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
   823
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
   824
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   826
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   827
lemma "P (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
   828
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   829
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   830
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   831
lemma "(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
   832
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   833
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   834
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   835
lemma "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
   836
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   837
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   838
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   839
datatype BitList = BitListNil | Bit0 BitList | Bit1 BitList
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   840
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   841
lemma "P (x::BitList)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   842
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   843
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   844
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   845
lemma "\<forall>x::BitList. P x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   846
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   847
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   848
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   849
lemma "P (Bit0 (Bit1 BitListNil))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   850
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   851
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   852
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   853
lemma "BitList_rec nil bit0 bit1 BitListNil = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   854
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   855
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   856
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   857
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   858
lemma "BitList_rec nil bit0 bit1 (Bit0 xs) = bit0 xs (BitList_rec nil bit0 bit1 xs)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   859
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   860
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   861
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   862
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   863
lemma "BitList_rec nil bit0 bit1 (Bit1 xs) = bit1 xs (BitList_rec nil bit0 bit1 xs)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   864
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   865
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   866
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   867
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   868
lemma "P (BitList_rec nil bit0 bit1 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   869
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   870
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   871
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   872
datatype 'a BinTree = Leaf 'a | Node "'a BinTree" "'a BinTree"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   873
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   874
lemma "P (x::'a BinTree)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   875
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   876
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   877
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   878
lemma "\<forall>x::'a BinTree. P x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   879
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   880
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   881
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   882
lemma "P (Node (Leaf x) (Leaf y))"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   883
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   884
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   885
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   886
lemma "BinTree_rec l n (Leaf x) = l x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   887
  refute [maxsize=1]  (* The "maxsize=1" tests are a bit pointless: for some formulae
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   888
                         below, refute will find no countermodel simply because this
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   889
                         size makes involved terms undefined.  Unfortunately, any
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   890
                         larger size already takes too long. *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   891
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   892
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   893
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   894
lemma "BinTree_rec l n (Node x y) = n x y (BinTree_rec l n x) (BinTree_rec l n y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   895
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   896
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   897
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   898
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
   899
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
   900
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   901
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   902
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   903
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
   904
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   905
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   906
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   907
text {* Mutually recursive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   908
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   909
datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   910
     and 'a bexp = Equal "'a aexp" "'a aexp"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   911
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   912
lemma "P (x::'a aexp)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   913
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   914
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   915
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   916
lemma "\<forall>x::'a aexp. P x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   917
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   918
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   919
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
   920
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
   921
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   922
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   923
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   924
lemma "P (x::'a bexp)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   925
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   926
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   927
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   928
lemma "\<forall>x::'a bexp. P x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   929
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   930
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   931
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   932
lemma "aexp_bexp_rec_1 number ite equal (Number x) = number x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   933
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   934
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   935
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   936
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   937
lemma "aexp_bexp_rec_1 number ite equal (ITE x y z) = ite x y z (aexp_bexp_rec_2 number ite equal x) (aexp_bexp_rec_1 number ite equal y) (aexp_bexp_rec_1 number ite equal z)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   938
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   939
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   940
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   941
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
   942
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
   943
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   944
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   945
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   946
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
   947
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   948
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   949
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   950
lemma "aexp_bexp_rec_2 number ite equal (Equal x y) = equal x y (aexp_bexp_rec_1 number ite equal x) (aexp_bexp_rec_1 number ite equal y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   951
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   952
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   953
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   954
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
   955
lemma "P (aexp_bexp_rec_2 number ite equal x)"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   956
  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
   957
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   958
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   959
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
   960
  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
   961
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   962
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   963
datatype X = A | B X | C Y
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   964
     and Y = D X | E Y | F
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   965
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   966
lemma "P (x::X)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   967
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   968
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   969
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   970
lemma "P (y::Y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   971
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   972
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   973
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   974
lemma "P (B (B A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   975
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   976
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   977
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   978
lemma "P (B (C F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   979
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   980
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   981
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   982
lemma "P (C (D A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   983
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   984
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   985
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   986
lemma "P (C (E F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   987
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   988
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   989
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   990
lemma "P (D (B A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   991
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   992
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   993
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   994
lemma "P (D (C F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   995
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   996
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   997
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   998
lemma "P (E (D A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   999
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1000
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1001
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1002
lemma "P (E (E F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1003
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1004
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1005
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1006
lemma "P (C (D (C F)))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1007
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1008
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1009
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1010
lemma "X_Y_rec_1 a b c d e f A = a"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1011
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1012
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1013
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1015
lemma "X_Y_rec_1 a b c d e f (B x) = b x (X_Y_rec_1 a b c d e f x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1016
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1017
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1018
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1019
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1020
lemma "X_Y_rec_1 a b c d e f (C y) = c y (X_Y_rec_2 a b c d e f y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1021
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1022
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1023
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1024
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1025
lemma "X_Y_rec_2 a b c d e f (D x) = d x (X_Y_rec_1 a b c d e f x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1026
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1027
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1028
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1029
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1030
lemma "X_Y_rec_2 a b c d e f (E y) = e y (X_Y_rec_2 a b c d e f y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1031
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1032
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1033
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1034
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1035
lemma "X_Y_rec_2 a b c d e f F = f"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1036
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1037
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1038
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1039
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1040
lemma "P (X_Y_rec_1 a b c d e f x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1041
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1042
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1043
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1044
lemma "P (X_Y_rec_2 a b c d e f y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1045
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1046
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1047
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1048
text {* Other datatype examples *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1049
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1050
text {* Indirect recursion is implemented via mutual recursion. *}
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1051
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1052
datatype XOpt = CX "XOpt option" | DX "bool \<Rightarrow> XOpt option"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1053
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1054
lemma "P (x::XOpt)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1055
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1056
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1057
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1058
lemma "P (CX None)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1059
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1060
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1061
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1062
lemma "P (CX (Some (CX None)))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1063
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1064
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1065
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1066
lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (CX x) = cx x (XOpt_rec_2 cx dx n1 s1 n2 s2 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1067
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1068
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1069
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1070
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1071
lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\<lambda>b. XOpt_rec_3 cx dx n1 s1 n2 s2 (x b))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1072
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1073
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1074
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1075
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1076
lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 None = n1"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1077
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1078
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1079
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1080
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1081
lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 (Some x) = s1 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1082
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1083
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1084
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1085
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1086
lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 None = n2"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1087
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1088
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1089
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1090
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1091
lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 (Some x) = s2 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1092
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1093
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1094
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1095
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1096
lemma "P (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1097
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1098
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1099
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1100
lemma "P (XOpt_rec_2 cx dx n1 s1 n2 s2 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1101
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1102
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1103
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1104
lemma "P (XOpt_rec_3 cx dx n1 s1 n2 s2 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1105
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1106
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1107
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1108
datatype 'a YOpt = CY "('a \<Rightarrow> 'a YOpt) option"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1109
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1110
lemma "P (x::'a YOpt)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1111
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1112
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1113
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1114
lemma "P (CY None)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1115
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1116
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1117
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1118
lemma "P (CY (Some (\<lambda>a. CY None)))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1119
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1120
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1121
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1122
lemma "YOpt_rec_1 cy n s (CY x) = cy x (YOpt_rec_2 cy n s x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1123
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1124
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1125
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1126
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1127
lemma "YOpt_rec_2 cy n s None = n"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1128
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1129
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1130
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1131
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1132
lemma "YOpt_rec_2 cy n s (Some x) = s x (\<lambda>a. YOpt_rec_1 cy n s (x a))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1133
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1134
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1135
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1136
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1137
lemma "P (YOpt_rec_1 cy n s x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1138
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1139
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1140
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1141
lemma "P (YOpt_rec_2 cy n s x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1142
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1143
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1144
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
  1145
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
  1146
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1147
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
  1148
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1149
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1150
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1151
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
  1152
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1153
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1154
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1155
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
  1156
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1157
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1158
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1159
lemma "Trie_rec_1 tr nil cons (TR x) = tr x (Trie_rec_2 tr nil cons x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1160
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1161
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1162
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1163
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1164
lemma "Trie_rec_2 tr nil cons [] = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1165
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1166
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1167
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1168
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1169
lemma "Trie_rec_2 tr nil cons (x#xs) = cons x xs (Trie_rec_1 tr nil cons x) (Trie_rec_2 tr nil cons xs)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1170
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1171
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1172
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1173
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1174
lemma "P (Trie_rec_1 tr nil cons x)"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1175
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1176
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1177
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1178
lemma "P (Trie_rec_2 tr nil cons 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
  1179
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1180
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1181
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1182
datatype InfTree = Leaf | Node "nat \<Rightarrow> InfTree"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1183
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1184
lemma "P (x::InfTree)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1185
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1186
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1187
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
  1188
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
  1189
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1190
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1191
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1192
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
  1193
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1194
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1195
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1196
lemma "InfTree_rec leaf node Leaf = leaf"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1197
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1198
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1199
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1200
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1201
lemma "InfTree_rec leaf node (Node x) = node x (\<lambda>n. InfTree_rec leaf node (x n))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1202
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1203
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1204
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1205
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
  1206
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
  1207
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1208
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1209
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1210
datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a \<Rightarrow> 'a lambda"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1211
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
  1212
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
  1213
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1214
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1215
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1216
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
  1217
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1218
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1219
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1220
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
  1221
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1222
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1223
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1224
lemma "lambda_rec var app lam (Var x) = var x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1225
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1226
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1227
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1228
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1229
lemma "lambda_rec var app lam (App x y) = app x y (lambda_rec var app lam x) (lambda_rec var app lam y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1230
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1231
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1232
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1233
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1234
lemma "lambda_rec var app lam (Lam x) = lam x (\<lambda>a. lambda_rec var app lam (x a))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1235
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1236
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1237
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1238
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
  1239
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
  1240
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1241
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1242
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1243
text {* Taken from "Inductive datatypes in HOL", p.8: *}
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1244
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1245
datatype ('a, 'b) T = C "'a \<Rightarrow> bool" | D "'b list"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1246
datatype 'c U = E "('c, 'c U) T"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1247
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1248
lemma "P (x::'c U)"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1249
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1250
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1251
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1252
lemma "\<forall>x::'c U. P x"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1253
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1254
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1255
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1256
lemma "P (E (C (\<lambda>a. True)))"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1257
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1258
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1259
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1260
lemma "U_rec_1 e c d nil cons (E x) = e x (U_rec_2 e c d nil cons x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1261
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1262
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1263
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1264
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1265
lemma "U_rec_2 e c d nil cons (C x) = c x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1266
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1267
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1268
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1269
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1270
lemma "U_rec_2 e c d nil cons (D x) = d x (U_rec_3 e c d nil cons x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1271
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1272
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1273
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1274
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1275
lemma "U_rec_3 e c d nil cons [] = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1276
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1277
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1278
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1279
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1280
lemma "U_rec_3 e c d nil cons (x#xs) = cons x xs (U_rec_1 e c d nil cons x) (U_rec_3 e c d nil cons xs)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1281
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1282
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1283
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1284
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1285
lemma "P (U_rec_1 e c d nil cons x)"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1286
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1287
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1288
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1289
lemma "P (U_rec_2 e c d nil cons x)"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1290
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1291
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1292
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1293
lemma "P (U_rec_3 e c d nil cons x)"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1294
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1295
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1296
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1297
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1298
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1299
subsubsection {* Records *}
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1300
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1301
(*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
  1302
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1303
record ('a, 'b) point =
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1304
  xpos :: 'a
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1305
  ypos :: 'b
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1306
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1307
lemma "(x::('a, 'b) point) = y"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1308
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1309
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1310
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1311
record ('a, 'b, 'c) extpoint = "('a, 'b) point" +
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1312
  ext :: 'c
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1313
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1314
lemma "(x::('a, 'b, 'c) extpoint) = y"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1315
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1316
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1317
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1318
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1319
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1320
subsubsection {* Inductively defined sets *}
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1321
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1322
inductive_set arbitrarySet :: "'a set"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1323
where
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
  1324
  "undefined : arbitrarySet"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1325
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1326
lemma "x : arbitrarySet"
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1327
  refute
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1328
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1329
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1330
inductive_set evenCard :: "'a set set"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1331
where
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1332
  "{} : evenCard"
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1333
| "\<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
  1334
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1335
lemma "S : evenCard"
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1336
  refute
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1337
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1338
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1339
inductive_set
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1340
  even :: "nat set"
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1341
  and odd  :: "nat set"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1342
where
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1343
  "0 : even"
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1344
| "n : even \<Longrightarrow> Suc n : odd"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1345
| "n : odd \<Longrightarrow> Suc n : even"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1346
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1347
lemma "n : odd"
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1348
  (*refute*)  (* TODO: there seems to be an issue here with undefined terms
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1349
                       because of the recursive datatype "nat" *)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1350
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1351
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1352
consts f :: "'a \<Rightarrow> 'a"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1353
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1354
inductive_set
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1355
  a_even :: "'a set"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1356
  and a_odd :: "'a set"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1357
where
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
  1358
  "undefined : a_even"
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1359
| "x : a_even \<Longrightarrow> f x : a_odd"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1360
| "x : a_odd \<Longrightarrow> f x : a_even"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1361
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1362
lemma "x : a_odd"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1363
  refute  -- {* finds a model of size 2, as expected *}
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1364
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1365
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1366
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1367
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1368
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
  1369
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1370
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
  1371
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1372
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1373
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1374
lemma "finite x"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1375
  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
  1376
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1377
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1378
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
  1379
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1380
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1381
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1382
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
  1383
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1384
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1385
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1386
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
  1387
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1388
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1389
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1390
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
  1391
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1392
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1393
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1394
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
  1395
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1396
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1397
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
  1398
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
  1399
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1400
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1401
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
  1402
lemma "xs @ ys = ys @ xs"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1403
  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
  1404
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1405
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1406
lemma "f (lfp f) = lfp f"
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1407
  refute
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1408
oops
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1409
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1410
lemma "f (gfp f) = gfp f"
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1411
  refute
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1412
oops
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1413
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1414
lemma "lfp f = gfp f"
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1415
  refute
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1416
oops
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1417
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1418
(*****************************************************************************)
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
  1419
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1420
subsubsection {* 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
  1421
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1422
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
  1423
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1424
class classA
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
  1425
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1426
lemma "P (x::'a::classA)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1427
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1428
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1429
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
  1430
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
  1431
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1432
class classC =
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1433
  assumes classC_ax: "\<exists>x y. x \<noteq> 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
  1434
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1435
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
  1436
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1437
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1438
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1439
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
  1440
  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
  1441
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1442
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1443
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
  1444
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1445
class classD =
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1446
  fixes classD_const :: "'a \<Rightarrow> 'a"
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1447
  assumes classD_ax: "classD_const (classD_const x) = classD_const 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
  1448
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1449
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
  1450
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1451
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1452
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1453
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
  1454
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1455
class classE = classC + classD
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
  1456
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1457
lemma "P (x::'a::classE)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1458
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1459
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1460
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
  1461
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
  1462
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1463
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
  1464
  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
  1465
  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
  1466
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
  1467
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1468
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
  1469
  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
  1470
  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
  1471
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
  1472
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1473
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
  1474
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1475
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1476
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1477
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
  1478
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1479
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
  1480
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1481
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
  1482
  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
  1483
  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
  1484
  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
  1485
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1486
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
  1487
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1488
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1489
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1490
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
  1491
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1492
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1493
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1494
lemma "P (inverse (p::'a\<times>'b))"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1495
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1496
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1497
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1498
text {* Structured proofs *}
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1499
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1500
lemma "x = y"
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1501
proof cases
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1502
  assume "x = y"
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1503
  show ?thesis
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1504
  refute
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1505
  refute [no_assms]
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1506
  refute [no_assms = false]
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1507
oops
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1508
18774
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
  1509
refute_params [satsolver="auto"]
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
  1510
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1511
end