author | wenzelm |
Mon, 22 May 2000 13:29:21 +0200 | |
changeset 8920 | af5e09b6c208 |
parent 8800 | e3688ef49f12 |
child 8940 | 55bc03d9f168 |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/HOL.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
2260 | 6 |
Higher-Order Logic. |
923 | 7 |
*) |
8 |
||
7357 | 9 |
theory HOL = CPure |
10 |
files ("HOL_lemmas.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML"): |
|
923 | 11 |
|
2260 | 12 |
|
13 |
(** Core syntax **) |
|
14 |
||
3947 | 15 |
global |
16 |
||
7357 | 17 |
classes "term" < logic |
18 |
defaultsort "term" |
|
923 | 19 |
|
7357 | 20 |
typedecl bool |
923 | 21 |
|
22 |
arities |
|
7357 | 23 |
bool :: "term" |
24 |
fun :: ("term", "term") "term" |
|
923 | 25 |
|
26 |
||
27 |
consts |
|
28 |
||
29 |
(* Constants *) |
|
30 |
||
7357 | 31 |
Trueprop :: "bool => prop" ("(_)" 5) |
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 |
|
38 |
(* Binders *) |
|
39 |
||
7357 | 40 |
Eps :: "('a => bool) => 'a" |
41 |
All :: "('a => bool) => bool" (binder "ALL " 10) |
|
42 |
Ex :: "('a => bool) => bool" (binder "EX " 10) |
|
43 |
Ex1 :: "('a => bool) => bool" (binder "EX! " 10) |
|
44 |
Let :: "['a, 'a => 'b] => 'b" |
|
923 | 45 |
|
46 |
(* Infixes *) |
|
47 |
||
7357 | 48 |
"=" :: "['a, 'a] => bool" (infixl 50) |
49 |
& :: "[bool, bool] => bool" (infixr 35) |
|
50 |
"|" :: "[bool, bool] => bool" (infixr 30) |
|
51 |
--> :: "[bool, bool] => bool" (infixr 25) |
|
923 | 52 |
|
2260 | 53 |
|
54 |
(* Overloaded Constants *) |
|
55 |
||
7357 | 56 |
axclass plus < "term" |
57 |
axclass minus < "term" |
|
58 |
axclass times < "term" |
|
59 |
axclass power < "term" |
|
3370
5c5fdce3a4e4
Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents:
3320
diff
changeset
|
60 |
|
2260 | 61 |
consts |
7357 | 62 |
"+" :: "['a::plus, 'a] => 'a" (infixl 65) |
63 |
- :: "['a::minus, 'a] => 'a" (infixl 65) |
|
64 |
uminus :: "['a::minus] => 'a" ("- _" [81] 80) |
|
8800 | 65 |
abs :: "('a::minus) => 'a" |
7426 | 66 |
* :: "['a::times, 'a] => 'a" (infixl 70) |
3370
5c5fdce3a4e4
Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents:
3320
diff
changeset
|
67 |
(*See Nat.thy for "^"*) |
2260 | 68 |
|
3820 | 69 |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
70 |
|
2260 | 71 |
(** Additional concrete syntax **) |
72 |
||
4868 | 73 |
nonterminals |
923 | 74 |
letbinds letbind |
75 |
case_syn cases_syn |
|
76 |
||
77 |
syntax |
|
7357 | 78 |
~= :: "['a, 'a] => bool" (infixl 50) |
79 |
"_Eps" :: "[pttrn, bool] => 'a" ("(3SOME _./ _)" [0, 10] 10) |
|
923 | 80 |
|
81 |
(* Let expressions *) |
|
82 |
||
7357 | 83 |
"_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10) |
84 |
"" :: "letbind => letbinds" ("_") |
|
85 |
"_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") |
|
86 |
"_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) |
|
923 | 87 |
|
88 |
(* Case expressions *) |
|
89 |
||
7357 | 90 |
"@case" :: "['a, cases_syn] => 'b" ("(case _ of/ _)" 10) |
91 |
"@case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10) |
|
92 |
"" :: "case_syn => cases_syn" ("_") |
|
93 |
"@case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _") |
|
923 | 94 |
|
95 |
translations |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
96 |
"x ~= y" == "~ (x = y)" |
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
97 |
"SOME x. P" == "Eps (%x. P)" |
923 | 98 |
"_Let (_binds b bs) e" == "_Let b (_Let bs e)" |
1114 | 99 |
"let x = a in e" == "Let a (%x. e)" |
923 | 100 |
|
3820 | 101 |
syntax ("" output) |
7357 | 102 |
"op =" :: "['a, 'a] => bool" ("(_ =/ _)" [51, 51] 50) |
103 |
"op ~=" :: "['a, 'a] => bool" ("(_ ~=/ _)" [51, 51] 50) |
|
2260 | 104 |
|
105 |
syntax (symbols) |
|
7357 | 106 |
Not :: "bool => bool" ("\\<not> _" [40] 40) |
107 |
"op &" :: "[bool, bool] => bool" (infixr "\\<and>" 35) |
|
108 |
"op |" :: "[bool, bool] => bool" (infixr "\\<or>" 30) |
|
109 |
"op -->" :: "[bool, bool] => bool" (infixr "\\<midarrow>\\<rightarrow>" 25) |
|
110 |
"op o" :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixl "\\<circ>" 55) |
|
111 |
"op ~=" :: "['a, 'a] => bool" (infixl "\\<noteq>" 50) |
|
112 |
"_Eps" :: "[pttrn, bool] => 'a" ("(3\\<epsilon>_./ _)" [0, 10] 10) |
|
113 |
"ALL " :: "[idts, bool] => bool" ("(3\\<forall>_./ _)" [0, 10] 10) |
|
114 |
"EX " :: "[idts, bool] => bool" ("(3\\<exists>_./ _)" [0, 10] 10) |
|
115 |
"EX! " :: "[idts, bool] => bool" ("(3\\<exists>!_./ _)" [0, 10] 10) |
|
116 |
"@case1" :: "['a, 'b] => case_syn" ("(2_ \\<Rightarrow>/ _)" 10) |
|
117 |
(*"@case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ \\<orelse> _")*) |
|
2372 | 118 |
|
3820 | 119 |
syntax (symbols output) |
7357 | 120 |
"op ~=" :: "['a, 'a] => bool" ("(_ \\<noteq>/ _)" [51, 51] 50) |
3820 | 121 |
|
6027
9dd06eeda95c
added new print_mode "xsymbols" for extended symbol support
oheimb
parents:
5786
diff
changeset
|
122 |
syntax (xsymbols) |
7357 | 123 |
"op -->" :: "[bool, bool] => bool" (infixr "\\<longrightarrow>" 25) |
2260 | 124 |
|
6340 | 125 |
syntax (HTML output) |
7357 | 126 |
Not :: "bool => bool" ("\\<not> _" [40] 40) |
6340 | 127 |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
128 |
syntax (HOL) |
7357 | 129 |
"_Eps" :: "[pttrn, bool] => 'a" ("(3@ _./ _)" [0, 10] 10) |
130 |
"ALL " :: "[idts, bool] => bool" ("(3! _./ _)" [0, 10] 10) |
|
131 |
"EX " :: "[idts, bool] => bool" ("(3? _./ _)" [0, 10] 10) |
|
132 |
"EX! " :: "[idts, bool] => bool" ("(3?! _./ _)" [0, 10] 10) |
|
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
133 |
|
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
134 |
|
6340 | 135 |
|
2260 | 136 |
(** Rules and definitions **) |
137 |
||
3947 | 138 |
local |
139 |
||
7357 | 140 |
axioms |
923 | 141 |
|
7357 | 142 |
eq_reflection: "(x=y) ==> (x==y)" |
923 | 143 |
|
144 |
(* Basic Rules *) |
|
145 |
||
7357 | 146 |
refl: "t = (t::'a)" |
147 |
subst: "[| s = t; P(s) |] ==> P(t::'a)" |
|
6289 | 148 |
|
149 |
(*Extensionality is built into the meta-logic, and this rule expresses |
|
150 |
a related property. It is an eta-expanded version of the traditional |
|
151 |
rule, and similar to the ABS rule of HOL.*) |
|
7357 | 152 |
ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" |
6289 | 153 |
|
7357 | 154 |
selectI: "P (x::'a) ==> P (@x. P x)" |
923 | 155 |
|
7357 | 156 |
impI: "(P ==> Q) ==> P-->Q" |
157 |
mp: "[| P-->Q; P |] ==> Q" |
|
923 | 158 |
|
159 |
defs |
|
160 |
||
7357 | 161 |
True_def: "True == ((%x::bool. x) = (%x. x))" |
162 |
All_def: "All(P) == (P = (%x. True))" |
|
163 |
Ex_def: "Ex(P) == P(@x. P(x))" |
|
164 |
False_def: "False == (!P. P)" |
|
165 |
not_def: "~ P == P-->False" |
|
166 |
and_def: "P & Q == !R. (P-->Q-->R) --> R" |
|
167 |
or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" |
|
168 |
Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" |
|
923 | 169 |
|
7357 | 170 |
axioms |
923 | 171 |
(* Axioms *) |
172 |
||
7357 | 173 |
iff: "(P-->Q) --> (Q-->P) --> (P=Q)" |
174 |
True_or_False: "(P=True) | (P=False)" |
|
923 | 175 |
|
176 |
defs |
|
5069 | 177 |
(*misc definitions*) |
7357 | 178 |
Let_def: "Let s f == f(s)" |
179 |
if_def: "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)" |
|
5069 | 180 |
|
181 |
(*arbitrary is completely unspecified, but is made to appear as a |
|
182 |
definition syntactically*) |
|
7357 | 183 |
arbitrary_def: "False ==> arbitrary == (@x. False)" |
923 | 184 |
|
3320 | 185 |
|
4868 | 186 |
|
7357 | 187 |
(* theory and package setup *) |
4868 | 188 |
|
7357 | 189 |
use "HOL_lemmas.ML" setup attrib_setup |
190 |
use "cladata.ML" setup Classical.setup setup clasetup |
|
191 |
use "blastdata.ML" setup Blast.setup |
|
8473 | 192 |
use "simpdata.ML" setup Simplifier.setup |
193 |
setup "Simplifier.method_setup Splitter.split_modifiers" |
|
8640 | 194 |
setup simpsetup setup cong_attrib_setup |
195 |
setup Splitter.setup setup Clasimp.setup setup iff_attrib_setup |
|
4868 | 196 |
|
197 |
||
923 | 198 |
end |