author | nipkow |
Thu, 02 Dec 2004 11:59:34 +0100 | |
changeset 15362 | a000b267be57 |
parent 15360 | 300e09825d8b |
child 15363 | 885a40edcdba |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/HOL.thy |
2 |
ID: $Id$ |
|
11750 | 3 |
Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson |
4 |
*) |
|
923 | 5 |
|
11750 | 6 |
header {* The basis of Higher-Order Logic *} |
923 | 7 |
|
15131 | 8 |
theory HOL |
15140 | 9 |
imports CPure |
15131 | 10 |
files ("HOL_lemmas.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML") |
15197 | 11 |
("antisym_setup.ML") |
15131 | 12 |
begin |
2260 | 13 |
|
11750 | 14 |
subsection {* Primitive logic *} |
15 |
||
16 |
subsubsection {* Core syntax *} |
|
2260 | 17 |
|
14854 | 18 |
classes type |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
19 |
defaultsort type |
3947 | 20 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
21 |
global |
923 | 22 |
|
7357 | 23 |
typedecl bool |
923 | 24 |
|
25 |
arities |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
26 |
bool :: type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
27 |
fun :: (type, type) type |
923 | 28 |
|
11750 | 29 |
judgment |
30 |
Trueprop :: "bool => prop" ("(_)" 5) |
|
923 | 31 |
|
11750 | 32 |
consts |
7357 | 33 |
Not :: "bool => bool" ("~ _" [40] 40) |
34 |
True :: bool |
|
35 |
False :: bool |
|
36 |
If :: "[bool, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) |
|
3947 | 37 |
arbitrary :: 'a |
923 | 38 |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
39 |
The :: "('a => bool) => 'a" |
7357 | 40 |
All :: "('a => bool) => bool" (binder "ALL " 10) |
41 |
Ex :: "('a => bool) => bool" (binder "EX " 10) |
|
42 |
Ex1 :: "('a => bool) => bool" (binder "EX! " 10) |
|
43 |
Let :: "['a, 'a => 'b] => 'b" |
|
923 | 44 |
|
7357 | 45 |
"=" :: "['a, 'a] => bool" (infixl 50) |
46 |
& :: "[bool, bool] => bool" (infixr 35) |
|
47 |
"|" :: "[bool, bool] => bool" (infixr 30) |
|
48 |
--> :: "[bool, bool] => bool" (infixr 25) |
|
923 | 49 |
|
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
50 |
local |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
51 |
|
2260 | 52 |
|
11750 | 53 |
subsubsection {* Additional concrete syntax *} |
2260 | 54 |
|
4868 | 55 |
nonterminals |
923 | 56 |
letbinds letbind |
57 |
case_syn cases_syn |
|
58 |
||
59 |
syntax |
|
12650 | 60 |
"_not_equal" :: "['a, 'a] => bool" (infixl "~=" 50) |
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
61 |
"_The" :: "[pttrn, bool] => 'a" ("(3THE _./ _)" [0, 10] 10) |
923 | 62 |
|
7357 | 63 |
"_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10) |
64 |
"" :: "letbind => letbinds" ("_") |
|
65 |
"_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") |
|
66 |
"_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) |
|
923 | 67 |
|
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
68 |
"_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
|
69 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10) |
7357 | 70 |
"" :: "case_syn => cases_syn" ("_") |
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
71 |
"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _") |
923 | 72 |
|
73 |
translations |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
74 |
"x ~= y" == "~ (x = y)" |
13764 | 75 |
"THE x. P" == "The (%x. P)" |
923 | 76 |
"_Let (_binds b bs) e" == "_Let b (_Let bs e)" |
1114 | 77 |
"let x = a in e" == "Let a (%x. e)" |
923 | 78 |
|
13763
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
79 |
print_translation {* |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
80 |
(* To avoid eta-contraction of body: *) |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
81 |
[("The", fn [Abs abs] => |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
82 |
let val (x,t) = atomic_abs_tr' abs |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
83 |
in Syntax.const "_The" $ x $ t end)] |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
84 |
*} |
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
85 |
|
12633 | 86 |
syntax (output) |
11687 | 87 |
"=" :: "['a, 'a] => bool" (infix 50) |
12650 | 88 |
"_not_equal" :: "['a, 'a] => bool" (infix "~=" 50) |
2260 | 89 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
90 |
syntax (xsymbols) |
11687 | 91 |
Not :: "bool => bool" ("\<not> _" [40] 40) |
92 |
"op &" :: "[bool, bool] => bool" (infixr "\<and>" 35) |
|
93 |
"op |" :: "[bool, bool] => bool" (infixr "\<or>" 30) |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
94 |
"op -->" :: "[bool, bool] => bool" (infixr "\<longrightarrow>" 25) |
12650 | 95 |
"_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) |
11687 | 96 |
"ALL " :: "[idts, bool] => bool" ("(3\<forall>_./ _)" [0, 10] 10) |
97 |
"EX " :: "[idts, bool] => bool" ("(3\<exists>_./ _)" [0, 10] 10) |
|
98 |
"EX! " :: "[idts, bool] => bool" ("(3\<exists>!_./ _)" [0, 10] 10) |
|
99 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ \<Rightarrow>/ _)" 10) |
|
14361
ad2f5da643b4
* Support for raw latex output in control symbols: \<^raw...>
schirmer
parents:
14357
diff
changeset
|
100 |
(*"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ \<orelse> _")*) |
2372 | 101 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
102 |
syntax (xsymbols output) |
12650 | 103 |
"_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) |
3820 | 104 |
|
6340 | 105 |
syntax (HTML output) |
14565 | 106 |
"_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) |
11687 | 107 |
Not :: "bool => bool" ("\<not> _" [40] 40) |
14565 | 108 |
"op &" :: "[bool, bool] => bool" (infixr "\<and>" 35) |
109 |
"op |" :: "[bool, bool] => bool" (infixr "\<or>" 30) |
|
110 |
"_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) |
|
111 |
"ALL " :: "[idts, bool] => bool" ("(3\<forall>_./ _)" [0, 10] 10) |
|
112 |
"EX " :: "[idts, bool] => bool" ("(3\<exists>_./ _)" [0, 10] 10) |
|
113 |
"EX! " :: "[idts, bool] => bool" ("(3\<exists>!_./ _)" [0, 10] 10) |
|
6340 | 114 |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
115 |
syntax (HOL) |
7357 | 116 |
"ALL " :: "[idts, bool] => bool" ("(3! _./ _)" [0, 10] 10) |
117 |
"EX " :: "[idts, bool] => bool" ("(3? _./ _)" [0, 10] 10) |
|
118 |
"EX! " :: "[idts, bool] => bool" ("(3?! _./ _)" [0, 10] 10) |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
119 |
|
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
120 |
|
11750 | 121 |
subsubsection {* Axioms and basic definitions *} |
2260 | 122 |
|
7357 | 123 |
axioms |
124 |
eq_reflection: "(x=y) ==> (x==y)" |
|
923 | 125 |
|
7357 | 126 |
refl: "t = (t::'a)" |
127 |
subst: "[| s = t; P(s) |] ==> P(t::'a)" |
|
6289 | 128 |
|
7357 | 129 |
ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" |
11750 | 130 |
-- {* Extensionality is built into the meta-logic, and this rule expresses *} |
131 |
-- {* a related property. It is an eta-expanded version of the traditional *} |
|
132 |
-- {* rule, and similar to the ABS rule of HOL *} |
|
6289 | 133 |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
134 |
the_eq_trivial: "(THE x. x = a) = (a::'a)" |
923 | 135 |
|
7357 | 136 |
impI: "(P ==> Q) ==> P-->Q" |
137 |
mp: "[| P-->Q; P |] ==> Q" |
|
923 | 138 |
|
139 |
defs |
|
7357 | 140 |
True_def: "True == ((%x::bool. x) = (%x. x))" |
141 |
All_def: "All(P) == (P = (%x. True))" |
|
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
142 |
Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" |
7357 | 143 |
False_def: "False == (!P. P)" |
144 |
not_def: "~ P == P-->False" |
|
145 |
and_def: "P & Q == !R. (P-->Q-->R) --> R" |
|
146 |
or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" |
|
147 |
Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" |
|
923 | 148 |
|
7357 | 149 |
axioms |
150 |
iff: "(P-->Q) --> (Q-->P) --> (P=Q)" |
|
151 |
True_or_False: "(P=True) | (P=False)" |
|
923 | 152 |
|
153 |
defs |
|
7357 | 154 |
Let_def: "Let s f == f(s)" |
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
155 |
if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" |
5069 | 156 |
|
14223
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
157 |
finalconsts |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
158 |
"op =" |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
159 |
"op -->" |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
160 |
The |
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14208
diff
changeset
|
161 |
arbitrary |
3320 | 162 |
|
11750 | 163 |
subsubsection {* Generic algebraic operations *} |
4868 | 164 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
165 |
axclass zero < type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
166 |
axclass one < type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
167 |
axclass plus < type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
168 |
axclass minus < type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
169 |
axclass times < type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
170 |
axclass inverse < type |
11750 | 171 |
|
172 |
global |
|
173 |
||
174 |
consts |
|
175 |
"0" :: "'a::zero" ("0") |
|
176 |
"1" :: "'a::one" ("1") |
|
177 |
"+" :: "['a::plus, 'a] => 'a" (infixl 65) |
|
178 |
- :: "['a::minus, 'a] => 'a" (infixl 65) |
|
179 |
uminus :: "['a::minus] => 'a" ("- _" [81] 80) |
|
180 |
* :: "['a::times, 'a] => 'a" (infixl 70) |
|
181 |
||
13456
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
182 |
syntax |
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
183 |
"_index1" :: index ("\<^sub>1") |
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
wenzelm
parents:
13438
diff
changeset
|
184 |
translations |
14690 | 185 |
(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
|
186 |
|
11750 | 187 |
local |
188 |
||
189 |
typed_print_translation {* |
|
190 |
let |
|
191 |
fun tr' c = (c, fn show_sorts => fn T => fn ts => |
|
192 |
if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match |
|
193 |
else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); |
|
194 |
in [tr' "0", tr' "1"] end; |
|
195 |
*} -- {* show types that are presumably too general *} |
|
196 |
||
197 |
||
198 |
consts |
|
199 |
abs :: "'a::minus => 'a" |
|
200 |
inverse :: "'a::inverse => 'a" |
|
201 |
divide :: "['a::inverse, 'a] => 'a" (infixl "'/" 70) |
|
202 |
||
203 |
syntax (xsymbols) |
|
204 |
abs :: "'a::minus => 'a" ("\<bar>_\<bar>") |
|
205 |
syntax (HTML output) |
|
206 |
abs :: "'a::minus => 'a" ("\<bar>_\<bar>") |
|
207 |
||
208 |
||
209 |
subsection {* Theory and package setup *} |
|
210 |
||
211 |
subsubsection {* Basic lemmas *} |
|
4868 | 212 |
|
9736 | 213 |
use "HOL_lemmas.ML" |
11687 | 214 |
theorems case_split = case_split_thm [case_names True False] |
9869 | 215 |
|
12386 | 216 |
|
217 |
subsubsection {* Intuitionistic Reasoning *} |
|
218 |
||
219 |
lemma impE': |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
220 |
assumes 1: "P --> Q" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
221 |
and 2: "Q ==> R" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
222 |
and 3: "P --> Q ==> P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
223 |
shows R |
12386 | 224 |
proof - |
225 |
from 3 and 1 have P . |
|
226 |
with 1 have Q by (rule impE) |
|
227 |
with 2 show R . |
|
228 |
qed |
|
229 |
||
230 |
lemma allE': |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
231 |
assumes 1: "ALL x. P x" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
232 |
and 2: "P x ==> ALL x. P x ==> Q" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
233 |
shows Q |
12386 | 234 |
proof - |
235 |
from 1 have "P x" by (rule spec) |
|
236 |
from this and 1 show Q by (rule 2) |
|
237 |
qed |
|
238 |
||
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
239 |
lemma notE': |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
240 |
assumes 1: "~ P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
241 |
and 2: "~ P ==> P" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
242 |
shows R |
12386 | 243 |
proof - |
244 |
from 2 and 1 have P . |
|
245 |
with 1 show R by (rule notE) |
|
246 |
qed |
|
247 |
||
248 |
lemmas [CPure.elim!] = disjE iffE FalseE conjE exE |
|
249 |
and [CPure.intro!] = iffI conjI impI TrueI notI allI refl |
|
250 |
and [CPure.elim 2] = allE notE' impE' |
|
251 |
and [CPure.intro] = exI disjI2 disjI1 |
|
252 |
||
253 |
lemmas [trans] = trans |
|
254 |
and [sym] = sym not_sym |
|
255 |
and [CPure.elim?] = iffD1 iffD2 impE |
|
11750 | 256 |
|
11438
3d9222b80989
declare trans [trans] (*overridden in theory Calculation*);
wenzelm
parents:
11432
diff
changeset
|
257 |
|
11750 | 258 |
subsubsection {* Atomizing meta-level connectives *} |
259 |
||
260 |
lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" |
|
12003 | 261 |
proof |
9488 | 262 |
assume "!!x. P x" |
10383 | 263 |
show "ALL x. P x" by (rule allI) |
9488 | 264 |
next |
265 |
assume "ALL x. P x" |
|
10383 | 266 |
thus "!!x. P x" by (rule allE) |
9488 | 267 |
qed |
268 |
||
11750 | 269 |
lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" |
12003 | 270 |
proof |
9488 | 271 |
assume r: "A ==> B" |
10383 | 272 |
show "A --> B" by (rule impI) (rule r) |
9488 | 273 |
next |
274 |
assume "A --> B" and A |
|
10383 | 275 |
thus B by (rule mp) |
9488 | 276 |
qed |
277 |
||
14749 | 278 |
lemma atomize_not: "(A ==> False) == Trueprop (~A)" |
279 |
proof |
|
280 |
assume r: "A ==> False" |
|
281 |
show "~A" by (rule notI) (rule r) |
|
282 |
next |
|
283 |
assume "~A" and A |
|
284 |
thus False by (rule notE) |
|
285 |
qed |
|
286 |
||
11750 | 287 |
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" |
12003 | 288 |
proof |
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
289 |
assume "x == y" |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
290 |
show "x = y" by (unfold prems) (rule refl) |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
291 |
next |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
292 |
assume "x = y" |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
293 |
thus "x == y" by (rule eq_reflection) |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
294 |
qed |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
295 |
|
12023 | 296 |
lemma atomize_conj [atomize]: |
297 |
"(!!C. (A ==> B ==> PROP C) ==> PROP C) == Trueprop (A & B)" |
|
12003 | 298 |
proof |
11953 | 299 |
assume "!!C. (A ==> B ==> PROP C) ==> PROP C" |
300 |
show "A & B" by (rule conjI) |
|
301 |
next |
|
302 |
fix C |
|
303 |
assume "A & B" |
|
304 |
assume "A ==> B ==> PROP C" |
|
305 |
thus "PROP C" |
|
306 |
proof this |
|
307 |
show A by (rule conjunct1) |
|
308 |
show B by (rule conjunct2) |
|
309 |
qed |
|
310 |
qed |
|
311 |
||
12386 | 312 |
lemmas [symmetric, rulify] = atomize_all atomize_imp |
313 |
||
11750 | 314 |
|
315 |
subsubsection {* Classical Reasoner setup *} |
|
9529 | 316 |
|
10383 | 317 |
use "cladata.ML" |
318 |
setup hypsubst_setup |
|
11977 | 319 |
|
12386 | 320 |
ML_setup {* |
321 |
Context.>> (ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)); |
|
322 |
*} |
|
11977 | 323 |
|
10383 | 324 |
setup Classical.setup |
325 |
setup clasetup |
|
326 |
||
12386 | 327 |
lemmas [intro?] = ext |
328 |
and [elim?] = ex1_implies_ex |
|
11977 | 329 |
|
9869 | 330 |
use "blastdata.ML" |
331 |
setup Blast.setup |
|
4868 | 332 |
|
11750 | 333 |
|
334 |
subsubsection {* Simplifier setup *} |
|
335 |
||
12281 | 336 |
lemma meta_eq_to_obj_eq: "x == y ==> x = y" |
337 |
proof - |
|
338 |
assume r: "x == y" |
|
339 |
show "x = y" by (unfold r) (rule refl) |
|
340 |
qed |
|
341 |
||
342 |
lemma eta_contract_eq: "(%s. f s) = f" .. |
|
343 |
||
344 |
lemma simp_thms: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
345 |
shows not_not: "(~ ~ P) = P" |
15354 | 346 |
and Not_eq_iff: "((~P) = (~Q)) = (P = Q)" |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
347 |
and |
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
348 |
"(P ~= Q) = (P = (~Q))" |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
349 |
"(P | ~P) = True" "(~P | P) = True" |
12281 | 350 |
"(x = x) = True" |
351 |
"(~True) = False" "(~False) = True" |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
352 |
"(~P) ~= P" "P ~= (~P)" |
12281 | 353 |
"(True=P) = P" "(P=True) = P" "(False=P) = (~P)" "(P=False) = (~P)" |
354 |
"(True --> P) = P" "(False --> P) = True" |
|
355 |
"(P --> True) = True" "(P --> P) = True" |
|
356 |
"(P --> False) = (~P)" "(P --> ~P) = (~P)" |
|
357 |
"(P & True) = P" "(True & P) = P" |
|
358 |
"(P & False) = False" "(False & P) = False" |
|
359 |
"(P & P) = P" "(P & (P & Q)) = (P & Q)" |
|
360 |
"(P & ~P) = False" "(~P & P) = False" |
|
361 |
"(P | True) = True" "(True | P) = True" |
|
362 |
"(P | False) = P" "(False | P) = P" |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
363 |
"(P | P) = P" "(P | (P | Q)) = (P | Q)" and |
12281 | 364 |
"(ALL x. P) = P" "(EX x. P) = P" "EX x. x=t" "EX x. t=x" |
365 |
-- {* needed for the one-point-rule quantifier simplification procs *} |
|
366 |
-- {* essential for termination!! *} and |
|
367 |
"!!P. (EX x. x=t & P(x)) = P(t)" |
|
368 |
"!!P. (EX x. t=x & P(x)) = P(t)" |
|
369 |
"!!P. (ALL x. x=t --> P(x)) = P(t)" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
370 |
"!!P. (ALL x. t=x --> P(x)) = P(t)" |
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
371 |
by (blast, blast, blast, blast, blast, rules+) |
13421 | 372 |
|
12281 | 373 |
lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))" |
12354 | 374 |
by rules |
12281 | 375 |
|
376 |
lemma ex_simps: |
|
377 |
"!!P Q. (EX x. P x & Q) = ((EX x. P x) & Q)" |
|
378 |
"!!P Q. (EX x. P & Q x) = (P & (EX x. Q x))" |
|
379 |
"!!P Q. (EX x. P x | Q) = ((EX x. P x) | Q)" |
|
380 |
"!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))" |
|
381 |
"!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)" |
|
382 |
"!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))" |
|
383 |
-- {* Miniscoping: pushing in existential quantifiers. *} |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
384 |
by (rules | blast)+ |
12281 | 385 |
|
386 |
lemma all_simps: |
|
387 |
"!!P Q. (ALL x. P x & Q) = ((ALL x. P x) & Q)" |
|
388 |
"!!P Q. (ALL x. P & Q x) = (P & (ALL x. Q x))" |
|
389 |
"!!P Q. (ALL x. P x | Q) = ((ALL x. P x) | Q)" |
|
390 |
"!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))" |
|
391 |
"!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)" |
|
392 |
"!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))" |
|
393 |
-- {* Miniscoping: pushing in universal quantifiers. *} |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
394 |
by (rules | blast)+ |
12281 | 395 |
|
14201 | 396 |
lemma disj_absorb: "(A | A) = A" |
397 |
by blast |
|
398 |
||
399 |
lemma disj_left_absorb: "(A | (A | B)) = (A | B)" |
|
400 |
by blast |
|
401 |
||
402 |
lemma conj_absorb: "(A & A) = A" |
|
403 |
by blast |
|
404 |
||
405 |
lemma conj_left_absorb: "(A & (A & B)) = (A & B)" |
|
406 |
by blast |
|
407 |
||
12281 | 408 |
lemma eq_ac: |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
409 |
shows eq_commute: "(a=b) = (b=a)" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
410 |
and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
411 |
and eq_assoc: "((P=Q)=R) = (P=(Q=R))" by (rules, blast+) |
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
412 |
lemma neq_commute: "(a~=b) = (b~=a)" by rules |
12281 | 413 |
|
414 |
lemma conj_comms: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
415 |
shows conj_commute: "(P&Q) = (Q&P)" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
416 |
and conj_left_commute: "(P&(Q&R)) = (Q&(P&R))" by rules+ |
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
417 |
lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by rules |
12281 | 418 |
|
419 |
lemma disj_comms: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
420 |
shows disj_commute: "(P|Q) = (Q|P)" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
421 |
and disj_left_commute: "(P|(Q|R)) = (Q|(P|R))" by rules+ |
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
422 |
lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by rules |
12281 | 423 |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
424 |
lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by rules |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
425 |
lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by rules |
12281 | 426 |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
427 |
lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by rules |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
428 |
lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by rules |
12281 | 429 |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
430 |
lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by rules |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
431 |
lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by rules |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
432 |
lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by rules |
12281 | 433 |
|
434 |
text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *} |
|
435 |
lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast |
|
436 |
lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast |
|
437 |
||
438 |
lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast |
|
439 |
lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast |
|
440 |
||
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
441 |
lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by rules |
12281 | 442 |
lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast |
443 |
lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast |
|
444 |
lemma not_iff: "(P~=Q) = (P = (~Q))" by blast |
|
445 |
lemma disj_not1: "(~P | Q) = (P --> Q)" by blast |
|
446 |
lemma disj_not2: "(P | ~Q) = (Q --> P)" -- {* changes orientation :-( *} |
|
447 |
by blast |
|
448 |
lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast |
|
449 |
||
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
450 |
lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by rules |
12281 | 451 |
|
452 |
||
453 |
lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q" |
|
454 |
-- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *} |
|
455 |
-- {* cases boil down to the same thing. *} |
|
456 |
by blast |
|
457 |
||
458 |
lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast |
|
459 |
lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
460 |
lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by rules |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
461 |
lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by rules |
12281 | 462 |
|
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
463 |
lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by rules |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
464 |
lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by rules |
12281 | 465 |
|
466 |
text {* |
|
467 |
\medskip The @{text "&"} congruence rule: not included by default! |
|
468 |
May slow rewrite proofs down by as much as 50\% *} |
|
469 |
||
470 |
lemma conj_cong: |
|
471 |
"(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))" |
|
12354 | 472 |
by rules |
12281 | 473 |
|
474 |
lemma rev_conj_cong: |
|
475 |
"(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))" |
|
12354 | 476 |
by rules |
12281 | 477 |
|
478 |
text {* The @{text "|"} congruence rule: not included by default! *} |
|
479 |
||
480 |
lemma disj_cong: |
|
481 |
"(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))" |
|
482 |
by blast |
|
483 |
||
484 |
lemma eq_sym_conv: "(x = y) = (y = x)" |
|
12354 | 485 |
by rules |
12281 | 486 |
|
487 |
||
488 |
text {* \medskip if-then-else rules *} |
|
489 |
||
490 |
lemma if_True: "(if True then x else y) = x" |
|
491 |
by (unfold if_def) blast |
|
492 |
||
493 |
lemma if_False: "(if False then x else y) = y" |
|
494 |
by (unfold if_def) blast |
|
495 |
||
496 |
lemma if_P: "P ==> (if P then x else y) = x" |
|
497 |
by (unfold if_def) blast |
|
498 |
||
499 |
lemma if_not_P: "~P ==> (if P then x else y) = y" |
|
500 |
by (unfold if_def) blast |
|
501 |
||
502 |
lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" |
|
503 |
apply (rule case_split [of Q]) |
|
504 |
apply (subst if_P) |
|
14208 | 505 |
prefer 3 apply (subst if_not_P, blast+) |
12281 | 506 |
done |
507 |
||
508 |
lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" |
|
14208 | 509 |
by (subst split_if, blast) |
12281 | 510 |
|
511 |
lemmas if_splits = split_if split_if_asm |
|
512 |
||
513 |
lemma if_def2: "(if Q then x else y) = ((Q --> x) & (~ Q --> y))" |
|
514 |
by (rule split_if) |
|
515 |
||
516 |
lemma if_cancel: "(if c then x else x) = x" |
|
14208 | 517 |
by (subst split_if, blast) |
12281 | 518 |
|
519 |
lemma if_eq_cancel: "(if x = y then y else x) = x" |
|
14208 | 520 |
by (subst split_if, blast) |
12281 | 521 |
|
522 |
lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))" |
|
523 |
-- {* This form is useful for expanding @{text if}s on the RIGHT of the @{text "==>"} symbol. *} |
|
524 |
by (rule split_if) |
|
525 |
||
526 |
lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" |
|
527 |
-- {* And this form is useful for expanding @{text if}s on the LEFT. *} |
|
14208 | 528 |
apply (subst split_if, blast) |
12281 | 529 |
done |
530 |
||
12436
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
531 |
lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) rules |
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
berghofe
parents:
12386
diff
changeset
|
532 |
lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) rules |
12281 | 533 |
|
14201 | 534 |
subsubsection {* Actual Installation of the Simplifier *} |
535 |
||
9869 | 536 |
use "simpdata.ML" |
537 |
setup Simplifier.setup |
|
538 |
setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup |
|
539 |
setup Splitter.setup setup Clasimp.setup |
|
540 |
||
14201 | 541 |
declare disj_absorb [simp] conj_absorb [simp] |
542 |
||
13723 | 543 |
lemma ex1_eq[iff]: "EX! x. x = t" "EX! x. t = x" |
544 |
by blast+ |
|
545 |
||
13638 | 546 |
theorem choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))" |
547 |
apply (rule iffI) |
|
548 |
apply (rule_tac a = "%x. THE y. P x y" in ex1I) |
|
549 |
apply (fast dest!: theI') |
|
550 |
apply (fast intro: ext the1_equality [symmetric]) |
|
551 |
apply (erule ex1E) |
|
552 |
apply (rule allI) |
|
553 |
apply (rule ex1I) |
|
554 |
apply (erule spec) |
|
555 |
apply (erule_tac x = "%z. if z = x then y else f z" in allE) |
|
556 |
apply (erule impE) |
|
557 |
apply (rule allI) |
|
558 |
apply (rule_tac P = "xa = x" in case_split_thm) |
|
14208 | 559 |
apply (drule_tac [3] x = x in fun_cong, simp_all) |
13638 | 560 |
done |
561 |
||
13438
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
562 |
text{*Needs only HOL-lemmas:*} |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
563 |
lemma mk_left_commute: |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
564 |
assumes a: "\<And>x y z. f (f x y) z = f x (f y z)" and |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
565 |
c: "\<And>x y. f x y = f y x" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
566 |
shows "f x (f y z) = f y (f x z)" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
567 |
by(rule trans[OF trans[OF c a] arg_cong[OF c, of "f y"]]) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
568 |
|
11750 | 569 |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
570 |
subsubsection {* Generic cases and induction *} |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
571 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
572 |
constdefs |
11989 | 573 |
induct_forall :: "('a => bool) => bool" |
574 |
"induct_forall P == \<forall>x. P x" |
|
575 |
induct_implies :: "bool => bool => bool" |
|
576 |
"induct_implies A B == A --> B" |
|
577 |
induct_equal :: "'a => 'a => bool" |
|
578 |
"induct_equal x y == x = y" |
|
579 |
induct_conj :: "bool => bool => bool" |
|
580 |
"induct_conj A B == A & B" |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
581 |
|
11989 | 582 |
lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))" |
583 |
by (simp only: atomize_all induct_forall_def) |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
584 |
|
11989 | 585 |
lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)" |
586 |
by (simp only: atomize_imp induct_implies_def) |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
587 |
|
11989 | 588 |
lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)" |
589 |
by (simp only: atomize_eq induct_equal_def) |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
590 |
|
11989 | 591 |
lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = |
592 |
induct_conj (induct_forall A) (induct_forall B)" |
|
12354 | 593 |
by (unfold induct_forall_def induct_conj_def) rules |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
594 |
|
11989 | 595 |
lemma induct_implies_conj: "induct_implies C (induct_conj A B) = |
596 |
induct_conj (induct_implies C A) (induct_implies C B)" |
|
12354 | 597 |
by (unfold induct_implies_def induct_conj_def) rules |
11989 | 598 |
|
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
599 |
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
|
600 |
proof |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
601 |
assume r: "induct_conj A B ==> PROP C" and A B |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
602 |
show "PROP C" by (rule r) (simp! add: induct_conj_def) |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
603 |
next |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
604 |
assume r: "A ==> B ==> PROP C" and "induct_conj A B" |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
605 |
show "PROP C" by (rule r) (simp! add: induct_conj_def)+ |
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
606 |
qed |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
607 |
|
11989 | 608 |
lemma induct_impliesI: "(A ==> B) ==> induct_implies A B" |
609 |
by (simp add: induct_implies_def) |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
610 |
|
12161 | 611 |
lemmas induct_atomize = atomize_conj induct_forall_eq induct_implies_eq induct_equal_eq |
612 |
lemmas induct_rulify1 [symmetric, standard] = induct_forall_eq induct_implies_eq induct_equal_eq |
|
613 |
lemmas induct_rulify2 = induct_forall_def induct_implies_def induct_equal_def induct_conj_def |
|
11989 | 614 |
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
|
615 |
|
11989 | 616 |
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
|
617 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
618 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
619 |
text {* Method setup. *} |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
620 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
621 |
ML {* |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
622 |
structure InductMethod = InductMethodFun |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
623 |
(struct |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
624 |
val dest_concls = HOLogic.dest_concls; |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
625 |
val cases_default = thm "case_split"; |
11989 | 626 |
val local_impI = thm "induct_impliesI"; |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
627 |
val conjI = thm "conjI"; |
11989 | 628 |
val atomize = thms "induct_atomize"; |
629 |
val rulify1 = thms "induct_rulify1"; |
|
630 |
val rulify2 = thms "induct_rulify2"; |
|
12240 | 631 |
val localize = [Thm.symmetric (thm "induct_implies_def")]; |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
632 |
end); |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
633 |
*} |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
634 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
635 |
setup InductMethod.setup |
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
636 |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
637 |
|
11750 | 638 |
subsection {* Order signatures and orders *} |
639 |
||
640 |
axclass |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12281
diff
changeset
|
641 |
ord < type |
11750 | 642 |
|
643 |
syntax |
|
644 |
"op <" :: "['a::ord, 'a] => bool" ("op <") |
|
645 |
"op <=" :: "['a::ord, 'a] => bool" ("op <=") |
|
646 |
||
647 |
global |
|
648 |
||
649 |
consts |
|
650 |
"op <" :: "['a::ord, 'a] => bool" ("(_/ < _)" [50, 51] 50) |
|
651 |
"op <=" :: "['a::ord, 'a] => bool" ("(_/ <= _)" [50, 51] 50) |
|
652 |
||
653 |
local |
|
654 |
||
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
655 |
syntax (xsymbols) |
11750 | 656 |
"op <=" :: "['a::ord, 'a] => bool" ("op \<le>") |
657 |
"op <=" :: "['a::ord, 'a] => bool" ("(_/ \<le> _)" [50, 51] 50) |
|
658 |
||
14565 | 659 |
syntax (HTML output) |
660 |
"op <=" :: "['a::ord, 'a] => bool" ("op \<le>") |
|
661 |
"op <=" :: "['a::ord, 'a] => bool" ("(_/ \<le> _)" [50, 51] 50) |
|
662 |
||
15354 | 663 |
text{* Syntactic sugar: *} |
11750 | 664 |
|
15354 | 665 |
consts |
666 |
"_gt" :: "'a::ord => 'a => bool" (infixl ">" 50) |
|
667 |
"_ge" :: "'a::ord => 'a => bool" (infixl ">=" 50) |
|
668 |
translations |
|
669 |
"x > y" => "y < x" |
|
670 |
"x >= y" => "y <= x" |
|
671 |
||
672 |
syntax (xsymbols) |
|
673 |
"_ge" :: "'a::ord => 'a => bool" (infixl "\<ge>" 50) |
|
674 |
||
675 |
syntax (HTML output) |
|
676 |
"_ge" :: "['a::ord, 'a] => bool" (infixl "\<ge>" 50) |
|
677 |
||
14295 | 678 |
|
11750 | 679 |
subsubsection {* Monotonicity *} |
680 |
||
13412 | 681 |
locale mono = |
682 |
fixes f |
|
683 |
assumes mono: "A <= B ==> f A <= f B" |
|
11750 | 684 |
|
13421 | 685 |
lemmas monoI [intro?] = mono.intro |
13412 | 686 |
and monoD [dest?] = mono.mono |
11750 | 687 |
|
688 |
constdefs |
|
689 |
min :: "['a::ord, 'a] => 'a" |
|
690 |
"min a b == (if a <= b then a else b)" |
|
691 |
max :: "['a::ord, 'a] => 'a" |
|
692 |
"max a b == (if a <= b then b else a)" |
|
693 |
||
694 |
lemma min_leastL: "(!!x. least <= x) ==> min least x = least" |
|
695 |
by (simp add: min_def) |
|
696 |
||
697 |
lemma min_of_mono: |
|
698 |
"ALL x y. (f x <= f y) = (x <= y) ==> min (f m) (f n) = f (min m n)" |
|
699 |
by (simp add: min_def) |
|
700 |
||
701 |
lemma max_leastL: "(!!x. least <= x) ==> max least x = x" |
|
702 |
by (simp add: max_def) |
|
703 |
||
704 |
lemma max_of_mono: |
|
705 |
"ALL x y. (f x <= f y) = (x <= y) ==> max (f m) (f n) = f (max m n)" |
|
706 |
by (simp add: max_def) |
|
707 |
||
708 |
||
709 |
subsubsection "Orders" |
|
710 |
||
711 |
axclass order < ord |
|
712 |
order_refl [iff]: "x <= x" |
|
713 |
order_trans: "x <= y ==> y <= z ==> x <= z" |
|
714 |
order_antisym: "x <= y ==> y <= x ==> x = y" |
|
715 |
order_less_le: "(x < y) = (x <= y & x ~= y)" |
|
716 |
||
717 |
||
718 |
text {* Reflexivity. *} |
|
719 |
||
720 |
lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y" |
|
721 |
-- {* This form is useful with the classical reasoner. *} |
|
722 |
apply (erule ssubst) |
|
723 |
apply (rule order_refl) |
|
724 |
done |
|
725 |
||
13553 | 726 |
lemma order_less_irrefl [iff]: "~ x < (x::'a::order)" |
11750 | 727 |
by (simp add: order_less_le) |
728 |
||
729 |
lemma order_le_less: "((x::'a::order) <= y) = (x < y | x = y)" |
|
730 |
-- {* NOT suitable for iff, since it can cause PROOF FAILED. *} |
|
14208 | 731 |
apply (simp add: order_less_le, blast) |
11750 | 732 |
done |
733 |
||
734 |
lemmas order_le_imp_less_or_eq = order_le_less [THEN iffD1, standard] |
|
735 |
||
736 |
lemma order_less_imp_le: "!!x::'a::order. x < y ==> x <= y" |
|
737 |
by (simp add: order_less_le) |
|
738 |
||
739 |
||
740 |
text {* Asymmetry. *} |
|
741 |
||
742 |
lemma order_less_not_sym: "(x::'a::order) < y ==> ~ (y < x)" |
|
743 |
by (simp add: order_less_le order_antisym) |
|
744 |
||
745 |
lemma order_less_asym: "x < (y::'a::order) ==> (~P ==> y < x) ==> P" |
|
746 |
apply (drule order_less_not_sym) |
|
14208 | 747 |
apply (erule contrapos_np, simp) |
11750 | 748 |
done |
749 |
||
14295 | 750 |
lemma order_eq_iff: "!!x::'a::order. (x = y) = (x \<le> y & y \<le> x)" |
751 |
by (blast intro: order_antisym) |
|
752 |
||
15197 | 753 |
lemma order_antisym_conv: "(y::'a::order) <= x ==> (x <= y) = (x = y)" |
754 |
by(blast intro:order_antisym) |
|
11750 | 755 |
|
756 |
text {* Transitivity. *} |
|
757 |
||
758 |
lemma order_less_trans: "!!x::'a::order. [| x < y; y < z |] ==> x < z" |
|
759 |
apply (simp add: order_less_le) |
|
760 |
apply (blast intro: order_trans order_antisym) |
|
761 |
done |
|
762 |
||
763 |
lemma order_le_less_trans: "!!x::'a::order. [| x <= y; y < z |] ==> x < z" |
|
764 |
apply (simp add: order_less_le) |
|
765 |
apply (blast intro: order_trans order_antisym) |
|
766 |
done |
|
767 |
||
768 |
lemma order_less_le_trans: "!!x::'a::order. [| x < y; y <= z |] ==> x < z" |
|
769 |
apply (simp add: order_less_le) |
|
770 |
apply (blast intro: order_trans order_antisym) |
|
771 |
done |
|
772 |
||
773 |
||
774 |
text {* Useful for simplification, but too risky to include by default. *} |
|
775 |
||
776 |
lemma order_less_imp_not_less: "(x::'a::order) < y ==> (~ y < x) = True" |
|
777 |
by (blast elim: order_less_asym) |
|
778 |
||
779 |
lemma order_less_imp_triv: "(x::'a::order) < y ==> (y < x --> P) = True" |
|
780 |
by (blast elim: order_less_asym) |
|
781 |
||
782 |
lemma order_less_imp_not_eq: "(x::'a::order) < y ==> (x = y) = False" |
|
783 |
by auto |
|
784 |
||
785 |
lemma order_less_imp_not_eq2: "(x::'a::order) < y ==> (y = x) = False" |
|
786 |
by auto |
|
787 |
||
788 |
||
789 |
text {* Other operators. *} |
|
790 |
||
791 |
lemma min_leastR: "(!!x::'a::order. least <= x) ==> min x least = least" |
|
792 |
apply (simp add: min_def) |
|
793 |
apply (blast intro: order_antisym) |
|
794 |
done |
|
795 |
||
796 |
lemma max_leastR: "(!!x::'a::order. least <= x) ==> max x least = x" |
|
797 |
apply (simp add: max_def) |
|
798 |
apply (blast intro: order_antisym) |
|
799 |
done |
|
800 |
||
801 |
||
802 |
subsubsection {* Least value operator *} |
|
803 |
||
804 |
constdefs |
|
805 |
Least :: "('a::ord => bool) => 'a" (binder "LEAST " 10) |
|
806 |
"Least P == THE x. P x & (ALL y. P y --> x <= y)" |
|
807 |
-- {* We can no longer use LeastM because the latter requires Hilbert-AC. *} |
|
808 |
||
809 |
lemma LeastI2: |
|
810 |
"[| P (x::'a::order); |
|
811 |
!!y. P y ==> x <= y; |
|
812 |
!!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |] |
|
12281 | 813 |
==> Q (Least P)" |
11750 | 814 |
apply (unfold Least_def) |
815 |
apply (rule theI2) |
|
816 |
apply (blast intro: order_antisym)+ |
|
817 |
done |
|
818 |
||
819 |
lemma Least_equality: |
|
12281 | 820 |
"[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k" |
11750 | 821 |
apply (simp add: Least_def) |
822 |
apply (rule the_equality) |
|
823 |
apply (auto intro!: order_antisym) |
|
824 |
done |
|
825 |
||
826 |
||
827 |
subsubsection "Linear / total orders" |
|
828 |
||
829 |
axclass linorder < order |
|
830 |
linorder_linear: "x <= y | y <= x" |
|
831 |
||
832 |
lemma linorder_less_linear: "!!x::'a::linorder. x<y | x=y | y<x" |
|
833 |
apply (simp add: order_less_le) |
|
14208 | 834 |
apply (insert linorder_linear, blast) |
11750 | 835 |
done |
836 |
||
15079
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents:
14981
diff
changeset
|
837 |
lemma linorder_le_less_linear: "!!x::'a::linorder. x\<le>y | y<x" |
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents:
14981
diff
changeset
|
838 |
by (simp add: order_le_less linorder_less_linear) |
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents:
14981
diff
changeset
|
839 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
840 |
lemma linorder_le_cases [case_names le ge]: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
841 |
"((x::'a::linorder) \<le> y ==> P) ==> (y \<le> x ==> P) ==> P" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
842 |
by (insert linorder_linear, blast) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
843 |
|
11750 | 844 |
lemma linorder_cases [case_names less equal greater]: |
845 |
"((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
846 |
by (insert linorder_less_linear, blast) |
11750 | 847 |
|
848 |
lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)" |
|
849 |
apply (simp add: order_less_le) |
|
850 |
apply (insert linorder_linear) |
|
851 |
apply (blast intro: order_antisym) |
|
852 |
done |
|
853 |
||
854 |
lemma linorder_not_le: "!!x::'a::linorder. (~ x <= y) = (y < x)" |
|
855 |
apply (simp add: order_less_le) |
|
856 |
apply (insert linorder_linear) |
|
857 |
apply (blast intro: order_antisym) |
|
858 |
done |
|
859 |
||
860 |
lemma linorder_neq_iff: "!!x::'a::linorder. (x ~= y) = (x<y | y<x)" |
|
14208 | 861 |
by (cut_tac x = x and y = y in linorder_less_linear, auto) |
11750 | 862 |
|
863 |
lemma linorder_neqE: "x ~= (y::'a::linorder) ==> (x < y ==> R) ==> (y < x ==> R) ==> R" |
|
14208 | 864 |
by (simp add: linorder_neq_iff, blast) |
11750 | 865 |
|
15197 | 866 |
lemma linorder_antisym_conv1: "~ (x::'a::linorder) < y ==> (x <= y) = (x = y)" |
867 |
by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1]) |
|
868 |
||
869 |
lemma linorder_antisym_conv2: "(x::'a::linorder) <= y ==> (~ x < y) = (x = y)" |
|
870 |
by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1]) |
|
871 |
||
872 |
lemma linorder_antisym_conv3: "~ (y::'a::linorder) < x ==> (~ x < y) = (x = y)" |
|
873 |
by(blast intro:order_antisym dest:linorder_not_less[THEN iffD1]) |
|
874 |
||
875 |
use "antisym_setup.ML"; |
|
876 |
setup antisym_setup |
|
11750 | 877 |
|
878 |
subsubsection "Min and max on (linear) orders" |
|
879 |
||
880 |
lemma min_same [simp]: "min (x::'a::order) x = x" |
|
881 |
by (simp add: min_def) |
|
882 |
||
883 |
lemma max_same [simp]: "max (x::'a::order) x = x" |
|
884 |
by (simp add: max_def) |
|
885 |
||
886 |
lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" |
|
887 |
apply (simp add: max_def) |
|
888 |
apply (insert linorder_linear) |
|
889 |
apply (blast intro: order_trans) |
|
890 |
done |
|
891 |
||
892 |
lemma le_maxI1: "(x::'a::linorder) <= max x y" |
|
893 |
by (simp add: le_max_iff_disj) |
|
894 |
||
895 |
lemma le_maxI2: "(y::'a::linorder) <= max x y" |
|
896 |
-- {* CANNOT use with @{text "[intro!]"} because blast will give PROOF FAILED. *} |
|
897 |
by (simp add: le_max_iff_disj) |
|
898 |
||
899 |
lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" |
|
900 |
apply (simp add: max_def order_le_less) |
|
901 |
apply (insert linorder_less_linear) |
|
902 |
apply (blast intro: order_less_trans) |
|
903 |
done |
|
904 |
||
905 |
lemma max_le_iff_conj [simp]: |
|
906 |
"!!z::'a::linorder. (max x y <= z) = (x <= z & y <= z)" |
|
907 |
apply (simp add: max_def) |
|
908 |
apply (insert linorder_linear) |
|
909 |
apply (blast intro: order_trans) |
|
910 |
done |
|
911 |
||
912 |
lemma max_less_iff_conj [simp]: |
|
913 |
"!!z::'a::linorder. (max x y < z) = (x < z & y < z)" |
|
914 |
apply (simp add: order_le_less max_def) |
|
915 |
apply (insert linorder_less_linear) |
|
916 |
apply (blast intro: order_less_trans) |
|
917 |
done |
|
918 |
||
919 |
lemma le_min_iff_conj [simp]: |
|
920 |
"!!z::'a::linorder. (z <= min x y) = (z <= x & z <= y)" |
|
12892 | 921 |
-- {* @{text "[iff]"} screws up a @{text blast} in MiniML *} |
11750 | 922 |
apply (simp add: min_def) |
923 |
apply (insert linorder_linear) |
|
924 |
apply (blast intro: order_trans) |
|
925 |
done |
|
926 |
||
927 |
lemma min_less_iff_conj [simp]: |
|
928 |
"!!z::'a::linorder. (z < min x y) = (z < x & z < y)" |
|
929 |
apply (simp add: order_le_less min_def) |
|
930 |
apply (insert linorder_less_linear) |
|
931 |
apply (blast intro: order_less_trans) |
|
932 |
done |
|
933 |
||
934 |
lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" |
|
935 |
apply (simp add: min_def) |
|
936 |
apply (insert linorder_linear) |
|
937 |
apply (blast intro: order_trans) |
|
938 |
done |
|
939 |
||
940 |
lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" |
|
941 |
apply (simp add: min_def order_le_less) |
|
942 |
apply (insert linorder_less_linear) |
|
943 |
apply (blast intro: order_less_trans) |
|
944 |
done |
|
945 |
||
13438
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
946 |
lemma max_assoc: "!!x::'a::linorder. max (max x y) z = max x (max y z)" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
947 |
apply(simp add:max_def) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
948 |
apply(rule conjI) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
949 |
apply(blast intro:order_trans) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
950 |
apply(simp add:linorder_not_le) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
951 |
apply(blast dest: order_less_trans order_le_less_trans) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
952 |
done |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
953 |
|
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
954 |
lemma max_commute: "!!x::'a::linorder. max x y = max y x" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
955 |
apply(simp add:max_def) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
956 |
apply(simp add:linorder_not_le) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
957 |
apply(blast dest: order_less_trans) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
958 |
done |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
959 |
|
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
960 |
lemmas max_ac = max_assoc max_commute |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
961 |
mk_left_commute[of max,OF max_assoc max_commute] |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
962 |
|
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
963 |
lemma min_assoc: "!!x::'a::linorder. min (min x y) z = min x (min y z)" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
964 |
apply(simp add:min_def) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
965 |
apply(rule conjI) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
966 |
apply(blast intro:order_trans) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
967 |
apply(simp add:linorder_not_le) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
968 |
apply(blast dest: order_less_trans order_le_less_trans) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
969 |
done |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
970 |
|
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
971 |
lemma min_commute: "!!x::'a::linorder. min x y = min y x" |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
972 |
apply(simp add:min_def) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
973 |
apply(simp add:linorder_not_le) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
974 |
apply(blast dest: order_less_trans) |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
975 |
done |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
976 |
|
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
977 |
lemmas min_ac = min_assoc min_commute |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
978 |
mk_left_commute[of min,OF min_assoc min_commute] |
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
nipkow
parents:
13421
diff
changeset
|
979 |
|
11750 | 980 |
lemma split_min: |
981 |
"P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" |
|
982 |
by (simp add: min_def) |
|
983 |
||
984 |
lemma split_max: |
|
985 |
"P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" |
|
986 |
by (simp add: max_def) |
|
987 |
||
988 |
||
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
989 |
subsubsection {* Transitivity rules for calculational reasoning *} |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
990 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
991 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
992 |
lemma order_neq_le_trans: "a ~= b ==> (a::'a::order) <= b ==> a < b" |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
993 |
by (simp add: order_less_le) |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
994 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
995 |
lemma order_le_neq_trans: "(a::'a::order) <= b ==> a ~= b ==> a < b" |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
996 |
by (simp add: order_less_le) |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
997 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
998 |
lemma order_less_asym': "(a::'a::order) < b ==> b < a ==> P" |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
999 |
by (rule order_less_asym) |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1000 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1001 |
|
14444
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1002 |
subsubsection {* Setup of transitivity reasoner as Solver *} |
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1003 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1004 |
lemma less_imp_neq: "[| (x::'a::order) < y |] ==> x ~= y" |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1005 |
by (erule contrapos_pn, erule subst, rule order_less_irrefl) |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1006 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1007 |
lemma eq_neq_eq_imp_neq: "[| x = a ; a ~= b; b = y |] ==> x ~= y" |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1008 |
by (erule subst, erule ssubst, assumption) |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1009 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1010 |
ML_setup {* |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1011 |
|
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1012 |
(* The setting up of Quasi_Tac serves as a demo. Since there is no |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1013 |
class for quasi orders, the tactics Quasi_Tac.trans_tac and |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1014 |
Quasi_Tac.quasi_tac are not of much use. *) |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1015 |
|
15288 | 1016 |
fun decomp_gen sort sign (Trueprop $ t) = |
1017 |
let fun of_sort t = Sign.of_sort sign (type_of t, sort) |
|
1018 |
fun dec (Const ("Not", _) $ t) = ( |
|
1019 |
case dec t of |
|
1020 |
None => None |
|
1021 |
| Some (t1, rel, t2) => Some (t1, "~" ^ rel, t2)) |
|
1022 |
| dec (Const ("op =", _) $ t1 $ t2) = |
|
1023 |
if of_sort t1 |
|
1024 |
then Some (t1, "=", t2) |
|
1025 |
else None |
|
1026 |
| dec (Const ("op <=", _) $ t1 $ t2) = |
|
1027 |
if of_sort t1 |
|
1028 |
then Some (t1, "<=", t2) |
|
1029 |
else None |
|
1030 |
| dec (Const ("op <", _) $ t1 $ t2) = |
|
1031 |
if of_sort t1 |
|
1032 |
then Some (t1, "<", t2) |
|
1033 |
else None |
|
1034 |
| dec _ = None |
|
1035 |
in dec t end; |
|
1036 |
||
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1037 |
structure Quasi_Tac = Quasi_Tac_Fun ( |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1038 |
struct |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1039 |
val le_trans = thm "order_trans"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1040 |
val le_refl = thm "order_refl"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1041 |
val eqD1 = thm "order_eq_refl"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1042 |
val eqD2 = thm "sym" RS thm "order_eq_refl"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1043 |
val less_reflE = thm "order_less_irrefl" RS thm "notE"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1044 |
val less_imp_le = thm "order_less_imp_le"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1045 |
val le_neq_trans = thm "order_le_neq_trans"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1046 |
val neq_le_trans = thm "order_neq_le_trans"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1047 |
val less_imp_neq = thm "less_imp_neq"; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1048 |
val decomp_trans = decomp_gen ["HOL.order"]; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1049 |
val decomp_quasi = decomp_gen ["HOL.order"]; |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1050 |
|
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1051 |
end); (* struct *) |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1052 |
|
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1053 |
structure Order_Tac = Order_Tac_Fun ( |
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1054 |
struct |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1055 |
val less_reflE = thm "order_less_irrefl" RS thm "notE"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1056 |
val le_refl = thm "order_refl"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1057 |
val less_imp_le = thm "order_less_imp_le"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1058 |
val not_lessI = thm "linorder_not_less" RS thm "iffD2"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1059 |
val not_leI = thm "linorder_not_le" RS thm "iffD2"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1060 |
val not_lessD = thm "linorder_not_less" RS thm "iffD1"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1061 |
val not_leD = thm "linorder_not_le" RS thm "iffD1"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1062 |
val eqI = thm "order_antisym"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1063 |
val eqD1 = thm "order_eq_refl"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1064 |
val eqD2 = thm "sym" RS thm "order_eq_refl"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1065 |
val less_trans = thm "order_less_trans"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1066 |
val less_le_trans = thm "order_less_le_trans"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1067 |
val le_less_trans = thm "order_le_less_trans"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1068 |
val le_trans = thm "order_trans"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1069 |
val le_neq_trans = thm "order_le_neq_trans"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1070 |
val neq_le_trans = thm "order_neq_le_trans"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1071 |
val less_imp_neq = thm "less_imp_neq"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1072 |
val eq_neq_eq_imp_neq = thm "eq_neq_eq_imp_neq"; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1073 |
val decomp_part = decomp_gen ["HOL.order"]; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1074 |
val decomp_lin = decomp_gen ["HOL.linorder"]; |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1075 |
|
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1076 |
end); (* struct *) |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1077 |
|
14590 | 1078 |
simpset_ref() := simpset () |
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1079 |
addSolver (mk_solver "Trans_linear" (fn _ => Order_Tac.linear_tac)) |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1080 |
addSolver (mk_solver "Trans_partial" (fn _ => Order_Tac.partial_tac)); |
14444
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1081 |
(* Adding the transitivity reasoners also as safe solvers showed a slight |
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1082 |
speed up, but the reasoning strength appears to be not higher (at least |
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1083 |
no breaking of additional proofs in the entire HOL distribution, as |
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1084 |
of 5 March 2004, was observed). *) |
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1085 |
*} |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1086 |
|
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1087 |
(* Optional setup of methods *) |
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1088 |
|
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1089 |
(* |
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1090 |
method_setup trans_partial = |
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1091 |
{* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (Order_Tac.partial_tac)) *} |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1092 |
{* transitivity reasoner for partial orders *} |
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1093 |
method_setup trans_linear = |
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1094 |
{* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (Order_Tac.linear_tac)) *} |
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1095 |
{* transitivity reasoner for linear orders *} |
14398
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1096 |
*) |
c5c47703f763
Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents:
14365
diff
changeset
|
1097 |
|
14444
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1098 |
(* |
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1099 |
declare order.order_refl [simp del] order_less_irrefl [simp del] |
15103
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1100 |
|
79846e8792eb
New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
15079
diff
changeset
|
1101 |
can currently not be removed, abel_cancel relies on it. |
14444
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1102 |
*) |
24724afce166
Added documentation for transitivity solver setup.
ballarin
parents:
14430
diff
changeset
|
1103 |
|
11750 | 1104 |
subsubsection "Bounded quantifiers" |
1105 |
||
1106 |
syntax |
|
1107 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) |
|
1108 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) |
|
1109 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) |
|
1110 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) |
|
1111 |
||
15360 | 1112 |
"_gtAll" :: "[idt, 'a, bool] => bool" ("(3ALL _>_./ _)" [0, 0, 10] 10) |
1113 |
"_gtEx" :: "[idt, 'a, bool] => bool" ("(3EX _>_./ _)" [0, 0, 10] 10) |
|
1114 |
"_geAll" :: "[idt, 'a, bool] => bool" ("(3ALL _>=_./ _)" [0, 0, 10] 10) |
|
1115 |
"_geEx" :: "[idt, 'a, bool] => bool" ("(3EX _>=_./ _)" [0, 0, 10] 10) |
|
1116 |
||
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
1117 |
syntax (xsymbols) |
11750 | 1118 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_<_./ _)" [0, 0, 10] 10) |
1119 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_<_./ _)" [0, 0, 10] 10) |
|
1120 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10) |
|
1121 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10) |
|
1122 |
||
15360 | 1123 |
"_gtAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_>_./ _)" [0, 0, 10] 10) |
1124 |
"_gtEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_>_./ _)" [0, 0, 10] 10) |
|
1125 |
"_geAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10) |
|
1126 |
"_geEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10) |
|
1127 |
||
11750 | 1128 |
syntax (HOL) |
1129 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10) |
|
1130 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10) |
|
1131 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10) |
|
1132 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10) |
|
1133 |
||
14565 | 1134 |
syntax (HTML output) |
1135 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_<_./ _)" [0, 0, 10] 10) |
|
1136 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_<_./ _)" [0, 0, 10] 10) |
|
1137 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10) |
|
1138 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10) |
|
1139 |
||
15360 | 1140 |
"_gtAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_>_./ _)" [0, 0, 10] 10) |
1141 |
"_gtEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_>_./ _)" [0, 0, 10] 10) |
|
1142 |
"_geAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10) |
|
1143 |
"_geEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10) |
|
1144 |
||
11750 | 1145 |
translations |
1146 |
"ALL x<y. P" => "ALL x. x < y --> P" |
|
1147 |
"EX x<y. P" => "EX x. x < y & P" |
|
1148 |
"ALL x<=y. P" => "ALL x. x <= y --> P" |
|
1149 |
"EX x<=y. P" => "EX x. x <= y & P" |
|
15360 | 1150 |
"ALL x>y. P" => "ALL x. x > y --> P" |
1151 |
"EX x>y. P" => "EX x. x > y & P" |
|
1152 |
"ALL x>=y. P" => "ALL x. x >= y --> P" |
|
1153 |
"EX x>=y. P" => "EX x. x >= y & P" |
|
11750 | 1154 |
|
14357 | 1155 |
print_translation {* |
1156 |
let |
|
1157 |
fun all_tr' [Const ("_bound",_) $ Free (v,_), |
|
1158 |
Const("op -->",_) $ (Const ("op <",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
|
1159 |
(if v=v' then Syntax.const "_lessAll" $ Syntax.mark_bound v' $ n $ P else raise Match) |
|
1160 |
||
1161 |
| all_tr' [Const ("_bound",_) $ Free (v,_), |
|
1162 |
Const("op -->",_) $ (Const ("op <=",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
|
15362 | 1163 |
(if v=v' then Syntax.const "_leAll" $ Syntax.mark_bound v' $ n $ P else raise Match) |
1164 |
||
1165 |
| all_tr' [Const ("_bound",_) $ Free (v,_), |
|
1166 |
Const("op -->",_) $ (Const ("op <",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
|
1167 |
(if v=v' then Syntax.const "_gtAll" $ Syntax.mark_bound v' $ n $ P else raise Match) |
|
1168 |
||
1169 |
| all_tr' [Const ("_bound",_) $ Free (v,_), |
|
1170 |
Const("op -->",_) $ (Const ("op <=",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
|
1171 |
(if v=v' then Syntax.const "_geAll" $ Syntax.mark_bound v' $ n $ P else raise Match); |
|
14357 | 1172 |
|
1173 |
fun ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
1174 |
Const("op &",_) $ (Const ("op <",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
|
1175 |
(if v=v' then Syntax.const "_lessEx" $ Syntax.mark_bound v' $ n $ P else raise Match) |
|
1176 |
||
1177 |
| ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
1178 |
Const("op &",_) $ (Const ("op <=",_) $ (Const ("_bound",_) $ Free (v',_)) $ n ) $ P] = |
|
15362 | 1179 |
(if v=v' then Syntax.const "_leEx" $ Syntax.mark_bound v' $ n $ P else |
1180 |
raise Match) |
|
1181 |
||
1182 |
| ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
1183 |
Const("op &",_) $ (Const ("op <",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
|
1184 |
(if v=v' then Syntax.const "_gtEx" $ Syntax.mark_bound v' $ n $ P else raise Match) |
|
1185 |
||
1186 |
| ex_tr' [Const ("_bound",_) $ Free (v,_), |
|
1187 |
Const("op &",_) $ (Const ("op <=",_) $ n $ (Const ("_bound",_) $ Free (v',_))) $ P] = |
|
1188 |
(if v=v' then Syntax.const "_geEx" $ Syntax.mark_bound v' $ n $ P else raise Match) |
|
14357 | 1189 |
in |
1190 |
[("ALL ", all_tr'), ("EX ", ex_tr')] |
|
923 | 1191 |
end |
14357 | 1192 |
*} |
1193 |
||
1194 |
end |