author | nipkow |
Tue, 08 Apr 1997 12:03:59 +0200 | |
changeset 2920 | feab36851df3 |
parent 2912 | 3fac3e8d5d3e |
child 3066 | 3c548f92e032 |
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 |
||
9 |
HOL = CPure + |
|
10 |
||
2260 | 11 |
|
12 |
(** Core syntax **) |
|
13 |
||
923 | 14 |
classes |
15 |
term < logic |
|
16 |
||
17 |
default |
|
18 |
term |
|
19 |
||
20 |
types |
|
21 |
bool |
|
22 |
||
23 |
arities |
|
24 |
fun :: (term, term) term |
|
25 |
bool :: term |
|
26 |
||
27 |
||
28 |
consts |
|
29 |
||
30 |
(* Constants *) |
|
31 |
||
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
32 |
Trueprop :: bool => prop ("(_)" 5) |
2720 | 33 |
Not :: bool => bool ("~ _" [40] 40) |
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
34 |
True, False :: bool |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
35 |
If :: [bool, 'a, 'a] => 'a ("(if (_)/ then (_)/ else (_))" 10) |
923 | 36 |
|
37 |
(* Binders *) |
|
38 |
||
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
39 |
Eps :: ('a => bool) => 'a |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
40 |
All :: ('a => bool) => bool (binder "! " 10) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
41 |
Ex :: ('a => bool) => bool (binder "? " 10) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
42 |
Ex1 :: ('a => bool) => bool (binder "?! " 10) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
43 |
Let :: ['a, 'a => 'b] => 'b |
923 | 44 |
|
45 |
(* Infixes *) |
|
46 |
||
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
47 |
o :: ['b => 'c, 'a => 'b, 'a] => 'c (infixl 55) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
48 |
"=" :: ['a, 'a] => bool (infixl 50) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
49 |
"&" :: [bool, bool] => bool (infixr 35) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
50 |
"|" :: [bool, bool] => bool (infixr 30) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
51 |
"-->" :: [bool, bool] => bool (infixr 25) |
923 | 52 |
|
2260 | 53 |
|
54 |
(* Overloaded Constants *) |
|
55 |
||
56 |
axclass |
|
57 |
plus < term |
|
923 | 58 |
|
2260 | 59 |
axclass |
60 |
minus < term |
|
61 |
||
62 |
axclass |
|
63 |
times < term |
|
64 |
||
65 |
consts |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
66 |
"+" :: ['a::plus, 'a] => 'a (infixl 65) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
67 |
"-" :: ['a::minus, 'a] => 'a (infixl 65) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
68 |
"*" :: ['a::times, 'a] => 'a (infixl 70) |
923 | 69 |
|
70 |
||
2260 | 71 |
|
72 |
(** Additional concrete syntax **) |
|
73 |
||
923 | 74 |
types |
75 |
letbinds letbind |
|
76 |
case_syn cases_syn |
|
77 |
||
78 |
syntax |
|
79 |
||
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
80 |
"~=" :: ['a, 'a] => bool (infixl 50) |
923 | 81 |
|
2368 | 82 |
"@Eps" :: [pttrn, bool] => 'a ("(3@ _./ _)" [0, 10] 10) |
1068 | 83 |
|
923 | 84 |
(* Alternative Quantifiers *) |
85 |
||
2368 | 86 |
"*All" :: [idts, bool] => bool ("(3ALL _./ _)" [0, 10] 10) |
87 |
"*Ex" :: [idts, bool] => bool ("(3EX _./ _)" [0, 10] 10) |
|
88 |
"*Ex1" :: [idts, bool] => bool ("(3EX! _./ _)" [0, 10] 10) |
|
923 | 89 |
|
90 |
(* Let expressions *) |
|
91 |
||
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
92 |
"_bind" :: [pttrn, 'a] => letbind ("(2_ =/ _)" 10) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
93 |
"" :: letbind => letbinds ("_") |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
94 |
"_binds" :: [letbind, letbinds] => letbinds ("_;/ _") |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
95 |
"_Let" :: [letbinds, 'a] => 'a ("(let (_)/ in (_))" 10) |
923 | 96 |
|
97 |
(* Case expressions *) |
|
98 |
||
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
99 |
"@case" :: ['a, cases_syn] => 'b ("(case _ of/ _)" 10) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
100 |
"@case1" :: ['a, 'b] => case_syn ("(2_ =>/ _)" 10) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
101 |
"" :: case_syn => cases_syn ("_") |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1273
diff
changeset
|
102 |
"@case2" :: [case_syn, cases_syn] => cases_syn ("_/ | _") |
923 | 103 |
|
104 |
translations |
|
105 |
"x ~= y" == "~ (x = y)" |
|
2260 | 106 |
"@ x.b" == "Eps (%x. b)" |
923 | 107 |
"ALL xs. P" => "! xs. P" |
108 |
"EX xs. P" => "? xs. P" |
|
109 |
"EX! xs. P" => "?! xs. P" |
|
110 |
"_Let (_binds b bs) e" == "_Let b (_Let bs e)" |
|
1114 | 111 |
"let x = a in e" == "Let a (%x. e)" |
923 | 112 |
|
2260 | 113 |
|
114 |
syntax (symbols) |
|
2762 | 115 |
Not :: bool => bool ("\\<not> _" [40] 40) |
2260 | 116 |
"op &" :: [bool, bool] => bool (infixr "\\<and>" 35) |
117 |
"op |" :: [bool, bool] => bool (infixr "\\<or>" 30) |
|
118 |
"op -->" :: [bool, bool] => bool (infixr "\\<midarrow>\\<rightarrow>" 25) |
|
119 |
"op o" :: ['b => 'c, 'a => 'b, 'a] => 'c (infixl "\\<circ>" 55) |
|
120 |
"op ~=" :: ['a, 'a] => bool (infixl "\\<noteq>" 50) |
|
2368 | 121 |
"@Eps" :: [pttrn, bool] => 'a ("(3\\<epsilon>_./ _)" [0, 10] 10) |
122 |
"! " :: [idts, bool] => bool ("(3\\<forall>_./ _)" [0, 10] 10) |
|
123 |
"? " :: [idts, bool] => bool ("(3\\<exists>_./ _)" [0, 10] 10) |
|
124 |
"?! " :: [idts, bool] => bool ("(3\\<exists>!_./ _)" [0, 10] 10) |
|
2552 | 125 |
"@case1" :: ['a, 'b] => case_syn ("(2_ \\<Rightarrow>/ _)" 10) |
2372 | 126 |
|
127 |
syntax (symbols output) |
|
2368 | 128 |
"*All" :: [idts, bool] => bool ("(3\\<forall>_./ _)" [0, 10] 10) |
129 |
"*Ex" :: [idts, bool] => bool ("(3\\<exists>_./ _)" [0, 10] 10) |
|
130 |
"*Ex1" :: [idts, bool] => bool ("(3\\<exists>!_./ _)" [0, 10] 10) |
|
2372 | 131 |
|
2260 | 132 |
|
133 |
||
134 |
||
135 |
(** Rules and definitions **) |
|
136 |
||
923 | 137 |
rules |
138 |
||
139 |
eq_reflection "(x=y) ==> (x==y)" |
|
140 |
||
141 |
(* Basic Rules *) |
|
142 |
||
143 |
refl "t = (t::'a)" |
|
144 |
subst "[| s = t; P(s) |] ==> P(t::'a)" |
|
145 |
ext "(!!x::'a. (f(x)::'b) = g(x)) ==> (%x.f(x)) = (%x.g(x))" |
|
146 |
selectI "P(x::'a) ==> P(@x.P(x))" |
|
147 |
||
148 |
impI "(P ==> Q) ==> P-->Q" |
|
149 |
mp "[| P-->Q; P |] ==> Q" |
|
150 |
||
151 |
defs |
|
152 |
||
153 |
True_def "True == ((%x::bool.x)=(%x.x))" |
|
154 |
All_def "All(P) == (P = (%x.True))" |
|
155 |
Ex_def "Ex(P) == P(@x.P(x))" |
|
156 |
False_def "False == (!P.P)" |
|
157 |
not_def "~ P == P-->False" |
|
158 |
and_def "P & Q == !R. (P-->Q-->R) --> R" |
|
159 |
or_def "P | Q == !R. (P-->R) --> (Q-->R) --> R" |
|
160 |
Ex1_def "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" |
|
161 |
||
162 |
rules |
|
163 |
(* Axioms *) |
|
164 |
||
165 |
iff "(P-->Q) --> (Q-->P) --> (P=Q)" |
|
166 |
True_or_False "(P=True) | (P=False)" |
|
167 |
||
168 |
defs |
|
169 |
(* Misc Definitions *) |
|
170 |
||
171 |
Let_def "Let s f == f(s)" |
|
172 |
o_def "(f::'b=>'c) o g == (%(x::'a). f(g(x)))" |
|
973 | 173 |
if_def "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)" |
923 | 174 |
|
175 |
end |
|
176 |
||
2260 | 177 |
|
923 | 178 |
ML |
179 |
||
180 |
(** Choice between the HOL and Isabelle style of quantifiers **) |
|
181 |
||
182 |
val HOL_quantifiers = ref true; |
|
183 |
||
184 |
fun alt_ast_tr' (name, alt_name) = |
|
185 |
let |
|
186 |
fun ast_tr' (*name*) args = |
|
187 |
if ! HOL_quantifiers then raise Match |
|
188 |
else Syntax.mk_appl (Syntax.Constant alt_name) args; |
|
189 |
in |
|
190 |
(name, ast_tr') |
|
191 |
end; |
|
192 |
||
193 |
||
194 |
val print_ast_translation = |
|
195 |
map alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")]; |