author | wenzelm |
Mon, 09 Oct 2006 02:20:01 +0200 | |
changeset 20907 | 9673c67dde9b |
parent 20833 | 4fcf8ddb54f5 |
child 20944 | 34b2c1bb7178 |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/HOL.thy |
2 |
ID: $Id$ |
|
11750 | 3 |
Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson |
4 |
*) |
|
923 | 5 |
|
11750 | 6 |
header {* The basis of Higher-Order Logic *} |
923 | 7 |
|
15131 | 8 |
theory HOL |
15140 | 9 |
imports CPure |
20766 | 10 |
uses ("cladata.ML") ("blastdata.ML") ("simpdata.ML") "Tools/res_atpset.ML" |
15131 | 11 |
begin |
2260 | 12 |
|
11750 | 13 |
subsection {* Primitive logic *} |
14 |
||
15 |
subsubsection {* Core syntax *} |
|
2260 | 16 |
|
14854 | 17 |
classes type |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
18 |
defaultsort type |
3947 | 19 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
20 |
global |
923 | 21 |
|
7357 | 22 |
typedecl bool |
923 | 23 |
|
24 |
arities |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
25 |
bool :: type |
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
26 |
"fun" :: (type, type) type |
923 | 27 |
|
11750 | 28 |
judgment |
29 |
Trueprop :: "bool => prop" ("(_)" 5) |
|
923 | 30 |
|
11750 | 31 |
consts |
7357 | 32 |
Not :: "bool => bool" ("~ _" [40] 40) |
33 |
True :: bool |
|
34 |
False :: bool |
|
3947 | 35 |
arbitrary :: 'a |
20172 | 36 |
undefined :: 'a |
923 | 37 |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
38 |
The :: "('a => bool) => 'a" |
7357 | 39 |
All :: "('a => bool) => bool" (binder "ALL " 10) |
40 |
Ex :: "('a => bool) => bool" (binder "EX " 10) |
|
41 |
Ex1 :: "('a => bool) => bool" (binder "EX! " 10) |
|
42 |
Let :: "['a, 'a => 'b] => 'b" |
|
923 | 43 |
|
7357 | 44 |
"=" :: "['a, 'a] => bool" (infixl 50) |
45 |
& :: "[bool, bool] => bool" (infixr 35) |
|
46 |
"|" :: "[bool, bool] => bool" (infixr 30) |
|
47 |
--> :: "[bool, bool] => bool" (infixr 25) |
|
923 | 48 |
|
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
49 |
local |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
50 |
|
16587 | 51 |
consts |
52 |
If :: "[bool, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) |
|
2260 | 53 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
54 |
|
11750 | 55 |
subsubsection {* Additional concrete syntax *} |
2260 | 56 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
57 |
const_syntax (output) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
58 |
"op =" (infix "=" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
59 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
60 |
abbreviation |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
61 |
not_equal :: "['a, 'a] => bool" (infixl "~=" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
62 |
"x ~= y == ~ (x = y)" |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
63 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
64 |
const_syntax (output) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
65 |
not_equal (infix "~=" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
66 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
67 |
const_syntax (xsymbols) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
68 |
Not ("\<not> _" [40] 40) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
69 |
"op &" (infixr "\<and>" 35) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
70 |
"op |" (infixr "\<or>" 30) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
71 |
"op -->" (infixr "\<longrightarrow>" 25) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
72 |
not_equal (infix "\<noteq>" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
73 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
74 |
const_syntax (HTML output) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
75 |
Not ("\<not> _" [40] 40) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
76 |
"op &" (infixr "\<and>" 35) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
77 |
"op |" (infixr "\<or>" 30) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
78 |
not_equal (infix "\<noteq>" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
79 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
80 |
abbreviation (iff) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
81 |
iff :: "[bool, bool] => bool" (infixr "<->" 25) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
82 |
"A <-> B == A = B" |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
83 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
84 |
const_syntax (xsymbols) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
85 |
iff (infixr "\<longleftrightarrow>" 25) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
86 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
87 |
|
4868 | 88 |
nonterminals |
923 | 89 |
letbinds letbind |
90 |
case_syn cases_syn |
|
91 |
||
92 |
syntax |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
93 |
"_The" :: "[pttrn, bool] => 'a" ("(3THE _./ _)" [0, 10] 10) |
923 | 94 |
|
7357 | 95 |
"_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10) |
96 |
"" :: "letbind => letbinds" ("_") |
|
97 |
"_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") |
|
98 |
"_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) |
|
923 | 99 |
|
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
100 |
"_case_syntax":: "['a, cases_syn] => 'b" ("(case _ of/ _)" 10) |
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
101 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10) |
7357 | 102 |
"" :: "case_syn => cases_syn" ("_") |
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
103 |
"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _") |
923 | 104 |
|
105 |
translations |
|
13764 | 106 |
"THE x. P" == "The (%x. P)" |
923 | 107 |
"_Let (_binds b bs) e" == "_Let b (_Let bs e)" |
1114 | 108 |
"let x = a in e" == "Let a (%x. e)" |
923 | 109 |
|
13763
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
110 |
print_translation {* |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
111 |
(* To avoid eta-contraction of body: *) |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
112 |
[("The", fn [Abs abs] => |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
113 |
let val (x,t) = atomic_abs_tr' abs |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
114 |
in Syntax.const "_The" $ x $ t end)] |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
115 |
*} |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
116 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
117 |
syntax (xsymbols) |
11687 | 118 |
"ALL " :: "[idts, bool] => bool" ("(3\<forall>_./ _)" [0, 10] 10) |
119 |
"EX " :: "[idts, bool] => bool" ("(3\<exists>_./ _)" [0, 10] 10) |
|
120 |
"EX! " :: "[idts, bool] => bool" ("(3\<exists>!_./ _)" [0, 10] 10) |
|
121 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ \<Rightarrow>/ _)" 10) |
|
14361
ad2f5da643b4
* Support for raw latex output in control symbols: \<^raw...>
schirmer
parents:
14357
diff
changeset
|
122 |
(*"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ \<orelse> _")*) |
2372 | 123 |
|
6340 | 124 |
syntax (HTML output) |
14565 | 125 |
"ALL " :: "[idts, bool] => bool" ("(3\<forall>_./ _)" [0, 10] 10) |
126 |
"EX " :: "[idts, bool] => bool" ("(3\<exists>_./ _)" [0, 10] 10) |
|
127 |
"EX! " :: "[idts, bool] => bool" ("(3\<exists>!_./ _)" [0, 10] 10) |
|
6340 | 128 |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
129 |
syntax (HOL) |
7357 | 130 |
"ALL " :: "[idts, bool] => bool" ("(3! _./ _)" [0, 10] 10) |
131 |
"EX " :: "[idts, bool] => bool" ("(3? _./ _)" [0, 10] 10) |
|
132 |
"EX! " :: "[idts, bool] => bool" ("(3?! _./ _)" [0, 10] 10) |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
133 |
|
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
134 |
|
11750 | 135 |
subsubsection {* Axioms and basic definitions *} |
2260 | 136 |
|
7357 | 137 |
axioms |
15380 | 138 |
eq_reflection: "(x=y) ==> (x==y)" |
923 | 139 |
|
15380 | 140 |
refl: "t = (t::'a)" |
6289 | 141 |
|
15380 | 142 |
ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" |
143 |
-- {*Extensionality is built into the meta-logic, and this rule expresses |
|
144 |
a related property. It is an eta-expanded version of the traditional |
|
145 |
rule, and similar to the ABS rule of HOL*} |
|
6289 | 146 |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
147 |
the_eq_trivial: "(THE x. x = a) = (a::'a)" |
923 | 148 |
|
15380 | 149 |
impI: "(P ==> Q) ==> P-->Q" |
150 |
mp: "[| P-->Q; P |] ==> Q" |
|
151 |
||
152 |
||
153 |
text{*Thanks to Stephan Merz*} |
|
154 |
theorem subst: |
|
155 |
assumes eq: "s = t" and p: "P(s)" |
|
156 |
shows "P(t::'a)" |
|
157 |
proof - |
|
158 |
from eq have meta: "s \<equiv> t" |
|
159 |
by (rule eq_reflection) |
|
160 |
from p show ?thesis |
|
161 |
by (unfold meta) |
|
162 |
qed |
|
923 | 163 |
|
164 |
defs |
|
7357 | 165 |
True_def: "True == ((%x::bool. x) = (%x. x))" |
166 |
All_def: "All(P) == (P = (%x. True))" |
|
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
167 |
Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" |
7357 | 168 |
False_def: "False == (!P. P)" |
169 |
not_def: "~ P == P-->False" |
|
170 |
and_def: "P & Q == !R. (P-->Q-->R) --> R" |
|
171 |
or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" |
|
172 |
Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" |
|
923 | 173 |
|
7357 | 174 |
axioms |
175 |
iff: "(P-->Q) --> (Q-->P) --> (P=Q)" |
|
176 |
True_or_False: "(P=True) | (P=False)" |
|
923 | 177 |
|
178 |
defs |
|
7357 | 179 |
Let_def: "Let s f == f(s)" |
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
180 |
if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" |
5069 | 181 |
|
14223
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
182 |
finalconsts |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
183 |
"op =" |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
184 |
"op -->" |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
185 |
The |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
186 |
arbitrary |
20172 | 187 |
undefined |
3320 | 188 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
189 |
|
11750 | 190 |
subsubsection {* Generic algebraic operations *} |
4868 | 191 |
|
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
192 |
class zero = |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
193 |
fixes zero :: "'a" ("\<^loc>0") |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
194 |
|
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
195 |
class one = |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
196 |
fixes one :: "'a" ("\<^loc>1") |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
197 |
|
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
198 |
hide (open) const zero one |
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
199 |
|
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
200 |
class plus = |
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
201 |
fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<^loc>+" 65) |
11750 | 202 |
|
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
203 |
class minus = |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
204 |
fixes uminus :: "'a \<Rightarrow> 'a" |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
205 |
fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<^loc>-" 65) |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
206 |
fixes abs :: "'a \<Rightarrow> 'a" |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
207 |
|
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
208 |
class times = |
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
209 |
fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<^loc>*" 70) |
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
210 |
|
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
211 |
class inverse = |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
212 |
fixes inverse :: "'a \<Rightarrow> 'a" |
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
213 |
fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<^loc>'/" 70) |
11750 | 214 |
|
13456
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
215 |
syntax |
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
216 |
"_index1" :: index ("\<^sub>1") |
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
217 |
translations |
14690 | 218 |
(index) "\<^sub>1" => (index) "\<^bsub>\<struct>\<^esub>" |
13456
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
219 |
|
11750 | 220 |
typed_print_translation {* |
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
221 |
let |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
222 |
fun tr' c = (c, fn show_sorts => fn T => fn ts => |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
223 |
if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
224 |
else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
225 |
in |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
226 |
map (tr' o prefix Syntax.constN) ["HOL.one", "HOL.zero"] |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
227 |
end; |
11750 | 228 |
*} -- {* show types that are presumably too general *} |
229 |
||
20741 | 230 |
const_syntax |
231 |
uminus ("- _" [81] 80) |
|
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
232 |
|
20741 | 233 |
const_syntax (xsymbols) |
234 |
abs ("\<bar>_\<bar>") |
|
235 |
const_syntax (HTML output) |
|
236 |
abs ("\<bar>_\<bar>") |
|
11750 | 237 |
|
238 |
||
15411 | 239 |
subsection {*Equality*} |
240 |
||
18457 | 241 |
lemma sym: "s = t ==> t = s" |
242 |
by (erule subst) (rule refl) |
|
15411 | 243 |
|
18457 | 244 |
lemma ssubst: "t = s ==> P s ==> P t" |
245 |
by (drule sym) (erule subst) |
|
15411 | 246 |
|
247 |
lemma trans: "[| r=s; s=t |] ==> r=t" |
|
18457 | 248 |
by (erule subst) |
15411 | 249 |
|
18457 | 250 |
lemma def_imp_eq: assumes meq: "A == B" shows "A = B" |
251 |
by (unfold meq) (rule refl) |
|
252 |
||
15411 | 253 |
|
254 |
(*Useful with eresolve_tac for proving equalties from known equalities. |
|
255 |
a = b |
|
256 |
| | |
|
257 |
c = d *) |
|
258 |
lemma box_equals: "[| a=b; a=c; b=d |] ==> c=d" |
|
259 |
apply (rule trans) |
|
260 |
apply (rule trans) |
|
261 |
apply (rule sym) |
|
262 |
apply assumption+ |
|
263 |
done |
|
264 |
||
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
265 |
text {* For calculational reasoning: *} |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
266 |
|
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
267 |
lemma forw_subst: "a = b ==> P b ==> P a" |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
268 |
by (rule ssubst) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
269 |
|
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
270 |
lemma back_subst: "P a ==> a = b ==> P b" |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
271 |
by (rule subst) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
272 |
|
15411 | 273 |
|
274 |
subsection {*Congruence rules for application*} |
|
275 |
||
276 |
(*similar to AP_THM in Gordon's HOL*) |
|
277 |
lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)" |
|
278 |
apply (erule subst) |
|
279 |
apply (rule refl) |
|
280 |
done |
|
281 |
||
282 |
(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*) |
|
283 |
lemma arg_cong: "x=y ==> f(x)=f(y)" |
|
284 |
apply (erule subst) |
|
285 |
apply (rule refl) |
|
286 |
done |
|
287 |
||
15655 | 288 |
lemma arg_cong2: "\<lbrakk> a = b; c = d \<rbrakk> \<Longrightarrow> f a c = f b d" |
289 |
apply (erule ssubst)+ |
|
290 |
apply (rule refl) |
|
291 |
done |
|
292 |
||
293 |
||
15411 | 294 |
lemma cong: "[| f = g; (x::'a) = y |] ==> f(x) = g(y)" |
295 |
apply (erule subst)+ |
|
296 |
apply (rule refl) |
|
297 |
done |
|
298 |
||
299 |
||
300 |
subsection {*Equality of booleans -- iff*} |
|
301 |
||
302 |
lemma iffI: assumes prems: "P ==> Q" "Q ==> P" shows "P=Q" |
|
18457 | 303 |
by (iprover intro: iff [THEN mp, THEN mp] impI prems) |
15411 | 304 |
|
305 |
lemma iffD2: "[| P=Q; Q |] ==> P" |
|
18457 | 306 |
by (erule ssubst) |
15411 | 307 |
|
308 |
lemma rev_iffD2: "[| Q; P=Q |] ==> P" |
|
18457 | 309 |
by (erule iffD2) |
15411 | 310 |
|
311 |
lemmas iffD1 = sym [THEN iffD2, standard] |
|
312 |
lemmas rev_iffD1 = sym [THEN [2] rev_iffD2, standard] |
|
313 |
||
314 |
lemma iffE: |
|
315 |
assumes major: "P=Q" |
|
316 |
and minor: "[| P --> Q; Q --> P |] ==> R" |
|
18457 | 317 |
shows R |
318 |
by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1]) |
|
15411 | 319 |
|
320 |
||
321 |
subsection {*True*} |
|
322 |
||
323 |
lemma TrueI: "True" |
|
18457 | 324 |
by (unfold True_def) (rule refl) |
15411 | 325 |
|
326 |
lemma eqTrueI: "P ==> P=True" |
|
18457 | 327 |
by (iprover intro: iffI TrueI) |
15411 | 328 |
|
329 |
lemma eqTrueE: "P=True ==> P" |
|
330 |
apply (erule iffD2) |
|
331 |
apply (rule TrueI) |
|
332 |
done |
|
333 |
||
334 |
||
335 |
subsection {*Universal quantifier*} |
|
336 |
||
337 |
lemma allI: assumes p: "!!x::'a. P(x)" shows "ALL x. P(x)" |
|
338 |
apply (unfold All_def) |
|
17589 | 339 |
apply (iprover intro: ext eqTrueI p) |
15411 | 340 |
done |
341 |
||
342 |
lemma spec: "ALL x::'a. P(x) ==> P(x)" |
|
343 |
apply (unfold All_def) |
|
344 |
apply (rule eqTrueE) |
|
345 |
apply (erule fun_cong) |
|
346 |
done |
|
347 |
||
348 |
lemma allE: |
|
349 |
assumes major: "ALL x. P(x)" |
|
350 |
and minor: "P(x) ==> R" |
|
351 |
shows "R" |
|
17589 | 352 |
by (iprover intro: minor major [THEN spec]) |
15411 | 353 |
|
354 |
lemma all_dupE: |
|
355 |
assumes major: "ALL x. P(x)" |
|
356 |
and minor: "[| P(x); ALL x. P(x) |] ==> R" |
|
357 |
shows "R" |
|
17589 | 358 |
by (iprover intro: minor major major [THEN spec]) |
15411 | 359 |
|
360 |
||
361 |
subsection {*False*} |
|
362 |
(*Depends upon spec; it is impossible to do propositional logic before quantifiers!*) |
|
363 |
||
364 |
lemma FalseE: "False ==> P" |
|
365 |
apply (unfold False_def) |
|
366 |
apply (erule spec) |
|
367 |
done |
|
368 |
||
369 |
lemma False_neq_True: "False=True ==> P" |
|
370 |
by (erule eqTrueE [THEN FalseE]) |
|
371 |
||
372 |
||
373 |
subsection {*Negation*} |
|
374 |
||
375 |
lemma notI: |
|
376 |
assumes p: "P ==> False" |
|
377 |
shows "~P" |
|
378 |
apply (unfold not_def) |
|
17589 | 379 |
apply (iprover intro: impI p) |
15411 | 380 |
done |
381 |
||
382 |
lemma False_not_True: "False ~= True" |
|
383 |
apply (rule notI) |
|
384 |
apply (erule False_neq_True) |
|
385 |
done |
|
386 |
||
387 |
lemma True_not_False: "True ~= False" |
|
388 |
apply (rule notI) |
|
389 |
apply (drule sym) |
|
390 |
apply (erule False_neq_True) |
|
391 |
done |
|
392 |
||
393 |
lemma notE: "[| ~P; P |] ==> R" |
|
394 |
apply (unfold not_def) |
|
395 |
apply (erule mp [THEN FalseE]) |
|
396 |
apply assumption |
|
397 |
done |
|
398 |
||
399 |
(* Alternative ~ introduction rule: [| P ==> ~ Pa; P ==> Pa |] ==> ~ P *) |
|
400 |
lemmas notI2 = notE [THEN notI, standard] |
|
401 |
||
402 |
||
403 |
subsection {*Implication*} |
|
404 |
||
405 |
lemma impE: |
|
406 |
assumes "P-->Q" "P" "Q ==> R" |
|
407 |
shows "R" |
|
17589 | 408 |
by (iprover intro: prems mp) |
15411 | 409 |
|
410 |
(* Reduces Q to P-->Q, allowing substitution in P. *) |
|
411 |
lemma rev_mp: "[| P; P --> Q |] ==> Q" |
|
17589 | 412 |
by (iprover intro: mp) |
15411 | 413 |
|
414 |
lemma contrapos_nn: |
|
415 |
assumes major: "~Q" |
|
416 |
and minor: "P==>Q" |
|
417 |
shows "~P" |
|
17589 | 418 |
by (iprover intro: notI minor major [THEN notE]) |
15411 | 419 |
|
420 |
(*not used at all, but we already have the other 3 combinations *) |
|
421 |
lemma contrapos_pn: |
|
422 |
assumes major: "Q" |
|
423 |
and minor: "P ==> ~Q" |
|
424 |
shows "~P" |
|
17589 | 425 |
by (iprover intro: notI minor major notE) |
15411 | 426 |
|
427 |
lemma not_sym: "t ~= s ==> s ~= t" |
|
428 |
apply (erule contrapos_nn) |
|
429 |
apply (erule sym) |
|
430 |
done |
|
431 |
||
432 |
(*still used in HOLCF*) |
|
433 |
lemma rev_contrapos: |
|
434 |
assumes pq: "P ==> Q" |
|
435 |
and nq: "~Q" |
|
436 |
shows "~P" |
|
437 |
apply (rule nq [THEN contrapos_nn]) |
|
438 |
apply (erule pq) |
|
439 |
done |
|
440 |
||
441 |
subsection {*Existential quantifier*} |
|
442 |
||
443 |
lemma exI: "P x ==> EX x::'a. P x" |
|
444 |
apply (unfold Ex_def) |
|
17589 | 445 |
apply (iprover intro: allI allE impI mp) |
15411 | 446 |
done |
447 |
||
448 |
lemma exE: |
|
449 |
assumes major: "EX x::'a. P(x)" |
|
450 |
and minor: "!!x. P(x) ==> Q" |
|
451 |
shows "Q" |
|
452 |
apply (rule major [unfolded Ex_def, THEN spec, THEN mp]) |
|
17589 | 453 |
apply (iprover intro: impI [THEN allI] minor) |
15411 | 454 |
done |
455 |
||
456 |
||
457 |
subsection {*Conjunction*} |
|
458 |
||
459 |
lemma conjI: "[| P; Q |] ==> P&Q" |
|
460 |
apply (unfold and_def) |
|
17589 | 461 |
apply (iprover intro: impI [THEN allI] mp) |
15411 | 462 |
done |
463 |
||
464 |
lemma conjunct1: "[| P & Q |] ==> P" |
|
465 |
apply (unfold and_def) |
|
17589 | 466 |
apply (iprover intro: impI dest: spec mp) |
15411 | 467 |
done |
468 |
||
469 |
lemma conjunct2: "[| P & Q |] ==> Q" |
|
470 |
apply (unfold and_def) |
|
17589 | 471 |
apply (iprover intro: impI dest: spec mp) |
15411 | 472 |
done |
473 |
||
474 |
lemma conjE: |
|
475 |
assumes major: "P&Q" |
|
476 |
and minor: "[| P; Q |] ==> R" |
|
477 |
shows "R" |
|
478 |
apply (rule minor) |
|
479 |
apply (rule major [THEN conjunct1]) |
|
480 |
apply (rule major [THEN conjunct2]) |
|
481 |
done |
|
482 |
||
483 |
lemma context_conjI: |
|
484 |
assumes prems: "P" "P ==> Q" shows "P & Q" |
|
17589 | 485 |
by (iprover intro: conjI prems) |
15411 | 486 |
|
487 |
||
488 |
subsection {*Disjunction*} |
|
489 |
||
490 |
lemma disjI1: "P ==> P|Q" |
|
491 |
apply (unfold or_def) |
|
17589 | 492 |
apply (iprover intro: allI impI mp) |
15411 | 493 |
done |
494 |
||
495 |
lemma disjI2: "Q ==> P|Q" |
|
496 |
apply (unfold or_def) |
|
17589 | 497 |
apply (iprover intro: allI impI mp) |
15411 | 498 |
done |
499 |
||
500 |
lemma disjE: |
|
501 |
assumes major: "P|Q" |
|
502 |
and minorP: "P ==> R" |
|
503 |
and minorQ: "Q ==> R" |
|
504 |
shows "R" |
|
17589 | 505 |
by (iprover intro: minorP minorQ impI |
15411 | 506 |
major [unfolded or_def, THEN spec, THEN mp, THEN mp]) |
507 |
||
508 |
||
509 |
subsection {*Classical logic*} |
|
510 |
||
511 |
||
512 |
lemma classical: |
|
513 |
assumes prem: "~P ==> P" |
|
514 |
shows "P" |
|
515 |
apply (rule True_or_False [THEN disjE, THEN eqTrueE]) |
|
516 |
apply assumption |
|
517 |
apply (rule notI [THEN prem, THEN eqTrueI]) |
|
518 |
apply (erule subst) |
|
519 |
apply assumption |
|
520 |
done |
|
521 |
||
522 |
lemmas ccontr = FalseE [THEN classical, standard] |
|
523 |
||
524 |
(*notE with premises exchanged; it discharges ~R so that it can be used to |
|
525 |
make elimination rules*) |
|
526 |
lemma rev_notE: |
|
527 |
assumes premp: "P" |
|
528 |
and premnot: "~R ==> ~P" |
|
529 |
shows "R" |
|
530 |
apply (rule ccontr) |
|
531 |
apply (erule notE [OF premnot premp]) |
|
532 |
done |
|
533 |
||
534 |
(*Double negation law*) |
|
535 |
lemma notnotD: "~~P ==> P" |
|
536 |
apply (rule classical) |
|
537 |
apply (erule notE) |
|
538 |
apply assumption |
|
539 |
done |
|
540 |
||
541 |
lemma contrapos_pp: |
|
542 |
assumes p1: "Q" |
|
543 |
and p2: "~P ==> ~Q" |
|
544 |
shows "P" |
|
17589 | 545 |
by (iprover intro: classical p1 p2 notE) |
15411 | 546 |
|
547 |
||
548 |
subsection {*Unique existence*} |
|
549 |
||
550 |
lemma ex1I: |
|
551 |
assumes prems: "P a" "!!x. P(x) ==> x=a" |
|
552 |
shows "EX! x. P(x)" |
|
17589 | 553 |
by (unfold Ex1_def, iprover intro: prems exI conjI allI impI) |
15411 | 554 |
|
555 |
text{*Sometimes easier to use: the premises have no shared variables. Safe!*} |
|
556 |
lemma ex_ex1I: |
|
557 |
assumes ex_prem: "EX x. P(x)" |
|
558 |
and eq: "!!x y. [| P(x); P(y) |] ==> x=y" |
|
559 |
shows "EX! x. P(x)" |
|
17589 | 560 |
by (iprover intro: ex_prem [THEN exE] ex1I eq) |
15411 | 561 |
|
562 |
lemma ex1E: |
|
563 |
assumes major: "EX! x. P(x)" |
|
564 |
and minor: "!!x. [| P(x); ALL y. P(y) --> y=x |] ==> R" |
|
565 |
shows "R" |
|
566 |
apply (rule major [unfolded Ex1_def, THEN exE]) |
|
567 |
apply (erule conjE) |
|
17589 | 568 |
apply (iprover intro: minor) |
15411 | 569 |
done |
570 |
||
571 |
lemma ex1_implies_ex: "EX! x. P x ==> EX x. P x" |
|
572 |
apply (erule ex1E) |
|
573 |
apply (rule exI) |
|
574 |
apply assumption |
|
575 |
done |
|
576 |
||
577 |
||
578 |
subsection {*THE: definite description operator*} |
|
579 |
||
580 |
lemma the_equality: |
|
581 |
assumes prema: "P a" |
|
582 |
and premx: "!!x. P x ==> x=a" |
|
583 |
shows "(THE x. P x) = a" |
|
584 |
apply (rule trans [OF _ the_eq_trivial]) |
|
585 |
apply (rule_tac f = "The" in arg_cong) |
|
586 |
apply (rule ext) |
|
587 |
apply (rule iffI) |
|
588 |
apply (erule premx) |
|
589 |
apply (erule ssubst, rule prema) |
|
590 |
done |
|
591 |
||
592 |
lemma theI: |
|
593 |
assumes "P a" and "!!x. P x ==> x=a" |
|
594 |
shows "P (THE x. P x)" |
|
17589 | 595 |
by (iprover intro: prems the_equality [THEN ssubst]) |
15411 | 596 |
|
597 |
lemma theI': "EX! x. P x ==> P (THE x. P x)" |
|
598 |
apply (erule ex1E) |
|
599 |
apply (erule theI) |
|
600 |
apply (erule allE) |
|
601 |
apply (erule mp) |
|
602 |
apply assumption |
|
603 |
done |
|
604 |
||
605 |
(*Easier to apply than theI: only one occurrence of P*) |
|
606 |
lemma theI2: |
|
607 |
assumes "P a" "!!x. P x ==> x=a" "!!x. P x ==> Q x" |
|
608 |
shows "Q (THE x. P x)" |
|
17589 | 609 |
by (iprover intro: prems theI) |
15411 | 610 |
|
18697 | 611 |
lemma the1_equality [elim?]: "[| EX!x. P x; P a |] ==> (THE x. P x) = a" |
15411 | 612 |
apply (rule the_equality) |
613 |
apply assumption |
|
614 |
apply (erule ex1E) |
|
615 |
apply (erule all_dupE) |
|
616 |
apply (drule mp) |
|
617 |
apply assumption |
|
618 |
apply (erule ssubst) |
|
619 |
apply (erule allE) |
|
620 |
apply (erule mp) |
|
621 |
apply assumption |
|
622 |
done |
|
623 |
||
624 |
lemma the_sym_eq_trivial: "(THE y. x=y) = x" |
|
625 |
apply (rule the_equality) |
|
626 |
apply (rule refl) |
|
627 |
apply (erule sym) |
|
628 |
done |
|
629 |
||
630 |
||
631 |
subsection {*Classical intro rules for disjunction and existential quantifiers*} |
|
632 |
||
633 |
lemma disjCI: |
|
634 |
assumes "~Q ==> P" shows "P|Q" |
|
635 |
apply (rule classical) |
|
17589 | 636 |
apply (iprover intro: prems disjI1 disjI2 notI elim: notE) |
15411 | 637 |
done |
638 |
||
639 |
lemma excluded_middle: "~P | P" |
|
17589 | 640 |
by (iprover intro: disjCI) |
15411 | 641 |
|
642 |
text{*case distinction as a natural deduction rule. Note that @{term "~P"} |
|
643 |
is the second case, not the first.*} |
|
644 |
lemma case_split_thm: |
|
645 |
assumes prem1: "P ==> Q" |
|
646 |
and prem2: "~P ==> Q" |
|
647 |
shows "Q" |
|
648 |
apply (rule excluded_middle [THEN disjE]) |
|
649 |
apply (erule prem2) |
|
650 |
apply (erule prem1) |
|
651 |
done |
|
652 |
||
653 |
(*Classical implies (-->) elimination. *) |
|
654 |
lemma impCE: |
|
655 |
assumes major: "P-->Q" |
|
656 |
and minor: "~P ==> R" "Q ==> R" |
|
657 |
shows "R" |
|
658 |
apply (rule excluded_middle [of P, THEN disjE]) |
|
17589 | 659 |
apply (iprover intro: minor major [THEN mp])+ |
15411 | 660 |
done |
661 |
||
662 |
(*This version of --> elimination works on Q before P. It works best for |
|
663 |
those cases in which P holds "almost everywhere". Can't install as |
|
664 |
default: would break old proofs.*) |
|
665 |
lemma impCE': |
|
666 |
assumes major: "P-->Q" |
|
667 |
and minor: "Q ==> R" "~P ==> R" |
|
668 |
shows "R" |
|
669 |
apply (rule excluded_middle [of P, THEN disjE]) |
|
17589 | 670 |
apply (iprover intro: minor major [THEN mp])+ |
15411 | 671 |
done |
672 |
||
673 |
(*Classical <-> elimination. *) |
|
674 |
lemma iffCE: |
|
675 |
assumes major: "P=Q" |
|
676 |
and minor: "[| P; Q |] ==> R" "[| ~P; ~Q |] ==> R" |
|
677 |
shows "R" |
|
678 |
apply (rule major [THEN iffE]) |
|
17589 | 679 |
apply (iprover intro: minor elim: impCE notE) |
15411 | 680 |
done |
681 |
||
682 |
lemma exCI: |
|
683 |
assumes "ALL x. ~P(x) ==> P(a)" |
|
684 |
shows "EX x. P(x)" |
|
685 |
apply (rule ccontr) |
|
17589 | 686 |
apply (iprover intro: prems exI allI notI notE [of "\<exists>x. P x"]) |
15411 | 687 |
done |
688 |
||
689 |
||
690 |
||
11750 | 691 |
subsection {* Theory and package setup *} |
692 |
||
15411 | 693 |
ML |
694 |
{* |
|
695 |
val eq_reflection = thm "eq_reflection" |
|
696 |
val refl = thm "refl" |
|
697 |
val subst = thm "subst" |
|
698 |
val ext = thm "ext" |
|
699 |
val impI = thm "impI" |
|
700 |
val mp = thm "mp" |
|
701 |
val True_def = thm "True_def" |
|
702 |
val All_def = thm "All_def" |
|
703 |
val Ex_def = thm "Ex_def" |
|
704 |
val False_def = thm "False_def" |
|
705 |
val not_def = thm "not_def" |
|
706 |
val and_def = thm "and_def" |
|
707 |
val or_def = thm "or_def" |
|
708 |
val Ex1_def = thm "Ex1_def" |
|
709 |
val iff = thm "iff" |
|
710 |
val True_or_False = thm "True_or_False" |
|
711 |
val Let_def = thm "Let_def" |
|
712 |
val if_def = thm "if_def" |
|
713 |
val sym = thm "sym" |
|
714 |
val ssubst = thm "ssubst" |
|
715 |
val trans = thm "trans" |
|
716 |
val def_imp_eq = thm "def_imp_eq" |
|
717 |
val box_equals = thm "box_equals" |
|
718 |
val fun_cong = thm "fun_cong" |
|
719 |
val arg_cong = thm "arg_cong" |
|
720 |
val cong = thm "cong" |
|
721 |
val iffI = thm "iffI" |
|
722 |
val iffD2 = thm "iffD2" |
|
723 |
val rev_iffD2 = thm "rev_iffD2" |
|
724 |
val iffD1 = thm "iffD1" |
|
725 |
val rev_iffD1 = thm "rev_iffD1" |
|
726 |
val iffE = thm "iffE" |
|
727 |
val TrueI = thm "TrueI" |
|
728 |
val eqTrueI = thm "eqTrueI" |
|
729 |
val eqTrueE = thm "eqTrueE" |
|
730 |
val allI = thm "allI" |
|
731 |
val spec = thm "spec" |
|
732 |
val allE = thm "allE" |
|
733 |
val all_dupE = thm "all_dupE" |
|
734 |
val FalseE = thm "FalseE" |
|
735 |
val False_neq_True = thm "False_neq_True" |
|
736 |
val notI = thm "notI" |
|
737 |
val False_not_True = thm "False_not_True" |
|
738 |
val True_not_False = thm "True_not_False" |
|
739 |
val notE = thm "notE" |
|
740 |
val notI2 = thm "notI2" |
|
741 |
val impE = thm "impE" |
|
742 |
val rev_mp = thm "rev_mp" |
|
743 |
val contrapos_nn = thm "contrapos_nn" |
|
744 |
val contrapos_pn = thm "contrapos_pn" |
|
745 |
val not_sym = thm "not_sym" |
|
746 |
val rev_contrapos = thm "rev_contrapos" |
|
747 |
val exI = thm "exI" |
|
748 |
val exE = thm "exE" |
|
749 |
val conjI = thm "conjI" |
|
750 |
val conjunct1 = thm "conjunct1" |
|
751 |
val conjunct2 = thm "conjunct2" |
|
752 |
val conjE = thm "conjE" |
|
753 |
val context_conjI = thm "context_conjI" |
|
754 |
val disjI1 = thm "disjI1" |
|
755 |
val disjI2 = thm "disjI2" |
|
756 |
val disjE = thm "disjE" |
|
757 |
val classical = thm "classical" |
|
758 |
val ccontr = thm "ccontr" |
|
759 |
val rev_notE = thm "rev_notE" |
|
760 |
val notnotD = thm "notnotD" |
|
761 |
val contrapos_pp = thm "contrapos_pp" |
|
762 |
val ex1I = thm "ex1I" |
|
763 |
val ex_ex1I = thm "ex_ex1I" |
|
764 |
val ex1E = thm "ex1E" |
|
765 |
val ex1_implies_ex = thm "ex1_implies_ex" |
|
766 |
val the_equality = thm "the_equality" |
|
767 |
val theI = thm "theI" |
|
768 |
val theI' = thm "theI'" |
|
769 |
val theI2 = thm "theI2" |
|
770 |
val the1_equality = thm "the1_equality" |
|
771 |
val the_sym_eq_trivial = thm "the_sym_eq_trivial" |
|
772 |
val disjCI = thm "disjCI" |
|
773 |
val excluded_middle = thm "excluded_middle" |
|
774 |
val case_split_thm = thm "case_split_thm" |
|
775 |
val impCE = thm "impCE" |
|
776 |
val impCE = thm "impCE" |
|
777 |
val iffCE = thm "iffCE" |
|
778 |
val exCI = thm "exCI" |
|
4868 | 779 |
|
15411 | 780 |
(* combination of (spec RS spec RS ...(j times) ... spec RS mp) *) |
781 |
local |
|
782 |
fun wrong_prem (Const ("All", _) $ (Abs (_, _, t))) = wrong_prem t |
|
783 |
| wrong_prem (Bound _) = true |
|
784 |
| wrong_prem _ = false |
|
15570 | 785 |
val filter_right = List.filter (fn t => not (wrong_prem (HOLogic.dest_Trueprop (hd (Thm.prems_of t))))) |
15411 | 786 |
in |
787 |
fun smp i = funpow i (fn m => filter_right ([spec] RL m)) ([mp]) |
|
788 |
fun smp_tac j = EVERY'[dresolve_tac (smp j), atac] |
|
789 |
end |
|
790 |
||
791 |
||
792 |
fun strip_tac i = REPEAT(resolve_tac [impI,allI] i) |
|
793 |
||
794 |
(*Obsolete form of disjunctive case analysis*) |
|
795 |
fun excluded_middle_tac sP = |
|
796 |
res_inst_tac [("Q",sP)] (excluded_middle RS disjE) |
|
797 |
||
798 |
fun case_tac a = res_inst_tac [("P",a)] case_split_thm |
|
799 |
*} |
|
800 |
||
11687 | 801 |
theorems case_split = case_split_thm [case_names True False] |
9869 | 802 |
|
18457 | 803 |
ML {* |
804 |
structure ProjectRule = ProjectRuleFun |
|
805 |
(struct |
|
806 |
val conjunct1 = thm "conjunct1"; |
|
807 |
val conjunct2 = thm "conjunct2"; |
|
808 |
val mp = thm "mp"; |
|
809 |
end) |
|
810 |
*} |
|
811 |
||
12386 | 812 |
|
813 |
subsubsection {* Intuitionistic Reasoning *} |
|
814 |
||
815 |
lemma impE': |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
816 |
assumes 1: "P --> Q" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
817 |
and 2: "Q ==> R" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
818 |
and 3: "P --> Q ==> P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
819 |
shows R |
12386 | 820 |
proof - |
821 |
from 3 and 1 have P . |
|
822 |
with 1 have Q by (rule impE) |
|
823 |
with 2 show R . |
|
824 |
qed |
|
825 |
||
826 |
lemma allE': |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
827 |
assumes 1: "ALL x. P x" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
828 |
and 2: "P x ==> ALL x. P x ==> Q" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
829 |
shows Q |
12386 | 830 |
proof - |
831 |
from 1 have "P x" by (rule spec) |
|
832 |
from this and 1 show Q by (rule 2) |
|
833 |
qed |
|
834 |
||
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
835 |
lemma notE': |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
836 |
assumes 1: "~ P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
837 |
and 2: "~ P ==> P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
838 |
shows R |
12386 | 839 |
proof - |
840 |
from 2 and 1 have P . |
|
841 |
with 1 show R by (rule notE) |
|
842 |
qed |
|
843 |
||
15801 | 844 |
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE |
845 |
and [Pure.intro!] = iffI conjI impI TrueI notI allI refl |
|
846 |
and [Pure.elim 2] = allE notE' impE' |
|
847 |
and [Pure.intro] = exI disjI2 disjI1 |
|
12386 | 848 |
|
849 |
lemmas [trans] = trans |
|
850 |
and [sym] = sym not_sym |
|
15801 | 851 |
and [Pure.elim?] = iffD1 iffD2 impE |
11750 | 852 |
|
11438
3d9222b80989
declare trans [trans] (*overridden in theory Calculation*);
wenzelm
parents:
11432
diff
changeset
|
853 |
|
11750 | 854 |
subsubsection {* Atomizing meta-level connectives *} |
855 |
||
856 |
lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" |
|
12003 | 857 |
proof |
9488 | 858 |
assume "!!x. P x" |
10383 | 859 |
show "ALL x. P x" by (rule allI) |
9488 | 860 |
next |
861 |
assume "ALL x. P x" |
|
10383 | 862 |
thus "!!x. P x" by (rule allE) |
9488 | 863 |
qed |
864 |
||
11750 | 865 |
lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" |
12003 | 866 |
proof |
9488 | 867 |
assume r: "A ==> B" |
10383 | 868 |
show "A --> B" by (rule impI) (rule r) |
9488 | 869 |
next |
870 |
assume "A --> B" and A |
|
10383 | 871 |
thus B by (rule mp) |
9488 | 872 |
qed |
873 |
||
14749 | 874 |
lemma atomize_not: "(A ==> False) == Trueprop (~A)" |
875 |
proof |
|
876 |
assume r: "A ==> False" |
|
877 |
show "~A" by (rule notI) (rule r) |
|
878 |
next |
|
879 |
assume "~A" and A |
|
880 |
thus False by (rule notE) |
|
881 |
qed |
|
882 |
||
11750 | 883 |
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" |
12003 | 884 |
proof |
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
885 |
assume "x == y" |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
886 |
show "x = y" by (unfold prems) (rule refl) |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
887 |
next |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
888 |
assume "x = y" |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
889 |
thus "x == y" by (rule eq_reflection) |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
890 |
qed |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
891 |
|
12023 | 892 |
lemma atomize_conj [atomize]: |
19121 | 893 |
includes meta_conjunction_syntax |
894 |
shows "(A && B) == Trueprop (A & B)" |
|
12003 | 895 |
proof |
19121 | 896 |
assume conj: "A && B" |
897 |
show "A & B" |
|
898 |
proof (rule conjI) |
|
899 |
from conj show A by (rule conjunctionD1) |
|
900 |
from conj show B by (rule conjunctionD2) |
|
901 |
qed |
|
11953 | 902 |
next |
19121 | 903 |
assume conj: "A & B" |
904 |
show "A && B" |
|
905 |
proof - |
|
906 |
from conj show A .. |
|
907 |
from conj show B .. |
|
11953 | 908 |
qed |
909 |
qed |
|
910 |
||
12386 | 911 |
lemmas [symmetric, rulify] = atomize_all atomize_imp |
18832 | 912 |
and [symmetric, defn] = atomize_all atomize_imp atomize_eq |
12386 | 913 |
|
11750 | 914 |
|
915 |
subsubsection {* Classical Reasoner setup *} |
|
9529 | 916 |
|
10383 | 917 |
use "cladata.ML" |
918 |
setup hypsubst_setup |
|
18708 | 919 |
setup {* ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac) *} |
10383 | 920 |
setup Classical.setup |
20766 | 921 |
setup ResAtpset.setup |
20223 | 922 |
setup clasetup |
19162 | 923 |
|
20223 | 924 |
lemma contrapos_np: "~ Q ==> (~ P ==> Q) ==> P" |
925 |
apply (erule swap) |
|
926 |
apply (erule (1) meta_mp) |
|
927 |
done |
|
10383 | 928 |
|
18689
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
929 |
declare ex_ex1I [rule del, intro! 2] |
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
930 |
and ex1I [intro] |
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
931 |
|
12386 | 932 |
lemmas [intro?] = ext |
933 |
and [elim?] = ex1_implies_ex |
|
11977 | 934 |
|
9869 | 935 |
use "blastdata.ML" |
936 |
setup Blast.setup |
|
4868 | 937 |
|
11750 | 938 |
|
17459 | 939 |
subsubsection {* Simplifier setup *} |
11750 | 940 |
|
12281 | 941 |
lemma meta_eq_to_obj_eq: "x == y ==> x = y" |
942 |
proof - |
|
943 |
assume r: "x == y" |
|
944 |
show "x = y" by (unfold r) (rule refl) |
|
945 |
qed |
|
946 |
||
947 |
lemma eta_contract_eq: "(%s. f s) = f" .. |
|
948 |
||
949 |
lemma simp_thms: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
950 |
shows not_not: "(~ ~ P) = P" |
15354 | 951 |
and Not_eq_iff: "((~P) = (~Q)) = (P = Q)" |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
952 |
and |
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
953 |
"(P ~= Q) = (P = (~Q))" |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
954 |
"(P | ~P) = True" "(~P | P) = True" |
12281 | 955 |
"(x = x) = True" |
956 |
"(~True) = False" "(~False) = True" |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
957 |
"(~P) ~= P" "P ~= (~P)" |
12281 | 958 |
"(True=P) = P" "(P=True) = P" "(False=P) = (~P)" "(P=False) = (~P)" |
959 |
"(True --> P) = P" "(False --> P) = True" |
|
960 |
"(P --> True) = True" "(P --> P) = True" |
|
961 |
"(P --> False) = (~P)" "(P --> ~P) = (~P)" |
|
962 |
"(P & True) = P" "(True & P) = P" |
|
963 |
"(P & False) = False" "(False & P) = False" |
|
964 |
"(P & P) = P" "(P & (P & Q)) = (P & Q)" |
|
965 |
"(P & ~P) = False" "(~P & P) = False" |
|
966 |
"(P | True) = True" "(True | P) = True" |
|
967 |
"(P | False) = P" "(False | P) = P" |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
968 |
"(P | P) = P" "(P | (P | Q)) = (P | Q)" and |
12281 | 969 |
"(ALL x. P) = P" "(EX x. P) = P" "EX x. x=t" "EX x. t=x" |
970 |
-- {* needed for the one-point-rule quantifier simplification procs *} |
|
971 |
-- {* essential for termination!! *} and |
|
972 |
"!!P. (EX x. x=t & P(x)) = P(t)" |
|
973 |
"!!P. (EX x. t=x & P(x)) = P(t)" |
|
974 |
"!!P. (ALL x. x=t --> P(x)) = P(t)" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
975 |
"!!P. (ALL x. t=x --> P(x)) = P(t)" |
17589 | 976 |
by (blast, blast, blast, blast, blast, iprover+) |
13421 | 977 |
|
12281 | 978 |
lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))" |
17589 | 979 |
by iprover |
12281 | 980 |
|
981 |
lemma ex_simps: |
|
982 |
"!!P Q. (EX x. P x & Q) = ((EX x. P x) & Q)" |
|
983 |
"!!P Q. (EX x. P & Q x) = (P & (EX x. Q x))" |
|
984 |
"!!P Q. (EX x. P x | Q) = ((EX x. P x) | Q)" |
|
985 |
"!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))" |
|
986 |
"!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)" |
|
987 |
"!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))" |
|
988 |
-- {* Miniscoping: pushing in existential quantifiers. *} |
|
17589 | 989 |
by (iprover | blast)+ |
12281 | 990 |
|
991 |
lemma all_simps: |
|
992 |
"!!P Q. (ALL x. P x & Q) = ((ALL x. P x) & Q)" |
|
993 |
"!!P Q. (ALL x. P & Q x) = (P & (ALL x. Q x))" |
|
994 |
"!!P Q. (ALL x. P x | Q) = ((ALL x. P x) | Q)" |
|
995 |
"!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))" |
|
996 |
"!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)" |
|
997 |
"!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))" |
|
998 |
-- {* Miniscoping: pushing in universal quantifiers. *} |
|
17589 | 999 |
by (iprover | blast)+ |
12281 | 1000 |
|
14201 | 1001 |
lemma disj_absorb: "(A | A) = A" |
1002 |
by blast |
|
1003 |
||
1004 |
lemma disj_left_absorb: "(A | (A | B)) = (A | B)" |
|
1005 |
by blast |
|
1006 |
||
1007 |
lemma conj_absorb: "(A & A) = A" |
|
1008 |
by blast |
|
1009 |
||
1010 |
lemma conj_left_absorb: "(A & (A & B)) = (A & B)" |
|
1011 |
by blast |
|
1012 |
||
12281 | 1013 |
lemma eq_ac: |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1014 |
shows eq_commute: "(a=b) = (b=a)" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1015 |
and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))" |
17589 | 1016 |
and eq_assoc: "((P=Q)=R) = (P=(Q=R))" by (iprover, blast+) |
1017 |
lemma neq_commute: "(a~=b) = (b~=a)" by iprover |
|
12281 | 1018 |
|
1019 |
lemma conj_comms: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1020 |
shows conj_commute: "(P&Q) = (Q&P)" |
17589 | 1021 |
and conj_left_commute: "(P&(Q&R)) = (Q&(P&R))" by iprover+ |
1022 |
lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by iprover |
|
12281 | 1023 |
|
19174 | 1024 |
lemmas conj_ac = conj_commute conj_left_commute conj_assoc |
1025 |
||
12281 | 1026 |
lemma disj_comms: |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1027 |
shows disj_commute: "(P|Q) = (Q|P)" |
17589 | 1028 |
and disj_left_commute: "(P|(Q|R)) = (Q|(P|R))" by iprover+ |
1029 |
lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by iprover |
|
12281 | 1030 |
|
19174 | 1031 |
lemmas disj_ac = disj_commute disj_left_commute disj_assoc |
1032 |
||
17589 | 1033 |
lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by iprover |
1034 |
lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by iprover |
|
12281 | 1035 |
|
17589 | 1036 |
lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by iprover |
1037 |
lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by iprover |
|
12281 | 1038 |
|
17589 | 1039 |
lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by iprover |
1040 |
lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by iprover |
|
1041 |
lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by iprover |
|
12281 | 1042 |
|
1043 |
text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *} |
|
1044 |
lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast |
|
1045 |
lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast |
|
1046 |
||
1047 |
lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast |
|
1048 |
lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast |
|
1049 |
||
17589 | 1050 |
lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by iprover |
12281 | 1051 |
lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast |
1052 |
lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast |
|
1053 |
lemma not_iff: "(P~=Q) = (P = (~Q))" by blast |
|
1054 |
lemma disj_not1: "(~P | Q) = (P --> Q)" by blast |
|
1055 |
lemma disj_not2: "(P | ~Q) = (Q --> P)" -- {* changes orientation :-( *} |
|
1056 |
by blast |
|
1057 |
lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast |
|
1058 |
||
17589 | 1059 |
lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by iprover |
12281 | 1060 |
|
1061 |
||
1062 |
lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q" |
|
1063 |
-- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *} |
|
1064 |
-- {* cases boil down to the same thing. *} |
|
1065 |
by blast |
|
1066 |
||
1067 |
lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast |
|
1068 |
lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast |
|
17589 | 1069 |
lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by iprover |
1070 |
lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by iprover |
|
12281 | 1071 |
|
17589 | 1072 |
lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by iprover |
1073 |
lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by iprover |
|
12281 | 1074 |
|
1075 |
text {* |
|
1076 |
\medskip The @{text "&"} congruence rule: not included by default! |
|
1077 |
May slow rewrite proofs down by as much as 50\% *} |
|
1078 |
||
1079 |
lemma conj_cong: |
|
1080 |
"(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))" |
|
17589 | 1081 |
by iprover |
12281 | 1082 |
|
1083 |
lemma rev_conj_cong: |
|
1084 |
"(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))" |
|
17589 | 1085 |
by iprover |
12281 | 1086 |
|
1087 |
text {* The @{text "|"} congruence rule: not included by default! *} |
|
1088 |
||
1089 |
lemma disj_cong: |
|
1090 |
"(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))" |
|
1091 |
by blast |
|
1092 |
||
1093 |
lemma eq_sym_conv: "(x = y) = (y = x)" |
|
17589 | 1094 |
by iprover |
12281 | 1095 |
|
1096 |
||
1097 |
text {* \medskip if-then-else rules *} |
|
1098 |
||
1099 |
lemma if_True: "(if True then x else y) = x" |
|
1100 |
by (unfold if_def) blast |
|
1101 |
||
1102 |
lemma if_False: "(if False then x else y) = y" |
|
1103 |
by (unfold if_def) blast |
|
1104 |
||
1105 |
lemma if_P: "P ==> (if P then x else y) = x" |
|
1106 |
by (unfold if_def) blast |
|
1107 |
||
1108 |
lemma if_not_P: "~P ==> (if P then x else y) = y" |
|
1109 |
by (unfold if_def) blast |
|
1110 |
||
1111 |
lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" |
|
1112 |
apply (rule case_split [of Q]) |
|
15481 | 1113 |
apply (simplesubst if_P) |
1114 |
prefer 3 apply (simplesubst if_not_P, blast+) |
|
12281 | 1115 |
done |
1116 |
||
1117 |
lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" |
|
15481 | 1118 |
by (simplesubst split_if, blast) |
12281 | 1119 |
|
1120 |
lemmas if_splits = split_if split_if_asm |
|
1121 |
||
1122 |
lemma if_def2: "(if Q then x else y) = ((Q --> x) & (~ Q --> y))" |
|
1123 |
by (rule split_if) |
|
1124 |
||
1125 |
lemma if_cancel: "(if c then x else x) = x" |
|
15481 | 1126 |
by (simplesubst split_if, blast) |
12281 | 1127 |
|
1128 |
lemma if_eq_cancel: "(if x = y then y else x) = x" |
|
15481 | 1129 |
by (simplesubst split_if, blast) |
12281 | 1130 |
|
1131 |
lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))" |
|
19796 | 1132 |
-- {* This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "==>"} symbol. *} |
12281 | 1133 |
by (rule split_if) |
1134 |
||
1135 |
lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" |
|
19796 | 1136 |
-- {* And this form is useful for expanding @{text "if"}s on the LEFT. *} |
15481 | 1137 |
apply (simplesubst split_if, blast) |
12281 | 1138 |
done |
1139 |
||
17589 | 1140 |
lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) iprover |
1141 |
lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) iprover |
|
12281 | 1142 |
|
15423 | 1143 |
text {* \medskip let rules for simproc *} |
1144 |
||
1145 |
lemma Let_folded: "f x \<equiv> g x \<Longrightarrow> Let x f \<equiv> Let x g" |
|
1146 |
by (unfold Let_def) |
|
1147 |
||
1148 |
lemma Let_unfold: "f x \<equiv> g \<Longrightarrow> Let x f \<equiv> g" |
|
1149 |
by (unfold Let_def) |
|
1150 |
||
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1151 |
text {* |
16999 | 1152 |
The following copy of the implication operator is useful for |
1153 |
fine-tuning congruence rules. It instructs the simplifier to simplify |
|
1154 |
its premise. |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1155 |
*} |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1156 |
|
17197 | 1157 |
constdefs |
1158 |
simp_implies :: "[prop, prop] => prop" (infixr "=simp=>" 1) |
|
1159 |
"simp_implies \<equiv> op ==>" |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1160 |
|
18457 | 1161 |
lemma simp_impliesI: |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1162 |
assumes PQ: "(PROP P \<Longrightarrow> PROP Q)" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1163 |
shows "PROP P =simp=> PROP Q" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1164 |
apply (unfold simp_implies_def) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1165 |
apply (rule PQ) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1166 |
apply assumption |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1167 |
done |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1168 |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1169 |
lemma simp_impliesE: |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1170 |
assumes PQ:"PROP P =simp=> PROP Q" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1171 |
and P: "PROP P" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1172 |
and QR: "PROP Q \<Longrightarrow> PROP R" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1173 |
shows "PROP R" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1174 |
apply (rule QR) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1175 |
apply (rule PQ [unfolded simp_implies_def]) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1176 |
apply (rule P) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1177 |
done |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1178 |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1179 |
lemma simp_implies_cong: |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1180 |
assumes PP' :"PROP P == PROP P'" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1181 |
and P'QQ': "PROP P' ==> (PROP Q == PROP Q')" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1182 |
shows "(PROP P =simp=> PROP Q) == (PROP P' =simp=> PROP Q')" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1183 |
proof (unfold simp_implies_def, rule equal_intr_rule) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1184 |
assume PQ: "PROP P \<Longrightarrow> PROP Q" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1185 |
and P': "PROP P'" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1186 |
from PP' [symmetric] and P' have "PROP P" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1187 |
by (rule equal_elim_rule1) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1188 |
hence "PROP Q" by (rule PQ) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1189 |
with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1190 |
next |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1191 |
assume P'Q': "PROP P' \<Longrightarrow> PROP Q'" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1192 |
and P: "PROP P" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1193 |
from PP' and P have P': "PROP P'" by (rule equal_elim_rule1) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1194 |
hence "PROP Q'" by (rule P'Q') |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1195 |
with P'QQ' [OF P', symmetric] show "PROP Q" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1196 |
by (rule equal_elim_rule1) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1197 |
qed |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1198 |
|
17459 | 1199 |
|
1200 |
text {* \medskip Actual Installation of the Simplifier. *} |
|
14201 | 1201 |
|
9869 | 1202 |
use "simpdata.ML" |
1203 |
setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup |
|
1204 |
setup Splitter.setup setup Clasimp.setup |
|
18591 | 1205 |
setup EqSubst.setup |
15481 | 1206 |
|
17459 | 1207 |
|
1208 |
subsubsection {* Code generator setup *} |
|
1209 |
||
1210 |
types_code |
|
1211 |
"bool" ("bool") |
|
1212 |
attach (term_of) {* |
|
1213 |
fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; |
|
1214 |
*} |
|
1215 |
attach (test) {* |
|
1216 |
fun gen_bool i = one_of [false, true]; |
|
1217 |
*} |
|
18887 | 1218 |
"prop" ("bool") |
1219 |
attach (term_of) {* |
|
1220 |
fun term_of_prop b = |
|
1221 |
HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const); |
|
1222 |
*} |
|
17459 | 1223 |
|
1224 |
consts_code |
|
18887 | 1225 |
"Trueprop" ("(_)") |
17459 | 1226 |
"True" ("true") |
1227 |
"False" ("false") |
|
1228 |
"Not" ("not") |
|
1229 |
"op |" ("(_ orelse/ _)") |
|
1230 |
"op &" ("(_ andalso/ _)") |
|
1231 |
"HOL.If" ("(if _/ then _/ else _)") |
|
1232 |
||
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1233 |
setup {* |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1234 |
let |
17459 | 1235 |
|
1236 |
fun eq_codegen thy defs gr dep thyname b t = |
|
1237 |
(case strip_comb t of |
|
1238 |
(Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE |
|
1239 |
| (Const ("op =", _), [t, u]) => |
|
1240 |
let |
|
1241 |
val (gr', pt) = Codegen.invoke_codegen thy defs dep thyname false (gr, t); |
|
17639
50878db27b94
eq_codegen now ensures that code for bool type is generated.
berghofe
parents:
17589
diff
changeset
|
1242 |
val (gr'', pu) = Codegen.invoke_codegen thy defs dep thyname false (gr', u); |
50878db27b94
eq_codegen now ensures that code for bool type is generated.
berghofe
parents:
17589
diff
changeset
|
1243 |
val (gr''', _) = Codegen.invoke_tycodegen thy defs dep thyname false (gr'', HOLogic.boolT) |
17459 | 1244 |
in |
17639
50878db27b94
eq_codegen now ensures that code for bool type is generated.
berghofe
parents:
17589
diff
changeset
|
1245 |
SOME (gr''', Codegen.parens |
17459 | 1246 |
(Pretty.block [pt, Pretty.str " =", Pretty.brk 1, pu])) |
1247 |
end |
|
1248 |
| (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen |
|
1249 |
thy defs dep thyname b (gr, Codegen.eta_expand t ts 2)) |
|
1250 |
| _ => NONE); |
|
1251 |
||
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1252 |
in |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1253 |
|
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1254 |
Codegen.add_codegen "eq_codegen" eq_codegen |
18887 | 1255 |
|
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1256 |
end |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1257 |
*} |
18887 | 1258 |
|
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1259 |
setup {* |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1260 |
let |
18887 | 1261 |
|
1262 |
fun evaluation_tac i = Tactical.PRIMITIVE (Drule.fconv_rule |
|
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1263 |
(Drule.goals_conv (equal i) Codegen.evaluation_conv)); |
18887 | 1264 |
|
1265 |
val evaluation_meth = |
|
1266 |
Method.no_args (Method.METHOD (fn _ => evaluation_tac 1 THEN rtac TrueI 1)); |
|
1267 |
||
17459 | 1268 |
in |
1269 |
||
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1270 |
Method.add_method ("evaluation", evaluation_meth, "solve goal by evaluation") |
18887 | 1271 |
|
17459 | 1272 |
end; |
1273 |
*} |
|
1274 |
||
19961 | 1275 |
|
15481 | 1276 |
subsection {* Other simple lemmas *} |
1277 |
||
15411 | 1278 |
declare disj_absorb [simp] conj_absorb [simp] |
14201 | 1279 |
|
13723 | 1280 |
lemma ex1_eq[iff]: "EX! x. x = t" "EX! x. t = x" |
1281 |
by blast+ |
|
1282 |
||
15481 | 1283 |
|
13638 | 1284 |
theorem choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))" |
1285 |
apply (rule iffI) |
|
1286 |
apply (rule_tac a = "%x. THE y. P x y" in ex1I) |
|
1287 |
apply (fast dest!: theI') |
|
1288 |
apply (fast intro: ext the1_equality [symmetric]) |
|
1289 |
apply (erule ex1E) |
|
1290 |
apply (rule allI) |
|
1291 |
apply (rule ex1I) |
|
1292 |
apply (erule spec) |
|
1293 |
apply (erule_tac x = "%z. if z = x then y else f z" in allE) |
|
1294 |
apply (erule impE) |
|
1295 |
apply (rule allI) |
|
1296 |
apply (rule_tac P = "xa = x" in case_split_thm) |
|
14208 | 1297 |
apply (drule_tac [3] x = x in fun_cong, simp_all) |
13638 | 1298 |
done |
1299 |
||
13438
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
1300 |
text{*Needs only HOL-lemmas:*} |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
1301 |
lemma mk_left_commute: |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
1302 |
assumes a: "\<And>x y z. f (f x y) z = f x (f y z)" and |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
1303 |
c: "\<And>x y. f x y = f y x" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
1304 |
shows "f x (f y z) = f y (f x z)" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
1305 |
by(rule trans[OF trans[OF c a] arg_cong[OF c, of "f y"]]) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
1306 |
|
11750 | 1307 |
|
15481 | 1308 |
subsection {* Generic cases and induction *} |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1309 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1310 |
constdefs |
18457 | 1311 |
induct_forall where "induct_forall P == \<forall>x. P x" |
1312 |
induct_implies where "induct_implies A B == A \<longrightarrow> B" |
|
1313 |
induct_equal where "induct_equal x y == x = y" |
|
1314 |
induct_conj where "induct_conj A B == A \<and> B" |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1315 |
|
11989 | 1316 |
lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))" |
18457 | 1317 |
by (unfold atomize_all induct_forall_def) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1318 |
|
11989 | 1319 |
lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)" |
18457 | 1320 |
by (unfold atomize_imp induct_implies_def) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1321 |
|
11989 | 1322 |
lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)" |
18457 | 1323 |
by (unfold atomize_eq induct_equal_def) |
1324 |
||
1325 |
lemma induct_conj_eq: |
|
1326 |
includes meta_conjunction_syntax |
|
1327 |
shows "(A && B) == Trueprop (induct_conj A B)" |
|
1328 |
by (unfold atomize_conj induct_conj_def) |
|
1329 |
||
1330 |
lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq |
|
1331 |
lemmas induct_rulify [symmetric, standard] = induct_atomize |
|
1332 |
lemmas induct_rulify_fallback = |
|
1333 |
induct_forall_def induct_implies_def induct_equal_def induct_conj_def |
|
1334 |
||
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1335 |
|
11989 | 1336 |
lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = |
1337 |
induct_conj (induct_forall A) (induct_forall B)" |
|
17589 | 1338 |
by (unfold induct_forall_def induct_conj_def) iprover |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1339 |
|
11989 | 1340 |
lemma induct_implies_conj: "induct_implies C (induct_conj A B) = |
1341 |
induct_conj (induct_implies C A) (induct_implies C B)" |
|
17589 | 1342 |
by (unfold induct_implies_def induct_conj_def) iprover |
11989 | 1343 |
|
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1344 |
lemma induct_conj_curry: "(induct_conj A B ==> PROP C) == (A ==> B ==> PROP C)" |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1345 |
proof |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1346 |
assume r: "induct_conj A B ==> PROP C" and A B |
18457 | 1347 |
show "PROP C" by (rule r) (simp add: induct_conj_def `A` `B`) |
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1348 |
next |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1349 |
assume r: "A ==> B ==> PROP C" and "induct_conj A B" |
18457 | 1350 |
show "PROP C" by (rule r) (simp_all add: `induct_conj A B` [unfolded induct_conj_def]) |
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1351 |
qed |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1352 |
|
11989 | 1353 |
lemmas induct_conj = induct_forall_conj induct_implies_conj induct_conj_curry |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1354 |
|
11989 | 1355 |
hide const induct_forall induct_implies induct_equal induct_conj |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1356 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1357 |
text {* Method setup. *} |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1358 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1359 |
ML {* |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1360 |
structure InductMethod = InductMethodFun |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1361 |
(struct |
15411 | 1362 |
val cases_default = thm "case_split" |
1363 |
val atomize = thms "induct_atomize" |
|
18457 | 1364 |
val rulify = thms "induct_rulify" |
1365 |
val rulify_fallback = thms "induct_rulify_fallback" |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1366 |
end); |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1367 |
*} |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1368 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1369 |
setup InductMethod.setup |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1370 |
|
18457 | 1371 |
|
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1372 |
text {* itself as a code generator datatype *} |
18702 | 1373 |
|
19598
d68dd20af31f
different object logic setup for CodegenTheorems
haftmann
parents:
19536
diff
changeset
|
1374 |
setup {* |
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1375 |
let fun add_itself thy = |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1376 |
let |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1377 |
val v = ("'a", []); |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1378 |
val t = Logic.mk_type (TFree v); |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1379 |
val Const (c, ty) = t; |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1380 |
val (_, Type (dtco, _)) = strip_type ty; |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1381 |
in |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1382 |
thy |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1383 |
|> CodegenData.add_datatype (dtco, (([v], [(c, [])]), CodegenData.lazy (fn () => []))) |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1384 |
end |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1385 |
in add_itself end; |
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
1386 |
*} |
18702 | 1387 |
|
20698 | 1388 |
text {* code generation for arbitrary as exception *} |
1389 |
||
1390 |
setup {* |
|
20833 | 1391 |
CodegenSerializer.add_undefined "SML" "arbitrary" "(raise Fail \"arbitrary\")" |
20698 | 1392 |
*} |
1393 |
code_const arbitrary |
|
1394 |
(Haskell target_atom "(error \"arbitrary\")") |
|
1395 |
||
14357 | 1396 |
end |