src/HOL/ex/Refute_Examples.thy
author krauss
Tue, 28 Sep 2010 09:54:07 +0200
changeset 39754 150f831ce4a3
parent 37388 793618618f78
child 45694 4a8743618257
permissions -rw-r--r--
no longer declare .psimps rules as [simp]. This regularly caused confusion (e.g., they show up in simp traces when the regular simp rules are disabled). In the few places where the rules are used, explicitly mentioning them actually clarifies the proof text.
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
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   252
definition "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   253
  "trans P == (ALL x y z. P x y \<longrightarrow> P y z \<longrightarrow> P x z)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   254
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   255
definition "subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   256
  "subset P Q == (ALL x y. P x y \<longrightarrow> Q x y)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   257
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   258
definition "trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   259
  "trans_closure P Q == (subset Q P) & (trans P) & (ALL R. subset Q R \<longrightarrow> trans R \<longrightarrow> subset P R)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   260
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   261
lemma "trans_closure T P \<longrightarrow> (\<exists>x y. T x y)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   262
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   263
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   264
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   265
text {* "The union of transitive closures is equal to the transitive closure of unions." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   266
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   267
lemma "(\<forall>x y. (P x y | R x y) \<longrightarrow> T x y) \<longrightarrow> trans T \<longrightarrow> (\<forall>Q. (\<forall>x y. (P x y | R x y) \<longrightarrow> Q x y) \<longrightarrow> trans Q \<longrightarrow> subset T Q)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   268
        \<longrightarrow> trans_closure TP P
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   269
        \<longrightarrow> trans_closure TR R
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   270
        \<longrightarrow> (T x y = (TP x y | TR x y))"
16910
19b4bf469fb2 minor parameter changes
webertj
parents: 16050
diff changeset
   271
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   272
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   273
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   274
text {* "Every surjective function is invertible." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   275
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   276
lemma "(\<forall>y. \<exists>x. y = f x) \<longrightarrow> (\<exists>g. \<forall>x. g (f x) = x)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   277
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   278
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   279
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   280
text {* "Every invertible function is surjective." *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   281
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   282
lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   283
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   284
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   285
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   286
text {* Every point is a fixed point of some function. *}
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   287
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   288
lemma "\<exists>f. f x = x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   289
  refute [maxsize=4]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   290
  apply (rule_tac x="\<lambda>x. x" in exI)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   291
  apply simp
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   292
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   293
14465
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
   294
text {* Axiom of Choice: first an incorrect version ... *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   295
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   296
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (EX!f. \<forall>x. P x (f x))"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   297
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   298
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   299
14465
8cc21ed7ef41 \<dots> replaced by ...
webertj
parents: 14462
diff changeset
   300
text {* ... and now two correct ones *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   301
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   302
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>f. \<forall>x. P x (f x))"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   303
  refute [maxsize=4]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   304
  apply (simp add: choice)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   305
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   306
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   307
lemma "(\<forall>x. EX!y. P x y) \<longrightarrow> (EX!f. \<forall>x. P x (f x))"
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   308
  refute [maxsize=2]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   309
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   310
    apply (simp add: ex1_implies_ex choice)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   311
  apply (fast intro: ext)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   312
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   313
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   314
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   315
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   316
subsubsection {* Meta-logic *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   317
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   318
lemma "!!x. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   319
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   320
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   321
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   322
lemma "f x == g x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   323
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   324
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   325
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   326
lemma "P \<Longrightarrow> Q"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   327
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   328
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   329
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   330
lemma "\<lbrakk> P; Q; R \<rbrakk> \<Longrightarrow> S"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   331
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   332
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   333
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   334
lemma "(x == all) \<Longrightarrow> False"
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   335
  refute
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   336
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   337
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   338
lemma "(x == (op ==)) \<Longrightarrow> False"
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   339
  refute
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   340
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   341
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   342
lemma "(x == (op \<Longrightarrow>)) \<Longrightarrow> False"
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   343
  refute
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   344
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   345
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   346
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   347
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   348
subsubsection {* Schematic variables *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   349
36319
8feb2c4bef1a mark schematic statements explicitly;
wenzelm
parents: 36131
diff changeset
   350
schematic_lemma "?P"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   351
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   352
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   353
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   354
36319
8feb2c4bef1a mark schematic statements explicitly;
wenzelm
parents: 36131
diff changeset
   355
schematic_lemma "x = ?y"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   356
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   357
  apply auto
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   358
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   359
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   360
(******************************************************************************)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   361
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   362
subsubsection {* Abstractions *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   363
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   364
lemma "(\<lambda>x. x) = (\<lambda>x. y)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   365
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   366
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   367
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   368
lemma "(\<lambda>f. f x) = (\<lambda>f. True)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   369
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   370
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   371
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   372
lemma "(\<lambda>x. x) = (\<lambda>y. y)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   373
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   374
  apply simp
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   375
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   376
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   377
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   378
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   379
subsubsection {* Sets *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   380
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   381
lemma "P (A::'a set)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   382
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   383
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   384
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   385
lemma "P (A::'a set set)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   386
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   387
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   388
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   389
lemma "{x. P x} = {y. P y}"
14489
3676def6b8b9 satsolver=dpll
webertj
parents: 14465
diff changeset
   390
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   391
  apply simp
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   392
done
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   393
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   394
lemma "x : {x. P x}"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   395
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   396
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   397
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   398
lemma "P op:"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   399
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   400
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   401
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   402
lemma "P (op: x)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   403
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   404
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   405
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   406
lemma "P Collect"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   407
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   408
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   409
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   410
lemma "A Un B = A Int B"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   411
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   412
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   413
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   414
lemma "(A Int B) Un C = (A Un C) Int B"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   415
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   416
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   417
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   418
lemma "Ball A P \<longrightarrow> Bex A P"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   419
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   420
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   421
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   422
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   423
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   424
subsubsection {* undefined *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   425
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   426
lemma "undefined"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   427
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   428
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   429
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   430
lemma "P undefined"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   431
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   432
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   433
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   434
lemma "undefined x"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   435
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   436
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   437
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   438
lemma "undefined undefined"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   439
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   440
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   441
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   442
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   443
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   444
subsubsection {* The *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   445
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   446
lemma "The P"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   447
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   448
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   449
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   450
lemma "P The"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   451
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   452
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   453
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   454
lemma "P (The P)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   455
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   456
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   457
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   458
lemma "(THE x. x=y) = z"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   459
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   460
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   461
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   462
lemma "Ex P \<longrightarrow> P (The P)"
14489
3676def6b8b9 satsolver=dpll
webertj
parents: 14465
diff changeset
   463
  refute
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   464
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   465
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   466
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   467
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   468
subsubsection {* Eps *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   469
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   470
lemma "Eps P"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   471
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   472
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   473
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   474
lemma "P Eps"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   475
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   476
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   477
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   478
lemma "P (Eps P)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   479
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   480
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   481
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   482
lemma "(SOME x. x=y) = z"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   483
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   484
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   485
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   486
lemma "Ex P \<longrightarrow> P (Eps P)"
14489
3676def6b8b9 satsolver=dpll
webertj
parents: 14465
diff changeset
   487
  refute [maxsize=3]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   488
  apply (auto simp add: someI)
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   489
done
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   490
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   491
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   492
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   493
subsubsection {* Subtypes (typedef), typedecl *}
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   494
15161
065ce5385a06 comment modified
webertj
parents: 14809
diff changeset
   495
text {* A completely unspecified non-empty subset of @{typ "'a"}: *}
065ce5385a06 comment modified
webertj
parents: 14809
diff changeset
   496
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   497
typedef 'a myTdef = "insert (undefined::'a) (undefined::'a set)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   498
  by auto
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   499
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   500
lemma "(x::'a myTdef) = y"
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   501
  refute
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   502
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   503
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   504
typedecl myTdecl
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   505
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   506
typedef 'a T_bij = "{(f::'a\<Rightarrow>'a). \<forall>y. \<exists>!x. f x = y}"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   507
  by auto
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   508
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   509
lemma "P (f::(myTdecl myTdef) T_bij)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   510
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   511
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   512
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   513
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   514
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   515
subsubsection {* Inductive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   516
21502
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 18789
diff changeset
   517
text {* With @{text quick_and_dirty} set, the datatype package does
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 18789
diff changeset
   518
  not generate certain axioms for recursion operators.  Without these
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 18789
diff changeset
   519
  axioms, refute may find spurious countermodels. *}
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   520
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   521
text {* unit *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   522
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   523
lemma "P (x::unit)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   524
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   525
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   526
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   527
lemma "\<forall>x::unit. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   528
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   529
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   530
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   531
lemma "P ()"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   532
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   533
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   534
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   535
lemma "unit_rec u x = u"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   536
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   537
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   538
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   539
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
   540
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
   541
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   545
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   547
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   548
text {* option *}
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   549
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   550
lemma "P (x::'a option)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   551
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   552
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   553
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   554
lemma "\<forall>x::'a option. P x"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   555
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   556
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   557
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   558
lemma "P None"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   559
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   560
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   561
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   562
lemma "P (Some x)"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   563
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   564
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   565
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   566
lemma "option_rec n s None = n"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   567
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   568
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   569
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   570
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   571
lemma "option_rec n s (Some x) = s x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   572
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   573
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   574
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   575
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
   576
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
   577
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   581
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   583
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   584
text {* * *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   585
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   586
lemma "P (x::'a*'b)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   587
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   588
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   589
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   590
lemma "\<forall>x::'a*'b. P x"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   591
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   592
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   593
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   594
lemma "P (x, y)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   595
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   596
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   597
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   598
lemma "P (fst x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   599
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   600
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   601
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   602
lemma "P (snd x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   603
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   604
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   605
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   606
lemma "P Pair"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   607
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   608
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   609
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   610
lemma "prod_rec p (a, b) = p a b"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   611
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   612
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   613
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   614
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   615
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
   616
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   617
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   618
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   619
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
   620
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   621
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   622
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   623
text {* + *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   624
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   625
lemma "P (x::'a+'b)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   626
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   627
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   628
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   629
lemma "\<forall>x::'a+'b. P x"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   630
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   631
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   632
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   633
lemma "P (Inl x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   634
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   635
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   636
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   637
lemma "P (Inr x)"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   638
  refute
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   639
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   640
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   641
lemma "P Inl"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   642
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   643
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   644
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   645
lemma "sum_rec l r (Inl x) = l x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   646
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   647
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   648
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   649
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   650
lemma "sum_rec l r (Inr x) = r x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   651
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   652
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   653
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   654
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
   655
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
   656
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   657
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   658
f08e2d83681e major code change: 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
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
   660
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   661
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   662
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   663
text {* Non-recursive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   664
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   665
datatype T1 = A | B
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   666
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   667
lemma "P (x::T1)"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   668
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   669
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   670
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   671
lemma "\<forall>x::T1. P x"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   672
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   673
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   674
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   675
lemma "P A"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   676
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   677
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   678
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   679
lemma "P B"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   680
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   681
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   682
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   683
lemma "T1_rec a b A = a"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   684
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   685
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   686
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   687
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   688
lemma "T1_rec a b B = b"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   689
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   690
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   691
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   692
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
   693
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
   694
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   695
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   696
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   697
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
   698
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   700
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   701
datatype 'a T2 = C T1 | D 'a
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   702
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   703
lemma "P (x::'a T2)"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   704
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   705
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   706
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   707
lemma "\<forall>x::'a T2. P x"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   708
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   709
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   710
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   711
lemma "P D"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   712
  refute
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   713
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   714
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   715
lemma "T2_rec c d (C x) = c x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   716
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   717
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   718
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   719
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   720
lemma "T2_rec c d (D x) = d x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   721
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   722
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   723
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   724
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   725
lemma "P (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
   726
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   727
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   728
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   729
lemma "P (case x of 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
   730
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   731
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   732
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   733
datatype ('a,'b) T3 = E "'a \<Rightarrow> 'b"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   734
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   735
lemma "P (x::('a,'b) T3)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   736
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   737
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   738
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   739
lemma "\<forall>x::('a,'b) T3. P x"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   740
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   741
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   742
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   743
lemma "P E"
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   744
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   745
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   746
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   747
lemma "T3_rec e (E x) = e x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   748
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   749
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   750
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   751
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
   752
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
   753
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   754
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   755
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   756
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
   757
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   758
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   759
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   760
text {* Recursive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   761
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   762
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
   763
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   764
lemma "P (x::nat)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   765
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   766
oops
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   767
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   768
lemma "\<forall>x::nat. P x"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   769
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   770
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   771
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   772
lemma "P (Suc 0)"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   773
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   774
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   775
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   776
lemma "P Suc"
37388
793618618f78 tuned quotes, antiquotations and whitespace
haftmann
parents: 36319
diff changeset
   777
  refute  -- {* @{term Suc} is a partial function (regardless of the size
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   778
                of the model), hence @{term "P Suc"} is undefined, hence no
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   779
                model will be found *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   780
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   781
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   782
lemma "nat_rec zero suc 0 = zero"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   783
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   784
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   785
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   786
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   787
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
   788
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   789
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   790
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   791
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   792
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
   793
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   797
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   801
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   802
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
   803
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   804
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   805
f08e2d83681e major code change: 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
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
   807
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   810
lemma "P [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
   811
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   812
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   813
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   814
lemma "list_rec nil cons [] = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   815
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   816
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   817
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   818
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   819
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
   820
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   821
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   822
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   823
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   824
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
   825
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   826
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   827
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   828
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
   829
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   830
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   831
f08e2d83681e major code change: 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
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
   833
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   837
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   839
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   840
datatype BitList = BitListNil | Bit0 BitList | Bit1 BitList
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   841
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   842
lemma "P (x::BitList)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   843
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   844
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   845
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   846
lemma "\<forall>x::BitList. P x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   847
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   848
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   849
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   850
lemma "P (Bit0 (Bit1 BitListNil))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   851
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   852
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   853
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   854
lemma "BitList_rec nil bit0 bit1 BitListNil = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   855
  refute [maxsize=4]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   856
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   857
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   858
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   859
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
   860
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   861
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   862
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   863
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   864
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
   865
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   866
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   867
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   868
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   869
lemma "P (BitList_rec nil bit0 bit1 x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   870
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   871
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   872
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   873
datatype 'a BinTree = Leaf 'a | Node "'a BinTree" "'a BinTree"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   874
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   875
lemma "P (x::'a BinTree)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   876
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   877
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   878
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   879
lemma "\<forall>x::'a BinTree. P x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   880
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   881
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   882
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   883
lemma "P (Node (Leaf x) (Leaf y))"
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   884
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   885
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   886
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   887
lemma "BinTree_rec l n (Leaf x) = l x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   888
  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
   889
                         below, refute will find no countermodel simply because this
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   890
                         size makes involved terms undefined.  Unfortunately, any
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   891
                         larger size already takes too long. *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   892
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   893
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   894
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   895
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
   896
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   897
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   898
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   899
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
   900
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
   901
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
   905
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   907
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
   908
text {* Mutually recursive datatypes *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   909
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   910
datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   911
     and 'a bexp = Equal "'a aexp" "'a aexp"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   912
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   913
lemma "P (x::'a aexp)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   914
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   915
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   916
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   917
lemma "\<forall>x::'a aexp. P x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   918
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   919
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   920
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
   921
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
   922
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   924
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   925
lemma "P (x::'a bexp)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   926
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   927
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   928
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   929
lemma "\<forall>x::'a bexp. P x"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   930
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   931
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   932
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   933
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
   934
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   935
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   936
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   937
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   938
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
   939
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   940
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   941
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   942
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   943
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
   944
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   947
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
   948
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   949
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   950
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   951
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
   952
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   953
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   954
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   955
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
   956
lemma "P (aexp_bexp_rec_2 number ite equal x)"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   957
  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
   958
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   960
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
   961
  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
   962
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   963
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   964
datatype X = A | B X | C Y
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   965
     and Y = D X | E Y | F
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   966
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   967
lemma "P (x::X)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   968
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   969
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   970
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   971
lemma "P (y::Y)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   972
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   973
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   974
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   975
lemma "P (B (B A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   976
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   977
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   978
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   979
lemma "P (B (C F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   980
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   981
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   982
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   983
lemma "P (C (D A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   984
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   985
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   986
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   987
lemma "P (C (E F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   988
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   989
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   990
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   991
lemma "P (D (B A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   992
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   993
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   994
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   995
lemma "P (D (C F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   996
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   997
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   998
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   999
lemma "P (E (D A))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1000
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1001
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1002
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1003
lemma "P (E (E F))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1004
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1005
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1006
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1007
lemma "P (C (D (C F)))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1008
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1009
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1010
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1011
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
  1012
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1013
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1014
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1015
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1016
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
  1017
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1018
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1019
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1020
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1021
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
  1022
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1023
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1024
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1025
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1026
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
  1027
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1028
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1029
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1030
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1031
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
  1032
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1033
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1034
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1035
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1036
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
  1037
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1038
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1039
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1040
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1041
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
  1042
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1043
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1044
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1045
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
  1046
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1047
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1048
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1049
text {* Other datatype examples *}
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1050
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1051
text {* Indirect recursion is implemented via mutual recursion. *}
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1052
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1053
datatype XOpt = CX "XOpt option" | DX "bool \<Rightarrow> XOpt option"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1054
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1055
lemma "P (x::XOpt)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1056
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1057
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1058
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1059
lemma "P (CX None)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1060
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1061
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1062
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1063
lemma "P (CX (Some (CX None)))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1064
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1065
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1066
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1067
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
  1068
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1069
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1070
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1071
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1072
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
  1073
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1074
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1075
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1076
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1077
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
  1078
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1079
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1080
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1081
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1082
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
  1083
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1084
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1085
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1086
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1087
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
  1088
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1089
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1090
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1091
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1092
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
  1093
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1094
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1095
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1096
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1097
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
  1098
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1099
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1100
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1101
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
  1102
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1103
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1104
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1105
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
  1106
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1107
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1108
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1109
datatype 'a YOpt = CY "('a \<Rightarrow> 'a YOpt) option"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1110
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1111
lemma "P (x::'a YOpt)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1112
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1113
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1114
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1115
lemma "P (CY None)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1116
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1117
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1118
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1119
lemma "P (CY (Some (\<lambda>a. CY None)))"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1120
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1121
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1122
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1123
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
  1124
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1125
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1126
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1127
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1128
lemma "YOpt_rec_2 cy n s None = n"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1129
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1130
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1131
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1132
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1133
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
  1134
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1135
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1136
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1137
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1138
lemma "P (YOpt_rec_1 cy n s x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1139
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1140
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1141
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1142
lemma "P (YOpt_rec_2 cy n s x)"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1143
  refute
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1144
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1145
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
  1146
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
  1147
f08e2d83681e major code change: 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
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
  1149
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1153
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1157
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1159
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1160
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
  1161
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1162
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1163
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1164
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1165
lemma "Trie_rec_2 tr nil cons [] = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1166
  refute [maxsize=3]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1167
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1168
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1169
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1170
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
  1171
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1172
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1173
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1174
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1175
lemma "P (Trie_rec_1 tr nil cons x)"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1176
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1177
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1178
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1179
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
  1180
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1182
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1183
datatype InfTree = Leaf | Node "nat \<Rightarrow> InfTree"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1184
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1185
lemma "P (x::InfTree)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1186
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1187
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1188
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
  1189
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
  1190
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1194
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1196
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1197
lemma "InfTree_rec leaf node Leaf = leaf"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1198
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1199
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1200
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1201
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1202
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
  1203
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1204
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1205
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1206
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
  1207
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
  1208
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1210
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1211
datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a \<Rightarrow> 'a lambda"
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1212
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
  1213
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
  1214
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1218
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1222
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1224
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1225
lemma "lambda_rec var app lam (Var x) = var x"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1226
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1227
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1228
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1229
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1230
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
  1231
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1232
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1233
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1234
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1235
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
  1236
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1237
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1238
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1239
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
  1240
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
  1241
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1243
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1244
text {* Taken from "Inductive datatypes in HOL", p.8: *}
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1245
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1246
datatype ('a, 'b) T = C "'a \<Rightarrow> bool" | D "'b list"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1247
datatype 'c U = E "('c, 'c U) T"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1248
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1249
lemma "P (x::'c U)"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1250
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1251
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1252
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1253
lemma "\<forall>x::'c U. P x"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1254
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1255
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1256
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1257
lemma "P (E (C (\<lambda>a. True)))"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1258
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1259
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1260
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1261
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
  1262
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1263
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1264
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1265
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1266
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
  1267
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1268
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1269
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1270
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1271
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
  1272
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1273
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1274
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1275
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1276
lemma "U_rec_3 e c d nil cons [] = nil"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1277
  refute [maxsize=2]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1278
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1279
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1280
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1281
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
  1282
  refute [maxsize=1]
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1283
  apply simp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1284
done
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1285
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1286
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
  1287
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1288
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1289
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1290
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
  1291
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1292
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1293
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1294
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
  1295
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1296
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1297
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1298
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1299
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1300
subsubsection {* Records *}
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1301
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1302
(*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
  1303
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1304
record ('a, 'b) point =
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1305
  xpos :: 'a
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1306
  ypos :: 'b
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1307
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1308
lemma "(x::('a, 'b) point) = y"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1309
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1310
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1311
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1312
record ('a, 'b, 'c) extpoint = "('a, 'b) point" +
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1313
  ext :: 'c
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1314
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1315
lemma "(x::('a, 'b, 'c) extpoint) = y"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1316
  refute
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1317
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1318
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1319
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1320
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1321
subsubsection {* Inductively defined sets *}
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1322
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1323
inductive_set arbitrarySet :: "'a set"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1324
where
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
  1325
  "undefined : arbitrarySet"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1326
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1327
lemma "x : arbitrarySet"
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1328
  refute
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1329
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1330
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1331
inductive_set evenCard :: "'a set set"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1332
where
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1333
  "{} : evenCard"
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1334
| "\<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
  1335
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1336
lemma "S : evenCard"
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1337
  refute
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1338
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1339
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1340
inductive_set
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1341
  even :: "nat set"
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1342
  and odd  :: "nat set"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1343
where
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1344
  "0 : even"
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1345
| "n : even \<Longrightarrow> Suc n : odd"
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 23219
diff changeset
  1346
| "n : odd \<Longrightarrow> Suc n : even"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1347
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1348
lemma "n : odd"
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1349
  (*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
  1350
                       because of the recursive datatype "nat" *)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1351
oops
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1352
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1353
consts f :: "'a \<Rightarrow> 'a"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1354
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1355
inductive_set
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1356
  a_even :: "'a set"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1357
  and a_odd :: "'a set"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1358
where
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
  1359
  "undefined : a_even"
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1360
| "x : a_even \<Longrightarrow> f x : a_odd"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1361
| "x : a_odd \<Longrightarrow> f x : a_even"
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1362
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1363
lemma "x : a_odd"
36131
64b6374a21a8 adapt Refute example to reflect latest soundness fix to Refute
blanchet
parents: 35416
diff changeset
  1364
  (* refute  -- {* finds a model of size 2, as expected *}
64b6374a21a8 adapt Refute example to reflect latest soundness fix to Refute
blanchet
parents: 35416
diff changeset
  1365
     NO LONGER WORKS since "lfp"'s interpreter is disabled *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1366
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1367
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1368
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1369
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 22284
diff changeset
  1370
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
  1371
f08e2d83681e major code change: 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
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
  1373
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1374
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1375
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1376
lemma "finite x"
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1377
  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
  1378
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1381
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1385
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1389
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1393
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1397
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1398
oops
f08e2d83681e major code change: 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
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
  1400
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
  1401
  refute
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1402
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1403
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
  1404
lemma "xs @ ys = ys @ xs"
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
  1405
  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
  1406
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1407
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1408
lemma "f (lfp f) = lfp f"
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1409
  refute
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1410
oops
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1411
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1412
lemma "f (gfp f) = gfp f"
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1413
  refute
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1414
oops
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1415
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1416
lemma "lfp f = gfp f"
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1417
  refute
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1418
oops
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15767
diff changeset
  1419
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
  1420
(*****************************************************************************)
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
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1422
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
  1423
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1424
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
  1425
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1426
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
  1427
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1428
lemma "P (x::'a::classA)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1429
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1430
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1431
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
  1432
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
  1433
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1434
class classC =
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1435
  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
  1436
f08e2d83681e major code change: 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
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
  1438
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1442
  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
  1443
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1445
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
  1446
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1447
class classD =
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1448
  fixes classD_const :: "'a \<Rightarrow> 'a"
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1449
  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
  1450
f08e2d83681e major code change: 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
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
  1452
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1455
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
  1456
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
  1457
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
  1458
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1459
lemma "P (x::'a::classE)"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1460
  refute
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1461
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1462
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
  1463
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
  1464
f08e2d83681e major code change: 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
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
  1466
  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
  1467
  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
  1468
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
  1469
f08e2d83681e major code change: 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
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
  1471
  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
  1472
  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
  1473
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
  1474
f08e2d83681e major code change: 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
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
  1476
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
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
  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
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
  1482
f08e2d83681e major code change: 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
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
  1484
  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
  1485
  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
  1486
  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
  1487
f08e2d83681e major code change: 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
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
  1489
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
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
  1493
  refute
f08e2d83681e major code change: 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
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1495
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
  1496
lemma "P (inverse (p::'a\<times>'b))"
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
  1497
  refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1498
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1499
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1500
text {* Structured proofs *}
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1501
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1502
lemma "x = y"
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1503
proof cases
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1504
  assume "x = y"
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1505
  show ?thesis
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1506
  refute
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1507
  refute [no_assms]
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1508
  refute [no_assms = false]
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1509
oops
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
  1510
18774
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
  1511
refute_params [satsolver="auto"]
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
  1512
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1513
end