author | huffman |
Sun, 04 Sep 2011 10:29:38 -0700 | |
changeset 44712 | 1e490e891c88 |
parent 44324 | d972b91ed09d |
child 44860 | 56101fa00193 |
permissions | -rw-r--r-- |
12024 | 1 |
header {* Main HOL *} |
2 |
||
15131 | 3 |
theory Main |
42695
a94ad372b2f5
adding creation of exhaustive generators for records; simplifying dependencies in Main theory
bulwahn
parents:
41919
diff
changeset
|
4 |
imports Plain Predicate_Compile Nitpick |
15131 | 5 |
begin |
9650
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
wenzelm
parents:
9619
diff
changeset
|
6 |
|
29304 | 7 |
text {* |
8 |
Classical Higher-order Logic -- only ``Main'', excluding real and |
|
9 |
complex numbers etc. |
|
10 |
*} |
|
11 |
||
27367 | 12 |
text {* See further \cite{Nipkow-et-al:2002:tutorial} *} |
25964 | 13 |
|
44324 | 14 |
text {* Compatibility layer -- to be dropped *} |
15 |
||
16 |
lemma Inf_bool_def: |
|
17 |
"Inf A \<longleftrightarrow> (\<forall>x\<in>A. x)" |
|
18 |
by (auto intro: bool_induct) |
|
19 |
||
20 |
lemma Sup_bool_def: |
|
21 |
"Sup A \<longleftrightarrow> (\<exists>x\<in>A. x)" |
|
22 |
by auto |
|
23 |
||
24 |
declare Complete_Lattice.Inf_bool_def [simp del] |
|
25 |
declare Complete_Lattice.Sup_bool_def [simp del] |
|
26 |
||
9650
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
wenzelm
parents:
9619
diff
changeset
|
27 |
end |