| author | wenzelm | 
| Fri, 26 Oct 2001 23:59:13 +0200 | |
| changeset 11953 | f98623fdf6ef | 
| parent 11824 | f4c1882dde2c | 
| child 11977 | 2e7c54b86763 | 
| 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  | 
|
81  | 
syntax (symbols)  | 
|
| 11687 | 82  | 
  Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
83  | 
"op &" :: "[bool, bool] => bool" (infixr "\<and>" 35)  | 
|
84  | 
"op |" :: "[bool, bool] => bool" (infixr "\<or>" 30)  | 
|
85  | 
"op -->" :: "[bool, bool] => bool" (infixr "\<midarrow>\<rightarrow>" 25)  | 
|
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  | 
|
| 3820 | 93  | 
syntax (symbols output)  | 
| 11687 | 94  | 
"op ~=" :: "['a, 'a] => bool" (infix "\<noteq>" 50)  | 
| 3820 | 95  | 
|
| 
6027
 
9dd06eeda95c
added new print_mode "xsymbols" for extended symbol support
 
oheimb 
parents: 
5786 
diff
changeset
 | 
96  | 
syntax (xsymbols)  | 
| 11687 | 97  | 
"op -->" :: "[bool, bool] => bool" (infixr "\<longrightarrow>" 25)  | 
| 2260 | 98  | 
|
| 6340 | 99  | 
syntax (HTML output)  | 
| 11687 | 100  | 
  Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
| 6340 | 101  | 
|
| 
7238
 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 
wenzelm 
parents: 
7220 
diff
changeset
 | 
102  | 
syntax (HOL)  | 
| 7357 | 103  | 
  "ALL "        :: "[idts, bool] => bool"                ("(3! _./ _)" [0, 10] 10)
 | 
104  | 
  "EX "         :: "[idts, bool] => bool"                ("(3? _./ _)" [0, 10] 10)
 | 
|
105  | 
  "EX! "        :: "[idts, bool] => bool"                ("(3?! _./ _)" [0, 10] 10)
 | 
|
| 
7238
 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 
wenzelm 
parents: 
7220 
diff
changeset
 | 
106  | 
|
| 
 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 
wenzelm 
parents: 
7220 
diff
changeset
 | 
107  | 
|
| 11750 | 108  | 
subsubsection {* Axioms and basic definitions *}
 | 
| 2260 | 109  | 
|
| 7357 | 110  | 
axioms  | 
111  | 
eq_reflection: "(x=y) ==> (x==y)"  | 
|
| 923 | 112  | 
|
| 7357 | 113  | 
refl: "t = (t::'a)"  | 
114  | 
subst: "[| s = t; P(s) |] ==> P(t::'a)"  | 
|
| 6289 | 115  | 
|
| 7357 | 116  | 
ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)"  | 
| 11750 | 117  | 
    -- {* Extensionality is built into the meta-logic, and this rule expresses *}
 | 
118  | 
    -- {* a related property.  It is an eta-expanded version of the traditional *}
 | 
|
119  | 
    -- {* rule, and similar to the ABS rule of HOL *}
 | 
|
| 6289 | 120  | 
|
| 
11432
 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 
wenzelm 
parents: 
10489 
diff
changeset
 | 
121  | 
the_eq_trivial: "(THE x. x = a) = (a::'a)"  | 
| 923 | 122  | 
|
| 7357 | 123  | 
impI: "(P ==> Q) ==> P-->Q"  | 
124  | 
mp: "[| P-->Q; P |] ==> Q"  | 
|
| 923 | 125  | 
|
126  | 
defs  | 
|
| 7357 | 127  | 
True_def: "True == ((%x::bool. x) = (%x. x))"  | 
128  | 
All_def: "All(P) == (P = (%x. True))"  | 
|
| 
11451
 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 
paulson 
parents: 
11438 
diff
changeset
 | 
129  | 
Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q"  | 
| 7357 | 130  | 
False_def: "False == (!P. P)"  | 
131  | 
not_def: "~ P == P-->False"  | 
|
132  | 
and_def: "P & Q == !R. (P-->Q-->R) --> R"  | 
|
133  | 
or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R"  | 
|
134  | 
Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)"  | 
|
| 923 | 135  | 
|
| 7357 | 136  | 
axioms  | 
137  | 
iff: "(P-->Q) --> (Q-->P) --> (P=Q)"  | 
|
138  | 
True_or_False: "(P=True) | (P=False)"  | 
|
| 923 | 139  | 
|
140  | 
defs  | 
|
| 7357 | 141  | 
Let_def: "Let s f == f(s)"  | 
| 
11451
 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 
paulson 
parents: 
11438 
diff
changeset
 | 
142  | 
if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)"  | 
| 5069 | 143  | 
|
| 
11451
 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 
paulson 
parents: 
11438 
diff
changeset
 | 
144  | 
arbitrary_def: "False ==> arbitrary == (THE x. False)"  | 
| 11750 | 145  | 
    -- {* @{term arbitrary} is completely unspecified, but is made to appear as a
 | 
146  | 
definition syntactically *}  | 
|
| 923 | 147  | 
|
| 3320 | 148  | 
|
| 11750 | 149  | 
subsubsection {* Generic algebraic operations *}
 | 
| 4868 | 150  | 
|
| 11750 | 151  | 
axclass zero < "term"  | 
152  | 
axclass one < "term"  | 
|
153  | 
axclass plus < "term"  | 
|
154  | 
axclass minus < "term"  | 
|
155  | 
axclass times < "term"  | 
|
156  | 
axclass inverse < "term"  | 
|
157  | 
||
158  | 
global  | 
|
159  | 
||
160  | 
consts  | 
|
161  | 
  "0"           :: "'a::zero"                       ("0")
 | 
|
162  | 
  "1"           :: "'a::one"                        ("1")
 | 
|
163  | 
"+" :: "['a::plus, 'a] => 'a" (infixl 65)  | 
|
164  | 
- :: "['a::minus, 'a] => 'a" (infixl 65)  | 
|
165  | 
  uminus        :: "['a::minus] => 'a"              ("- _" [81] 80)
 | 
|
166  | 
* :: "['a::times, 'a] => 'a" (infixl 70)  | 
|
167  | 
||
168  | 
local  | 
|
169  | 
||
170  | 
typed_print_translation {*
 | 
|
171  | 
let  | 
|
172  | 
fun tr' c = (c, fn show_sorts => fn T => fn ts =>  | 
|
173  | 
if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match  | 
|
174  | 
else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T);  | 
|
175  | 
in [tr' "0", tr' "1"] end;  | 
|
176  | 
*} -- {* show types that are presumably too general *}
 | 
|
177  | 
||
178  | 
||
179  | 
consts  | 
|
180  | 
abs :: "'a::minus => 'a"  | 
|
181  | 
inverse :: "'a::inverse => 'a"  | 
|
182  | 
divide :: "['a::inverse, 'a] => 'a" (infixl "'/" 70)  | 
|
183  | 
||
184  | 
syntax (xsymbols)  | 
|
185  | 
  abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | 
|
186  | 
syntax (HTML output)  | 
|
187  | 
  abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | 
|
188  | 
||
189  | 
axclass plus_ac0 < plus, zero  | 
|
190  | 
commute: "x + y = y + x"  | 
|
191  | 
assoc: "(x + y) + z = x + (y + z)"  | 
|
192  | 
zero: "0 + x = x"  | 
|
193  | 
||
194  | 
||
195  | 
subsection {* Theory and package setup *}
 | 
|
196  | 
||
197  | 
subsubsection {* Basic lemmas *}
 | 
|
| 4868 | 198  | 
|
| 9736 | 199  | 
use "HOL_lemmas.ML"  | 
| 11687 | 200  | 
theorems case_split = case_split_thm [case_names True False]  | 
| 9869 | 201  | 
|
| 11750 | 202  | 
declare trans [trans]  | 
203  | 
declare impE [CPure.elim] iffD1 [CPure.elim] iffD2 [CPure.elim]  | 
|
204  | 
||
| 
11438
 
3d9222b80989
declare trans [trans]  (*overridden in theory Calculation*);
 
wenzelm 
parents: 
11432 
diff
changeset
 | 
205  | 
|
| 11750 | 206  | 
subsubsection {* Atomizing meta-level connectives *}
 | 
207  | 
||
208  | 
lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)"  | 
|
| 9488 | 209  | 
proof (rule equal_intr_rule)  | 
210  | 
assume "!!x. P x"  | 
|
| 10383 | 211  | 
show "ALL x. P x" by (rule allI)  | 
| 9488 | 212  | 
next  | 
213  | 
assume "ALL x. P x"  | 
|
| 10383 | 214  | 
thus "!!x. P x" by (rule allE)  | 
| 9488 | 215  | 
qed  | 
216  | 
||
| 11750 | 217  | 
lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)"  | 
| 9488 | 218  | 
proof (rule equal_intr_rule)  | 
219  | 
assume r: "A ==> B"  | 
|
| 10383 | 220  | 
show "A --> B" by (rule impI) (rule r)  | 
| 9488 | 221  | 
next  | 
222  | 
assume "A --> B" and A  | 
|
| 10383 | 223  | 
thus B by (rule mp)  | 
| 9488 | 224  | 
qed  | 
225  | 
||
| 11750 | 226  | 
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)"  | 
| 
10432
 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 
wenzelm 
parents: 
10383 
diff
changeset
 | 
227  | 
proof (rule equal_intr_rule)  | 
| 
 
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  | 
show "x = y" by (unfold prems) (rule refl)  | 
| 
 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 
wenzelm 
parents: 
10383 
diff
changeset
 | 
230  | 
next  | 
| 
 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 
wenzelm 
parents: 
10383 
diff
changeset
 | 
231  | 
assume "x = y"  | 
| 
 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 
wenzelm 
parents: 
10383 
diff
changeset
 | 
232  | 
thus "x == y" by (rule eq_reflection)  | 
| 
 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 
wenzelm 
parents: 
10383 
diff
changeset
 | 
233  | 
qed  | 
| 
 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 
wenzelm 
parents: 
10383 
diff
changeset
 | 
234  | 
|
| 11953 | 235  | 
lemma atomize_conj [atomize]: "(!!C. (A ==> B ==> PROP C) ==> PROP C) == Trueprop (A & B)"  | 
236  | 
proof (rule equal_intr_rule)  | 
|
237  | 
assume "!!C. (A ==> B ==> PROP C) ==> PROP C"  | 
|
238  | 
show "A & B" by (rule conjI)  | 
|
239  | 
next  | 
|
240  | 
fix C  | 
|
241  | 
assume "A & B"  | 
|
242  | 
assume "A ==> B ==> PROP C"  | 
|
243  | 
thus "PROP C"  | 
|
244  | 
proof this  | 
|
245  | 
show A by (rule conjunct1)  | 
|
246  | 
show B by (rule conjunct2)  | 
|
247  | 
qed  | 
|
248  | 
qed  | 
|
249  | 
||
| 11750 | 250  | 
|
251  | 
subsubsection {* Classical Reasoner setup *}
 | 
|
| 9529 | 252  | 
|
| 10383 | 253  | 
use "cladata.ML"  | 
254  | 
setup hypsubst_setup  | 
|
| 11770 | 255  | 
declare atomize_all [symmetric, rulify] atomize_imp [symmetric, rulify]  | 
| 10383 | 256  | 
setup Classical.setup  | 
257  | 
setup clasetup  | 
|
258  | 
||
| 9869 | 259  | 
use "blastdata.ML"  | 
260  | 
setup Blast.setup  | 
|
| 4868 | 261  | 
|
| 11750 | 262  | 
|
263  | 
subsubsection {* Simplifier setup *}
 | 
|
264  | 
||
| 9869 | 265  | 
use "simpdata.ML"  | 
266  | 
setup Simplifier.setup  | 
|
267  | 
setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup  | 
|
268  | 
setup Splitter.setup setup Clasimp.setup  | 
|
269  | 
||
| 11750 | 270  | 
|
| 
11824
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
271  | 
subsubsection {* Generic cases and induction *}
 | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
272  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
273  | 
constdefs  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
274  | 
  inductive_forall :: "('a => bool) => bool"
 | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
275  | 
"inductive_forall P == \<forall>x. P x"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
276  | 
inductive_implies :: "bool => bool => bool"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
277  | 
"inductive_implies A B == A --> B"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
278  | 
inductive_equal :: "'a => 'a => bool"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
279  | 
"inductive_equal x y == x = y"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
280  | 
inductive_conj :: "bool => bool => bool"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
281  | 
"inductive_conj A B == A & B"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
282  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
283  | 
lemma inductive_forall_eq: "(!!x. P x) == Trueprop (inductive_forall (\<lambda>x. P x))"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
284  | 
by (simp only: atomize_all inductive_forall_def)  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
285  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
286  | 
lemma inductive_implies_eq: "(A ==> B) == Trueprop (inductive_implies A B)"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
287  | 
by (simp only: atomize_imp inductive_implies_def)  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
288  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
289  | 
lemma inductive_equal_eq: "(x == y) == Trueprop (inductive_equal x y)"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
290  | 
by (simp only: atomize_eq inductive_equal_def)  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
291  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
292  | 
lemma inductive_forall_conj: "inductive_forall (\<lambda>x. inductive_conj (A x) (B x)) =  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
293  | 
inductive_conj (inductive_forall A) (inductive_forall B)"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
294  | 
by (unfold inductive_forall_def inductive_conj_def) blast  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
295  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
296  | 
lemma inductive_implies_conj: "inductive_implies C (inductive_conj A B) =  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
297  | 
inductive_conj (inductive_implies C A) (inductive_implies C B)"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
298  | 
by (unfold inductive_implies_def inductive_conj_def) blast  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
299  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
300  | 
lemma inductive_conj_curry: "(inductive_conj A B ==> C) == (A ==> B ==> C)"  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
301  | 
by (simp only: atomize_imp atomize_eq inductive_conj_def) (rule equal_intr_rule, blast+)  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
302  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
303  | 
lemmas inductive_atomize = inductive_forall_eq inductive_implies_eq inductive_equal_eq  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
304  | 
lemmas inductive_rulify1 = inductive_atomize [symmetric, standard]  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
305  | 
lemmas inductive_rulify2 =  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
306  | 
inductive_forall_def inductive_implies_def inductive_equal_def inductive_conj_def  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
307  | 
lemmas inductive_conj = inductive_forall_conj inductive_implies_conj inductive_conj_curry  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
308  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
309  | 
hide const inductive_forall inductive_implies inductive_equal inductive_conj  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
310  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
311  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
312  | 
text {* Method setup. *}
 | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
313  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
314  | 
ML {*
 | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
315  | 
structure InductMethod = InductMethodFun  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
316  | 
(struct  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
317  | 
val dest_concls = HOLogic.dest_concls;  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
318  | 
val cases_default = thm "case_split";  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
319  | 
val conjI = thm "conjI";  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
320  | 
val atomize = thms "inductive_atomize";  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
321  | 
val rulify1 = thms "inductive_rulify1";  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
322  | 
val rulify2 = thms "inductive_rulify2";  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
323  | 
end);  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
324  | 
*}  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
325  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
326  | 
setup InductMethod.setup  | 
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
327  | 
|
| 
 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 
wenzelm 
parents: 
11770 
diff
changeset
 | 
328  | 
|
| 11750 | 329  | 
subsection {* Order signatures and orders *}
 | 
330  | 
||
331  | 
axclass  | 
|
332  | 
ord < "term"  | 
|
333  | 
||
334  | 
syntax  | 
|
335  | 
  "op <"        :: "['a::ord, 'a] => bool"             ("op <")
 | 
|
336  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("op <=")
 | 
|
337  | 
||
338  | 
global  | 
|
339  | 
||
340  | 
consts  | 
|
341  | 
  "op <"        :: "['a::ord, 'a] => bool"             ("(_/ < _)"  [50, 51] 50)
 | 
|
342  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("(_/ <= _)" [50, 51] 50)
 | 
|
343  | 
||
344  | 
local  | 
|
345  | 
||
346  | 
syntax (symbols)  | 
|
347  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("op \<le>")
 | 
|
348  | 
  "op <="       :: "['a::ord, 'a] => bool"             ("(_/ \<le> _)"  [50, 51] 50)
 | 
|
349  | 
||
350  | 
(*Tell blast about overloading of < and <= to reduce the risk of  | 
|
351  | 
its applying a rule for the wrong type*)  | 
|
352  | 
ML {*
 | 
|
353  | 
Blast.overloaded ("op <" , domain_type);
 | 
|
354  | 
Blast.overloaded ("op <=", domain_type);
 | 
|
355  | 
*}  | 
|
356  | 
||
357  | 
||
358  | 
subsubsection {* Monotonicity *}
 | 
|
359  | 
||
360  | 
constdefs  | 
|
361  | 
mono :: "['a::ord => 'b::ord] => bool"  | 
|
362  | 
"mono f == ALL A B. A <= B --> f A <= f B"  | 
|
363  | 
||
364  | 
lemma monoI [intro?]: "(!!A B. A <= B ==> f A <= f B) ==> mono f"  | 
|
365  | 
by (unfold mono_def) blast  | 
|
366  | 
||
367  | 
lemma monoD [dest?]: "mono f ==> A <= B ==> f A <= f B"  | 
|
368  | 
by (unfold mono_def) blast  | 
|
369  | 
||
370  | 
constdefs  | 
|
371  | 
min :: "['a::ord, 'a] => 'a"  | 
|
372  | 
"min a b == (if a <= b then a else b)"  | 
|
373  | 
max :: "['a::ord, 'a] => 'a"  | 
|
374  | 
"max a b == (if a <= b then b else a)"  | 
|
375  | 
||
376  | 
lemma min_leastL: "(!!x. least <= x) ==> min least x = least"  | 
|
377  | 
by (simp add: min_def)  | 
|
378  | 
||
379  | 
lemma min_of_mono:  | 
|
380  | 
"ALL x y. (f x <= f y) = (x <= y) ==> min (f m) (f n) = f (min m n)"  | 
|
381  | 
by (simp add: min_def)  | 
|
382  | 
||
383  | 
lemma max_leastL: "(!!x. least <= x) ==> max least x = x"  | 
|
384  | 
by (simp add: max_def)  | 
|
385  | 
||
386  | 
lemma max_of_mono:  | 
|
387  | 
"ALL x y. (f x <= f y) = (x <= y) ==> max (f m) (f n) = f (max m n)"  | 
|
388  | 
by (simp add: max_def)  | 
|
389  | 
||
390  | 
||
391  | 
subsubsection "Orders"  | 
|
392  | 
||
393  | 
axclass order < ord  | 
|
394  | 
order_refl [iff]: "x <= x"  | 
|
395  | 
order_trans: "x <= y ==> y <= z ==> x <= z"  | 
|
396  | 
order_antisym: "x <= y ==> y <= x ==> x = y"  | 
|
397  | 
order_less_le: "(x < y) = (x <= y & x ~= y)"  | 
|
398  | 
||
399  | 
||
400  | 
text {* Reflexivity. *}
 | 
|
401  | 
||
402  | 
lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y"  | 
|
403  | 
    -- {* This form is useful with the classical reasoner. *}
 | 
|
404  | 
apply (erule ssubst)  | 
|
405  | 
apply (rule order_refl)  | 
|
406  | 
done  | 
|
407  | 
||
408  | 
lemma order_less_irrefl [simp]: "~ x < (x::'a::order)"  | 
|
409  | 
by (simp add: order_less_le)  | 
|
410  | 
||
411  | 
lemma order_le_less: "((x::'a::order) <= y) = (x < y | x = y)"  | 
|
412  | 
    -- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
 | 
|
413  | 
apply (simp add: order_less_le)  | 
|
414  | 
apply (blast intro!: order_refl)  | 
|
415  | 
done  | 
|
416  | 
||
417  | 
lemmas order_le_imp_less_or_eq = order_le_less [THEN iffD1, standard]  | 
|
418  | 
||
419  | 
lemma order_less_imp_le: "!!x::'a::order. x < y ==> x <= y"  | 
|
420  | 
by (simp add: order_less_le)  | 
|
421  | 
||
422  | 
||
423  | 
text {* Asymmetry. *}
 | 
|
424  | 
||
425  | 
lemma order_less_not_sym: "(x::'a::order) < y ==> ~ (y < x)"  | 
|
426  | 
by (simp add: order_less_le order_antisym)  | 
|
427  | 
||
428  | 
lemma order_less_asym: "x < (y::'a::order) ==> (~P ==> y < x) ==> P"  | 
|
429  | 
apply (drule order_less_not_sym)  | 
|
430  | 
apply (erule contrapos_np)  | 
|
431  | 
apply simp  | 
|
432  | 
done  | 
|
433  | 
||
434  | 
||
435  | 
text {* Transitivity. *}
 | 
|
436  | 
||
437  | 
lemma order_less_trans: "!!x::'a::order. [| x < y; y < z |] ==> x < z"  | 
|
438  | 
apply (simp add: order_less_le)  | 
|
439  | 
apply (blast intro: order_trans order_antisym)  | 
|
440  | 
done  | 
|
441  | 
||
442  | 
lemma order_le_less_trans: "!!x::'a::order. [| x <= y; y < z |] ==> x < z"  | 
|
443  | 
apply (simp add: order_less_le)  | 
|
444  | 
apply (blast intro: order_trans order_antisym)  | 
|
445  | 
done  | 
|
446  | 
||
447  | 
lemma order_less_le_trans: "!!x::'a::order. [| x < y; y <= z |] ==> x < z"  | 
|
448  | 
apply (simp add: order_less_le)  | 
|
449  | 
apply (blast intro: order_trans order_antisym)  | 
|
450  | 
done  | 
|
451  | 
||
452  | 
||
453  | 
text {* Useful for simplification, but too risky to include by default. *}
 | 
|
454  | 
||
455  | 
lemma order_less_imp_not_less: "(x::'a::order) < y ==> (~ y < x) = True"  | 
|
456  | 
by (blast elim: order_less_asym)  | 
|
457  | 
||
458  | 
lemma order_less_imp_triv: "(x::'a::order) < y ==> (y < x --> P) = True"  | 
|
459  | 
by (blast elim: order_less_asym)  | 
|
460  | 
||
461  | 
lemma order_less_imp_not_eq: "(x::'a::order) < y ==> (x = y) = False"  | 
|
462  | 
by auto  | 
|
463  | 
||
464  | 
lemma order_less_imp_not_eq2: "(x::'a::order) < y ==> (y = x) = False"  | 
|
465  | 
by auto  | 
|
466  | 
||
467  | 
||
468  | 
text {* Other operators. *}
 | 
|
469  | 
||
470  | 
lemma min_leastR: "(!!x::'a::order. least <= x) ==> min x least = least"  | 
|
471  | 
apply (simp add: min_def)  | 
|
472  | 
apply (blast intro: order_antisym)  | 
|
473  | 
done  | 
|
474  | 
||
475  | 
lemma max_leastR: "(!!x::'a::order. least <= x) ==> max x least = x"  | 
|
476  | 
apply (simp add: max_def)  | 
|
477  | 
apply (blast intro: order_antisym)  | 
|
478  | 
done  | 
|
479  | 
||
480  | 
||
481  | 
subsubsection {* Least value operator *}
 | 
|
482  | 
||
483  | 
constdefs  | 
|
484  | 
  Least :: "('a::ord => bool) => 'a"               (binder "LEAST " 10)
 | 
|
485  | 
"Least P == THE x. P x & (ALL y. P y --> x <= y)"  | 
|
486  | 
    -- {* We can no longer use LeastM because the latter requires Hilbert-AC. *}
 | 
|
487  | 
||
488  | 
lemma LeastI2:  | 
|
489  | 
"[| P (x::'a::order);  | 
|
490  | 
!!y. P y ==> x <= y;  | 
|
491  | 
!!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |]  | 
|
492  | 
==> Q (Least P)";  | 
|
493  | 
apply (unfold Least_def)  | 
|
494  | 
apply (rule theI2)  | 
|
495  | 
apply (blast intro: order_antisym)+  | 
|
496  | 
done  | 
|
497  | 
||
498  | 
lemma Least_equality:  | 
|
499  | 
"[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k";  | 
|
500  | 
apply (simp add: Least_def)  | 
|
501  | 
apply (rule the_equality)  | 
|
502  | 
apply (auto intro!: order_antisym)  | 
|
503  | 
done  | 
|
504  | 
||
505  | 
||
506  | 
subsubsection "Linear / total orders"  | 
|
507  | 
||
508  | 
axclass linorder < order  | 
|
509  | 
linorder_linear: "x <= y | y <= x"  | 
|
510  | 
||
511  | 
lemma linorder_less_linear: "!!x::'a::linorder. x<y | x=y | y<x"  | 
|
512  | 
apply (simp add: order_less_le)  | 
|
513  | 
apply (insert linorder_linear)  | 
|
514  | 
apply blast  | 
|
515  | 
done  | 
|
516  | 
||
517  | 
lemma linorder_cases [case_names less equal greater]:  | 
|
518  | 
"((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P"  | 
|
519  | 
apply (insert linorder_less_linear)  | 
|
520  | 
apply blast  | 
|
521  | 
done  | 
|
522  | 
||
523  | 
lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)"  | 
|
524  | 
apply (simp add: order_less_le)  | 
|
525  | 
apply (insert linorder_linear)  | 
|
526  | 
apply (blast intro: order_antisym)  | 
|
527  | 
done  | 
|
528  | 
||
529  | 
lemma linorder_not_le: "!!x::'a::linorder. (~ x <= y) = (y < x)"  | 
|
530  | 
apply (simp add: order_less_le)  | 
|
531  | 
apply (insert linorder_linear)  | 
|
532  | 
apply (blast intro: order_antisym)  | 
|
533  | 
done  | 
|
534  | 
||
535  | 
lemma linorder_neq_iff: "!!x::'a::linorder. (x ~= y) = (x<y | y<x)"  | 
|
536  | 
apply (cut_tac x = x and y = y in linorder_less_linear)  | 
|
537  | 
apply auto  | 
|
538  | 
done  | 
|
539  | 
||
540  | 
lemma linorder_neqE: "x ~= (y::'a::linorder) ==> (x < y ==> R) ==> (y < x ==> R) ==> R"  | 
|
541  | 
apply (simp add: linorder_neq_iff)  | 
|
542  | 
apply blast  | 
|
543  | 
done  | 
|
544  | 
||
545  | 
||
546  | 
subsubsection "Min and max on (linear) orders"  | 
|
547  | 
||
548  | 
lemma min_same [simp]: "min (x::'a::order) x = x"  | 
|
549  | 
by (simp add: min_def)  | 
|
550  | 
||
551  | 
lemma max_same [simp]: "max (x::'a::order) x = x"  | 
|
552  | 
by (simp add: max_def)  | 
|
553  | 
||
554  | 
lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)"  | 
|
555  | 
apply (simp add: max_def)  | 
|
556  | 
apply (insert linorder_linear)  | 
|
557  | 
apply (blast intro: order_trans)  | 
|
558  | 
done  | 
|
559  | 
||
560  | 
lemma le_maxI1: "(x::'a::linorder) <= max x y"  | 
|
561  | 
by (simp add: le_max_iff_disj)  | 
|
562  | 
||
563  | 
lemma le_maxI2: "(y::'a::linorder) <= max x y"  | 
|
564  | 
    -- {* CANNOT use with @{text "[intro!]"} because blast will give PROOF FAILED. *}
 | 
|
565  | 
by (simp add: le_max_iff_disj)  | 
|
566  | 
||
567  | 
lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)"  | 
|
568  | 
apply (simp add: max_def order_le_less)  | 
|
569  | 
apply (insert linorder_less_linear)  | 
|
570  | 
apply (blast intro: order_less_trans)  | 
|
571  | 
done  | 
|
572  | 
||
573  | 
lemma max_le_iff_conj [simp]:  | 
|
574  | 
"!!z::'a::linorder. (max x y <= z) = (x <= z & y <= z)"  | 
|
575  | 
apply (simp add: max_def)  | 
|
576  | 
apply (insert linorder_linear)  | 
|
577  | 
apply (blast intro: order_trans)  | 
|
578  | 
done  | 
|
579  | 
||
580  | 
lemma max_less_iff_conj [simp]:  | 
|
581  | 
"!!z::'a::linorder. (max x y < z) = (x < z & y < z)"  | 
|
582  | 
apply (simp add: order_le_less max_def)  | 
|
583  | 
apply (insert linorder_less_linear)  | 
|
584  | 
apply (blast intro: order_less_trans)  | 
|
585  | 
done  | 
|
586  | 
||
587  | 
lemma le_min_iff_conj [simp]:  | 
|
588  | 
"!!z::'a::linorder. (z <= min x y) = (z <= x & z <= y)"  | 
|
589  | 
    -- {* @{text "[iff]"} screws up a Q{text blast} in MiniML *}
 | 
|
590  | 
apply (simp add: min_def)  | 
|
591  | 
apply (insert linorder_linear)  | 
|
592  | 
apply (blast intro: order_trans)  | 
|
593  | 
done  | 
|
594  | 
||
595  | 
lemma min_less_iff_conj [simp]:  | 
|
596  | 
"!!z::'a::linorder. (z < min x y) = (z < x & z < y)"  | 
|
597  | 
apply (simp add: order_le_less min_def)  | 
|
598  | 
apply (insert linorder_less_linear)  | 
|
599  | 
apply (blast intro: order_less_trans)  | 
|
600  | 
done  | 
|
601  | 
||
602  | 
lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)"  | 
|
603  | 
apply (simp add: min_def)  | 
|
604  | 
apply (insert linorder_linear)  | 
|
605  | 
apply (blast intro: order_trans)  | 
|
606  | 
done  | 
|
607  | 
||
608  | 
lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)"  | 
|
609  | 
apply (simp add: min_def order_le_less)  | 
|
610  | 
apply (insert linorder_less_linear)  | 
|
611  | 
apply (blast intro: order_less_trans)  | 
|
612  | 
done  | 
|
613  | 
||
614  | 
lemma split_min:  | 
|
615  | 
"P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))"  | 
|
616  | 
by (simp add: min_def)  | 
|
617  | 
||
618  | 
lemma split_max:  | 
|
619  | 
"P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))"  | 
|
620  | 
by (simp add: max_def)  | 
|
621  | 
||
622  | 
||
623  | 
subsubsection "Bounded quantifiers"  | 
|
624  | 
||
625  | 
syntax  | 
|
626  | 
  "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3ALL _<_./ _)"  [0, 0, 10] 10)
 | 
|
627  | 
  "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3EX _<_./ _)"  [0, 0, 10] 10)
 | 
|
628  | 
  "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3ALL _<=_./ _)" [0, 0, 10] 10)
 | 
|
629  | 
  "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3EX _<=_./ _)" [0, 0, 10] 10)
 | 
|
630  | 
||
631  | 
syntax (symbols)  | 
|
632  | 
  "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
 | 
|
633  | 
  "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
 | 
|
634  | 
  "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
 | 
|
635  | 
  "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
 | 
|
636  | 
||
637  | 
syntax (HOL)  | 
|
638  | 
  "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
 | 
|
639  | 
  "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
 | 
|
640  | 
  "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
 | 
|
641  | 
  "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
 | 
|
642  | 
||
643  | 
translations  | 
|
644  | 
"ALL x<y. P" => "ALL x. x < y --> P"  | 
|
645  | 
"EX x<y. P" => "EX x. x < y & P"  | 
|
646  | 
"ALL x<=y. P" => "ALL x. x <= y --> P"  | 
|
647  | 
"EX x<=y. P" => "EX x. x <= y & P"  | 
|
648  | 
||
| 923 | 649  | 
end  |