|
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 = Pure + |
|
10 |
|
11 classes |
|
12 term < logic |
|
13 plus < term |
|
14 minus < term |
|
15 times < term |
|
16 |
|
17 default term |
|
18 |
|
19 types |
|
20 bool 0 |
|
21 |
|
22 arities |
|
23 fun :: (term, term) term |
|
24 bool :: term |
|
25 |
|
26 |
|
27 consts |
|
28 |
|
29 (* Constants *) |
|
30 |
|
31 Trueprop :: "bool => prop" ("(_)" [0] 5) |
|
32 not :: "bool => bool" ("~ _" [40] 40) |
|
33 True, False :: "bool" |
|
34 if :: "[bool, 'a, 'a] => 'a" |
|
35 Inv :: "('a => 'b) => ('b => 'a)" |
|
36 |
|
37 (* Binders *) |
|
38 |
|
39 Eps :: "('a => bool) => 'a" (binder "@" 10) |
|
40 All :: "('a => bool) => bool" (binder "! " 10) |
|
41 Ex :: "('a => bool) => bool" (binder "? " 10) |
|
42 Ex1 :: "('a => bool) => bool" (binder "?! " 10) |
|
43 |
|
44 Let :: "['a, 'a=>'b] => 'b" |
|
45 "@let" :: "[idt,'a,'b] => 'b" ("(let _ = (2_)/ in (2_))" 10) |
|
46 |
|
47 (* Alternative Quantifiers *) |
|
48 |
|
49 "*All" :: "[idts, bool] => bool" ("(3ALL _./ _)" [0,0] 10) |
|
50 "*Ex" :: "[idts, bool] => bool" ("(3EX _./ _)" [0,0] 10) |
|
51 "*Ex1" :: "[idts, bool] => bool" ("(3EX! _./ _)" [0,0] 10) |
|
52 |
|
53 (* Infixes *) |
|
54 |
|
55 o :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixr 50) |
|
56 "=" :: "['a, 'a] => bool" (infixl 50) |
|
57 "&" :: "[bool, bool] => bool" (infixr 35) |
|
58 "|" :: "[bool, bool] => bool" (infixr 30) |
|
59 "-->" :: "[bool, bool] => bool" (infixr 25) |
|
60 |
|
61 (* Overloaded Constants *) |
|
62 |
|
63 "+" :: "['a::plus, 'a] => 'a" (infixl 65) |
|
64 "-" :: "['a::minus, 'a] => 'a" (infixl 65) |
|
65 "*" :: "['a::times, 'a] => 'a" (infixl 70) |
|
66 |
|
67 (* Rewriting Gadget *) |
|
68 |
|
69 NORM :: "'a => 'a" |
|
70 |
|
71 |
|
72 translations |
|
73 "ALL xs. P" => "! xs. P" |
|
74 "EX xs. P" => "? xs. P" |
|
75 "EX! xs. P" => "?! xs. P" |
|
76 |
|
77 "let x = s in t" == "Let(s,%x.t)" |
|
78 |
|
79 rules |
|
80 |
|
81 eq_reflection "(x=y) ==> (x==y)" |
|
82 |
|
83 (* Basic Rules *) |
|
84 |
|
85 refl "t = t::'a" |
|
86 subst "[| s = t; P(s) |] ==> P(t::'a)" |
|
87 ext "(!!x::'a. f(x)::'b = g(x)) ==> (%x.f(x)) = (%x.g(x))" |
|
88 selectI "P(x::'a) ==> P(@x.P(x))" |
|
89 |
|
90 impI "(P ==> Q) ==> P-->Q" |
|
91 mp "[| P-->Q; P |] ==> Q" |
|
92 |
|
93 (* Definitions *) |
|
94 |
|
95 True_def "True = ((%x.x)=(%x.x))" |
|
96 All_def "All = (%P. P = (%x.True))" |
|
97 Ex_def "Ex = (%P. P(@x.P(x)))" |
|
98 False_def "False = (!P.P)" |
|
99 not_def "not = (%P. P-->False)" |
|
100 and_def "op & = (%P Q. !R. (P-->Q-->R) --> R)" |
|
101 or_def "op | = (%P Q. !R. (P-->R) --> (Q-->R) --> R)" |
|
102 Ex1_def "Ex1 = (%P. ? x. P(x) & (! y. P(y) --> y=x))" |
|
103 Let_def "Let(s,f) = f(s)" |
|
104 |
|
105 (* Axioms *) |
|
106 |
|
107 iff "(P-->Q) --> (Q-->P) --> (P=Q)" |
|
108 True_or_False "(P=True) | (P=False)" |
|
109 |
|
110 (* Misc Definitions *) |
|
111 |
|
112 Inv_def "Inv = (%(f::'a=>'b) y. @x. f(x)=y)" |
|
113 o_def "op o = (%(f::'b=>'c) g (x::'a). f(g(x)))" |
|
114 |
|
115 if_def "if = (%P x y.@z::'a. (P=True --> z=x) & (P=False --> z=y))" |
|
116 |
|
117 (* Rewriting -- special constant to flag normalized terms *) |
|
118 |
|
119 NORM_def "NORM(x) = x" |
|
120 |
|
121 end |
|
122 |
|
123 |
|
124 ML |
|
125 |
|
126 (** Choice between the HOL and Isabelle style of quantifiers **) |
|
127 |
|
128 val HOL_quantifiers = ref true; |
|
129 |
|
130 fun mk_alt_ast_tr' (name, alt_name) = |
|
131 let |
|
132 fun ast_tr' (*name*) args = |
|
133 if ! HOL_quantifiers then raise Match |
|
134 else Ast.mk_appl (Ast.Constant alt_name) args; |
|
135 in |
|
136 (name, ast_tr') |
|
137 end; |
|
138 |
|
139 |
|
140 val print_ast_translation = |
|
141 map mk_alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")]; |
|
142 |