author | haftmann |
Mon, 06 Jul 2009 14:19:13 +0200 | |
changeset 31949 | 3f933687fae9 |
parent 31902 | 862ae16a799d |
child 31956 | c3844c4d0c2c |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/HOL.thy |
11750 | 2 |
Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson |
3 |
*) |
|
923 | 4 |
|
11750 | 5 |
header {* The basis of Higher-Order Logic *} |
923 | 6 |
|
15131 | 7 |
theory HOL |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
8 |
imports Pure "~~/src/Tools/Code_Generator" |
23163 | 9 |
uses |
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28856
diff
changeset
|
10 |
("Tools/hologic.ML") |
30980 | 11 |
"~~/src/Tools/auto_solve.ML" |
23171 | 12 |
"~~/src/Tools/IsaPlanner/zipper.ML" |
13 |
"~~/src/Tools/IsaPlanner/isand.ML" |
|
14 |
"~~/src/Tools/IsaPlanner/rw_tools.ML" |
|
15 |
"~~/src/Tools/IsaPlanner/rw_inst.ML" |
|
30165
6ee87f67d9cd
moved generic intuitionistic prover to src/Tools/intuitionistic.ML;
wenzelm
parents:
30160
diff
changeset
|
16 |
"~~/src/Tools/intuitionistic.ML" |
30160
5f7b17941730
moved some generic tools to src/Tools/ -- src/Provers is essentially obsolete;
wenzelm
parents:
30063
diff
changeset
|
17 |
"~~/src/Tools/project_rule.ML" |
23263 | 18 |
"~~/src/Provers/hypsubst.ML" |
19 |
"~~/src/Provers/splitter.ML" |
|
23163 | 20 |
"~~/src/Provers/classical.ML" |
21 |
"~~/src/Provers/blast.ML" |
|
22 |
"~~/src/Provers/clasimp.ML" |
|
30160
5f7b17941730
moved some generic tools to src/Tools/ -- src/Provers is essentially obsolete;
wenzelm
parents:
30063
diff
changeset
|
23 |
"~~/src/Tools/coherent.ML" |
5f7b17941730
moved some generic tools to src/Tools/ -- src/Provers is essentially obsolete;
wenzelm
parents:
30063
diff
changeset
|
24 |
"~~/src/Tools/eqsubst.ML" |
23163 | 25 |
"~~/src/Provers/quantifier1.ML" |
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28856
diff
changeset
|
26 |
("Tools/simpdata.ML") |
25741 | 27 |
"~~/src/Tools/random_word.ML" |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
28 |
"~~/src/Tools/atomize_elim.ML" |
24901
d3cbf79769b9
added first version of user-space type system for class target
haftmann
parents:
24844
diff
changeset
|
29 |
"~~/src/Tools/induct.ML" |
27326
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
wenzelm
parents:
27212
diff
changeset
|
30 |
("~~/src/Tools/induct_tacs.ML") |
29505 | 31 |
("Tools/recfun_codegen.ML") |
15131 | 32 |
begin |
2260 | 33 |
|
31299 | 34 |
setup {* Intuitionistic.method_setup @{binding iprover} *} |
30165
6ee87f67d9cd
moved generic intuitionistic prover to src/Tools/intuitionistic.ML;
wenzelm
parents:
30160
diff
changeset
|
35 |
|
6ee87f67d9cd
moved generic intuitionistic prover to src/Tools/intuitionistic.ML;
wenzelm
parents:
30160
diff
changeset
|
36 |
|
11750 | 37 |
subsection {* Primitive logic *} |
38 |
||
39 |
subsubsection {* Core syntax *} |
|
2260 | 40 |
|
14854 | 41 |
classes type |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
42 |
defaultsort type |
25494
b2484a7912ac
replaced typedecl interpretation by ObjectLogic.typedecl (based on base_sort);
wenzelm
parents:
25460
diff
changeset
|
43 |
setup {* ObjectLogic.add_base_sort @{sort type} *} |
25460
b80087af2274
interpretation of typedecls: instantiation to class type
haftmann
parents:
25388
diff
changeset
|
44 |
|
b80087af2274
interpretation of typedecls: instantiation to class type
haftmann
parents:
25388
diff
changeset
|
45 |
arities |
b80087af2274
interpretation of typedecls: instantiation to class type
haftmann
parents:
25388
diff
changeset
|
46 |
"fun" :: (type, type) type |
b80087af2274
interpretation of typedecls: instantiation to class type
haftmann
parents:
25388
diff
changeset
|
47 |
itself :: (type) type |
b80087af2274
interpretation of typedecls: instantiation to class type
haftmann
parents:
25388
diff
changeset
|
48 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
49 |
global |
923 | 50 |
|
7357 | 51 |
typedecl bool |
923 | 52 |
|
11750 | 53 |
judgment |
54 |
Trueprop :: "bool => prop" ("(_)" 5) |
|
923 | 55 |
|
11750 | 56 |
consts |
7357 | 57 |
Not :: "bool => bool" ("~ _" [40] 40) |
58 |
True :: bool |
|
59 |
False :: bool |
|
923 | 60 |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
61 |
The :: "('a => bool) => 'a" |
7357 | 62 |
All :: "('a => bool) => bool" (binder "ALL " 10) |
63 |
Ex :: "('a => bool) => bool" (binder "EX " 10) |
|
64 |
Ex1 :: "('a => bool) => bool" (binder "EX! " 10) |
|
65 |
Let :: "['a, 'a => 'b] => 'b" |
|
923 | 66 |
|
22839 | 67 |
"op =" :: "['a, 'a] => bool" (infixl "=" 50) |
68 |
"op &" :: "[bool, bool] => bool" (infixr "&" 35) |
|
69 |
"op |" :: "[bool, bool] => bool" (infixr "|" 30) |
|
70 |
"op -->" :: "[bool, bool] => bool" (infixr "-->" 25) |
|
923 | 71 |
|
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
72 |
local |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
73 |
|
16587 | 74 |
consts |
75 |
If :: "[bool, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) |
|
2260 | 76 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
77 |
|
11750 | 78 |
subsubsection {* Additional concrete syntax *} |
2260 | 79 |
|
21210 | 80 |
notation (output) |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
81 |
"op =" (infix "=" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
82 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
83 |
abbreviation |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
84 |
not_equal :: "['a, 'a] => bool" (infixl "~=" 50) where |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
85 |
"x ~= y == ~ (x = y)" |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
86 |
|
21210 | 87 |
notation (output) |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
88 |
not_equal (infix "~=" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
89 |
|
21210 | 90 |
notation (xsymbols) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
91 |
Not ("\<not> _" [40] 40) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
92 |
"op &" (infixr "\<and>" 35) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
93 |
"op |" (infixr "\<or>" 30) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
94 |
"op -->" (infixr "\<longrightarrow>" 25) and |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
95 |
not_equal (infix "\<noteq>" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
96 |
|
21210 | 97 |
notation (HTML output) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
98 |
Not ("\<not> _" [40] 40) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
99 |
"op &" (infixr "\<and>" 35) and |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
100 |
"op |" (infixr "\<or>" 30) and |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
101 |
not_equal (infix "\<noteq>" 50) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
102 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
103 |
abbreviation (iff) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21250
diff
changeset
|
104 |
iff :: "[bool, bool] => bool" (infixr "<->" 25) where |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
105 |
"A <-> B == A = B" |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
106 |
|
21210 | 107 |
notation (xsymbols) |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
108 |
iff (infixr "\<longleftrightarrow>" 25) |
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
109 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
110 |
|
4868 | 111 |
nonterminals |
923 | 112 |
letbinds letbind |
113 |
case_syn cases_syn |
|
114 |
||
115 |
syntax |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
116 |
"_The" :: "[pttrn, bool] => 'a" ("(3THE _./ _)" [0, 10] 10) |
923 | 117 |
|
7357 | 118 |
"_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10) |
119 |
"" :: "letbind => letbinds" ("_") |
|
120 |
"_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") |
|
121 |
"_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) |
|
923 | 122 |
|
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
123 |
"_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
|
124 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10) |
7357 | 125 |
"" :: "case_syn => cases_syn" ("_") |
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
126 |
"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _") |
923 | 127 |
|
128 |
translations |
|
13764 | 129 |
"THE x. P" == "The (%x. P)" |
923 | 130 |
"_Let (_binds b bs) e" == "_Let b (_Let bs e)" |
1114 | 131 |
"let x = a in e" == "Let a (%x. e)" |
923 | 132 |
|
13763
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
133 |
print_translation {* |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
134 |
(* To avoid eta-contraction of body: *) |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
135 |
[("The", fn [Abs abs] => |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
136 |
let val (x,t) = atomic_abs_tr' abs |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
137 |
in Syntax.const "_The" $ x $ t end)] |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
138 |
*} |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
139 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
140 |
syntax (xsymbols) |
11687 | 141 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ \<Rightarrow>/ _)" 10) |
21524 | 142 |
|
143 |
notation (xsymbols) |
|
144 |
All (binder "\<forall>" 10) and |
|
145 |
Ex (binder "\<exists>" 10) and |
|
146 |
Ex1 (binder "\<exists>!" 10) |
|
2372 | 147 |
|
21524 | 148 |
notation (HTML output) |
149 |
All (binder "\<forall>" 10) and |
|
150 |
Ex (binder "\<exists>" 10) and |
|
151 |
Ex1 (binder "\<exists>!" 10) |
|
6340 | 152 |
|
21524 | 153 |
notation (HOL) |
154 |
All (binder "! " 10) and |
|
155 |
Ex (binder "? " 10) and |
|
156 |
Ex1 (binder "?! " 10) |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
157 |
|
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
158 |
|
11750 | 159 |
subsubsection {* Axioms and basic definitions *} |
2260 | 160 |
|
7357 | 161 |
axioms |
15380 | 162 |
refl: "t = (t::'a)" |
28513 | 163 |
subst: "s = t \<Longrightarrow> P s \<Longrightarrow> P t" |
15380 | 164 |
ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" |
165 |
-- {*Extensionality is built into the meta-logic, and this rule expresses |
|
166 |
a related property. It is an eta-expanded version of the traditional |
|
167 |
rule, and similar to the ABS rule of HOL*} |
|
6289 | 168 |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
169 |
the_eq_trivial: "(THE x. x = a) = (a::'a)" |
923 | 170 |
|
15380 | 171 |
impI: "(P ==> Q) ==> P-->Q" |
172 |
mp: "[| P-->Q; P |] ==> Q" |
|
173 |
||
174 |
||
923 | 175 |
defs |
7357 | 176 |
True_def: "True == ((%x::bool. x) = (%x. x))" |
177 |
All_def: "All(P) == (P = (%x. True))" |
|
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
178 |
Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" |
7357 | 179 |
False_def: "False == (!P. P)" |
180 |
not_def: "~ P == P-->False" |
|
181 |
and_def: "P & Q == !R. (P-->Q-->R) --> R" |
|
182 |
or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" |
|
183 |
Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" |
|
923 | 184 |
|
7357 | 185 |
axioms |
186 |
iff: "(P-->Q) --> (Q-->P) --> (P=Q)" |
|
187 |
True_or_False: "(P=True) | (P=False)" |
|
923 | 188 |
|
189 |
defs |
|
24219 | 190 |
Let_def: "Let s f == f(s)" |
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
191 |
if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" |
5069 | 192 |
|
14223
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
193 |
finalconsts |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
194 |
"op =" |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
195 |
"op -->" |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
196 |
The |
22481
79c2724c36b5
added class "default" and expansion axioms for undefined
haftmann
parents:
22473
diff
changeset
|
197 |
|
79c2724c36b5
added class "default" and expansion axioms for undefined
haftmann
parents:
22473
diff
changeset
|
198 |
axiomatization |
79c2724c36b5
added class "default" and expansion axioms for undefined
haftmann
parents:
22473
diff
changeset
|
199 |
undefined :: 'a |
79c2724c36b5
added class "default" and expansion axioms for undefined
haftmann
parents:
22473
diff
changeset
|
200 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
201 |
|
22481
79c2724c36b5
added class "default" and expansion axioms for undefined
haftmann
parents:
22473
diff
changeset
|
202 |
subsubsection {* Generic classes and algebraic operations *} |
79c2724c36b5
added class "default" and expansion axioms for undefined
haftmann
parents:
22473
diff
changeset
|
203 |
|
29608 | 204 |
class default = |
24901
d3cbf79769b9
added first version of user-space type system for class target
haftmann
parents:
24844
diff
changeset
|
205 |
fixes default :: 'a |
4868 | 206 |
|
29608 | 207 |
class zero = |
25062 | 208 |
fixes zero :: 'a ("0") |
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
209 |
|
29608 | 210 |
class one = |
25062 | 211 |
fixes one :: 'a ("1") |
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
212 |
|
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
213 |
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
|
214 |
|
29608 | 215 |
class plus = |
25062 | 216 |
fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "+" 65) |
11750 | 217 |
|
29608 | 218 |
class minus = |
25762 | 219 |
fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "-" 65) |
220 |
||
29608 | 221 |
class uminus = |
25062 | 222 |
fixes uminus :: "'a \<Rightarrow> 'a" ("- _" [81] 80) |
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
223 |
|
29608 | 224 |
class times = |
25062 | 225 |
fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70) |
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
226 |
|
29608 | 227 |
class inverse = |
20590
bf92900995f8
introduced syntactic classes; moved some setup to Pure/codegen, Pure/nbe or OperationalEquality.thy
haftmann
parents:
20453
diff
changeset
|
228 |
fixes inverse :: "'a \<Rightarrow> 'a" |
25062 | 229 |
and divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "'/" 70) |
21524 | 230 |
|
29608 | 231 |
class abs = |
23878 | 232 |
fixes abs :: "'a \<Rightarrow> 'a" |
25388 | 233 |
begin |
23878 | 234 |
|
21524 | 235 |
notation (xsymbols) |
236 |
abs ("\<bar>_\<bar>") |
|
25388 | 237 |
|
21524 | 238 |
notation (HTML output) |
239 |
abs ("\<bar>_\<bar>") |
|
11750 | 240 |
|
25388 | 241 |
end |
242 |
||
29608 | 243 |
class sgn = |
25062 | 244 |
fixes sgn :: "'a \<Rightarrow> 'a" |
245 |
||
29608 | 246 |
class ord = |
24748 | 247 |
fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
248 |
and less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
23878 | 249 |
begin |
250 |
||
251 |
notation |
|
252 |
less_eq ("op <=") and |
|
253 |
less_eq ("(_/ <= _)" [51, 51] 50) and |
|
254 |
less ("op <") and |
|
255 |
less ("(_/ < _)" [51, 51] 50) |
|
256 |
||
257 |
notation (xsymbols) |
|
258 |
less_eq ("op \<le>") and |
|
259 |
less_eq ("(_/ \<le> _)" [51, 51] 50) |
|
260 |
||
261 |
notation (HTML output) |
|
262 |
less_eq ("op \<le>") and |
|
263 |
less_eq ("(_/ \<le> _)" [51, 51] 50) |
|
264 |
||
25388 | 265 |
abbreviation (input) |
266 |
greater_eq (infix ">=" 50) where |
|
267 |
"x >= y \<equiv> y <= x" |
|
268 |
||
24842 | 269 |
notation (input) |
23878 | 270 |
greater_eq (infix "\<ge>" 50) |
271 |
||
25388 | 272 |
abbreviation (input) |
273 |
greater (infix ">" 50) where |
|
274 |
"x > y \<equiv> y < x" |
|
275 |
||
276 |
end |
|
277 |
||
13456
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
278 |
syntax |
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
279 |
"_index1" :: index ("\<^sub>1") |
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
280 |
translations |
14690 | 281 |
(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
|
282 |
|
11750 | 283 |
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
|
284 |
let |
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
285 |
fun tr' c = (c, fn show_sorts => fn T => fn ts => |
29968 | 286 |
if (not o null) ts orelse T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match |
20713
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents:
20698
diff
changeset
|
287 |
else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); |
22993 | 288 |
in map tr' [@{const_syntax HOL.one}, @{const_syntax HOL.zero}] end; |
11750 | 289 |
*} -- {* show types that are presumably too general *} |
290 |
||
291 |
||
20944 | 292 |
subsection {* Fundamental rules *} |
293 |
||
20973 | 294 |
subsubsection {* Equality *} |
20944 | 295 |
|
18457 | 296 |
lemma sym: "s = t ==> t = s" |
297 |
by (erule subst) (rule refl) |
|
15411 | 298 |
|
18457 | 299 |
lemma ssubst: "t = s ==> P s ==> P t" |
300 |
by (drule sym) (erule subst) |
|
15411 | 301 |
|
302 |
lemma trans: "[| r=s; s=t |] ==> r=t" |
|
18457 | 303 |
by (erule subst) |
15411 | 304 |
|
20944 | 305 |
lemma meta_eq_to_obj_eq: |
306 |
assumes meq: "A == B" |
|
307 |
shows "A = B" |
|
308 |
by (unfold meq) (rule refl) |
|
15411 | 309 |
|
21502 | 310 |
text {* Useful with @{text erule} for proving equalities from known equalities. *} |
20944 | 311 |
(* a = b |
15411 | 312 |
| | |
313 |
c = d *) |
|
314 |
lemma box_equals: "[| a=b; a=c; b=d |] ==> c=d" |
|
315 |
apply (rule trans) |
|
316 |
apply (rule trans) |
|
317 |
apply (rule sym) |
|
318 |
apply assumption+ |
|
319 |
done |
|
320 |
||
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
321 |
text {* For calculational reasoning: *} |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
322 |
|
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
323 |
lemma forw_subst: "a = b ==> P b ==> P a" |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
324 |
by (rule ssubst) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
325 |
|
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
326 |
lemma back_subst: "P a ==> a = b ==> P b" |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
327 |
by (rule subst) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
328 |
|
15411 | 329 |
|
20944 | 330 |
subsubsection {*Congruence rules for application*} |
15411 | 331 |
|
332 |
(*similar to AP_THM in Gordon's HOL*) |
|
333 |
lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)" |
|
334 |
apply (erule subst) |
|
335 |
apply (rule refl) |
|
336 |
done |
|
337 |
||
338 |
(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*) |
|
339 |
lemma arg_cong: "x=y ==> f(x)=f(y)" |
|
340 |
apply (erule subst) |
|
341 |
apply (rule refl) |
|
342 |
done |
|
343 |
||
15655 | 344 |
lemma arg_cong2: "\<lbrakk> a = b; c = d \<rbrakk> \<Longrightarrow> f a c = f b d" |
345 |
apply (erule ssubst)+ |
|
346 |
apply (rule refl) |
|
347 |
done |
|
348 |
||
15411 | 349 |
lemma cong: "[| f = g; (x::'a) = y |] ==> f(x) = g(y)" |
350 |
apply (erule subst)+ |
|
351 |
apply (rule refl) |
|
352 |
done |
|
353 |
||
354 |
||
20944 | 355 |
subsubsection {*Equality of booleans -- iff*} |
15411 | 356 |
|
21504 | 357 |
lemma iffI: assumes "P ==> Q" and "Q ==> P" shows "P=Q" |
358 |
by (iprover intro: iff [THEN mp, THEN mp] impI assms) |
|
15411 | 359 |
|
360 |
lemma iffD2: "[| P=Q; Q |] ==> P" |
|
18457 | 361 |
by (erule ssubst) |
15411 | 362 |
|
363 |
lemma rev_iffD2: "[| Q; P=Q |] ==> P" |
|
18457 | 364 |
by (erule iffD2) |
15411 | 365 |
|
21504 | 366 |
lemma iffD1: "Q = P \<Longrightarrow> Q \<Longrightarrow> P" |
367 |
by (drule sym) (rule iffD2) |
|
368 |
||
369 |
lemma rev_iffD1: "Q \<Longrightarrow> Q = P \<Longrightarrow> P" |
|
370 |
by (drule sym) (rule rev_iffD2) |
|
15411 | 371 |
|
372 |
lemma iffE: |
|
373 |
assumes major: "P=Q" |
|
21504 | 374 |
and minor: "[| P --> Q; Q --> P |] ==> R" |
18457 | 375 |
shows R |
376 |
by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1]) |
|
15411 | 377 |
|
378 |
||
20944 | 379 |
subsubsection {*True*} |
15411 | 380 |
|
381 |
lemma TrueI: "True" |
|
21504 | 382 |
unfolding True_def by (rule refl) |
15411 | 383 |
|
21504 | 384 |
lemma eqTrueI: "P ==> P = True" |
18457 | 385 |
by (iprover intro: iffI TrueI) |
15411 | 386 |
|
21504 | 387 |
lemma eqTrueE: "P = True ==> P" |
388 |
by (erule iffD2) (rule TrueI) |
|
15411 | 389 |
|
390 |
||
20944 | 391 |
subsubsection {*Universal quantifier*} |
15411 | 392 |
|
21504 | 393 |
lemma allI: assumes "!!x::'a. P(x)" shows "ALL x. P(x)" |
394 |
unfolding All_def by (iprover intro: ext eqTrueI assms) |
|
15411 | 395 |
|
396 |
lemma spec: "ALL x::'a. P(x) ==> P(x)" |
|
397 |
apply (unfold All_def) |
|
398 |
apply (rule eqTrueE) |
|
399 |
apply (erule fun_cong) |
|
400 |
done |
|
401 |
||
402 |
lemma allE: |
|
403 |
assumes major: "ALL x. P(x)" |
|
21504 | 404 |
and minor: "P(x) ==> R" |
405 |
shows R |
|
406 |
by (iprover intro: minor major [THEN spec]) |
|
15411 | 407 |
|
408 |
lemma all_dupE: |
|
409 |
assumes major: "ALL x. P(x)" |
|
21504 | 410 |
and minor: "[| P(x); ALL x. P(x) |] ==> R" |
411 |
shows R |
|
412 |
by (iprover intro: minor major major [THEN spec]) |
|
15411 | 413 |
|
414 |
||
21504 | 415 |
subsubsection {* False *} |
416 |
||
417 |
text {* |
|
418 |
Depends upon @{text spec}; it is impossible to do propositional |
|
419 |
logic before quantifiers! |
|
420 |
*} |
|
15411 | 421 |
|
422 |
lemma FalseE: "False ==> P" |
|
21504 | 423 |
apply (unfold False_def) |
424 |
apply (erule spec) |
|
425 |
done |
|
15411 | 426 |
|
21504 | 427 |
lemma False_neq_True: "False = True ==> P" |
428 |
by (erule eqTrueE [THEN FalseE]) |
|
15411 | 429 |
|
430 |
||
21504 | 431 |
subsubsection {* Negation *} |
15411 | 432 |
|
433 |
lemma notI: |
|
21504 | 434 |
assumes "P ==> False" |
15411 | 435 |
shows "~P" |
21504 | 436 |
apply (unfold not_def) |
437 |
apply (iprover intro: impI assms) |
|
438 |
done |
|
15411 | 439 |
|
440 |
lemma False_not_True: "False ~= True" |
|
21504 | 441 |
apply (rule notI) |
442 |
apply (erule False_neq_True) |
|
443 |
done |
|
15411 | 444 |
|
445 |
lemma True_not_False: "True ~= False" |
|
21504 | 446 |
apply (rule notI) |
447 |
apply (drule sym) |
|
448 |
apply (erule False_neq_True) |
|
449 |
done |
|
15411 | 450 |
|
451 |
lemma notE: "[| ~P; P |] ==> R" |
|
21504 | 452 |
apply (unfold not_def) |
453 |
apply (erule mp [THEN FalseE]) |
|
454 |
apply assumption |
|
455 |
done |
|
15411 | 456 |
|
21504 | 457 |
lemma notI2: "(P \<Longrightarrow> \<not> Pa) \<Longrightarrow> (P \<Longrightarrow> Pa) \<Longrightarrow> \<not> P" |
458 |
by (erule notE [THEN notI]) (erule meta_mp) |
|
15411 | 459 |
|
460 |
||
20944 | 461 |
subsubsection {*Implication*} |
15411 | 462 |
|
463 |
lemma impE: |
|
464 |
assumes "P-->Q" "P" "Q ==> R" |
|
465 |
shows "R" |
|
23553 | 466 |
by (iprover intro: assms mp) |
15411 | 467 |
|
468 |
(* Reduces Q to P-->Q, allowing substitution in P. *) |
|
469 |
lemma rev_mp: "[| P; P --> Q |] ==> Q" |
|
17589 | 470 |
by (iprover intro: mp) |
15411 | 471 |
|
472 |
lemma contrapos_nn: |
|
473 |
assumes major: "~Q" |
|
474 |
and minor: "P==>Q" |
|
475 |
shows "~P" |
|
17589 | 476 |
by (iprover intro: notI minor major [THEN notE]) |
15411 | 477 |
|
478 |
(*not used at all, but we already have the other 3 combinations *) |
|
479 |
lemma contrapos_pn: |
|
480 |
assumes major: "Q" |
|
481 |
and minor: "P ==> ~Q" |
|
482 |
shows "~P" |
|
17589 | 483 |
by (iprover intro: notI minor major notE) |
15411 | 484 |
|
485 |
lemma not_sym: "t ~= s ==> s ~= t" |
|
21250 | 486 |
by (erule contrapos_nn) (erule sym) |
487 |
||
488 |
lemma eq_neq_eq_imp_neq: "[| x = a ; a ~= b; b = y |] ==> x ~= y" |
|
489 |
by (erule subst, erule ssubst, assumption) |
|
15411 | 490 |
|
491 |
(*still used in HOLCF*) |
|
492 |
lemma rev_contrapos: |
|
493 |
assumes pq: "P ==> Q" |
|
494 |
and nq: "~Q" |
|
495 |
shows "~P" |
|
496 |
apply (rule nq [THEN contrapos_nn]) |
|
497 |
apply (erule pq) |
|
498 |
done |
|
499 |
||
20944 | 500 |
subsubsection {*Existential quantifier*} |
15411 | 501 |
|
502 |
lemma exI: "P x ==> EX x::'a. P x" |
|
503 |
apply (unfold Ex_def) |
|
17589 | 504 |
apply (iprover intro: allI allE impI mp) |
15411 | 505 |
done |
506 |
||
507 |
lemma exE: |
|
508 |
assumes major: "EX x::'a. P(x)" |
|
509 |
and minor: "!!x. P(x) ==> Q" |
|
510 |
shows "Q" |
|
511 |
apply (rule major [unfolded Ex_def, THEN spec, THEN mp]) |
|
17589 | 512 |
apply (iprover intro: impI [THEN allI] minor) |
15411 | 513 |
done |
514 |
||
515 |
||
20944 | 516 |
subsubsection {*Conjunction*} |
15411 | 517 |
|
518 |
lemma conjI: "[| P; Q |] ==> P&Q" |
|
519 |
apply (unfold and_def) |
|
17589 | 520 |
apply (iprover intro: impI [THEN allI] mp) |
15411 | 521 |
done |
522 |
||
523 |
lemma conjunct1: "[| P & Q |] ==> P" |
|
524 |
apply (unfold and_def) |
|
17589 | 525 |
apply (iprover intro: impI dest: spec mp) |
15411 | 526 |
done |
527 |
||
528 |
lemma conjunct2: "[| P & Q |] ==> Q" |
|
529 |
apply (unfold and_def) |
|
17589 | 530 |
apply (iprover intro: impI dest: spec mp) |
15411 | 531 |
done |
532 |
||
533 |
lemma conjE: |
|
534 |
assumes major: "P&Q" |
|
535 |
and minor: "[| P; Q |] ==> R" |
|
536 |
shows "R" |
|
537 |
apply (rule minor) |
|
538 |
apply (rule major [THEN conjunct1]) |
|
539 |
apply (rule major [THEN conjunct2]) |
|
540 |
done |
|
541 |
||
542 |
lemma context_conjI: |
|
23553 | 543 |
assumes "P" "P ==> Q" shows "P & Q" |
544 |
by (iprover intro: conjI assms) |
|
15411 | 545 |
|
546 |
||
20944 | 547 |
subsubsection {*Disjunction*} |
15411 | 548 |
|
549 |
lemma disjI1: "P ==> P|Q" |
|
550 |
apply (unfold or_def) |
|
17589 | 551 |
apply (iprover intro: allI impI mp) |
15411 | 552 |
done |
553 |
||
554 |
lemma disjI2: "Q ==> P|Q" |
|
555 |
apply (unfold or_def) |
|
17589 | 556 |
apply (iprover intro: allI impI mp) |
15411 | 557 |
done |
558 |
||
559 |
lemma disjE: |
|
560 |
assumes major: "P|Q" |
|
561 |
and minorP: "P ==> R" |
|
562 |
and minorQ: "Q ==> R" |
|
563 |
shows "R" |
|
17589 | 564 |
by (iprover intro: minorP minorQ impI |
15411 | 565 |
major [unfolded or_def, THEN spec, THEN mp, THEN mp]) |
566 |
||
567 |
||
20944 | 568 |
subsubsection {*Classical logic*} |
15411 | 569 |
|
570 |
lemma classical: |
|
571 |
assumes prem: "~P ==> P" |
|
572 |
shows "P" |
|
573 |
apply (rule True_or_False [THEN disjE, THEN eqTrueE]) |
|
574 |
apply assumption |
|
575 |
apply (rule notI [THEN prem, THEN eqTrueI]) |
|
576 |
apply (erule subst) |
|
577 |
apply assumption |
|
578 |
done |
|
579 |
||
580 |
lemmas ccontr = FalseE [THEN classical, standard] |
|
581 |
||
582 |
(*notE with premises exchanged; it discharges ~R so that it can be used to |
|
583 |
make elimination rules*) |
|
584 |
lemma rev_notE: |
|
585 |
assumes premp: "P" |
|
586 |
and premnot: "~R ==> ~P" |
|
587 |
shows "R" |
|
588 |
apply (rule ccontr) |
|
589 |
apply (erule notE [OF premnot premp]) |
|
590 |
done |
|
591 |
||
592 |
(*Double negation law*) |
|
593 |
lemma notnotD: "~~P ==> P" |
|
594 |
apply (rule classical) |
|
595 |
apply (erule notE) |
|
596 |
apply assumption |
|
597 |
done |
|
598 |
||
599 |
lemma contrapos_pp: |
|
600 |
assumes p1: "Q" |
|
601 |
and p2: "~P ==> ~Q" |
|
602 |
shows "P" |
|
17589 | 603 |
by (iprover intro: classical p1 p2 notE) |
15411 | 604 |
|
605 |
||
20944 | 606 |
subsubsection {*Unique existence*} |
15411 | 607 |
|
608 |
lemma ex1I: |
|
23553 | 609 |
assumes "P a" "!!x. P(x) ==> x=a" |
15411 | 610 |
shows "EX! x. P(x)" |
23553 | 611 |
by (unfold Ex1_def, iprover intro: assms exI conjI allI impI) |
15411 | 612 |
|
613 |
text{*Sometimes easier to use: the premises have no shared variables. Safe!*} |
|
614 |
lemma ex_ex1I: |
|
615 |
assumes ex_prem: "EX x. P(x)" |
|
616 |
and eq: "!!x y. [| P(x); P(y) |] ==> x=y" |
|
617 |
shows "EX! x. P(x)" |
|
17589 | 618 |
by (iprover intro: ex_prem [THEN exE] ex1I eq) |
15411 | 619 |
|
620 |
lemma ex1E: |
|
621 |
assumes major: "EX! x. P(x)" |
|
622 |
and minor: "!!x. [| P(x); ALL y. P(y) --> y=x |] ==> R" |
|
623 |
shows "R" |
|
624 |
apply (rule major [unfolded Ex1_def, THEN exE]) |
|
625 |
apply (erule conjE) |
|
17589 | 626 |
apply (iprover intro: minor) |
15411 | 627 |
done |
628 |
||
629 |
lemma ex1_implies_ex: "EX! x. P x ==> EX x. P x" |
|
630 |
apply (erule ex1E) |
|
631 |
apply (rule exI) |
|
632 |
apply assumption |
|
633 |
done |
|
634 |
||
635 |
||
20944 | 636 |
subsubsection {*THE: definite description operator*} |
15411 | 637 |
|
638 |
lemma the_equality: |
|
639 |
assumes prema: "P a" |
|
640 |
and premx: "!!x. P x ==> x=a" |
|
641 |
shows "(THE x. P x) = a" |
|
642 |
apply (rule trans [OF _ the_eq_trivial]) |
|
643 |
apply (rule_tac f = "The" in arg_cong) |
|
644 |
apply (rule ext) |
|
645 |
apply (rule iffI) |
|
646 |
apply (erule premx) |
|
647 |
apply (erule ssubst, rule prema) |
|
648 |
done |
|
649 |
||
650 |
lemma theI: |
|
651 |
assumes "P a" and "!!x. P x ==> x=a" |
|
652 |
shows "P (THE x. P x)" |
|
23553 | 653 |
by (iprover intro: assms the_equality [THEN ssubst]) |
15411 | 654 |
|
655 |
lemma theI': "EX! x. P x ==> P (THE x. P x)" |
|
656 |
apply (erule ex1E) |
|
657 |
apply (erule theI) |
|
658 |
apply (erule allE) |
|
659 |
apply (erule mp) |
|
660 |
apply assumption |
|
661 |
done |
|
662 |
||
663 |
(*Easier to apply than theI: only one occurrence of P*) |
|
664 |
lemma theI2: |
|
665 |
assumes "P a" "!!x. P x ==> x=a" "!!x. P x ==> Q x" |
|
666 |
shows "Q (THE x. P x)" |
|
23553 | 667 |
by (iprover intro: assms theI) |
15411 | 668 |
|
24553 | 669 |
lemma the1I2: assumes "EX! x. P x" "\<And>x. P x \<Longrightarrow> Q x" shows "Q (THE x. P x)" |
670 |
by(iprover intro:assms(2) theI2[where P=P and Q=Q] ex1E[OF assms(1)] |
|
671 |
elim:allE impE) |
|
672 |
||
18697 | 673 |
lemma the1_equality [elim?]: "[| EX!x. P x; P a |] ==> (THE x. P x) = a" |
15411 | 674 |
apply (rule the_equality) |
675 |
apply assumption |
|
676 |
apply (erule ex1E) |
|
677 |
apply (erule all_dupE) |
|
678 |
apply (drule mp) |
|
679 |
apply assumption |
|
680 |
apply (erule ssubst) |
|
681 |
apply (erule allE) |
|
682 |
apply (erule mp) |
|
683 |
apply assumption |
|
684 |
done |
|
685 |
||
686 |
lemma the_sym_eq_trivial: "(THE y. x=y) = x" |
|
687 |
apply (rule the_equality) |
|
688 |
apply (rule refl) |
|
689 |
apply (erule sym) |
|
690 |
done |
|
691 |
||
692 |
||
20944 | 693 |
subsubsection {*Classical intro rules for disjunction and existential quantifiers*} |
15411 | 694 |
|
695 |
lemma disjCI: |
|
696 |
assumes "~Q ==> P" shows "P|Q" |
|
697 |
apply (rule classical) |
|
23553 | 698 |
apply (iprover intro: assms disjI1 disjI2 notI elim: notE) |
15411 | 699 |
done |
700 |
||
701 |
lemma excluded_middle: "~P | P" |
|
17589 | 702 |
by (iprover intro: disjCI) |
15411 | 703 |
|
20944 | 704 |
text {* |
705 |
case distinction as a natural deduction rule. |
|
706 |
Note that @{term "~P"} is the second case, not the first |
|
707 |
*} |
|
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
708 |
lemma case_split [case_names True False]: |
15411 | 709 |
assumes prem1: "P ==> Q" |
710 |
and prem2: "~P ==> Q" |
|
711 |
shows "Q" |
|
712 |
apply (rule excluded_middle [THEN disjE]) |
|
713 |
apply (erule prem2) |
|
714 |
apply (erule prem1) |
|
715 |
done |
|
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
716 |
|
15411 | 717 |
(*Classical implies (-->) elimination. *) |
718 |
lemma impCE: |
|
719 |
assumes major: "P-->Q" |
|
720 |
and minor: "~P ==> R" "Q ==> R" |
|
721 |
shows "R" |
|
722 |
apply (rule excluded_middle [of P, THEN disjE]) |
|
17589 | 723 |
apply (iprover intro: minor major [THEN mp])+ |
15411 | 724 |
done |
725 |
||
726 |
(*This version of --> elimination works on Q before P. It works best for |
|
727 |
those cases in which P holds "almost everywhere". Can't install as |
|
728 |
default: would break old proofs.*) |
|
729 |
lemma impCE': |
|
730 |
assumes major: "P-->Q" |
|
731 |
and minor: "Q ==> R" "~P ==> R" |
|
732 |
shows "R" |
|
733 |
apply (rule excluded_middle [of P, THEN disjE]) |
|
17589 | 734 |
apply (iprover intro: minor major [THEN mp])+ |
15411 | 735 |
done |
736 |
||
737 |
(*Classical <-> elimination. *) |
|
738 |
lemma iffCE: |
|
739 |
assumes major: "P=Q" |
|
740 |
and minor: "[| P; Q |] ==> R" "[| ~P; ~Q |] ==> R" |
|
741 |
shows "R" |
|
742 |
apply (rule major [THEN iffE]) |
|
17589 | 743 |
apply (iprover intro: minor elim: impCE notE) |
15411 | 744 |
done |
745 |
||
746 |
lemma exCI: |
|
747 |
assumes "ALL x. ~P(x) ==> P(a)" |
|
748 |
shows "EX x. P(x)" |
|
749 |
apply (rule ccontr) |
|
23553 | 750 |
apply (iprover intro: assms exI allI notI notE [of "\<exists>x. P x"]) |
15411 | 751 |
done |
752 |
||
753 |
||
12386 | 754 |
subsubsection {* Intuitionistic Reasoning *} |
755 |
||
756 |
lemma impE': |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
757 |
assumes 1: "P --> Q" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
758 |
and 2: "Q ==> R" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
759 |
and 3: "P --> Q ==> P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
760 |
shows R |
12386 | 761 |
proof - |
762 |
from 3 and 1 have P . |
|
763 |
with 1 have Q by (rule impE) |
|
764 |
with 2 show R . |
|
765 |
qed |
|
766 |
||
767 |
lemma allE': |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
768 |
assumes 1: "ALL x. P x" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
769 |
and 2: "P x ==> ALL x. P x ==> Q" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
770 |
shows Q |
12386 | 771 |
proof - |
772 |
from 1 have "P x" by (rule spec) |
|
773 |
from this and 1 show Q by (rule 2) |
|
774 |
qed |
|
775 |
||
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
776 |
lemma notE': |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
777 |
assumes 1: "~ P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
778 |
and 2: "~ P ==> P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
779 |
shows R |
12386 | 780 |
proof - |
781 |
from 2 and 1 have P . |
|
782 |
with 1 show R by (rule notE) |
|
783 |
qed |
|
784 |
||
22444
fb80fedd192d
added safe intro rules for removing "True" subgoals as well as "~ False" ones.
dixon
parents:
22377
diff
changeset
|
785 |
lemma TrueE: "True ==> P ==> P" . |
fb80fedd192d
added safe intro rules for removing "True" subgoals as well as "~ False" ones.
dixon
parents:
22377
diff
changeset
|
786 |
lemma notFalseE: "~ False ==> P ==> P" . |
fb80fedd192d
added safe intro rules for removing "True" subgoals as well as "~ False" ones.
dixon
parents:
22377
diff
changeset
|
787 |
|
22467
c9357ef01168
TrueElim and notTrueElim tested and added as safe elim rules.
dixon
parents:
22445
diff
changeset
|
788 |
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE TrueE notFalseE |
15801 | 789 |
and [Pure.intro!] = iffI conjI impI TrueI notI allI refl |
790 |
and [Pure.elim 2] = allE notE' impE' |
|
791 |
and [Pure.intro] = exI disjI2 disjI1 |
|
12386 | 792 |
|
793 |
lemmas [trans] = trans |
|
794 |
and [sym] = sym not_sym |
|
15801 | 795 |
and [Pure.elim?] = iffD1 iffD2 impE |
11750 | 796 |
|
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28856
diff
changeset
|
797 |
use "Tools/hologic.ML" |
23553 | 798 |
|
11438
3d9222b80989
declare trans [trans] (*overridden in theory Calculation*);
wenzelm
parents:
11432
diff
changeset
|
799 |
|
11750 | 800 |
subsubsection {* Atomizing meta-level connectives *} |
801 |
||
28513 | 802 |
axiomatization where |
803 |
eq_reflection: "x = y \<Longrightarrow> x \<equiv> y" (*admissible axiom*) |
|
804 |
||
11750 | 805 |
lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" |
12003 | 806 |
proof |
9488 | 807 |
assume "!!x. P x" |
23389 | 808 |
then show "ALL x. P x" .. |
9488 | 809 |
next |
810 |
assume "ALL x. P x" |
|
23553 | 811 |
then show "!!x. P x" by (rule allE) |
9488 | 812 |
qed |
813 |
||
11750 | 814 |
lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" |
12003 | 815 |
proof |
9488 | 816 |
assume r: "A ==> B" |
10383 | 817 |
show "A --> B" by (rule impI) (rule r) |
9488 | 818 |
next |
819 |
assume "A --> B" and A |
|
23553 | 820 |
then show B by (rule mp) |
9488 | 821 |
qed |
822 |
||
14749 | 823 |
lemma atomize_not: "(A ==> False) == Trueprop (~A)" |
824 |
proof |
|
825 |
assume r: "A ==> False" |
|
826 |
show "~A" by (rule notI) (rule r) |
|
827 |
next |
|
828 |
assume "~A" and A |
|
23553 | 829 |
then show False by (rule notE) |
14749 | 830 |
qed |
831 |
||
11750 | 832 |
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" |
12003 | 833 |
proof |
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
834 |
assume "x == y" |
23553 | 835 |
show "x = y" by (unfold `x == y`) (rule refl) |
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
836 |
next |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
837 |
assume "x = y" |
23553 | 838 |
then show "x == y" by (rule eq_reflection) |
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
839 |
qed |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
840 |
|
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28741
diff
changeset
|
841 |
lemma atomize_conj [atomize]: "(A &&& B) == Trueprop (A & B)" |
12003 | 842 |
proof |
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28741
diff
changeset
|
843 |
assume conj: "A &&& B" |
19121 | 844 |
show "A & B" |
845 |
proof (rule conjI) |
|
846 |
from conj show A by (rule conjunctionD1) |
|
847 |
from conj show B by (rule conjunctionD2) |
|
848 |
qed |
|
11953 | 849 |
next |
19121 | 850 |
assume conj: "A & B" |
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28741
diff
changeset
|
851 |
show "A &&& B" |
19121 | 852 |
proof - |
853 |
from conj show A .. |
|
854 |
from conj show B .. |
|
11953 | 855 |
qed |
856 |
qed |
|
857 |
||
12386 | 858 |
lemmas [symmetric, rulify] = atomize_all atomize_imp |
18832 | 859 |
and [symmetric, defn] = atomize_all atomize_imp atomize_eq |
12386 | 860 |
|
11750 | 861 |
|
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
862 |
subsubsection {* Atomizing elimination rules *} |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
863 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
864 |
setup AtomizeElim.setup |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
865 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
866 |
lemma atomize_exL[atomize_elim]: "(!!x. P x ==> Q) == ((EX x. P x) ==> Q)" |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
867 |
by rule iprover+ |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
868 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
869 |
lemma atomize_conjL[atomize_elim]: "(A ==> B ==> C) == (A & B ==> C)" |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
870 |
by rule iprover+ |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
871 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
872 |
lemma atomize_disjL[atomize_elim]: "((A ==> C) ==> (B ==> C) ==> C) == ((A | B ==> C) ==> C)" |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
873 |
by rule iprover+ |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
874 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
875 |
lemma atomize_elimL[atomize_elim]: "(!!B. (A ==> B) ==> B) == Trueprop A" .. |
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
876 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
877 |
|
20944 | 878 |
subsection {* Package setup *} |
879 |
||
11750 | 880 |
subsubsection {* Classical Reasoner setup *} |
9529 | 881 |
|
26411 | 882 |
lemma imp_elim: "P --> Q ==> (~ R ==> P) ==> (Q ==> R) ==> R" |
883 |
by (rule classical) iprover |
|
884 |
||
885 |
lemma swap: "~ P ==> (~ R ==> P) ==> R" |
|
886 |
by (rule classical) iprover |
|
887 |
||
20944 | 888 |
lemma thin_refl: |
889 |
"\<And>X. \<lbrakk> x=x; PROP W \<rbrakk> \<Longrightarrow> PROP W" . |
|
890 |
||
21151 | 891 |
ML {* |
892 |
structure Hypsubst = HypsubstFun( |
|
893 |
struct |
|
894 |
structure Simplifier = Simplifier |
|
21218 | 895 |
val dest_eq = HOLogic.dest_eq |
21151 | 896 |
val dest_Trueprop = HOLogic.dest_Trueprop |
897 |
val dest_imp = HOLogic.dest_imp |
|
26411 | 898 |
val eq_reflection = @{thm eq_reflection} |
899 |
val rev_eq_reflection = @{thm meta_eq_to_obj_eq} |
|
900 |
val imp_intr = @{thm impI} |
|
901 |
val rev_mp = @{thm rev_mp} |
|
902 |
val subst = @{thm subst} |
|
903 |
val sym = @{thm sym} |
|
22129 | 904 |
val thin_refl = @{thm thin_refl}; |
27572
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
krauss
parents:
27338
diff
changeset
|
905 |
val prop_subst = @{lemma "PROP P t ==> PROP prop (x = t ==> PROP P x)" |
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
krauss
parents:
27338
diff
changeset
|
906 |
by (unfold prop_def) (drule eq_reflection, unfold)} |
21151 | 907 |
end); |
21671 | 908 |
open Hypsubst; |
21151 | 909 |
|
910 |
structure Classical = ClassicalFun( |
|
911 |
struct |
|
26411 | 912 |
val imp_elim = @{thm imp_elim} |
913 |
val not_elim = @{thm notE} |
|
914 |
val swap = @{thm swap} |
|
915 |
val classical = @{thm classical} |
|
21151 | 916 |
val sizef = Drule.size_of_thm |
917 |
val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] |
|
918 |
end); |
|
919 |
||
920 |
structure BasicClassical: BASIC_CLASSICAL = Classical; |
|
21671 | 921 |
open BasicClassical; |
22129 | 922 |
|
27338 | 923 |
ML_Antiquote.value "claset" |
924 |
(Scan.succeed "Classical.local_claset_of (ML_Context.the_local_context ())"); |
|
24035 | 925 |
|
31902 | 926 |
structure ResAtpset = Named_Thms |
927 |
(val name = "atp" val description = "ATP rules"); |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
928 |
|
31902 | 929 |
structure ResBlacklist = Named_Thms |
930 |
(val name = "noatp" val description = "theorems blacklisted for ATP"); |
|
21151 | 931 |
*} |
932 |
||
25388 | 933 |
text {*ResBlacklist holds theorems blacklisted to sledgehammer. |
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
934 |
These theorems typically produce clauses that are prolific (match too many equality or |
25388 | 935 |
membership literals) and relate to seldom-used facts. Some duplicate other rules.*} |
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
936 |
|
21009 | 937 |
setup {* |
938 |
let |
|
939 |
(*prevent substitution on bool*) |
|
940 |
fun hyp_subst_tac' i thm = if i <= Thm.nprems_of thm andalso |
|
941 |
Term.exists_Const (fn ("op =", Type (_, [T, _])) => T <> Type ("bool", []) | _ => false) |
|
942 |
(nth (Thm.prems_of thm) (i - 1)) then Hypsubst.hyp_subst_tac i thm else no_tac thm; |
|
943 |
in |
|
21151 | 944 |
Hypsubst.hypsubst_setup |
945 |
#> ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac) |
|
946 |
#> Classical.setup |
|
947 |
#> ResAtpset.setup |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
948 |
#> ResBlacklist.setup |
21009 | 949 |
end |
950 |
*} |
|
951 |
||
952 |
declare iffI [intro!] |
|
953 |
and notI [intro!] |
|
954 |
and impI [intro!] |
|
955 |
and disjCI [intro!] |
|
956 |
and conjI [intro!] |
|
957 |
and TrueI [intro!] |
|
958 |
and refl [intro!] |
|
959 |
||
960 |
declare iffCE [elim!] |
|
961 |
and FalseE [elim!] |
|
962 |
and impCE [elim!] |
|
963 |
and disjE [elim!] |
|
964 |
and conjE [elim!] |
|
965 |
and conjE [elim!] |
|
966 |
||
967 |
declare ex_ex1I [intro!] |
|
968 |
and allI [intro!] |
|
969 |
and the_equality [intro] |
|
970 |
and exI [intro] |
|
971 |
||
972 |
declare exE [elim!] |
|
973 |
allE [elim] |
|
974 |
||
22377 | 975 |
ML {* val HOL_cs = @{claset} *} |
19162 | 976 |
|
20223 | 977 |
lemma contrapos_np: "~ Q ==> (~ P ==> Q) ==> P" |
978 |
apply (erule swap) |
|
979 |
apply (erule (1) meta_mp) |
|
980 |
done |
|
10383 | 981 |
|
18689
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
982 |
declare ex_ex1I [rule del, intro! 2] |
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
983 |
and ex1I [intro] |
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
984 |
|
12386 | 985 |
lemmas [intro?] = ext |
986 |
and [elim?] = ex1_implies_ex |
|
11977 | 987 |
|
20944 | 988 |
(*Better then ex1E for classical reasoner: needs no quantifier duplication!*) |
20973 | 989 |
lemma alt_ex1E [elim!]: |
20944 | 990 |
assumes major: "\<exists>!x. P x" |
991 |
and prem: "\<And>x. \<lbrakk> P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y' \<rbrakk> \<Longrightarrow> R" |
|
992 |
shows R |
|
993 |
apply (rule ex1E [OF major]) |
|
994 |
apply (rule prem) |
|
22129 | 995 |
apply (tactic {* ares_tac @{thms allI} 1 *})+ |
996 |
apply (tactic {* etac (Classical.dup_elim @{thm allE}) 1 *}) |
|
997 |
apply iprover |
|
998 |
done |
|
20944 | 999 |
|
21151 | 1000 |
ML {* |
25388 | 1001 |
structure Blast = BlastFun |
1002 |
( |
|
21151 | 1003 |
type claset = Classical.claset |
22744
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22481
diff
changeset
|
1004 |
val equality_name = @{const_name "op ="} |
22993 | 1005 |
val not_name = @{const_name Not} |
26411 | 1006 |
val notE = @{thm notE} |
1007 |
val ccontr = @{thm ccontr} |
|
21151 | 1008 |
val contr_tac = Classical.contr_tac |
1009 |
val dup_intr = Classical.dup_intr |
|
1010 |
val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac |
|
1011 |
val rep_cs = Classical.rep_cs |
|
1012 |
val cla_modifiers = Classical.cla_modifiers |
|
1013 |
val cla_meth' = Classical.cla_meth' |
|
25388 | 1014 |
); |
21671 | 1015 |
val blast_tac = Blast.blast_tac; |
20944 | 1016 |
*} |
1017 |
||
21151 | 1018 |
setup Blast.setup |
1019 |
||
20944 | 1020 |
|
1021 |
subsubsection {* Simplifier *} |
|
12281 | 1022 |
|
1023 |
lemma eta_contract_eq: "(%s. f s) = f" .. |
|
1024 |
||
1025 |
lemma simp_thms: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1026 |
shows not_not: "(~ ~ P) = P" |
15354 | 1027 |
and Not_eq_iff: "((~P) = (~Q)) = (P = Q)" |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1028 |
and |
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
1029 |
"(P ~= Q) = (P = (~Q))" |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
1030 |
"(P | ~P) = True" "(~P | P) = True" |
12281 | 1031 |
"(x = x) = True" |
20944 | 1032 |
and not_True_eq_False: "(\<not> True) = False" |
1033 |
and not_False_eq_True: "(\<not> False) = True" |
|
1034 |
and |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
1035 |
"(~P) ~= P" "P ~= (~P)" |
20944 | 1036 |
"(True=P) = P" |
1037 |
and eq_True: "(P = True) = P" |
|
1038 |
and "(False=P) = (~P)" |
|
1039 |
and eq_False: "(P = False) = (\<not> P)" |
|
1040 |
and |
|
12281 | 1041 |
"(True --> P) = P" "(False --> P) = True" |
1042 |
"(P --> True) = True" "(P --> P) = True" |
|
1043 |
"(P --> False) = (~P)" "(P --> ~P) = (~P)" |
|
1044 |
"(P & True) = P" "(True & P) = P" |
|
1045 |
"(P & False) = False" "(False & P) = False" |
|
1046 |
"(P & P) = P" "(P & (P & Q)) = (P & Q)" |
|
1047 |
"(P & ~P) = False" "(~P & P) = False" |
|
1048 |
"(P | True) = True" "(True | P) = True" |
|
1049 |
"(P | False) = P" "(False | P) = P" |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
1050 |
"(P | P) = P" "(P | (P | Q)) = (P | Q)" and |
12281 | 1051 |
"(ALL x. P) = P" "(EX x. P) = P" "EX x. x=t" "EX x. t=x" |
31166
a90fe83f58ea
"{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents:
31156
diff
changeset
|
1052 |
and |
12281 | 1053 |
"!!P. (EX x. x=t & P(x)) = P(t)" |
1054 |
"!!P. (EX x. t=x & P(x)) = P(t)" |
|
1055 |
"!!P. (ALL x. x=t --> P(x)) = P(t)" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1056 |
"!!P. (ALL x. t=x --> P(x)) = P(t)" |
17589 | 1057 |
by (blast, blast, blast, blast, blast, iprover+) |
13421 | 1058 |
|
14201 | 1059 |
lemma disj_absorb: "(A | A) = A" |
1060 |
by blast |
|
1061 |
||
1062 |
lemma disj_left_absorb: "(A | (A | B)) = (A | B)" |
|
1063 |
by blast |
|
1064 |
||
1065 |
lemma conj_absorb: "(A & A) = A" |
|
1066 |
by blast |
|
1067 |
||
1068 |
lemma conj_left_absorb: "(A & (A & B)) = (A & B)" |
|
1069 |
by blast |
|
1070 |
||
12281 | 1071 |
lemma eq_ac: |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1072 |
shows eq_commute: "(a=b) = (b=a)" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1073 |
and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))" |
17589 | 1074 |
and eq_assoc: "((P=Q)=R) = (P=(Q=R))" by (iprover, blast+) |
1075 |
lemma neq_commute: "(a~=b) = (b~=a)" by iprover |
|
12281 | 1076 |
|
1077 |
lemma conj_comms: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1078 |
shows conj_commute: "(P&Q) = (Q&P)" |
17589 | 1079 |
and conj_left_commute: "(P&(Q&R)) = (Q&(P&R))" by iprover+ |
1080 |
lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by iprover |
|
12281 | 1081 |
|
19174 | 1082 |
lemmas conj_ac = conj_commute conj_left_commute conj_assoc |
1083 |
||
12281 | 1084 |
lemma disj_comms: |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1085 |
shows disj_commute: "(P|Q) = (Q|P)" |
17589 | 1086 |
and disj_left_commute: "(P|(Q|R)) = (Q|(P|R))" by iprover+ |
1087 |
lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by iprover |
|
12281 | 1088 |
|
19174 | 1089 |
lemmas disj_ac = disj_commute disj_left_commute disj_assoc |
1090 |
||
17589 | 1091 |
lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by iprover |
1092 |
lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by iprover |
|
12281 | 1093 |
|
17589 | 1094 |
lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by iprover |
1095 |
lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by iprover |
|
12281 | 1096 |
|
17589 | 1097 |
lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by iprover |
1098 |
lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by iprover |
|
1099 |
lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by iprover |
|
12281 | 1100 |
|
1101 |
text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *} |
|
1102 |
lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast |
|
1103 |
lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast |
|
1104 |
||
1105 |
lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast |
|
1106 |
lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast |
|
1107 |
||
21151 | 1108 |
lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))" |
1109 |
by iprover |
|
1110 |
||
17589 | 1111 |
lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by iprover |
12281 | 1112 |
lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast |
1113 |
lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast |
|
1114 |
lemma not_iff: "(P~=Q) = (P = (~Q))" by blast |
|
1115 |
lemma disj_not1: "(~P | Q) = (P --> Q)" by blast |
|
1116 |
lemma disj_not2: "(P | ~Q) = (Q --> P)" -- {* changes orientation :-( *} |
|
1117 |
by blast |
|
1118 |
lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast |
|
1119 |
||
17589 | 1120 |
lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by iprover |
12281 | 1121 |
|
1122 |
||
1123 |
lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q" |
|
1124 |
-- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *} |
|
1125 |
-- {* cases boil down to the same thing. *} |
|
1126 |
by blast |
|
1127 |
||
1128 |
lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast |
|
1129 |
lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast |
|
17589 | 1130 |
lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by iprover |
1131 |
lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by iprover |
|
23403 | 1132 |
lemma all_not_ex: "(ALL x. P x) = (~ (EX x. ~ P x ))" by blast |
12281 | 1133 |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
1134 |
declare All_def [noatp] |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
1135 |
|
17589 | 1136 |
lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by iprover |
1137 |
lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by iprover |
|
12281 | 1138 |
|
1139 |
text {* |
|
1140 |
\medskip The @{text "&"} congruence rule: not included by default! |
|
1141 |
May slow rewrite proofs down by as much as 50\% *} |
|
1142 |
||
1143 |
lemma conj_cong: |
|
1144 |
"(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))" |
|
17589 | 1145 |
by iprover |
12281 | 1146 |
|
1147 |
lemma rev_conj_cong: |
|
1148 |
"(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))" |
|
17589 | 1149 |
by iprover |
12281 | 1150 |
|
1151 |
text {* The @{text "|"} congruence rule: not included by default! *} |
|
1152 |
||
1153 |
lemma disj_cong: |
|
1154 |
"(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))" |
|
1155 |
by blast |
|
1156 |
||
1157 |
||
1158 |
text {* \medskip if-then-else rules *} |
|
1159 |
||
1160 |
lemma if_True: "(if True then x else y) = x" |
|
1161 |
by (unfold if_def) blast |
|
1162 |
||
1163 |
lemma if_False: "(if False then x else y) = y" |
|
1164 |
by (unfold if_def) blast |
|
1165 |
||
1166 |
lemma if_P: "P ==> (if P then x else y) = x" |
|
1167 |
by (unfold if_def) blast |
|
1168 |
||
1169 |
lemma if_not_P: "~P ==> (if P then x else y) = y" |
|
1170 |
by (unfold if_def) blast |
|
1171 |
||
1172 |
lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" |
|
1173 |
apply (rule case_split [of Q]) |
|
15481 | 1174 |
apply (simplesubst if_P) |
1175 |
prefer 3 apply (simplesubst if_not_P, blast+) |
|
12281 | 1176 |
done |
1177 |
||
1178 |
lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" |
|
15481 | 1179 |
by (simplesubst split_if, blast) |
12281 | 1180 |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
1181 |
lemmas if_splits [noatp] = split_if split_if_asm |
12281 | 1182 |
|
1183 |
lemma if_cancel: "(if c then x else x) = x" |
|
15481 | 1184 |
by (simplesubst split_if, blast) |
12281 | 1185 |
|
1186 |
lemma if_eq_cancel: "(if x = y then y else x) = x" |
|
15481 | 1187 |
by (simplesubst split_if, blast) |
12281 | 1188 |
|
1189 |
lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))" |
|
19796 | 1190 |
-- {* This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "==>"} symbol. *} |
12281 | 1191 |
by (rule split_if) |
1192 |
||
1193 |
lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" |
|
19796 | 1194 |
-- {* And this form is useful for expanding @{text "if"}s on the LEFT. *} |
15481 | 1195 |
apply (simplesubst split_if, blast) |
12281 | 1196 |
done |
1197 |
||
17589 | 1198 |
lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) iprover |
1199 |
lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) iprover |
|
12281 | 1200 |
|
15423 | 1201 |
text {* \medskip let rules for simproc *} |
1202 |
||
1203 |
lemma Let_folded: "f x \<equiv> g x \<Longrightarrow> Let x f \<equiv> Let x g" |
|
1204 |
by (unfold Let_def) |
|
1205 |
||
1206 |
lemma Let_unfold: "f x \<equiv> g \<Longrightarrow> Let x f \<equiv> g" |
|
1207 |
by (unfold Let_def) |
|
1208 |
||
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1209 |
text {* |
16999 | 1210 |
The following copy of the implication operator is useful for |
1211 |
fine-tuning congruence rules. It instructs the simplifier to simplify |
|
1212 |
its premise. |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1213 |
*} |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1214 |
|
17197 | 1215 |
constdefs |
1216 |
simp_implies :: "[prop, prop] => prop" (infixr "=simp=>" 1) |
|
28562 | 1217 |
[code del]: "simp_implies \<equiv> op ==>" |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1218 |
|
18457 | 1219 |
lemma simp_impliesI: |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1220 |
assumes PQ: "(PROP P \<Longrightarrow> PROP Q)" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1221 |
shows "PROP P =simp=> PROP Q" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1222 |
apply (unfold simp_implies_def) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1223 |
apply (rule PQ) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1224 |
apply assumption |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1225 |
done |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1226 |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1227 |
lemma simp_impliesE: |
25388 | 1228 |
assumes PQ: "PROP P =simp=> PROP Q" |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1229 |
and P: "PROP P" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1230 |
and QR: "PROP Q \<Longrightarrow> PROP R" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1231 |
shows "PROP R" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1232 |
apply (rule QR) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1233 |
apply (rule PQ [unfolded simp_implies_def]) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1234 |
apply (rule P) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1235 |
done |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1236 |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1237 |
lemma simp_implies_cong: |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1238 |
assumes PP' :"PROP P == PROP P'" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1239 |
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
|
1240 |
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
|
1241 |
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
|
1242 |
assume PQ: "PROP P \<Longrightarrow> PROP Q" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1243 |
and P': "PROP P'" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1244 |
from PP' [symmetric] and P' have "PROP P" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1245 |
by (rule equal_elim_rule1) |
23553 | 1246 |
then have "PROP Q" by (rule PQ) |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1247 |
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
|
1248 |
next |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1249 |
assume P'Q': "PROP P' \<Longrightarrow> PROP Q'" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1250 |
and P: "PROP P" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1251 |
from PP' and P have P': "PROP P'" by (rule equal_elim_rule1) |
23553 | 1252 |
then have "PROP Q'" by (rule P'Q') |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1253 |
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
|
1254 |
by (rule equal_elim_rule1) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1255 |
qed |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1256 |
|
20944 | 1257 |
lemma uncurry: |
1258 |
assumes "P \<longrightarrow> Q \<longrightarrow> R" |
|
1259 |
shows "P \<and> Q \<longrightarrow> R" |
|
23553 | 1260 |
using assms by blast |
20944 | 1261 |
|
1262 |
lemma iff_allI: |
|
1263 |
assumes "\<And>x. P x = Q x" |
|
1264 |
shows "(\<forall>x. P x) = (\<forall>x. Q x)" |
|
23553 | 1265 |
using assms by blast |
20944 | 1266 |
|
1267 |
lemma iff_exI: |
|
1268 |
assumes "\<And>x. P x = Q x" |
|
1269 |
shows "(\<exists>x. P x) = (\<exists>x. Q x)" |
|
23553 | 1270 |
using assms by blast |
20944 | 1271 |
|
1272 |
lemma all_comm: |
|
1273 |
"(\<forall>x y. P x y) = (\<forall>y x. P x y)" |
|
1274 |
by blast |
|
1275 |
||
1276 |
lemma ex_comm: |
|
1277 |
"(\<exists>x y. P x y) = (\<exists>y x. P x y)" |
|
1278 |
by blast |
|
1279 |
||
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28856
diff
changeset
|
1280 |
use "Tools/simpdata.ML" |
21671 | 1281 |
ML {* open Simpdata *} |
1282 |
||
21151 | 1283 |
setup {* |
1284 |
Simplifier.method_setup Splitter.split_modifiers |
|
26496
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
wenzelm
parents:
26411
diff
changeset
|
1285 |
#> Simplifier.map_simpset (K Simpdata.simpset_simprocs) |
21151 | 1286 |
#> Splitter.setup |
26496
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
wenzelm
parents:
26411
diff
changeset
|
1287 |
#> clasimp_setup |
21151 | 1288 |
#> EqSubst.setup |
1289 |
*} |
|
1290 |
||
24035 | 1291 |
text {* Simproc for proving @{text "(y = x) == False"} from premise @{text "~(x = y)"}: *} |
1292 |
||
1293 |
simproc_setup neq ("x = y") = {* fn _ => |
|
1294 |
let |
|
1295 |
val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI}; |
|
1296 |
fun is_neq eq lhs rhs thm = |
|
1297 |
(case Thm.prop_of thm of |
|
1298 |
_ $ (Not $ (eq' $ l' $ r')) => |
|
1299 |
Not = HOLogic.Not andalso eq' = eq andalso |
|
1300 |
r' aconv lhs andalso l' aconv rhs |
|
1301 |
| _ => false); |
|
1302 |
fun proc ss ct = |
|
1303 |
(case Thm.term_of ct of |
|
1304 |
eq $ lhs $ rhs => |
|
1305 |
(case find_first (is_neq eq lhs rhs) (Simplifier.prems_of_ss ss) of |
|
1306 |
SOME thm => SOME (thm RS neq_to_EQ_False) |
|
1307 |
| NONE => NONE) |
|
1308 |
| _ => NONE); |
|
1309 |
in proc end; |
|
1310 |
*} |
|
1311 |
||
1312 |
simproc_setup let_simp ("Let x f") = {* |
|
1313 |
let |
|
1314 |
val (f_Let_unfold, x_Let_unfold) = |
|
28741 | 1315 |
let val [(_ $ (f $ x) $ _)] = prems_of @{thm Let_unfold} |
24035 | 1316 |
in (cterm_of @{theory} f, cterm_of @{theory} x) end |
1317 |
val (f_Let_folded, x_Let_folded) = |
|
28741 | 1318 |
let val [(_ $ (f $ x) $ _)] = prems_of @{thm Let_folded} |
24035 | 1319 |
in (cterm_of @{theory} f, cterm_of @{theory} x) end; |
1320 |
val g_Let_folded = |
|
28741 | 1321 |
let val [(_ $ _ $ (g $ _))] = prems_of @{thm Let_folded} |
1322 |
in cterm_of @{theory} g end; |
|
1323 |
fun count_loose (Bound i) k = if i >= k then 1 else 0 |
|
1324 |
| count_loose (s $ t) k = count_loose s k + count_loose t k |
|
1325 |
| count_loose (Abs (_, _, t)) k = count_loose t (k + 1) |
|
1326 |
| count_loose _ _ = 0; |
|
1327 |
fun is_trivial_let (Const (@{const_name Let}, _) $ x $ t) = |
|
1328 |
case t |
|
1329 |
of Abs (_, _, t') => count_loose t' 0 <= 1 |
|
1330 |
| _ => true; |
|
1331 |
in fn _ => fn ss => fn ct => if is_trivial_let (Thm.term_of ct) |
|
31151 | 1332 |
then SOME @{thm Let_def} (*no or one ocurrence of bound variable*) |
28741 | 1333 |
else let (*Norbert Schirmer's case*) |
1334 |
val ctxt = Simplifier.the_context ss; |
|
1335 |
val thy = ProofContext.theory_of ctxt; |
|
1336 |
val t = Thm.term_of ct; |
|
1337 |
val ([t'], ctxt') = Variable.import_terms false [t] ctxt; |
|
1338 |
in Option.map (hd o Variable.export ctxt' ctxt o single) |
|
1339 |
(case t' of Const (@{const_name Let},_) $ x $ f => (* x and f are already in normal form *) |
|
1340 |
if is_Free x orelse is_Bound x orelse is_Const x |
|
1341 |
then SOME @{thm Let_def} |
|
1342 |
else |
|
1343 |
let |
|
1344 |
val n = case f of (Abs (x, _, _)) => x | _ => "x"; |
|
1345 |
val cx = cterm_of thy x; |
|
1346 |
val {T = xT, ...} = rep_cterm cx; |
|
1347 |
val cf = cterm_of thy f; |
|
1348 |
val fx_g = Simplifier.rewrite ss (Thm.capply cf cx); |
|
1349 |
val (_ $ _ $ g) = prop_of fx_g; |
|
1350 |
val g' = abstract_over (x,g); |
|
1351 |
in (if (g aconv g') |
|
1352 |
then |
|
1353 |
let |
|
1354 |
val rl = |
|
1355 |
cterm_instantiate [(f_Let_unfold, cf), (x_Let_unfold, cx)] @{thm Let_unfold}; |
|
1356 |
in SOME (rl OF [fx_g]) end |
|
1357 |
else if Term.betapply (f, x) aconv g then NONE (*avoid identity conversion*) |
|
1358 |
else let |
|
1359 |
val abs_g'= Abs (n,xT,g'); |
|
1360 |
val g'x = abs_g'$x; |
|
1361 |
val g_g'x = symmetric (beta_conversion false (cterm_of thy g'x)); |
|
1362 |
val rl = cterm_instantiate |
|
1363 |
[(f_Let_folded, cterm_of thy f), (x_Let_folded, cx), |
|
1364 |
(g_Let_folded, cterm_of thy abs_g')] |
|
1365 |
@{thm Let_folded}; |
|
1366 |
in SOME (rl OF [transitive fx_g g_g'x]) |
|
1367 |
end) |
|
1368 |
end |
|
1369 |
| _ => NONE) |
|
1370 |
end |
|
1371 |
end *} |
|
24035 | 1372 |
|
21151 | 1373 |
lemma True_implies_equals: "(True \<Longrightarrow> PROP P) \<equiv> PROP P" |
1374 |
proof |
|
23389 | 1375 |
assume "True \<Longrightarrow> PROP P" |
1376 |
from this [OF TrueI] show "PROP P" . |
|
21151 | 1377 |
next |
1378 |
assume "PROP P" |
|
23389 | 1379 |
then show "PROP P" . |
21151 | 1380 |
qed |
1381 |
||
1382 |
lemma ex_simps: |
|
1383 |
"!!P Q. (EX x. P x & Q) = ((EX x. P x) & Q)" |
|
1384 |
"!!P Q. (EX x. P & Q x) = (P & (EX x. Q x))" |
|
1385 |
"!!P Q. (EX x. P x | Q) = ((EX x. P x) | Q)" |
|
1386 |
"!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))" |
|
1387 |
"!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)" |
|
1388 |
"!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))" |
|
1389 |
-- {* Miniscoping: pushing in existential quantifiers. *} |
|
1390 |
by (iprover | blast)+ |
|
1391 |
||
1392 |
lemma all_simps: |
|
1393 |
"!!P Q. (ALL x. P x & Q) = ((ALL x. P x) & Q)" |
|
1394 |
"!!P Q. (ALL x. P & Q x) = (P & (ALL x. Q x))" |
|
1395 |
"!!P Q. (ALL x. P x | Q) = ((ALL x. P x) | Q)" |
|
1396 |
"!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))" |
|
1397 |
"!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)" |
|
1398 |
"!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))" |
|
1399 |
-- {* Miniscoping: pushing in universal quantifiers. *} |
|
1400 |
by (iprover | blast)+ |
|
15481 | 1401 |
|
21671 | 1402 |
lemmas [simp] = |
1403 |
triv_forall_equality (*prunes params*) |
|
1404 |
True_implies_equals (*prune asms `True'*) |
|
1405 |
if_True |
|
1406 |
if_False |
|
1407 |
if_cancel |
|
1408 |
if_eq_cancel |
|
1409 |
imp_disjL |
|
20973 | 1410 |
(*In general it seems wrong to add distributive laws by default: they |
1411 |
might cause exponential blow-up. But imp_disjL has been in for a while |
|
1412 |
and cannot be removed without affecting existing proofs. Moreover, |
|
1413 |
rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the |
|
1414 |
grounds that it allows simplification of R in the two cases.*) |
|
21671 | 1415 |
conj_assoc |
1416 |
disj_assoc |
|
1417 |
de_Morgan_conj |
|
1418 |
de_Morgan_disj |
|
1419 |
imp_disj1 |
|
1420 |
imp_disj2 |
|
1421 |
not_imp |
|
1422 |
disj_not1 |
|
1423 |
not_all |
|
1424 |
not_ex |
|
1425 |
cases_simp |
|
1426 |
the_eq_trivial |
|
1427 |
the_sym_eq_trivial |
|
1428 |
ex_simps |
|
1429 |
all_simps |
|
1430 |
simp_thms |
|
1431 |
||
1432 |
lemmas [cong] = imp_cong simp_implies_cong |
|
1433 |
lemmas [split] = split_if |
|
20973 | 1434 |
|
22377 | 1435 |
ML {* val HOL_ss = @{simpset} *} |
20973 | 1436 |
|
20944 | 1437 |
text {* Simplifies x assuming c and y assuming ~c *} |
1438 |
lemma if_cong: |
|
1439 |
assumes "b = c" |
|
1440 |
and "c \<Longrightarrow> x = u" |
|
1441 |
and "\<not> c \<Longrightarrow> y = v" |
|
1442 |
shows "(if b then x else y) = (if c then u else v)" |
|
23553 | 1443 |
unfolding if_def using assms by simp |
20944 | 1444 |
|
1445 |
text {* Prevents simplification of x and y: |
|
1446 |
faster and allows the execution of functional programs. *} |
|
1447 |
lemma if_weak_cong [cong]: |
|
1448 |
assumes "b = c" |
|
1449 |
shows "(if b then x else y) = (if c then x else y)" |
|
23553 | 1450 |
using assms by (rule arg_cong) |
20944 | 1451 |
|
1452 |
text {* Prevents simplification of t: much faster *} |
|
1453 |
lemma let_weak_cong: |
|
1454 |
assumes "a = b" |
|
1455 |
shows "(let x = a in t x) = (let x = b in t x)" |
|
23553 | 1456 |
using assms by (rule arg_cong) |
20944 | 1457 |
|
1458 |
text {* To tidy up the result of a simproc. Only the RHS will be simplified. *} |
|
1459 |
lemma eq_cong2: |
|
1460 |
assumes "u = u'" |
|
1461 |
shows "(t \<equiv> u) \<equiv> (t \<equiv> u')" |
|
23553 | 1462 |
using assms by simp |
20944 | 1463 |
|
1464 |
lemma if_distrib: |
|
1465 |
"f (if c then x else y) = (if c then f x else f y)" |
|
1466 |
by simp |
|
1467 |
||
1468 |
text {* This lemma restricts the effect of the rewrite rule u=v to the left-hand |
|
21502 | 1469 |
side of an equality. Used in @{text "{Integ,Real}/simproc.ML"} *} |
20944 | 1470 |
lemma restrict_to_left: |
1471 |
assumes "x = y" |
|
1472 |
shows "(x = z) = (y = z)" |
|
23553 | 1473 |
using assms by simp |
20944 | 1474 |
|
17459 | 1475 |
|
20944 | 1476 |
subsubsection {* Generic cases and induction *} |
17459 | 1477 |
|
20944 | 1478 |
text {* Rule projections: *} |
18887 | 1479 |
|
20944 | 1480 |
ML {* |
1481 |
structure ProjectRule = ProjectRuleFun |
|
25388 | 1482 |
( |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1483 |
val conjunct1 = @{thm conjunct1} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1484 |
val conjunct2 = @{thm conjunct2} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1485 |
val mp = @{thm mp} |
25388 | 1486 |
) |
17459 | 1487 |
*} |
1488 |
||
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1489 |
constdefs |
18457 | 1490 |
induct_forall where "induct_forall P == \<forall>x. P x" |
1491 |
induct_implies where "induct_implies A B == A \<longrightarrow> B" |
|
1492 |
induct_equal where "induct_equal x y == x = y" |
|
1493 |
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
|
1494 |
|
11989 | 1495 |
lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))" |
18457 | 1496 |
by (unfold atomize_all induct_forall_def) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1497 |
|
11989 | 1498 |
lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)" |
18457 | 1499 |
by (unfold atomize_imp induct_implies_def) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1500 |
|
11989 | 1501 |
lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)" |
18457 | 1502 |
by (unfold atomize_eq induct_equal_def) |
1503 |
||
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28741
diff
changeset
|
1504 |
lemma induct_conj_eq: "(A &&& B) == Trueprop (induct_conj A B)" |
18457 | 1505 |
by (unfold atomize_conj induct_conj_def) |
1506 |
||
1507 |
lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq |
|
1508 |
lemmas induct_rulify [symmetric, standard] = induct_atomize |
|
1509 |
lemmas induct_rulify_fallback = |
|
1510 |
induct_forall_def induct_implies_def induct_equal_def induct_conj_def |
|
1511 |
||
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1512 |
|
11989 | 1513 |
lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = |
1514 |
induct_conj (induct_forall A) (induct_forall B)" |
|
17589 | 1515 |
by (unfold induct_forall_def induct_conj_def) iprover |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1516 |
|
11989 | 1517 |
lemma induct_implies_conj: "induct_implies C (induct_conj A B) = |
1518 |
induct_conj (induct_implies C A) (induct_implies C B)" |
|
17589 | 1519 |
by (unfold induct_implies_def induct_conj_def) iprover |
11989 | 1520 |
|
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1521 |
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
|
1522 |
proof |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1523 |
assume r: "induct_conj A B ==> PROP C" and A B |
18457 | 1524 |
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
|
1525 |
next |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1526 |
assume r: "A ==> B ==> PROP C" and "induct_conj A B" |
18457 | 1527 |
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
|
1528 |
qed |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1529 |
|
11989 | 1530 |
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
|
1531 |
|
11989 | 1532 |
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
|
1533 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1534 |
text {* Method setup. *} |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1535 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1536 |
ML {* |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1537 |
structure Induct = InductFun |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1538 |
( |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1539 |
val cases_default = @{thm case_split} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1540 |
val atomize = @{thms induct_atomize} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1541 |
val rulify = @{thms induct_rulify} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1542 |
val rulify_fallback = @{thms induct_rulify_fallback} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1543 |
) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1544 |
*} |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1545 |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
24748
diff
changeset
|
1546 |
setup Induct.setup |
18457 | 1547 |
|
27326
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
wenzelm
parents:
27212
diff
changeset
|
1548 |
use "~~/src/Tools/induct_tacs.ML" |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1549 |
setup InductTacs.setup |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1550 |
|
20944 | 1551 |
|
28325 | 1552 |
subsubsection {* Coherent logic *} |
1553 |
||
1554 |
ML {* |
|
1555 |
structure Coherent = CoherentFun |
|
1556 |
( |
|
1557 |
val atomize_elimL = @{thm atomize_elimL} |
|
1558 |
val atomize_exL = @{thm atomize_exL} |
|
1559 |
val atomize_conjL = @{thm atomize_conjL} |
|
1560 |
val atomize_disjL = @{thm atomize_disjL} |
|
1561 |
val operator_names = |
|
1562 |
[@{const_name "op |"}, @{const_name "op &"}, @{const_name "Ex"}] |
|
1563 |
); |
|
1564 |
*} |
|
1565 |
||
1566 |
setup Coherent.setup |
|
1567 |
||
1568 |
||
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1569 |
subsubsection {* Reorienting equalities *} |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1570 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1571 |
ML {* |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1572 |
signature REORIENT_PROC = |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1573 |
sig |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1574 |
val init : theory -> theory |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1575 |
val add : (term -> bool) -> theory -> theory |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1576 |
val proc : morphism -> simpset -> cterm -> thm option |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1577 |
end; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1578 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1579 |
structure ReorientProc : REORIENT_PROC = |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1580 |
struct |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1581 |
structure Data = TheoryDataFun |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1582 |
( |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1583 |
type T = term -> bool; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1584 |
val empty = (fn _ => false); |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1585 |
val copy = I; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1586 |
val extend = I; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1587 |
fun merge _ (m1, m2) = (fn t => m1 t orelse m2 t); |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1588 |
) |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1589 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1590 |
val init = Data.init; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1591 |
fun add m = Data.map (fn matches => fn t => matches t orelse m t); |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1592 |
val meta_reorient = @{thm eq_commute [THEN eq_reflection]}; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1593 |
fun proc phi ss ct = |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1594 |
let |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1595 |
val ctxt = Simplifier.the_context ss; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1596 |
val thy = ProofContext.theory_of ctxt; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1597 |
val matches = Data.get thy; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1598 |
in |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1599 |
case Thm.term_of ct of |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1600 |
(_ $ t $ u) => if matches u then NONE else SOME meta_reorient |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1601 |
| _ => NONE |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1602 |
end; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1603 |
end; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1604 |
*} |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1605 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1606 |
setup ReorientProc.init |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1607 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1608 |
setup {* |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1609 |
ReorientProc.add |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1610 |
(fn Const(@{const_name HOL.zero}, _) => true |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1611 |
| Const(@{const_name HOL.one}, _) => true |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1612 |
| _ => false) |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1613 |
*} |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1614 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1615 |
simproc_setup reorient_zero ("0 = x") = ReorientProc.proc |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1616 |
simproc_setup reorient_one ("1 = x") = ReorientProc.proc |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1617 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1618 |
|
20944 | 1619 |
subsection {* Other simple lemmas and lemma duplicates *} |
1620 |
||
24166 | 1621 |
lemma Let_0 [simp]: "Let 0 f = f 0" |
1622 |
unfolding Let_def .. |
|
1623 |
||
1624 |
lemma Let_1 [simp]: "Let 1 f = f 1" |
|
1625 |
unfolding Let_def .. |
|
1626 |
||
20944 | 1627 |
lemma ex1_eq [iff]: "EX! x. x = t" "EX! x. t = x" |
1628 |
by blast+ |
|
1629 |
||
1630 |
lemma choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))" |
|
1631 |
apply (rule iffI) |
|
1632 |
apply (rule_tac a = "%x. THE y. P x y" in ex1I) |
|
1633 |
apply (fast dest!: theI') |
|
1634 |
apply (fast intro: ext the1_equality [symmetric]) |
|
1635 |
apply (erule ex1E) |
|
1636 |
apply (rule allI) |
|
1637 |
apply (rule ex1I) |
|
1638 |
apply (erule spec) |
|
1639 |
apply (erule_tac x = "%z. if z = x then y else f z" in allE) |
|
1640 |
apply (erule impE) |
|
1641 |
apply (rule allI) |
|
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1642 |
apply (case_tac "xa = x") |
20944 | 1643 |
apply (drule_tac [3] x = x in fun_cong, simp_all) |
1644 |
done |
|
1645 |
||
1646 |
lemma mk_left_commute: |
|
21547
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
haftmann
parents:
21524
diff
changeset
|
1647 |
fixes f (infix "\<otimes>" 60) |
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
haftmann
parents:
21524
diff
changeset
|
1648 |
assumes a: "\<And>x y z. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" and |
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
haftmann
parents:
21524
diff
changeset
|
1649 |
c: "\<And>x y. x \<otimes> y = y \<otimes> x" |
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
haftmann
parents:
21524
diff
changeset
|
1650 |
shows "x \<otimes> (y \<otimes> z) = y \<otimes> (x \<otimes> z)" |
20944 | 1651 |
by (rule trans [OF trans [OF c a] arg_cong [OF c, of "f y"]]) |
1652 |
||
22218 | 1653 |
lemmas eq_sym_conv = eq_commute |
1654 |
||
23037
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1655 |
lemma nnf_simps: |
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1656 |
"(\<not>(P \<and> Q)) = (\<not> P \<or> \<not> Q)" "(\<not> (P \<or> Q)) = (\<not> P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)" |
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1657 |
"(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not>(P = Q)) = ((P \<and> \<not> Q) \<or> (\<not>P \<and> Q))" |
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1658 |
"(\<not> \<not>(P)) = P" |
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1659 |
by blast+ |
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1660 |
|
21671 | 1661 |
|
1662 |
subsection {* Basic ML bindings *} |
|
1663 |
||
1664 |
ML {* |
|
22129 | 1665 |
val FalseE = @{thm FalseE} |
1666 |
val Let_def = @{thm Let_def} |
|
1667 |
val TrueI = @{thm TrueI} |
|
1668 |
val allE = @{thm allE} |
|
1669 |
val allI = @{thm allI} |
|
1670 |
val all_dupE = @{thm all_dupE} |
|
1671 |
val arg_cong = @{thm arg_cong} |
|
1672 |
val box_equals = @{thm box_equals} |
|
1673 |
val ccontr = @{thm ccontr} |
|
1674 |
val classical = @{thm classical} |
|
1675 |
val conjE = @{thm conjE} |
|
1676 |
val conjI = @{thm conjI} |
|
1677 |
val conjunct1 = @{thm conjunct1} |
|
1678 |
val conjunct2 = @{thm conjunct2} |
|
1679 |
val disjCI = @{thm disjCI} |
|
1680 |
val disjE = @{thm disjE} |
|
1681 |
val disjI1 = @{thm disjI1} |
|
1682 |
val disjI2 = @{thm disjI2} |
|
1683 |
val eq_reflection = @{thm eq_reflection} |
|
1684 |
val ex1E = @{thm ex1E} |
|
1685 |
val ex1I = @{thm ex1I} |
|
1686 |
val ex1_implies_ex = @{thm ex1_implies_ex} |
|
1687 |
val exE = @{thm exE} |
|
1688 |
val exI = @{thm exI} |
|
1689 |
val excluded_middle = @{thm excluded_middle} |
|
1690 |
val ext = @{thm ext} |
|
1691 |
val fun_cong = @{thm fun_cong} |
|
1692 |
val iffD1 = @{thm iffD1} |
|
1693 |
val iffD2 = @{thm iffD2} |
|
1694 |
val iffI = @{thm iffI} |
|
1695 |
val impE = @{thm impE} |
|
1696 |
val impI = @{thm impI} |
|
1697 |
val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq} |
|
1698 |
val mp = @{thm mp} |
|
1699 |
val notE = @{thm notE} |
|
1700 |
val notI = @{thm notI} |
|
1701 |
val not_all = @{thm not_all} |
|
1702 |
val not_ex = @{thm not_ex} |
|
1703 |
val not_iff = @{thm not_iff} |
|
1704 |
val not_not = @{thm not_not} |
|
1705 |
val not_sym = @{thm not_sym} |
|
1706 |
val refl = @{thm refl} |
|
1707 |
val rev_mp = @{thm rev_mp} |
|
1708 |
val spec = @{thm spec} |
|
1709 |
val ssubst = @{thm ssubst} |
|
1710 |
val subst = @{thm subst} |
|
1711 |
val sym = @{thm sym} |
|
1712 |
val trans = @{thm trans} |
|
21671 | 1713 |
*} |
1714 |
||
1715 |
||
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1716 |
subsection {* Code generator setup *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1717 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1718 |
subsubsection {* SML code generator setup *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1719 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1720 |
use "Tools/recfun_codegen.ML" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1721 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1722 |
setup {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1723 |
Codegen.setup |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1724 |
#> RecfunCodegen.setup |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1725 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1726 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1727 |
types_code |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1728 |
"bool" ("bool") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1729 |
attach (term_of) {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1730 |
fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1731 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1732 |
attach (test) {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1733 |
fun gen_bool i = |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1734 |
let val b = one_of [false, true] |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1735 |
in (b, fn () => term_of_bool b) end; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1736 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1737 |
"prop" ("bool") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1738 |
attach (term_of) {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1739 |
fun term_of_prop b = |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1740 |
HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const); |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1741 |
*} |
28400 | 1742 |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1743 |
consts_code |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1744 |
"Trueprop" ("(_)") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1745 |
"True" ("true") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1746 |
"False" ("false") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1747 |
"Not" ("Bool.not") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1748 |
"op |" ("(_ orelse/ _)") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1749 |
"op &" ("(_ andalso/ _)") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1750 |
"If" ("(if _/ then _/ else _)") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1751 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1752 |
setup {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1753 |
let |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1754 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1755 |
fun eq_codegen thy defs dep thyname b t gr = |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1756 |
(case strip_comb t of |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1757 |
(Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1758 |
| (Const ("op =", _), [t, u]) => |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1759 |
let |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1760 |
val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1761 |
val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr'; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1762 |
val (_, gr''') = Codegen.invoke_tycodegen thy defs dep thyname false HOLogic.boolT gr''; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1763 |
in |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1764 |
SOME (Codegen.parens |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1765 |
(Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''') |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1766 |
end |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1767 |
| (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1768 |
thy defs dep thyname b (Codegen.eta_expand t ts 2) gr) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1769 |
| _ => NONE); |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1770 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1771 |
in |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1772 |
Codegen.add_codegen "eq_codegen" eq_codegen |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1773 |
end |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1774 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1775 |
|
31151 | 1776 |
subsubsection {* Generic code generator preprocessor setup *} |
1777 |
||
1778 |
setup {* |
|
1779 |
Code_Preproc.map_pre (K HOL_basic_ss) |
|
1780 |
#> Code_Preproc.map_post (K HOL_basic_ss) |
|
1781 |
*} |
|
1782 |
||
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1783 |
subsubsection {* Equality *} |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24842
diff
changeset
|
1784 |
|
29608 | 1785 |
class eq = |
26513 | 1786 |
fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
28400 | 1787 |
assumes eq_equals: "eq x y \<longleftrightarrow> x = y" |
26513 | 1788 |
begin |
1789 |
||
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1790 |
lemma eq [code unfold, code inline del]: "eq = (op =)" |
28346
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
haftmann
parents:
28325
diff
changeset
|
1791 |
by (rule ext eq_equals)+ |
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
haftmann
parents:
28325
diff
changeset
|
1792 |
|
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
haftmann
parents:
28325
diff
changeset
|
1793 |
lemma eq_refl: "eq x x \<longleftrightarrow> True" |
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
haftmann
parents:
28325
diff
changeset
|
1794 |
unfolding eq by rule+ |
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
haftmann
parents:
28325
diff
changeset
|
1795 |
|
31151 | 1796 |
lemma equals_eq: "(op =) \<equiv> eq" |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1797 |
by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1798 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1799 |
declare equals_eq [symmetric, code post] |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1800 |
|
26513 | 1801 |
end |
1802 |
||
30966 | 1803 |
declare equals_eq [code] |
1804 |
||
31151 | 1805 |
setup {* |
1806 |
Code_Preproc.map_pre (fn simpset => |
|
1807 |
simpset addsimprocs [Simplifier.simproc_i @{theory} "eq" [@{term "op ="}] |
|
1808 |
(fn thy => fn _ => fn t as Const (_, T) => case strip_type T |
|
1809 |
of ((T as Type _) :: _, _) => SOME @{thm equals_eq} |
|
1810 |
| _ => NONE)]) |
|
1811 |
*} |
|
1812 |
||
30966 | 1813 |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1814 |
subsubsection {* Generic code generator foundation *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1815 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1816 |
text {* Datatypes *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1817 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1818 |
code_datatype True False |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1819 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1820 |
code_datatype "TYPE('a\<Colon>{})" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1821 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1822 |
code_datatype Trueprop "prop" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1823 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1824 |
text {* Code equations *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1825 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1826 |
lemma [code]: |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1827 |
shows "(True \<Longrightarrow> PROP P) \<equiv> PROP P" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1828 |
and "(False \<Longrightarrow> Q) \<equiv> Trueprop True" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1829 |
and "(PROP P \<Longrightarrow> True) \<equiv> Trueprop True" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1830 |
and "(Q \<Longrightarrow> False) \<equiv> Trueprop (\<not> Q)" by (auto intro!: equal_intr_rule) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1831 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1832 |
lemma [code]: |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1833 |
shows "False \<and> x \<longleftrightarrow> False" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1834 |
and "True \<and> x \<longleftrightarrow> x" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1835 |
and "x \<and> False \<longleftrightarrow> False" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1836 |
and "x \<and> True \<longleftrightarrow> x" by simp_all |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1837 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1838 |
lemma [code]: |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1839 |
shows "False \<or> x \<longleftrightarrow> x" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1840 |
and "True \<or> x \<longleftrightarrow> True" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1841 |
and "x \<or> False \<longleftrightarrow> x" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1842 |
and "x \<or> True \<longleftrightarrow> True" by simp_all |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1843 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1844 |
lemma [code]: |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1845 |
shows "\<not> True \<longleftrightarrow> False" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1846 |
and "\<not> False \<longleftrightarrow> True" by (rule HOL.simp_thms)+ |
28513 | 1847 |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1848 |
lemmas [code] = Let_def if_True if_False |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1849 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1850 |
lemmas [code, code unfold, symmetric, code post] = imp_conv_disj |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1851 |
|
31132 | 1852 |
instantiation itself :: (type) eq |
1853 |
begin |
|
1854 |
||
1855 |
definition eq_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where |
|
1856 |
"eq_itself x y \<longleftrightarrow> x = y" |
|
1857 |
||
1858 |
instance proof |
|
1859 |
qed (fact eq_itself_def) |
|
1860 |
||
1861 |
end |
|
1862 |
||
1863 |
lemma eq_itself_code [code]: |
|
1864 |
"eq_class.eq TYPE('a) TYPE('a) \<longleftrightarrow> True" |
|
1865 |
by (simp add: eq) |
|
1866 |
||
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1867 |
text {* Equality *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1868 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1869 |
declare simp_thms(6) [code nbe] |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1870 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1871 |
setup {* |
31156 | 1872 |
Code.add_const_alias @{thm equals_eq} |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1873 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1874 |
|
31151 | 1875 |
hide (open) const eq |
1876 |
hide const eq |
|
1877 |
||
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1878 |
text {* Cases *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1879 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1880 |
lemma Let_case_cert: |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1881 |
assumes "CASE \<equiv> (\<lambda>x. Let x f)" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1882 |
shows "CASE x \<equiv> f x" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1883 |
using assms by simp_all |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1884 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1885 |
lemma If_case_cert: |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1886 |
assumes "CASE \<equiv> (\<lambda>b. If b f g)" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1887 |
shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1888 |
using assms by simp_all |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1889 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1890 |
setup {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1891 |
Code.add_case @{thm Let_case_cert} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1892 |
#> Code.add_case @{thm If_case_cert} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1893 |
#> Code.add_undefined @{const_name undefined} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1894 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1895 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1896 |
code_abort undefined |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1897 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1898 |
subsubsection {* Generic code generator target languages *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1899 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1900 |
text {* type bool *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1901 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1902 |
code_type bool |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1903 |
(SML "bool") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1904 |
(OCaml "bool") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1905 |
(Haskell "Bool") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1906 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1907 |
code_const True and False and Not and "op &" and "op |" and If |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1908 |
(SML "true" and "false" and "not" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1909 |
and infixl 1 "andalso" and infixl 0 "orelse" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1910 |
and "!(if (_)/ then (_)/ else (_))") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1911 |
(OCaml "true" and "false" and "not" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1912 |
and infixl 4 "&&" and infixl 2 "||" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1913 |
and "!(if (_)/ then (_)/ else (_))") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1914 |
(Haskell "True" and "False" and "not" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1915 |
and infixl 3 "&&" and infixl 2 "||" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1916 |
and "!(if (_)/ then (_)/ else (_))") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1917 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1918 |
code_reserved SML |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1919 |
bool true false not |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1920 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1921 |
code_reserved OCaml |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1922 |
bool not |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1923 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1924 |
text {* using built-in Haskell equality *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1925 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1926 |
code_class eq |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1927 |
(Haskell "Eq") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1928 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1929 |
code_const "eq_class.eq" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1930 |
(Haskell infixl 4 "==") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1931 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1932 |
code_const "op =" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1933 |
(Haskell infixl 4 "==") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1934 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1935 |
text {* undefined *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1936 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1937 |
code_const undefined |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1938 |
(SML "!(raise/ Fail/ \"undefined\")") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1939 |
(OCaml "failwith/ \"undefined\"") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1940 |
(Haskell "error/ \"undefined\"") |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1941 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1942 |
subsubsection {* Evaluation and normalization by evaluation *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1943 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1944 |
setup {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1945 |
Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1946 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1947 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1948 |
ML {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1949 |
structure Eval_Method = |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1950 |
struct |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1951 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1952 |
val eval_ref : (unit -> bool) option ref = ref NONE; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1953 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1954 |
end; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1955 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1956 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1957 |
oracle eval_oracle = {* fn ct => |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1958 |
let |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1959 |
val thy = Thm.theory_of_cterm ct; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1960 |
val t = Thm.term_of ct; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1961 |
val dummy = @{cprop True}; |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1962 |
in case try HOLogic.dest_Trueprop t |
30947 | 1963 |
of SOME t' => if Code_ML.eval NONE |
30970
3fe2e418a071
generic postprocessing scheme for term evaluations
haftmann
parents:
30966
diff
changeset
|
1964 |
("Eval_Method.eval_ref", Eval_Method.eval_ref) (K I) thy t' [] |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1965 |
then Thm.capply (Thm.capply @{cterm "op \<equiv> \<Colon> prop \<Rightarrow> prop \<Rightarrow> prop"} ct) dummy |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1966 |
else dummy |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1967 |
| NONE => dummy |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1968 |
end |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1969 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1970 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1971 |
ML {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1972 |
fun gen_eval_method conv ctxt = SIMPLE_METHOD' |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1973 |
(CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1974 |
THEN' rtac TrueI) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1975 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1976 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1977 |
method_setup eval = {* Scan.succeed (gen_eval_method eval_oracle) *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1978 |
"solve goal by evaluation" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1979 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1980 |
method_setup evaluation = {* Scan.succeed (gen_eval_method Codegen.evaluation_conv) *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1981 |
"solve goal by evaluation" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1982 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1983 |
method_setup normalization = {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1984 |
Scan.succeed (K (SIMPLE_METHOD' (CONVERSION Nbe.norm_conv THEN' (fn k => TRY (rtac TrueI k))))) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1985 |
*} "solve goal by normalization" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1986 |
|
31902 | 1987 |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1988 |
subsubsection {* Quickcheck *} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1989 |
|
31172
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1990 |
ML {* |
31902 | 1991 |
structure Quickcheck_RecFun_Simps = Named_Thms |
31172
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1992 |
( |
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1993 |
val name = "quickcheck_recfun_simp" |
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1994 |
val description = "simplification rules of recursive functions as needed by Quickcheck" |
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1995 |
) |
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1996 |
*} |
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1997 |
|
31902 | 1998 |
setup Quickcheck_RecFun_Simps.setup |
31172
74d72ba262fb
added collection of simplification rules of recursive functions for quickcheck
bulwahn
parents:
31156
diff
changeset
|
1999 |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2000 |
setup {* |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2001 |
Quickcheck.add_generator ("SML", Codegen.test_term) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2002 |
*} |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2003 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2004 |
quickcheck_params [size = 5, iterations = 50] |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2005 |
|
23247 | 2006 |
|
30980 | 2007 |
subsection {* Nitpick setup *} |
30309
188f0658af9f
Added a "nitpick_maybe" symbol, which is used by Nitpick. This will go away once Nitpick is part of HOL.
blanchet
parents:
30254
diff
changeset
|
2008 |
|
188f0658af9f
Added a "nitpick_maybe" symbol, which is used by Nitpick. This will go away once Nitpick is part of HOL.
blanchet
parents:
30254
diff
changeset
|
2009 |
text {* This will be relocated once Nitpick is moved to HOL. *} |
188f0658af9f
Added a "nitpick_maybe" symbol, which is used by Nitpick. This will go away once Nitpick is part of HOL.
blanchet
parents:
30254
diff
changeset
|
2010 |
|
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2011 |
ML {* |
31902 | 2012 |
structure Nitpick_Const_Defs = Named_Thms |
30254
7b8afdfa2f83
Second try at adding "nitpick_const_def" attribute.
blanchet
parents:
30242
diff
changeset
|
2013 |
( |
7b8afdfa2f83
Second try at adding "nitpick_const_def" attribute.
blanchet
parents:
30242
diff
changeset
|
2014 |
val name = "nitpick_const_def" |
7b8afdfa2f83
Second try at adding "nitpick_const_def" attribute.
blanchet
parents:
30242
diff
changeset
|
2015 |
val description = "alternative definitions of constants as needed by Nitpick" |
7b8afdfa2f83
Second try at adding "nitpick_const_def" attribute.
blanchet
parents:
30242
diff
changeset
|
2016 |
) |
31902 | 2017 |
structure Nitpick_Const_Simps = Named_Thms |
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2018 |
( |
29866
6e93ae65c678
Added Nitpick_Const_Psimp attribute, dropped the 's' in Nitpick_Const_Simps, and killed the Nitpick_Ind_Intros attribute.
blanchet
parents:
29863
diff
changeset
|
2019 |
val name = "nitpick_const_simp" |
29869
a7a8b90cd882
Renamed descriptions of Nitpick (and ATP) attributes, so that they fit well with the rest of the sentence in ProofGeneral.
blanchet
parents:
29868
diff
changeset
|
2020 |
val description = "equational specification of constants as needed by Nitpick" |
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2021 |
) |
31902 | 2022 |
structure Nitpick_Const_Psimps = Named_Thms |
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2023 |
( |
29866
6e93ae65c678
Added Nitpick_Const_Psimp attribute, dropped the 's' in Nitpick_Const_Simps, and killed the Nitpick_Ind_Intros attribute.
blanchet
parents:
29863
diff
changeset
|
2024 |
val name = "nitpick_const_psimp" |
29869
a7a8b90cd882
Renamed descriptions of Nitpick (and ATP) attributes, so that they fit well with the rest of the sentence in ProofGeneral.
blanchet
parents:
29868
diff
changeset
|
2025 |
val description = "partial equational specification of constants as needed by Nitpick" |
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2026 |
) |
31902 | 2027 |
structure Nitpick_Ind_Intros = Named_Thms |
29868 | 2028 |
( |
2029 |
val name = "nitpick_ind_intro" |
|
29869
a7a8b90cd882
Renamed descriptions of Nitpick (and ATP) attributes, so that they fit well with the rest of the sentence in ProofGeneral.
blanchet
parents:
29868
diff
changeset
|
2030 |
val description = "introduction rules for (co)inductive predicates as needed by Nitpick" |
29868 | 2031 |
) |
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2032 |
*} |
30980 | 2033 |
|
2034 |
setup {* |
|
31902 | 2035 |
Nitpick_Const_Defs.setup |
2036 |
#> Nitpick_Const_Simps.setup |
|
2037 |
#> Nitpick_Const_Psimps.setup |
|
2038 |
#> Nitpick_Ind_Intros.setup |
|
30980 | 2039 |
*} |
2040 |
||
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2041 |
|
22839 | 2042 |
subsection {* Legacy tactics and ML bindings *} |
21671 | 2043 |
|
2044 |
ML {* |
|
2045 |
fun strip_tac i = REPEAT (resolve_tac [impI, allI] i); |
|
2046 |
||
2047 |
(* combination of (spec RS spec RS ...(j times) ... spec RS mp) *) |
|
2048 |
local |
|
2049 |
fun wrong_prem (Const ("All", _) $ (Abs (_, _, t))) = wrong_prem t |
|
2050 |
| wrong_prem (Bound _) = true |
|
2051 |
| wrong_prem _ = false; |
|
2052 |
val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.prems_of); |
|
2053 |
in |
|
2054 |
fun smp i = funpow i (fn m => filter_right ([spec] RL m)) ([mp]); |
|
2055 |
fun smp_tac j = EVERY'[dresolve_tac (smp j), atac]; |
|
2056 |
end; |
|
22839 | 2057 |
|
2058 |
val all_conj_distrib = thm "all_conj_distrib"; |
|
2059 |
val all_simps = thms "all_simps"; |
|
2060 |
val atomize_not = thm "atomize_not"; |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
24748
diff
changeset
|
2061 |
val case_split = thm "case_split"; |
22839 | 2062 |
val cases_simp = thm "cases_simp"; |
2063 |
val choice_eq = thm "choice_eq" |
|
2064 |
val cong = thm "cong" |
|
2065 |
val conj_comms = thms "conj_comms"; |
|
2066 |
val conj_cong = thm "conj_cong"; |
|
2067 |
val de_Morgan_conj = thm "de_Morgan_conj"; |
|
2068 |
val de_Morgan_disj = thm "de_Morgan_disj"; |
|
2069 |
val disj_assoc = thm "disj_assoc"; |
|
2070 |
val disj_comms = thms "disj_comms"; |
|
2071 |
val disj_cong = thm "disj_cong"; |
|
2072 |
val eq_ac = thms "eq_ac"; |
|
2073 |
val eq_cong2 = thm "eq_cong2" |
|
2074 |
val Eq_FalseI = thm "Eq_FalseI"; |
|
2075 |
val Eq_TrueI = thm "Eq_TrueI"; |
|
2076 |
val Ex1_def = thm "Ex1_def" |
|
2077 |
val ex_disj_distrib = thm "ex_disj_distrib"; |
|
2078 |
val ex_simps = thms "ex_simps"; |
|
2079 |
val if_cancel = thm "if_cancel"; |
|
2080 |
val if_eq_cancel = thm "if_eq_cancel"; |
|
2081 |
val if_False = thm "if_False"; |
|
2082 |
val iff_conv_conj_imp = thm "iff_conv_conj_imp"; |
|
2083 |
val iff = thm "iff" |
|
2084 |
val if_splits = thms "if_splits"; |
|
2085 |
val if_True = thm "if_True"; |
|
2086 |
val if_weak_cong = thm "if_weak_cong" |
|
2087 |
val imp_all = thm "imp_all"; |
|
2088 |
val imp_cong = thm "imp_cong"; |
|
2089 |
val imp_conjL = thm "imp_conjL"; |
|
2090 |
val imp_conjR = thm "imp_conjR"; |
|
2091 |
val imp_conv_disj = thm "imp_conv_disj"; |
|
2092 |
val simp_implies_def = thm "simp_implies_def"; |
|
2093 |
val simp_thms = thms "simp_thms"; |
|
2094 |
val split_if = thm "split_if"; |
|
2095 |
val the1_equality = thm "the1_equality" |
|
2096 |
val theI = thm "theI" |
|
2097 |
val theI' = thm "theI'" |
|
2098 |
val True_implies_equals = thm "True_implies_equals"; |
|
23037
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
2099 |
val nnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps simp_thms @ @{thms "nnf_simps"}) |
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
2100 |
|
21671 | 2101 |
*} |
2102 |
||
14357 | 2103 |
end |