src/HOL/ex/Refute_Examples.thy
author Fabian Huch <huch@in.tum.de>
Fri, 17 Jan 2025 12:17:37 +0100
changeset 81822 e7be7c4b871c
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
isabelle_id: report sync id, if available;
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
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
     8
section \<open>Examples for the 'refute' command\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     9
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    10
theory Refute_Examples
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63901
diff changeset
    11
imports "HOL-Library.Refute"
15297
0aff5d912422 imports (new syntax for theory headers)
webertj
parents: 15161
diff changeset
    12
begin
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    13
56851
35ff4ede3409 renamed 'dpll_p' to 'cdclite', to avoid confusion with the old 'dpll' and to reflect the idea that the new prover implements some ideas from CDCL not in DPLL -- this follows its author's, Sascha B.'s, wish
blanchet
parents: 56845
diff changeset
    14
refute_params [satsolver = "cdclite"]
18774
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
    15
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    16
lemma "P \<and> Q"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    17
apply (rule conjI)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    18
refute [expect = genuine] 1  \<comment> \<open>refutes \<^term>\<open>P\<close>\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    19
refute [expect = genuine] 2  \<comment> \<open>refutes \<^term>\<open>Q\<close>\<close>
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    20
refute [expect = genuine]    \<comment> \<open>equivalent to 'refute 1'\<close>
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    21
  \<comment> \<open>here 'refute 3' would cause an exception, since we only have 2 subgoals\<close>
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    22
refute [maxsize = 5, expect = genuine]   \<comment> \<open>we can override parameters ...\<close>
56851
35ff4ede3409 renamed 'dpll_p' to 'cdclite', to avoid confusion with the old 'dpll' and to reflect the idea that the new prover implements some ideas from CDCL not in DPLL -- this follows its author's, Sascha B.'s, wish
blanchet
parents: 56845
diff changeset
    23
refute [satsolver = "cdclite", expect = genuine] 2
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    24
  \<comment> \<open>... and specify a subgoal at the same time\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    25
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    26
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
    27
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
    28
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    29
subsection \<open>Examples and Test Cases\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    30
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    31
subsubsection \<open>Propositional logic\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    32
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    33
lemma "True"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    34
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    35
by auto
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    36
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    37
lemma "False"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    38
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    39
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    40
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    41
lemma "P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    42
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    43
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    44
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    45
lemma "~ P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    46
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    47
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    48
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    49
lemma "P & Q"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    50
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    51
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    52
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    53
lemma "P | Q"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    54
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    55
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    56
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    57
lemma "P \<longrightarrow> Q"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    58
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    59
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    60
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    61
lemma "(P::bool) = Q"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    62
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    63
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    64
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    65
lemma "(P | Q) \<longrightarrow> (P & Q)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    66
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    67
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    68
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
    69
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
    70
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    71
subsubsection \<open>Predicate logic\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    72
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    73
lemma "P x y z"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    74
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    75
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    76
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    77
lemma "P x y \<longrightarrow> P y x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    78
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    79
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    80
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    81
lemma "P (f (f x)) \<longrightarrow> P x \<longrightarrow> P (f x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    82
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    83
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
    84
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
    85
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
    86
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    87
subsubsection \<open>Equality\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    88
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    89
lemma "P = True"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    90
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    91
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    92
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    93
lemma "P = False"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    94
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    95
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    96
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    97
lemma "x = y"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
    98
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    99
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   100
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   101
lemma "f x = g x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   102
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   103
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   104
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   105
lemma "(f::'a\<Rightarrow>'b) = g"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   106
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   107
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   108
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   109
lemma "(f::('d\<Rightarrow>'d)\<Rightarrow>('c\<Rightarrow>'d)) = g"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   110
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   111
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   112
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   113
lemma "distinct [a, b]"
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   114
(* refute *)
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   115
apply simp
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   116
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   117
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   118
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   119
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   120
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   121
subsubsection \<open>First-Order Logic\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   122
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   123
lemma "\<exists>x. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   124
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   125
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   126
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   127
lemma "\<forall>x. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   128
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   129
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   130
63901
4ce989e962e0 more symbols;
wenzelm
parents: 62148
diff changeset
   131
lemma "\<exists>!x. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   132
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   133
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   134
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   135
lemma "Ex P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   136
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   137
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   138
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   139
lemma "All P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   140
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   141
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   142
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   143
lemma "Ex1 P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   144
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   145
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   146
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   147
lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   148
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   149
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   150
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   151
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>y. \<forall>x. P x y)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   152
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   153
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   154
63901
4ce989e962e0 more symbols;
wenzelm
parents: 62148
diff changeset
   155
lemma "(\<exists>x. P x) \<longrightarrow> (\<exists>!x. P x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   156
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   157
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   158
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   159
text \<open>A true statement (also testing names of free and bound variables being identical)\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   160
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   161
lemma "(\<forall>x y. P x y \<longrightarrow> P y x) \<longrightarrow> (\<forall>x. P x y) \<longrightarrow> P y x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   162
refute [maxsize = 4, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   163
by fast
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   164
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   165
text \<open>"A type has at most 4 elements."\<close>
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"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   168
refute [expect = genuine]
14455
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"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   172
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   173
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   174
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   175
text \<open>"Every reflexive and symmetric relation is transitive."\<close>
14350
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"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   178
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   179
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   180
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   181
text \<open>The "Drinker's theorem" ...\<close>
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"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   184
refute [maxsize = 4, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   185
by (auto simp add: ext)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   186
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   187
text \<open>... and an incorrect version of it\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   188
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   189
lemma "(\<exists>x. f x = g x) \<longrightarrow> f = g"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   190
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   191
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   192
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   193
text \<open>"Every function has a fixed point."\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   194
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   195
lemma "\<exists>x. f x = x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   196
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   197
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   198
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   199
text \<open>"Function composition is commutative."\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   200
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   201
lemma "f (g x) = g (f x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   202
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   203
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   204
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   205
text \<open>"Two functions that are equivalent wrt.\ the same predicate 'P' are equal."\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   206
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   207
lemma "((P::('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   208
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   209
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   210
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   211
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   212
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   213
subsubsection \<open>Higher-Order Logic\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   214
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   215
lemma "\<exists>P. P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   216
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   217
by auto
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   218
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   219
lemma "\<forall>P. P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   220
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   221
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   222
63901
4ce989e962e0 more symbols;
wenzelm
parents: 62148
diff changeset
   223
lemma "\<exists>!P. P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   224
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   225
by auto
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   226
63901
4ce989e962e0 more symbols;
wenzelm
parents: 62148
diff changeset
   227
lemma "\<exists>!P. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   228
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   229
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   230
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   231
lemma "P Q | Q x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   232
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   233
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   234
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   235
lemma "x \<noteq> All"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   236
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   237
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   238
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   239
lemma "x \<noteq> Ex"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   240
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   241
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   242
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   243
lemma "x \<noteq> Ex1"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   244
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   245
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   246
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   247
text \<open>"The transitive closure 'T' of an arbitrary relation 'P' is non-empty."\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   248
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   249
definition "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   250
  "trans P \<equiv> (\<forall>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
   251
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   252
definition "subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   253
  "subset P Q \<equiv> (\<forall>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
   254
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35315
diff changeset
   255
definition "trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   256
  "trans_closure P Q \<equiv> (subset Q P) \<and> (trans P) \<and> (\<forall>R. subset Q R \<longrightarrow> trans R \<longrightarrow> subset P R)"
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   257
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   258
lemma "trans_closure T P \<longrightarrow> (\<exists>x y. T x y)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   259
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   260
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   261
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   262
text \<open>"Every surjective function is invertible."\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   263
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   264
lemma "(\<forall>y. \<exists>x. y = f x) \<longrightarrow> (\<exists>g. \<forall>x. g (f x) = x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   265
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   266
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   267
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   268
text \<open>"Every invertible function is surjective."\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   269
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   270
lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   271
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   272
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   273
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   274
text \<open>Every point is a fixed point of some function.\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   275
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   276
lemma "\<exists>f. f x = x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   277
refute [maxsize = 4, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   278
apply (rule_tac x="\<lambda>x. x" in exI)
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   279
by simp
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   280
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   281
text \<open>Axiom of Choice: first an incorrect version ...\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   282
63901
4ce989e962e0 more symbols;
wenzelm
parents: 62148
diff changeset
   283
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>!f. \<forall>x. P x (f x))"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   284
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   285
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   286
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   287
text \<open>... and now two correct ones\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   288
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   289
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>f. \<forall>x. P x (f x))"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   290
refute [maxsize = 4, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   291
by (simp add: choice)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   292
63901
4ce989e962e0 more symbols;
wenzelm
parents: 62148
diff changeset
   293
lemma "(\<forall>x. \<exists>!y. P x y) \<longrightarrow> (\<exists>!f. \<forall>x. P x (f x))"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   294
refute [maxsize = 2, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   295
apply auto
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   296
  apply (simp add: ex1_implies_ex choice)
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   297
by (fast intro: ext)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   298
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   299
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   300
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   301
subsubsection \<open>Meta-logic\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   302
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   303
lemma "!!x. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   304
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   305
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   306
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   307
lemma "f x == g x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   308
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   309
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   310
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   311
lemma "P \<Longrightarrow> Q"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   312
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   313
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   314
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   315
lemma "\<lbrakk> P; Q; R \<rbrakk> \<Longrightarrow> S"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   316
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   317
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   318
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55867
diff changeset
   319
lemma "(x == Pure.all) \<Longrightarrow> False"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   320
refute [expect = genuine]
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   321
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   322
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   323
lemma "(x == (==)) \<Longrightarrow> False"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   324
refute [expect = genuine]
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   325
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   326
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   327
lemma "(x == (\<Longrightarrow>)) \<Longrightarrow> False"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   328
refute [expect = genuine]
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   329
oops
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   330
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   331
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   332
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   333
subsubsection \<open>Schematic variables\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   334
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 59987
diff changeset
   335
schematic_goal "?P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   336
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   337
by auto
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   338
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 59987
diff changeset
   339
schematic_goal "x = ?y"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   340
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   341
by auto
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   342
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   343
(******************************************************************************)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   344
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   345
subsubsection \<open>Abstractions\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   346
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   347
lemma "(\<lambda>x. x) = (\<lambda>x. y)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   348
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   349
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   351
lemma "(\<lambda>f. f x) = (\<lambda>f. True)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   352
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   353
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   354
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   355
lemma "(\<lambda>x. x) = (\<lambda>y. y)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   356
refute
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   357
by simp
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   358
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   359
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   360
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   361
subsubsection \<open>Sets\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   362
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   363
lemma "P (A::'a set)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   364
refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   365
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   366
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   367
lemma "P (A::'a set set)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   368
refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   369
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   370
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   371
lemma "{x. P x} = {y. P y}"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   372
refute
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   373
by simp
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   374
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   375
lemma "x \<in> {x. P x}"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   376
refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   377
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   378
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   379
lemma "P (\<in>)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   380
refute
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   381
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   382
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   383
lemma "P ((\<in>) x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   384
refute
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   385
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   386
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   387
lemma "P Collect"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   388
refute
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   389
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   390
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   391
lemma "A \<union> B = A \<inter> B"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   392
refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   393
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   394
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   395
lemma "(A \<inter> B) \<union> C = (A \<union> C) \<inter> B"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   396
refute
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   397
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   398
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   399
lemma "Ball A P \<longrightarrow> Bex A P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   400
refute
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   401
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   402
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   403
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   404
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   405
subsubsection \<open>undefined\<close>
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   406
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   407
lemma "undefined"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   408
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   409
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   410
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   411
lemma "P undefined"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   412
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   413
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   414
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   415
lemma "undefined x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   416
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   417
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   418
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 25032
diff changeset
   419
lemma "undefined undefined"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   420
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   421
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   422
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   423
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   424
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   425
subsubsection \<open>The\<close>
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   426
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   427
lemma "The P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   428
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   429
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   430
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   431
lemma "P The"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   432
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   433
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   434
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   435
lemma "P (The P)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   436
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   437
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   438
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   439
lemma "(THE x. x=y) = z"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   440
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   441
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   442
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   443
lemma "Ex P \<longrightarrow> P (The P)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   444
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   445
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   446
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   447
(*****************************************************************************)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   448
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   449
subsubsection \<open>Eps\<close>
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   450
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   451
lemma "Eps P"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   452
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   453
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   454
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   455
lemma "P Eps"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   456
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   457
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   458
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   459
lemma "P (Eps P)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   460
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   461
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   462
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   463
lemma "(SOME x. x=y) = z"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   464
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   465
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   466
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   467
lemma "Ex P \<longrightarrow> P (Eps P)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   468
refute [maxsize = 3, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   469
by (auto simp add: someI)
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   470
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   471
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   472
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   473
subsubsection \<open>Subtypes (typedef), typedecl\<close>
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   474
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   475
text \<open>A completely unspecified non-empty subset of \<^typ>\<open>'a\<close>:\<close>
15161
065ce5385a06 comment modified
webertj
parents: 14809
diff changeset
   476
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 37388
diff changeset
   477
definition "myTdef = insert (undefined::'a) (undefined::'a set)"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 37388
diff changeset
   478
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 46099
diff changeset
   479
typedef 'a myTdef = "myTdef :: 'a set"
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 37388
diff changeset
   480
  unfolding myTdef_def by auto
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   481
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   482
lemma "(x::'a myTdef) = y"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   483
refute
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   484
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   485
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   486
typedecl myTdecl
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   487
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 37388
diff changeset
   488
definition "T_bij = {(f::'a\<Rightarrow>'a). \<forall>y. \<exists>!x. f x = y}"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 37388
diff changeset
   489
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 46099
diff changeset
   490
typedef 'a T_bij = "T_bij :: ('a \<Rightarrow> 'a) set"
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 37388
diff changeset
   491
  unfolding T_bij_def by auto
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   492
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   493
lemma "P (f::(myTdecl myTdef) T_bij)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   494
refute
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   495
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   496
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   497
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   498
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   499
subsubsection \<open>Inductive datatypes\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   500
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   501
text \<open>unit\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   502
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   503
lemma "P (x::unit)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   504
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   505
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   506
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   507
lemma "\<forall>x::unit. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   508
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   509
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   510
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   511
lemma "P ()"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   512
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   513
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   514
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
   515
lemma "P (case x of () \<Rightarrow> u)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   516
refute [expect = genuine]
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
   517
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   518
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   519
text \<open>option\<close>
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   520
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   521
lemma "P (x::'a option)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   522
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   523
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   524
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   525
lemma "\<forall>x::'a option. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   526
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   527
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   528
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   529
lemma "P None"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   530
refute [expect = genuine]
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   531
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   532
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   533
lemma "P (Some x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   534
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   535
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   536
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   537
lemma "P (case x of None \<Rightarrow> n | Some u \<Rightarrow> s u)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   538
refute [expect = genuine]
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   539
oops
f08e2d83681e major code change: 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
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   541
text \<open>*\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   542
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   543
lemma "P (x::'a*'b)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   544
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   545
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   546
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   547
lemma "\<forall>x::'a*'b. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   548
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   549
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   550
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   551
lemma "P (x, y)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   552
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   553
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   554
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   555
lemma "P (fst x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   556
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   557
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   558
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   559
lemma "P (snd x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   560
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   561
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   562
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   563
lemma "P Pair"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   564
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   565
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   566
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
   567
lemma "P (case x of Pair a b \<Rightarrow> p a b)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   568
refute [expect = genuine]
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
   569
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   570
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   571
text \<open>+\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   572
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   573
lemma "P (x::'a+'b)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   574
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   575
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   576
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   577
lemma "\<forall>x::'a+'b. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   578
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   579
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   580
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   581
lemma "P (Inl x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   582
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   583
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   584
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   585
lemma "P (Inr x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   586
refute [expect = genuine]
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   587
oops
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   588
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   589
lemma "P Inl"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   590
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   591
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   592
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
   593
lemma "P (case x of Inl a \<Rightarrow> l a | Inr b \<Rightarrow> r b)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   594
refute [expect = genuine]
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
   595
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   596
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   597
text \<open>Non-recursive datatypes\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   598
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   599
datatype T1 = A | B
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   600
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   601
lemma "P (x::T1)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   602
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   603
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   604
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   605
lemma "\<forall>x::T1. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   606
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   607
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   608
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   609
lemma "P A"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   610
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   611
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   612
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   613
lemma "P B"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   614
refute [expect = genuine]
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   615
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   616
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   617
lemma "rec_T1 a b A = a"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   618
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   619
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   620
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   621
lemma "rec_T1 a b B = b"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   622
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   623
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   624
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   625
lemma "P (rec_T1 a b x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   626
refute [expect = genuine]
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   627
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   628
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   629
lemma "P (case x of A \<Rightarrow> a | B \<Rightarrow> b)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   630
refute [expect = genuine]
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
   631
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   632
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   633
datatype 'a T2 = C T1 | D 'a
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   634
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   635
lemma "P (x::'a T2)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   636
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   637
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   638
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   639
lemma "\<forall>x::'a T2. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   640
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   641
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   642
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   643
lemma "P D"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   644
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   645
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   646
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   647
lemma "rec_T2 c d (C x) = c x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   648
refute [maxsize = 4, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   649
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   650
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   651
lemma "rec_T2 c d (D x) = d x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   652
refute [maxsize = 4, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   653
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   654
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   655
lemma "P (rec_T2 c d x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   656
refute [expect = genuine]
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
   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 C u \<Rightarrow> c u | D v \<Rightarrow> d v)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   660
refute [expect = genuine]
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
   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
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   663
datatype ('a,'b) T3 = E "'a \<Rightarrow> 'b"
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   664
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   665
lemma "P (x::('a,'b) T3)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   666
refute [expect = genuine]
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   667
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   668
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   669
lemma "\<forall>x::('a,'b) T3. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   670
refute [expect = genuine]
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   671
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   672
14455
5c4a1e96efd6 Updated examples
webertj
parents: 14350
diff changeset
   673
lemma "P E"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   674
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   675
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   676
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   677
lemma "rec_T3 e (E x) = e x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   678
refute [maxsize = 2, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   679
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   680
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   681
lemma "P (rec_T3 e x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   682
refute [expect = genuine]
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
   683
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   684
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   685
lemma "P (case x of E f \<Rightarrow> e f)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   686
refute [expect = genuine]
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
   687
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   688
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   689
text \<open>Recursive datatypes\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   690
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   691
text \<open>nat\<close>
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   692
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   693
lemma "P (x::nat)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   694
refute [expect = potential]
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   695
oops
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   696
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   697
lemma "\<forall>x::nat. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   698
refute [expect = potential]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   699
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   700
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   701
lemma "P (Suc 0)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   702
refute [expect = potential]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   703
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   704
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   705
lemma "P Suc"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   706
refute [maxsize = 3, expect = none]
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   707
\<comment> \<open>\<^term>\<open>Suc\<close> is a partial function (regardless of the size
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   708
      of the model), hence \<^term>\<open>P Suc\<close> is undefined and no
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   709
      model will be found\<close>
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   710
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   711
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55413
diff changeset
   712
lemma "rec_nat zero suc 0 = zero"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   713
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   714
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   715
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55413
diff changeset
   716
lemma "rec_nat zero suc (Suc x) = suc x (rec_nat zero suc x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   717
refute [maxsize = 2, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   718
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   719
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55413
diff changeset
   720
lemma "P (rec_nat zero suc x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   721
refute [expect = potential]
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
   722
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   723
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   724
lemma "P (case x of 0 \<Rightarrow> zero | Suc n \<Rightarrow> suc n)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   725
refute [expect = potential]
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
   726
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   727
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   728
text \<open>'a list\<close>
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
   729
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   730
lemma "P (xs::'a list)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   731
refute [expect = potential]
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
   732
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   733
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   734
lemma "\<forall>xs::'a list. P xs"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   735
refute [expect = potential]
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
   736
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   737
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   738
lemma "P [x, y]"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   739
refute [expect = potential]
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
   740
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   741
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55395
diff changeset
   742
lemma "rec_list nil cons [] = nil"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   743
refute [maxsize = 3, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   744
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   745
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55395
diff changeset
   746
lemma "rec_list nil cons (x#xs) = cons x xs (rec_list nil cons xs)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   747
refute [maxsize = 2, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   748
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   749
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55395
diff changeset
   750
lemma "P (rec_list nil cons xs)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   751
refute [expect = potential]
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
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
lemma "P (case x of Nil \<Rightarrow> nil | Cons a b \<Rightarrow> cons a b)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   755
refute [expect = potential]
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
   756
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
lemma "(xs::'a list) = ys"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   759
refute [expect = potential]
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
   760
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   761
f08e2d83681e major code change: 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
lemma "a # xs = b # xs"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   763
refute [expect = potential]
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
   764
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   765
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   766
datatype BitList = BitListNil | Bit0 BitList | Bit1 BitList
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   767
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   768
lemma "P (x::BitList)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   769
refute [expect = potential]
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   770
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   771
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   772
lemma "\<forall>x::BitList. P x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   773
refute [expect = potential]
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   774
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   775
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   776
lemma "P (Bit0 (Bit1 BitListNil))"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   777
refute [expect = potential]
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   778
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   779
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   780
lemma "rec_BitList nil bit0 bit1 BitListNil = nil"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   781
refute [maxsize = 4, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   782
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   783
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   784
lemma "rec_BitList nil bit0 bit1 (Bit0 xs) = bit0 xs (rec_BitList nil bit0 bit1 xs)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   785
refute [maxsize = 2, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   786
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   787
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   788
lemma "rec_BitList nil bit0 bit1 (Bit1 xs) = bit1 xs (rec_BitList nil bit0 bit1 xs)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   789
refute [maxsize = 2, expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   790
by simp
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   791
55416
dd7992d4a61a adapted theories to 'xxx_case' to 'case_xxx'
blanchet
parents: 55415
diff changeset
   792
lemma "P (rec_BitList nil bit0 bit1 x)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   793
refute [expect = potential]
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   794
oops
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   795
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   796
(*****************************************************************************)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   797
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   798
subsubsection \<open>Examples involving special functions\<close>
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
   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
lemma "card x = 0"
58143
7f7026ae9dc5 tuned Refute example
blanchet
parents: 58129
diff changeset
   801
refute [expect = potential]
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
   802
oops
f08e2d83681e major code change: 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
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15547
diff changeset
   804
lemma "finite x"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   805
refute \<comment> \<open>no finite countermodel exists\<close>
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
   806
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
lemma "(x::nat) + y = 0"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   809
refute [expect = potential]
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   810
oops
f08e2d83681e major code change: 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
f08e2d83681e major code change: 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
lemma "(x::nat) = x + x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   813
refute [expect = potential]
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
   814
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   815
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   816
lemma "(x::nat) - y + y = x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   817
refute [expect = potential]
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   818
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   819
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   820
lemma "(x::nat) = x * x"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   821
refute [expect = potential]
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
   822
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   823
f08e2d83681e major code change: 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 "(x::nat) < x + y"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   825
refute [expect = potential]
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
   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
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   828
lemma "xs @ [] = ys @ []"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   829
refute [expect = potential]
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
   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
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21559
diff changeset
   832
lemma "xs @ ys = ys @ xs"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   833
refute [expect = potential]
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
   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
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24447
diff changeset
   836
(*****************************************************************************)
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
   837
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   838
subsubsection \<open>Type classes and overloading\<close>
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
   839
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   840
text \<open>A type class without axioms:\<close>
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
   841
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
   842
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
   843
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   844
lemma "P (x::'a::classA)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   845
refute [expect = genuine]
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   846
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   847
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   848
text \<open>An axiom with a type variable (denoting types which have at least two elements):\<close>
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
   849
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
   850
class classC =
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
   851
  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
   852
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   853
lemma "P (x::'a::classC)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   854
refute [expect = genuine]
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
   855
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   856
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   857
lemma "\<exists>x y. (x::'a::classC) \<noteq> y"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   858
(* refute [expect = none] FIXME *)
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
   859
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   860
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   861
text \<open>A type class for which a constant is defined:\<close>
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
   862
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
   863
class classD =
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
   864
  fixes classD_const :: "'a \<Rightarrow> 'a"
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
   865
  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
   866
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   867
lemma "P (x::'a::classD)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   868
refute [expect = genuine]
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
   869
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   870
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   871
text \<open>A type class with multiple superclasses:\<close>
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
   872
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 34120
diff changeset
   873
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
   874
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   875
lemma "P (x::'a::classE)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   876
refute [expect = genuine]
14809
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   877
oops
eaa5d6987ba2 mainly new/different datatype examples
webertj
parents: 14489
diff changeset
   878
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   879
text \<open>OFCLASS:\<close>
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
   880
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   881
lemma "OFCLASS('a::type, type_class)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   882
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   883
by intro_classes
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
   884
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   885
lemma "OFCLASS('a::classC, type_class)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   886
refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   887
by intro_classes
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
   888
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   889
lemma "OFCLASS('a::type, classC_class)"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   890
refute [expect = genuine]
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
   891
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   892
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   893
text \<open>Overloading:\<close>
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
   894
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   895
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
   896
62148
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   897
overloading inverse_bool \<equiv> "inverse :: bool \<Rightarrow> bool"
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   898
begin
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   899
  definition "inverse (b::bool) \<equiv> \<not> b"
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   900
end
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   901
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   902
overloading inverse_set \<equiv> "inverse :: 'a set \<Rightarrow> 'a set"
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   903
begin
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   904
  definition "inverse (S::'a set) \<equiv> -S"
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   905
end
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   906
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   907
overloading inverse_pair \<equiv> "inverse :: 'a \<times> 'b \<Rightarrow> 'a \<times> 'b"
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   908
begin
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   909
  definition "inverse_pair p \<equiv> (inverse (fst p), inverse (snd p))"
e410c6287103 eliminated old defs;
wenzelm
parents: 61933
diff changeset
   910
end
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
   911
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   912
lemma "inverse b"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   913
refute [expect = genuine]
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
   914
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   915
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   916
lemma "P (inverse (S::'a set))"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   917
refute [expect = genuine]
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
   918
oops
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   919
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15297
diff changeset
   920
lemma "P (inverse (p::'a\<times>'b))"
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   921
refute [expect = genuine]
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   922
oops
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   923
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   924
text \<open>Structured proofs\<close>
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
   925
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
   926
lemma "x = y"
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
   927
proof cases
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
   928
  assume "x = y"
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
   929
  show ?thesis
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   930
  refute [expect = none]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   931
  refute [no_assms, expect = genuine]
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   932
  refute [no_assms = false, expect = none]
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
   933
oops
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 32968
diff changeset
   934
46099
40ac5ae6d16f reintroduced 'refute' calls taken out after reintroducing the "set" constructor, and use "expect" feature
blanchet
parents: 45972
diff changeset
   935
refute_params [satsolver = "auto"]
18774
7cf74a743c32 works with DPLL solver now
webertj
parents: 16910
diff changeset
   936
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   937
end