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