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