| author | oheimb | 
| Tue, 23 Apr 1996 17:01:51 +0200 | |
| changeset 1674 | 33aff4d854e4 | 
| parent 1672 | 2c109cd2fdd0 | 
| child 2260 | b59781f2b809 | 
| permissions | -rw-r--r-- | 
| 923 | 1  | 
(* Title: HOL/HOL.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Tobias Nipkow  | 
|
4  | 
Copyright 1993 University of Cambridge  | 
|
5  | 
||
6  | 
Higher-Order Logic  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
HOL = CPure +  | 
|
10  | 
||
11  | 
classes  | 
|
12  | 
term < logic  | 
|
13  | 
||
14  | 
axclass  | 
|
15  | 
plus < term  | 
|
16  | 
||
17  | 
axclass  | 
|
18  | 
minus < term  | 
|
19  | 
||
20  | 
axclass  | 
|
21  | 
times < term  | 
|
22  | 
||
23  | 
default  | 
|
24  | 
term  | 
|
25  | 
||
26  | 
types  | 
|
27  | 
bool  | 
|
28  | 
||
29  | 
arities  | 
|
30  | 
fun :: (term, term) term  | 
|
31  | 
bool :: term  | 
|
32  | 
||
33  | 
||
34  | 
consts  | 
|
35  | 
||
36  | 
(* Constants *)  | 
|
37  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
38  | 
  Trueprop      :: bool => prop                     ("(_)" 5)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
39  | 
  not           :: bool => bool                     ("~ _" [40] 40)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
40  | 
True, False :: bool  | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
41  | 
  If            :: [bool, 'a, 'a] => 'a   ("(if (_)/ then (_)/ else (_))" 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
42  | 
  Inv           :: ('a => 'b) => ('b => 'a)
 | 
| 923 | 43  | 
|
44  | 
(* Binders *)  | 
|
45  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
46  | 
  Eps           :: ('a => bool) => 'a
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
47  | 
  All           :: ('a => bool) => bool             (binder "! " 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
48  | 
  Ex            :: ('a => bool) => bool             (binder "? " 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
49  | 
  Ex1           :: ('a => bool) => bool             (binder "?! " 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
50  | 
Let :: ['a, 'a => 'b] => 'b  | 
| 923 | 51  | 
|
52  | 
(* Infixes *)  | 
|
53  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
54  | 
o :: ['b => 'c, 'a => 'b, 'a] => 'c (infixl 55)  | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
55  | 
"=" :: ['a, 'a] => bool (infixl 50)  | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
56  | 
"&" :: [bool, bool] => bool (infixr 35)  | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
57  | 
"|" :: [bool, bool] => bool (infixr 30)  | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
58  | 
"-->" :: [bool, bool] => bool (infixr 25)  | 
| 923 | 59  | 
|
60  | 
(* Overloaded Constants *)  | 
|
61  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
62  | 
"+" :: ['a::plus, 'a] => 'a (infixl 65)  | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
63  | 
"-" :: ['a::minus, 'a] => 'a (infixl 65)  | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
64  | 
"*" :: ['a::times, 'a] => 'a (infixl 70)  | 
| 923 | 65  | 
|
66  | 
||
67  | 
types  | 
|
68  | 
letbinds letbind  | 
|
69  | 
case_syn cases_syn  | 
|
70  | 
||
71  | 
syntax  | 
|
72  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
73  | 
"~=" :: ['a, 'a] => bool (infixl 50)  | 
| 923 | 74  | 
|
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
75  | 
  "@Eps"        :: [pttrn,bool] => 'a               ("(3@ _./ _)" 10)
 | 
| 1068 | 76  | 
|
| 923 | 77  | 
(* Alternative Quantifiers *)  | 
78  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
79  | 
  "*All"        :: [idts, bool] => bool             ("(3ALL _./ _)" 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
80  | 
  "*Ex"         :: [idts, bool] => bool             ("(3EX _./ _)" 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
81  | 
  "*Ex1"        :: [idts, bool] => bool             ("(3EX! _./ _)" 10)
 | 
| 923 | 82  | 
|
83  | 
(* Let expressions *)  | 
|
84  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
85  | 
  "_bind"       :: [pttrn, 'a] => letbind           ("(2_ =/ _)" 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
86  | 
  ""            :: letbind => letbinds              ("_")
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
87  | 
  "_binds"      :: [letbind, letbinds] => letbinds  ("_;/ _")
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
88  | 
  "_Let"        :: [letbinds, 'a] => 'a             ("(let (_)/ in (_))" 10)
 | 
| 923 | 89  | 
|
90  | 
(* Case expressions *)  | 
|
91  | 
||
| 
1370
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
92  | 
  "@case"       :: ['a, cases_syn] => 'b            ("(case _ of/ _)" 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
93  | 
  "@case1"      :: ['a, 'b] => case_syn             ("(2_ =>/ _)" 10)
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
94  | 
  ""            :: case_syn => cases_syn            ("_")
 | 
| 
 
7361ac9b024d
removed quotes from types in consts and syntax sections
 
clasohm 
parents: 
1273 
diff
changeset
 | 
95  | 
  "@case2"      :: [case_syn, cases_syn] => cases_syn   ("_/ | _")
 | 
| 923 | 96  | 
|
97  | 
translations  | 
|
98  | 
"x ~= y" == "~ (x = y)"  | 
|
| 1068 | 99  | 
"@ x.b" == "Eps(%x.b)"  | 
| 923 | 100  | 
"ALL xs. P" => "! xs. P"  | 
101  | 
"EX xs. P" => "? xs. P"  | 
|
102  | 
"EX! xs. P" => "?! xs. P"  | 
|
103  | 
"_Let (_binds b bs) e" == "_Let b (_Let bs e)"  | 
|
| 1114 | 104  | 
"let x = a in e" == "Let a (%x. e)"  | 
| 923 | 105  | 
|
106  | 
rules  | 
|
107  | 
||
108  | 
eq_reflection "(x=y) ==> (x==y)"  | 
|
109  | 
||
110  | 
(* Basic Rules *)  | 
|
111  | 
||
112  | 
refl "t = (t::'a)"  | 
|
113  | 
subst "[| s = t; P(s) |] ==> P(t::'a)"  | 
|
114  | 
ext "(!!x::'a. (f(x)::'b) = g(x)) ==> (%x.f(x)) = (%x.g(x))"  | 
|
115  | 
selectI "P(x::'a) ==> P(@x.P(x))"  | 
|
116  | 
||
117  | 
impI "(P ==> Q) ==> P-->Q"  | 
|
118  | 
mp "[| P-->Q; P |] ==> Q"  | 
|
119  | 
||
120  | 
defs  | 
|
121  | 
||
122  | 
True_def "True == ((%x::bool.x)=(%x.x))"  | 
|
123  | 
All_def "All(P) == (P = (%x.True))"  | 
|
124  | 
Ex_def "Ex(P) == P(@x.P(x))"  | 
|
125  | 
False_def "False == (!P.P)"  | 
|
126  | 
not_def "~ P == P-->False"  | 
|
127  | 
and_def "P & Q == !R. (P-->Q-->R) --> R"  | 
|
128  | 
or_def "P | Q == !R. (P-->R) --> (Q-->R) --> R"  | 
|
129  | 
Ex1_def "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)"  | 
|
130  | 
||
131  | 
rules  | 
|
132  | 
(* Axioms *)  | 
|
133  | 
||
134  | 
iff "(P-->Q) --> (Q-->P) --> (P=Q)"  | 
|
135  | 
True_or_False "(P=True) | (P=False)"  | 
|
136  | 
||
137  | 
defs  | 
|
138  | 
(* Misc Definitions *)  | 
|
139  | 
||
140  | 
Let_def "Let s f == f(s)"  | 
|
141  | 
Inv_def "Inv(f::'a=>'b) == (% y. @x. f(x)=y)"  | 
|
142  | 
o_def "(f::'b=>'c) o g == (%(x::'a). f(g(x)))"  | 
|
| 973 | 143  | 
if_def "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)"  | 
| 923 | 144  | 
|
| 1273 | 145  | 
(* start 8bit 1 *)  | 
146  | 
(* end 8bit 1 *)  | 
|
147  | 
||
| 923 | 148  | 
end  | 
149  | 
||
150  | 
ML  | 
|
151  | 
||
152  | 
(** Choice between the HOL and Isabelle style of quantifiers **)  | 
|
153  | 
||
154  | 
val HOL_quantifiers = ref true;  | 
|
155  | 
||
156  | 
fun alt_ast_tr' (name, alt_name) =  | 
|
157  | 
let  | 
|
158  | 
fun ast_tr' (*name*) args =  | 
|
159  | 
if ! HOL_quantifiers then raise Match  | 
|
160  | 
else Syntax.mk_appl (Syntax.Constant alt_name) args;  | 
|
161  | 
in  | 
|
162  | 
(name, ast_tr')  | 
|
163  | 
end;  | 
|
164  | 
||
165  | 
||
166  | 
val print_ast_translation =  | 
|
167  | 
  map alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")];
 | 
|
| 1273 | 168  | 
|
169  | 
(* start 8bit 2 *)  | 
|
170  | 
(* end 8bit 2 *)  | 
|
171  | 
||
172  | 
||
173  |