| author | nipkow |
| Thu, 29 Nov 2001 13:33:45 +0100 | |
| changeset 12327 | 5a4d78204492 |
| parent 12281 | 3bd113b8f7a6 |
| child 12338 | de0f4a63baa5 |
| 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 |
|
| 7357 | 8 |
theory HOL = CPure |
|
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
9 |
files ("HOL_lemmas.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML"):
|
| 923 | 10 |
|
| 2260 | 11 |
|
| 11750 | 12 |
subsection {* Primitive logic *}
|
13 |
||
14 |
subsubsection {* Core syntax *}
|
|
| 2260 | 15 |
|
| 3947 | 16 |
global |
17 |
||
| 7357 | 18 |
classes "term" < logic |
19 |
defaultsort "term" |
|
| 923 | 20 |
|
| 7357 | 21 |
typedecl bool |
| 923 | 22 |
|
23 |
arities |
|
| 7357 | 24 |
bool :: "term" |
25 |
fun :: ("term", "term") "term"
|
|
| 923 | 26 |
|
| 11750 | 27 |
judgment |
28 |
Trueprop :: "bool => prop" ("(_)" 5)
|
|
| 923 | 29 |
|
| 11750 | 30 |
consts |
| 7357 | 31 |
Not :: "bool => bool" ("~ _" [40] 40)
|
32 |
True :: bool |
|
33 |
False :: bool |
|
34 |
If :: "[bool, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10)
|
|
| 3947 | 35 |
arbitrary :: 'a |
| 923 | 36 |
|
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
37 |
The :: "('a => bool) => 'a"
|
| 7357 | 38 |
All :: "('a => bool) => bool" (binder "ALL " 10)
|
39 |
Ex :: "('a => bool) => bool" (binder "EX " 10)
|
|
40 |
Ex1 :: "('a => bool) => bool" (binder "EX! " 10)
|
|
41 |
Let :: "['a, 'a => 'b] => 'b" |
|
| 923 | 42 |
|
| 7357 | 43 |
"=" :: "['a, 'a] => bool" (infixl 50) |
44 |
& :: "[bool, bool] => bool" (infixr 35) |
|
45 |
"|" :: "[bool, bool] => bool" (infixr 30) |
|
46 |
--> :: "[bool, bool] => bool" (infixr 25) |
|
| 923 | 47 |
|
|
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
48 |
local |
|
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
49 |
|
| 2260 | 50 |
|
| 11750 | 51 |
subsubsection {* Additional concrete syntax *}
|
| 2260 | 52 |
|
| 4868 | 53 |
nonterminals |
| 923 | 54 |
letbinds letbind |
55 |
case_syn cases_syn |
|
56 |
||
57 |
syntax |
|
| 7357 | 58 |
~= :: "['a, 'a] => bool" (infixl 50) |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
59 |
"_The" :: "[pttrn, bool] => 'a" ("(3THE _./ _)" [0, 10] 10)
|
| 923 | 60 |
|
| 7357 | 61 |
"_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10)
|
62 |
"" :: "letbind => letbinds" ("_")
|
|
63 |
"_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _")
|
|
64 |
"_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10)
|
|
| 923 | 65 |
|
|
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
66 |
"_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
|
67 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10)
|
| 7357 | 68 |
"" :: "case_syn => cases_syn" ("_")
|
|
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
69 |
"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _")
|
| 923 | 70 |
|
71 |
translations |
|
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
72 |
"x ~= y" == "~ (x = y)" |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
73 |
"THE x. P" == "The (%x. P)" |
| 923 | 74 |
"_Let (_binds b bs) e" == "_Let b (_Let bs e)" |
| 1114 | 75 |
"let x = a in e" == "Let a (%x. e)" |
| 923 | 76 |
|
| 3820 | 77 |
syntax ("" output)
|
| 11687 | 78 |
"=" :: "['a, 'a] => bool" (infix 50) |
79 |
"~=" :: "['a, 'a] => bool" (infix 50) |
|
| 2260 | 80 |
|
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
81 |
syntax (xsymbols) |
| 11687 | 82 |
Not :: "bool => bool" ("\<not> _" [40] 40)
|
83 |
"op &" :: "[bool, bool] => bool" (infixr "\<and>" 35) |
|
84 |
"op |" :: "[bool, bool] => bool" (infixr "\<or>" 30) |
|
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
85 |
"op -->" :: "[bool, bool] => bool" (infixr "\<longrightarrow>" 25) |
| 11687 | 86 |
"op ~=" :: "['a, 'a] => bool" (infix "\<noteq>" 50) |
87 |
"ALL " :: "[idts, bool] => bool" ("(3\<forall>_./ _)" [0, 10] 10)
|
|
88 |
"EX " :: "[idts, bool] => bool" ("(3\<exists>_./ _)" [0, 10] 10)
|
|
89 |
"EX! " :: "[idts, bool] => bool" ("(3\<exists>!_./ _)" [0, 10] 10)
|
|
90 |
"_case1" :: "['a, 'b] => case_syn" ("(2_ \<Rightarrow>/ _)" 10)
|
|
|
9060
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents:
8959
diff
changeset
|
91 |
(*"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ \\<orelse> _")*)
|
| 2372 | 92 |
|
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
93 |
syntax (xsymbols output) |
| 11687 | 94 |
"op ~=" :: "['a, 'a] => bool" (infix "\<noteq>" 50) |
| 3820 | 95 |
|
| 6340 | 96 |
syntax (HTML output) |
| 11687 | 97 |
Not :: "bool => bool" ("\<not> _" [40] 40)
|
| 6340 | 98 |
|
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
99 |
syntax (HOL) |
| 7357 | 100 |
"ALL " :: "[idts, bool] => bool" ("(3! _./ _)" [0, 10] 10)
|
101 |
"EX " :: "[idts, bool] => bool" ("(3? _./ _)" [0, 10] 10)
|
|
102 |
"EX! " :: "[idts, bool] => bool" ("(3?! _./ _)" [0, 10] 10)
|
|
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
103 |
|
|
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
104 |
|
| 11750 | 105 |
subsubsection {* Axioms and basic definitions *}
|
| 2260 | 106 |
|
| 7357 | 107 |
axioms |
108 |
eq_reflection: "(x=y) ==> (x==y)" |
|
| 923 | 109 |
|
| 7357 | 110 |
refl: "t = (t::'a)" |
111 |
subst: "[| s = t; P(s) |] ==> P(t::'a)" |
|
| 6289 | 112 |
|
| 7357 | 113 |
ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" |
| 11750 | 114 |
-- {* Extensionality is built into the meta-logic, and this rule expresses *}
|
115 |
-- {* a related property. It is an eta-expanded version of the traditional *}
|
|
116 |
-- {* rule, and similar to the ABS rule of HOL *}
|
|
| 6289 | 117 |
|
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
118 |
the_eq_trivial: "(THE x. x = a) = (a::'a)" |
| 923 | 119 |
|
| 7357 | 120 |
impI: "(P ==> Q) ==> P-->Q" |
121 |
mp: "[| P-->Q; P |] ==> Q" |
|
| 923 | 122 |
|
123 |
defs |
|
| 7357 | 124 |
True_def: "True == ((%x::bool. x) = (%x. x))" |
125 |
All_def: "All(P) == (P = (%x. True))" |
|
|
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
126 |
Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" |
| 7357 | 127 |
False_def: "False == (!P. P)" |
128 |
not_def: "~ P == P-->False" |
|
129 |
and_def: "P & Q == !R. (P-->Q-->R) --> R" |
|
130 |
or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" |
|
131 |
Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" |
|
| 923 | 132 |
|
| 7357 | 133 |
axioms |
134 |
iff: "(P-->Q) --> (Q-->P) --> (P=Q)" |
|
135 |
True_or_False: "(P=True) | (P=False)" |
|
| 923 | 136 |
|
137 |
defs |
|
| 7357 | 138 |
Let_def: "Let s f == f(s)" |
|
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
139 |
if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" |
| 5069 | 140 |
|
|
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11438
diff
changeset
|
141 |
arbitrary_def: "False ==> arbitrary == (THE x. False)" |
| 11750 | 142 |
-- {* @{term arbitrary} is completely unspecified, but is made to appear as a
|
143 |
definition syntactically *} |
|
| 923 | 144 |
|
| 3320 | 145 |
|
| 11750 | 146 |
subsubsection {* Generic algebraic operations *}
|
| 4868 | 147 |
|
| 11750 | 148 |
axclass zero < "term" |
149 |
axclass one < "term" |
|
150 |
axclass plus < "term" |
|
151 |
axclass minus < "term" |
|
152 |
axclass times < "term" |
|
153 |
axclass inverse < "term" |
|
154 |
||
155 |
global |
|
156 |
||
157 |
consts |
|
158 |
"0" :: "'a::zero" ("0")
|
|
159 |
"1" :: "'a::one" ("1")
|
|
160 |
"+" :: "['a::plus, 'a] => 'a" (infixl 65) |
|
161 |
- :: "['a::minus, 'a] => 'a" (infixl 65) |
|
162 |
uminus :: "['a::minus] => 'a" ("- _" [81] 80)
|
|
163 |
* :: "['a::times, 'a] => 'a" (infixl 70) |
|
164 |
||
165 |
local |
|
166 |
||
167 |
typed_print_translation {*
|
|
168 |
let |
|
169 |
fun tr' c = (c, fn show_sorts => fn T => fn ts => |
|
170 |
if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match |
|
171 |
else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); |
|
172 |
in [tr' "0", tr' "1"] end; |
|
173 |
*} -- {* show types that are presumably too general *}
|
|
174 |
||
175 |
||
176 |
consts |
|
177 |
abs :: "'a::minus => 'a" |
|
178 |
inverse :: "'a::inverse => 'a" |
|
179 |
divide :: "['a::inverse, 'a] => 'a" (infixl "'/" 70) |
|
180 |
||
181 |
syntax (xsymbols) |
|
182 |
abs :: "'a::minus => 'a" ("\<bar>_\<bar>")
|
|
183 |
syntax (HTML output) |
|
184 |
abs :: "'a::minus => 'a" ("\<bar>_\<bar>")
|
|
185 |
||
186 |
axclass plus_ac0 < plus, zero |
|
187 |
commute: "x + y = y + x" |
|
188 |
assoc: "(x + y) + z = x + (y + z)" |
|
189 |
zero: "0 + x = x" |
|
190 |
||
191 |
||
192 |
subsection {* Theory and package setup *}
|
|
193 |
||
194 |
subsubsection {* Basic lemmas *}
|
|
| 4868 | 195 |
|
| 9736 | 196 |
use "HOL_lemmas.ML" |
| 11687 | 197 |
theorems case_split = case_split_thm [case_names True False] |
| 9869 | 198 |
|
| 11750 | 199 |
declare trans [trans] |
200 |
declare impE [CPure.elim] iffD1 [CPure.elim] iffD2 [CPure.elim] |
|
201 |
||
|
11438
3d9222b80989
declare trans [trans] (*overridden in theory Calculation*);
wenzelm
parents:
11432
diff
changeset
|
202 |
|
| 11750 | 203 |
subsubsection {* Atomizing meta-level connectives *}
|
204 |
||
205 |
lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" |
|
| 12003 | 206 |
proof |
| 9488 | 207 |
assume "!!x. P x" |
| 10383 | 208 |
show "ALL x. P x" by (rule allI) |
| 9488 | 209 |
next |
210 |
assume "ALL x. P x" |
|
| 10383 | 211 |
thus "!!x. P x" by (rule allE) |
| 9488 | 212 |
qed |
213 |
||
| 11750 | 214 |
lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" |
| 12003 | 215 |
proof |
| 9488 | 216 |
assume r: "A ==> B" |
| 10383 | 217 |
show "A --> B" by (rule impI) (rule r) |
| 9488 | 218 |
next |
219 |
assume "A --> B" and A |
|
| 10383 | 220 |
thus B by (rule mp) |
| 9488 | 221 |
qed |
222 |
||
| 11750 | 223 |
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" |
| 12003 | 224 |
proof |
|
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
225 |
assume "x == y" |
|
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
226 |
show "x = y" by (unfold prems) (rule refl) |
|
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
227 |
next |
|
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
228 |
assume "x = y" |
|
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
229 |
thus "x == y" by (rule eq_reflection) |
|
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
230 |
qed |
|
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
231 |
|
| 12023 | 232 |
lemma atomize_conj [atomize]: |
233 |
"(!!C. (A ==> B ==> PROP C) ==> PROP C) == Trueprop (A & B)" |
|
| 12003 | 234 |
proof |
| 11953 | 235 |
assume "!!C. (A ==> B ==> PROP C) ==> PROP C" |
236 |
show "A & B" by (rule conjI) |
|
237 |
next |
|
238 |
fix C |
|
239 |
assume "A & B" |
|
240 |
assume "A ==> B ==> PROP C" |
|
241 |
thus "PROP C" |
|
242 |
proof this |
|
243 |
show A by (rule conjunct1) |
|
244 |
show B by (rule conjunct2) |
|
245 |
qed |
|
246 |
qed |
|
247 |
||
| 11750 | 248 |
|
249 |
subsubsection {* Classical Reasoner setup *}
|
|
| 9529 | 250 |
|
| 10383 | 251 |
use "cladata.ML" |
252 |
setup hypsubst_setup |
|
| 11977 | 253 |
|
| 11770 | 254 |
declare atomize_all [symmetric, rulify] atomize_imp [symmetric, rulify] |
| 11977 | 255 |
|
| 10383 | 256 |
setup Classical.setup |
257 |
setup clasetup |
|
258 |
||
| 11977 | 259 |
declare ext [intro?] |
260 |
declare disjI1 [elim?] disjI2 [elim?] ex1_implies_ex [elim?] sym [elim?] |
|
261 |
||
| 9869 | 262 |
use "blastdata.ML" |
263 |
setup Blast.setup |
|
| 4868 | 264 |
|
| 11750 | 265 |
|
266 |
subsubsection {* Simplifier setup *}
|
|
267 |
||
| 12281 | 268 |
lemma meta_eq_to_obj_eq: "x == y ==> x = y" |
269 |
proof - |
|
270 |
assume r: "x == y" |
|
271 |
show "x = y" by (unfold r) (rule refl) |
|
272 |
qed |
|
273 |
||
274 |
lemma eta_contract_eq: "(%s. f s) = f" .. |
|
275 |
||
276 |
lemma simp_thms: |
|
277 |
(not_not: "(~ ~ P) = P" and |
|
278 |
"(x = x) = True" |
|
279 |
"(~True) = False" "(~False) = True" |
|
280 |
"(~P) ~= P" "P ~= (~P)" "(P ~= Q) = (P = (~Q))" |
|
281 |
"(True=P) = P" "(P=True) = P" "(False=P) = (~P)" "(P=False) = (~P)" |
|
282 |
"(True --> P) = P" "(False --> P) = True" |
|
283 |
"(P --> True) = True" "(P --> P) = True" |
|
284 |
"(P --> False) = (~P)" "(P --> ~P) = (~P)" |
|
285 |
"(P & True) = P" "(True & P) = P" |
|
286 |
"(P & False) = False" "(False & P) = False" |
|
287 |
"(P & P) = P" "(P & (P & Q)) = (P & Q)" |
|
288 |
"(P & ~P) = False" "(~P & P) = False" |
|
289 |
"(P | True) = True" "(True | P) = True" |
|
290 |
"(P | False) = P" "(False | P) = P" |
|
291 |
"(P | P) = P" "(P | (P | Q)) = (P | Q)" |
|
292 |
"(P | ~P) = True" "(~P | P) = True" |
|
293 |
"((~P) = (~Q)) = (P=Q)" and |
|
294 |
"(ALL x. P) = P" "(EX x. P) = P" "EX x. x=t" "EX x. t=x" |
|
295 |
-- {* needed for the one-point-rule quantifier simplification procs *}
|
|
296 |
-- {* essential for termination!! *} and
|
|
297 |
"!!P. (EX x. x=t & P(x)) = P(t)" |
|
298 |
"!!P. (EX x. t=x & P(x)) = P(t)" |
|
299 |
"!!P. (ALL x. x=t --> P(x)) = P(t)" |
|
300 |
"!!P. (ALL x. t=x --> P(x)) = P(t)") |
|
301 |
by blast+ |
|
302 |
||
303 |
lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))" |
|
304 |
by blast |
|
305 |
||
306 |
lemma ex_simps: |
|
307 |
"!!P Q. (EX x. P x & Q) = ((EX x. P x) & Q)" |
|
308 |
"!!P Q. (EX x. P & Q x) = (P & (EX x. Q x))" |
|
309 |
"!!P Q. (EX x. P x | Q) = ((EX x. P x) | Q)" |
|
310 |
"!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))" |
|
311 |
"!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)" |
|
312 |
"!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))" |
|
313 |
-- {* Miniscoping: pushing in existential quantifiers. *}
|
|
314 |
by blast+ |
|
315 |
||
316 |
lemma all_simps: |
|
317 |
"!!P Q. (ALL x. P x & Q) = ((ALL x. P x) & Q)" |
|
318 |
"!!P Q. (ALL x. P & Q x) = (P & (ALL x. Q x))" |
|
319 |
"!!P Q. (ALL x. P x | Q) = ((ALL x. P x) | Q)" |
|
320 |
"!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))" |
|
321 |
"!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)" |
|
322 |
"!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))" |
|
323 |
-- {* Miniscoping: pushing in universal quantifiers. *}
|
|
324 |
by blast+ |
|
325 |
||
326 |
lemma eq_ac: |
|
327 |
(eq_commute: "(a=b) = (b=a)" and |
|
328 |
eq_left_commute: "(P=(Q=R)) = (Q=(P=R))" and |
|
329 |
eq_assoc: "((P=Q)=R) = (P=(Q=R))") by blast+ |
|
330 |
lemma neq_commute: "(a~=b) = (b~=a)" by blast |
|
331 |
||
332 |
lemma conj_comms: |
|
333 |
(conj_commute: "(P&Q) = (Q&P)" and |
|
334 |
conj_left_commute: "(P&(Q&R)) = (Q&(P&R))") by blast+ |
|
335 |
lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by blast |
|
336 |
||
337 |
lemma disj_comms: |
|
338 |
(disj_commute: "(P|Q) = (Q|P)" and |
|
339 |
disj_left_commute: "(P|(Q|R)) = (Q|(P|R))") by blast+ |
|
340 |
lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by blast |
|
341 |
||
342 |
lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by blast |
|
343 |
lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by blast |
|
344 |
||
345 |
lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by blast |
|
346 |
lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by blast |
|
347 |
||
348 |
lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by blast |
|
349 |
lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by blast |
|
350 |
lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by blast |
|
351 |
||
352 |
text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *}
|
|
353 |
lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast |
|
354 |
lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast |
|
355 |
||
356 |
lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast |
|
357 |
lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast |
|
358 |
||
359 |
lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by blast |
|
360 |
lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast |
|
361 |
lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast |
|
362 |
lemma not_iff: "(P~=Q) = (P = (~Q))" by blast |
|
363 |
lemma disj_not1: "(~P | Q) = (P --> Q)" by blast |
|
364 |
lemma disj_not2: "(P | ~Q) = (Q --> P)" -- {* changes orientation :-( *}
|
|
365 |
by blast |
|
366 |
lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast |
|
367 |
||
368 |
lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by blast |
|
369 |
||
370 |
||
371 |
lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q" |
|
372 |
-- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *}
|
|
373 |
-- {* cases boil down to the same thing. *}
|
|
374 |
by blast |
|
375 |
||
376 |
lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast |
|
377 |
lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast |
|
378 |
lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by blast |
|
379 |
lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by blast |
|
380 |
||
381 |
lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by blast |
|
382 |
lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by blast |
|
383 |
||
384 |
text {*
|
|
385 |
\medskip The @{text "&"} congruence rule: not included by default!
|
|
386 |
May slow rewrite proofs down by as much as 50\% *} |
|
387 |
||
388 |
lemma conj_cong: |
|
389 |
"(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))" |
|
390 |
by blast |
|
391 |
||
392 |
lemma rev_conj_cong: |
|
393 |
"(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))" |
|
394 |
by blast |
|
395 |
||
396 |
text {* The @{text "|"} congruence rule: not included by default! *}
|
|
397 |
||
398 |
lemma disj_cong: |
|
399 |
"(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))" |
|
400 |
by blast |
|
401 |
||
402 |
lemma eq_sym_conv: "(x = y) = (y = x)" |
|
403 |
by blast |
|
404 |
||
405 |
||
406 |
text {* \medskip if-then-else rules *}
|
|
407 |
||
408 |
lemma if_True: "(if True then x else y) = x" |
|
409 |
by (unfold if_def) blast |
|
410 |
||
411 |
lemma if_False: "(if False then x else y) = y" |
|
412 |
by (unfold if_def) blast |
|
413 |
||
414 |
lemma if_P: "P ==> (if P then x else y) = x" |
|
415 |
by (unfold if_def) blast |
|
416 |
||
417 |
lemma if_not_P: "~P ==> (if P then x else y) = y" |
|
418 |
by (unfold if_def) blast |
|
419 |
||
420 |
lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" |
|
421 |
apply (rule case_split [of Q]) |
|
422 |
apply (subst if_P) |
|
423 |
prefer 3 apply (subst if_not_P) |
|
424 |
apply blast+ |
|
425 |
done |
|
426 |
||
427 |
lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" |
|
428 |
apply (subst split_if) |
|
429 |
apply blast |
|
430 |
done |
|
431 |
||
432 |
lemmas if_splits = split_if split_if_asm |
|
433 |
||
434 |
lemma if_def2: "(if Q then x else y) = ((Q --> x) & (~ Q --> y))" |
|
435 |
by (rule split_if) |
|
436 |
||
437 |
lemma if_cancel: "(if c then x else x) = x" |
|
438 |
apply (subst split_if) |
|
439 |
apply blast |
|
440 |
done |
|
441 |
||
442 |
lemma if_eq_cancel: "(if x = y then y else x) = x" |
|
443 |
apply (subst split_if) |
|
444 |
apply blast |
|
445 |
done |
|
446 |
||
447 |
lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))" |
|
448 |
-- {* This form is useful for expanding @{text if}s on the RIGHT of the @{text "==>"} symbol. *}
|
|
449 |
by (rule split_if) |
|
450 |
||
451 |
lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" |
|
452 |
-- {* And this form is useful for expanding @{text if}s on the LEFT. *}
|
|
453 |
apply (subst split_if) |
|
454 |
apply blast |
|
455 |
done |
|
456 |
||
457 |
lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) blast |
|
458 |
lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) blast |
|
459 |
||
| 9869 | 460 |
use "simpdata.ML" |
461 |
setup Simplifier.setup |
|
462 |
setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup |
|
463 |
setup Splitter.setup setup Clasimp.setup |
|
464 |
||
| 11750 | 465 |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
466 |
subsubsection {* Generic cases and induction *}
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
467 |
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
468 |
constdefs |
| 11989 | 469 |
induct_forall :: "('a => bool) => bool"
|
470 |
"induct_forall P == \<forall>x. P x" |
|
471 |
induct_implies :: "bool => bool => bool" |
|
472 |
"induct_implies A B == A --> B" |
|
473 |
induct_equal :: "'a => 'a => bool" |
|
474 |
"induct_equal x y == x = y" |
|
475 |
induct_conj :: "bool => bool => bool" |
|
476 |
"induct_conj A B == A & B" |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
477 |
|
| 11989 | 478 |
lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))" |
479 |
by (simp only: atomize_all induct_forall_def) |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
480 |
|
| 11989 | 481 |
lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)" |
482 |
by (simp only: atomize_imp induct_implies_def) |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
483 |
|
| 11989 | 484 |
lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)" |
485 |
by (simp only: atomize_eq induct_equal_def) |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
486 |
|
| 11989 | 487 |
lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = |
488 |
induct_conj (induct_forall A) (induct_forall B)" |
|
489 |
by (unfold induct_forall_def induct_conj_def) blast |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
490 |
|
| 11989 | 491 |
lemma induct_implies_conj: "induct_implies C (induct_conj A B) = |
492 |
induct_conj (induct_implies C A) (induct_implies C B)" |
|
493 |
by (unfold induct_implies_def induct_conj_def) blast |
|
494 |
||
495 |
lemma induct_conj_curry: "(induct_conj A B ==> C) == (A ==> B ==> C)" |
|
496 |
by (simp only: atomize_imp atomize_eq induct_conj_def) (rule equal_intr_rule, blast+) |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
497 |
|
| 11989 | 498 |
lemma induct_impliesI: "(A ==> B) ==> induct_implies A B" |
499 |
by (simp add: induct_implies_def) |
|
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
500 |
|
| 12161 | 501 |
lemmas induct_atomize = atomize_conj induct_forall_eq induct_implies_eq induct_equal_eq |
502 |
lemmas induct_rulify1 [symmetric, standard] = induct_forall_eq induct_implies_eq induct_equal_eq |
|
503 |
lemmas induct_rulify2 = induct_forall_def induct_implies_def induct_equal_def induct_conj_def |
|
| 11989 | 504 |
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
|
505 |
|
| 11989 | 506 |
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
|
507 |
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
508 |
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
509 |
text {* Method setup. *}
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
510 |
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
511 |
ML {*
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
512 |
structure InductMethod = InductMethodFun |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
513 |
(struct |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
514 |
val dest_concls = HOLogic.dest_concls; |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
515 |
val cases_default = thm "case_split"; |
| 11989 | 516 |
val local_impI = thm "induct_impliesI"; |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
517 |
val conjI = thm "conjI"; |
| 11989 | 518 |
val atomize = thms "induct_atomize"; |
519 |
val rulify1 = thms "induct_rulify1"; |
|
520 |
val rulify2 = thms "induct_rulify2"; |
|
| 12240 | 521 |
val localize = [Thm.symmetric (thm "induct_implies_def")]; |
|
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
522 |
end); |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
523 |
*} |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
524 |
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
525 |
setup InductMethod.setup |
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
526 |
|
|
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
527 |
|
| 11750 | 528 |
subsection {* Order signatures and orders *}
|
529 |
||
530 |
axclass |
|
531 |
ord < "term" |
|
532 |
||
533 |
syntax |
|
534 |
"op <" :: "['a::ord, 'a] => bool" ("op <")
|
|
535 |
"op <=" :: "['a::ord, 'a] => bool" ("op <=")
|
|
536 |
||
537 |
global |
|
538 |
||
539 |
consts |
|
540 |
"op <" :: "['a::ord, 'a] => bool" ("(_/ < _)" [50, 51] 50)
|
|
541 |
"op <=" :: "['a::ord, 'a] => bool" ("(_/ <= _)" [50, 51] 50)
|
|
542 |
||
543 |
local |
|
544 |
||
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
545 |
syntax (xsymbols) |
| 11750 | 546 |
"op <=" :: "['a::ord, 'a] => bool" ("op \<le>")
|
547 |
"op <=" :: "['a::ord, 'a] => bool" ("(_/ \<le> _)" [50, 51] 50)
|
|
548 |
||
549 |
(*Tell blast about overloading of < and <= to reduce the risk of |
|
550 |
its applying a rule for the wrong type*) |
|
551 |
ML {*
|
|
552 |
Blast.overloaded ("op <" , domain_type);
|
|
553 |
Blast.overloaded ("op <=", domain_type);
|
|
554 |
*} |
|
555 |
||
556 |
||
557 |
subsubsection {* Monotonicity *}
|
|
558 |
||
559 |
constdefs |
|
560 |
mono :: "['a::ord => 'b::ord] => bool" |
|
561 |
"mono f == ALL A B. A <= B --> f A <= f B" |
|
562 |
||
563 |
lemma monoI [intro?]: "(!!A B. A <= B ==> f A <= f B) ==> mono f" |
|
564 |
by (unfold mono_def) blast |
|
565 |
||
566 |
lemma monoD [dest?]: "mono f ==> A <= B ==> f A <= f B" |
|
567 |
by (unfold mono_def) blast |
|
568 |
||
569 |
constdefs |
|
570 |
min :: "['a::ord, 'a] => 'a" |
|
571 |
"min a b == (if a <= b then a else b)" |
|
572 |
max :: "['a::ord, 'a] => 'a" |
|
573 |
"max a b == (if a <= b then b else a)" |
|
574 |
||
575 |
lemma min_leastL: "(!!x. least <= x) ==> min least x = least" |
|
576 |
by (simp add: min_def) |
|
577 |
||
578 |
lemma min_of_mono: |
|
579 |
"ALL x y. (f x <= f y) = (x <= y) ==> min (f m) (f n) = f (min m n)" |
|
580 |
by (simp add: min_def) |
|
581 |
||
582 |
lemma max_leastL: "(!!x. least <= x) ==> max least x = x" |
|
583 |
by (simp add: max_def) |
|
584 |
||
585 |
lemma max_of_mono: |
|
586 |
"ALL x y. (f x <= f y) = (x <= y) ==> max (f m) (f n) = f (max m n)" |
|
587 |
by (simp add: max_def) |
|
588 |
||
589 |
||
590 |
subsubsection "Orders" |
|
591 |
||
592 |
axclass order < ord |
|
593 |
order_refl [iff]: "x <= x" |
|
594 |
order_trans: "x <= y ==> y <= z ==> x <= z" |
|
595 |
order_antisym: "x <= y ==> y <= x ==> x = y" |
|
596 |
order_less_le: "(x < y) = (x <= y & x ~= y)" |
|
597 |
||
598 |
||
599 |
text {* Reflexivity. *}
|
|
600 |
||
601 |
lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y" |
|
602 |
-- {* This form is useful with the classical reasoner. *}
|
|
603 |
apply (erule ssubst) |
|
604 |
apply (rule order_refl) |
|
605 |
done |
|
606 |
||
607 |
lemma order_less_irrefl [simp]: "~ x < (x::'a::order)" |
|
608 |
by (simp add: order_less_le) |
|
609 |
||
610 |
lemma order_le_less: "((x::'a::order) <= y) = (x < y | x = y)" |
|
611 |
-- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
|
|
612 |
apply (simp add: order_less_le) |
|
| 12256 | 613 |
apply blast |
| 11750 | 614 |
done |
615 |
||
616 |
lemmas order_le_imp_less_or_eq = order_le_less [THEN iffD1, standard] |
|
617 |
||
618 |
lemma order_less_imp_le: "!!x::'a::order. x < y ==> x <= y" |
|
619 |
by (simp add: order_less_le) |
|
620 |
||
621 |
||
622 |
text {* Asymmetry. *}
|
|
623 |
||
624 |
lemma order_less_not_sym: "(x::'a::order) < y ==> ~ (y < x)" |
|
625 |
by (simp add: order_less_le order_antisym) |
|
626 |
||
627 |
lemma order_less_asym: "x < (y::'a::order) ==> (~P ==> y < x) ==> P" |
|
628 |
apply (drule order_less_not_sym) |
|
629 |
apply (erule contrapos_np) |
|
630 |
apply simp |
|
631 |
done |
|
632 |
||
633 |
||
634 |
text {* Transitivity. *}
|
|
635 |
||
636 |
lemma order_less_trans: "!!x::'a::order. [| x < y; y < z |] ==> x < z" |
|
637 |
apply (simp add: order_less_le) |
|
638 |
apply (blast intro: order_trans order_antisym) |
|
639 |
done |
|
640 |
||
641 |
lemma order_le_less_trans: "!!x::'a::order. [| x <= y; y < z |] ==> x < z" |
|
642 |
apply (simp add: order_less_le) |
|
643 |
apply (blast intro: order_trans order_antisym) |
|
644 |
done |
|
645 |
||
646 |
lemma order_less_le_trans: "!!x::'a::order. [| x < y; y <= z |] ==> x < z" |
|
647 |
apply (simp add: order_less_le) |
|
648 |
apply (blast intro: order_trans order_antisym) |
|
649 |
done |
|
650 |
||
651 |
||
652 |
text {* Useful for simplification, but too risky to include by default. *}
|
|
653 |
||
654 |
lemma order_less_imp_not_less: "(x::'a::order) < y ==> (~ y < x) = True" |
|
655 |
by (blast elim: order_less_asym) |
|
656 |
||
657 |
lemma order_less_imp_triv: "(x::'a::order) < y ==> (y < x --> P) = True" |
|
658 |
by (blast elim: order_less_asym) |
|
659 |
||
660 |
lemma order_less_imp_not_eq: "(x::'a::order) < y ==> (x = y) = False" |
|
661 |
by auto |
|
662 |
||
663 |
lemma order_less_imp_not_eq2: "(x::'a::order) < y ==> (y = x) = False" |
|
664 |
by auto |
|
665 |
||
666 |
||
667 |
text {* Other operators. *}
|
|
668 |
||
669 |
lemma min_leastR: "(!!x::'a::order. least <= x) ==> min x least = least" |
|
670 |
apply (simp add: min_def) |
|
671 |
apply (blast intro: order_antisym) |
|
672 |
done |
|
673 |
||
674 |
lemma max_leastR: "(!!x::'a::order. least <= x) ==> max x least = x" |
|
675 |
apply (simp add: max_def) |
|
676 |
apply (blast intro: order_antisym) |
|
677 |
done |
|
678 |
||
679 |
||
680 |
subsubsection {* Least value operator *}
|
|
681 |
||
682 |
constdefs |
|
683 |
Least :: "('a::ord => bool) => 'a" (binder "LEAST " 10)
|
|
684 |
"Least P == THE x. P x & (ALL y. P y --> x <= y)" |
|
685 |
-- {* We can no longer use LeastM because the latter requires Hilbert-AC. *}
|
|
686 |
||
687 |
lemma LeastI2: |
|
688 |
"[| P (x::'a::order); |
|
689 |
!!y. P y ==> x <= y; |
|
690 |
!!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |] |
|
| 12281 | 691 |
==> Q (Least P)" |
| 11750 | 692 |
apply (unfold Least_def) |
693 |
apply (rule theI2) |
|
694 |
apply (blast intro: order_antisym)+ |
|
695 |
done |
|
696 |
||
697 |
lemma Least_equality: |
|
| 12281 | 698 |
"[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k" |
| 11750 | 699 |
apply (simp add: Least_def) |
700 |
apply (rule the_equality) |
|
701 |
apply (auto intro!: order_antisym) |
|
702 |
done |
|
703 |
||
704 |
||
705 |
subsubsection "Linear / total orders" |
|
706 |
||
707 |
axclass linorder < order |
|
708 |
linorder_linear: "x <= y | y <= x" |
|
709 |
||
710 |
lemma linorder_less_linear: "!!x::'a::linorder. x<y | x=y | y<x" |
|
711 |
apply (simp add: order_less_le) |
|
712 |
apply (insert linorder_linear) |
|
713 |
apply blast |
|
714 |
done |
|
715 |
||
716 |
lemma linorder_cases [case_names less equal greater]: |
|
717 |
"((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P" |
|
718 |
apply (insert linorder_less_linear) |
|
719 |
apply blast |
|
720 |
done |
|
721 |
||
722 |
lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)" |
|
723 |
apply (simp add: order_less_le) |
|
724 |
apply (insert linorder_linear) |
|
725 |
apply (blast intro: order_antisym) |
|
726 |
done |
|
727 |
||
728 |
lemma linorder_not_le: "!!x::'a::linorder. (~ x <= y) = (y < x)" |
|
729 |
apply (simp add: order_less_le) |
|
730 |
apply (insert linorder_linear) |
|
731 |
apply (blast intro: order_antisym) |
|
732 |
done |
|
733 |
||
734 |
lemma linorder_neq_iff: "!!x::'a::linorder. (x ~= y) = (x<y | y<x)" |
|
735 |
apply (cut_tac x = x and y = y in linorder_less_linear) |
|
736 |
apply auto |
|
737 |
done |
|
738 |
||
739 |
lemma linorder_neqE: "x ~= (y::'a::linorder) ==> (x < y ==> R) ==> (y < x ==> R) ==> R" |
|
740 |
apply (simp add: linorder_neq_iff) |
|
741 |
apply blast |
|
742 |
done |
|
743 |
||
744 |
||
745 |
subsubsection "Min and max on (linear) orders" |
|
746 |
||
747 |
lemma min_same [simp]: "min (x::'a::order) x = x" |
|
748 |
by (simp add: min_def) |
|
749 |
||
750 |
lemma max_same [simp]: "max (x::'a::order) x = x" |
|
751 |
by (simp add: max_def) |
|
752 |
||
753 |
lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" |
|
754 |
apply (simp add: max_def) |
|
755 |
apply (insert linorder_linear) |
|
756 |
apply (blast intro: order_trans) |
|
757 |
done |
|
758 |
||
759 |
lemma le_maxI1: "(x::'a::linorder) <= max x y" |
|
760 |
by (simp add: le_max_iff_disj) |
|
761 |
||
762 |
lemma le_maxI2: "(y::'a::linorder) <= max x y" |
|
763 |
-- {* CANNOT use with @{text "[intro!]"} because blast will give PROOF FAILED. *}
|
|
764 |
by (simp add: le_max_iff_disj) |
|
765 |
||
766 |
lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" |
|
767 |
apply (simp add: max_def order_le_less) |
|
768 |
apply (insert linorder_less_linear) |
|
769 |
apply (blast intro: order_less_trans) |
|
770 |
done |
|
771 |
||
772 |
lemma max_le_iff_conj [simp]: |
|
773 |
"!!z::'a::linorder. (max x y <= z) = (x <= z & y <= z)" |
|
774 |
apply (simp add: max_def) |
|
775 |
apply (insert linorder_linear) |
|
776 |
apply (blast intro: order_trans) |
|
777 |
done |
|
778 |
||
779 |
lemma max_less_iff_conj [simp]: |
|
780 |
"!!z::'a::linorder. (max x y < z) = (x < z & y < z)" |
|
781 |
apply (simp add: order_le_less max_def) |
|
782 |
apply (insert linorder_less_linear) |
|
783 |
apply (blast intro: order_less_trans) |
|
784 |
done |
|
785 |
||
786 |
lemma le_min_iff_conj [simp]: |
|
787 |
"!!z::'a::linorder. (z <= min x y) = (z <= x & z <= y)" |
|
788 |
-- {* @{text "[iff]"} screws up a Q{text blast} in MiniML *}
|
|
789 |
apply (simp add: min_def) |
|
790 |
apply (insert linorder_linear) |
|
791 |
apply (blast intro: order_trans) |
|
792 |
done |
|
793 |
||
794 |
lemma min_less_iff_conj [simp]: |
|
795 |
"!!z::'a::linorder. (z < min x y) = (z < x & z < y)" |
|
796 |
apply (simp add: order_le_less min_def) |
|
797 |
apply (insert linorder_less_linear) |
|
798 |
apply (blast intro: order_less_trans) |
|
799 |
done |
|
800 |
||
801 |
lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" |
|
802 |
apply (simp add: min_def) |
|
803 |
apply (insert linorder_linear) |
|
804 |
apply (blast intro: order_trans) |
|
805 |
done |
|
806 |
||
807 |
lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" |
|
808 |
apply (simp add: min_def order_le_less) |
|
809 |
apply (insert linorder_less_linear) |
|
810 |
apply (blast intro: order_less_trans) |
|
811 |
done |
|
812 |
||
813 |
lemma split_min: |
|
814 |
"P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" |
|
815 |
by (simp add: min_def) |
|
816 |
||
817 |
lemma split_max: |
|
818 |
"P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" |
|
819 |
by (simp add: max_def) |
|
820 |
||
821 |
||
822 |
subsubsection "Bounded quantifiers" |
|
823 |
||
824 |
syntax |
|
825 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10)
|
|
826 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10)
|
|
827 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10)
|
|
828 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10)
|
|
829 |
||
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12023
diff
changeset
|
830 |
syntax (xsymbols) |
| 11750 | 831 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_<_./ _)" [0, 0, 10] 10)
|
832 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_<_./ _)" [0, 0, 10] 10)
|
|
833 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
|
|
834 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
|
|
835 |
||
836 |
syntax (HOL) |
|
837 |
"_lessAll" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10)
|
|
838 |
"_lessEx" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10)
|
|
839 |
"_leAll" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10)
|
|
840 |
"_leEx" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10)
|
|
841 |
||
842 |
translations |
|
843 |
"ALL x<y. P" => "ALL x. x < y --> P" |
|
844 |
"EX x<y. P" => "EX x. x < y & P" |
|
845 |
"ALL x<=y. P" => "ALL x. x <= y --> P" |
|
846 |
"EX x<=y. P" => "EX x. x <= y & P" |
|
847 |
||
| 923 | 848 |
end |