author | kleing |
Sun, 09 Dec 2001 14:36:14 +0100 | |
changeset 12432 | 90b0fc84f8d9 |
parent 3836 | f1a1817659e6 |
child 17480 | fd19f77dcf60 |
permissions | -rw-r--r-- |
1477 | 1 |
(* Title: FOLP/FOLP.thy |
1142 | 2 |
ID: $Id$ |
1477 | 3 |
Author: Martin D Coen, Cambridge University Computer Laboratory |
1142 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
Classical First-Order Logic with Proofs |
|
7 |
*) |
|
8 |
||
0 | 9 |
FOLP = IFOLP + |
10 |
consts |
|
11 |
cla :: "[p=>p]=>p" |
|
12 |
rules |
|
3836 | 13 |
classical "(!!x. x:~P ==> f(x):P) ==> cla(f):P" |
0 | 14 |
end |