39348
|
1 |
(* ========================================================================= *)
|
|
2 |
(* PROOFS IN FIRST ORDER LOGIC *)
|
72004
|
3 |
(* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *)
|
39348
|
4 |
(* ========================================================================= *)
|
|
5 |
|
|
6 |
signature Proof =
|
|
7 |
sig
|
|
8 |
|
|
9 |
(* ------------------------------------------------------------------------- *)
|
|
10 |
(* A type of first order logic proofs. *)
|
|
11 |
(* ------------------------------------------------------------------------- *)
|
|
12 |
|
|
13 |
datatype inference =
|
|
14 |
Axiom of LiteralSet.set
|
|
15 |
| Assume of Atom.atom
|
|
16 |
| Subst of Subst.subst * Thm.thm
|
|
17 |
| Resolve of Atom.atom * Thm.thm * Thm.thm
|
|
18 |
| Refl of Term.term
|
|
19 |
| Equality of Literal.literal * Term.path * Term.term
|
|
20 |
|
|
21 |
type proof = (Thm.thm * inference) list
|
|
22 |
|
|
23 |
(* ------------------------------------------------------------------------- *)
|
|
24 |
(* Reconstructing single inferences. *)
|
|
25 |
(* ------------------------------------------------------------------------- *)
|
|
26 |
|
|
27 |
val inferenceType : inference -> Thm.inferenceType
|
|
28 |
|
|
29 |
val parents : inference -> Thm.thm list
|
|
30 |
|
|
31 |
val inferenceToThm : inference -> Thm.thm
|
|
32 |
|
|
33 |
val thmToInference : Thm.thm -> inference
|
|
34 |
|
|
35 |
(* ------------------------------------------------------------------------- *)
|
|
36 |
(* Reconstructing whole proofs. *)
|
|
37 |
(* ------------------------------------------------------------------------- *)
|
|
38 |
|
|
39 |
val proof : Thm.thm -> proof
|
|
40 |
|
|
41 |
(* ------------------------------------------------------------------------- *)
|
|
42 |
(* Free variables. *)
|
|
43 |
(* ------------------------------------------------------------------------- *)
|
|
44 |
|
|
45 |
val freeIn : Term.var -> proof -> bool
|
|
46 |
|
|
47 |
val freeVars : proof -> NameSet.set
|
|
48 |
|
|
49 |
(* ------------------------------------------------------------------------- *)
|
|
50 |
(* Printing. *)
|
|
51 |
(* ------------------------------------------------------------------------- *)
|
|
52 |
|
|
53 |
val ppInference : inference Print.pp
|
|
54 |
|
|
55 |
val inferenceToString : inference -> string
|
|
56 |
|
|
57 |
val pp : proof Print.pp
|
|
58 |
|
|
59 |
val toString : proof -> string
|
|
60 |
|
|
61 |
end
|