src/Tools/Metis/src/Proof.sml
author blanchet
Thu, 16 Sep 2010 07:30:15 +0200
changeset 39444 beabb8443ee4
parent 39443 e330437cd22a
child 39501 aaa7078fff55
permissions -rw-r--r--
MIT license -> BSD License
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     1
(* ========================================================================= *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     2
(* PROOFS IN FIRST ORDER LOGIC                                               *)
39444
beabb8443ee4 MIT license -> BSD License
blanchet
parents: 39443
diff changeset
     3
(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     4
(* ========================================================================= *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     5
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     6
structure Proof :> Proof =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     7
struct
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     8
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     9
open Useful;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    10
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    11
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    12
(* A type of first order logic proofs.                                       *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    13
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    14
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    15
datatype inference =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    16
    Axiom of LiteralSet.set
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    17
  | Assume of Atom.atom
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    18
  | Subst of Subst.subst * Thm.thm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    19
  | Resolve of Atom.atom * Thm.thm * Thm.thm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    20
  | Refl of Term.term
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    21
  | Equality of Literal.literal * Term.path * Term.term;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    22
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    23
type proof = (Thm.thm * inference) list;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    24
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    25
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    26
(* Printing.                                                                 *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    27
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    28
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    29
fun inferenceType (Axiom _) = Thm.Axiom
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    30
  | inferenceType (Assume _) = Thm.Assume
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    31
  | inferenceType (Subst _) = Thm.Subst
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    32
  | inferenceType (Resolve _) = Thm.Resolve
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    33
  | inferenceType (Refl _) = Thm.Refl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    34
  | inferenceType (Equality _) = Thm.Equality;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    35
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    36
local
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    37
  fun ppAssume atm = Print.sequence (Print.addBreak 1) (Atom.pp atm);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    38
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    39
  fun ppSubst ppThm (sub,thm) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    40
      Print.sequence (Print.addBreak 1)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    41
        (Print.blockProgram Print.Inconsistent 1
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    42
           [Print.ppString "{",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    43
            Print.ppOp2 " =" Print.ppString Subst.pp ("sub",sub),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    44
            Print.ppString ",",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    45
            Print.addBreak 1,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    46
            Print.ppOp2 " =" Print.ppString ppThm ("thm",thm),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    47
            Print.ppString "}"]);
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    48
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    49
  fun ppResolve ppThm (res,pos,neg) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    50
      Print.sequence (Print.addBreak 1)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    51
        (Print.blockProgram Print.Inconsistent 1
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    52
           [Print.ppString "{",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    53
            Print.ppOp2 " =" Print.ppString Atom.pp ("res",res),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    54
            Print.ppString ",",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    55
            Print.addBreak 1,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    56
            Print.ppOp2 " =" Print.ppString ppThm ("pos",pos),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    57
            Print.ppString ",",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    58
            Print.addBreak 1,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    59
            Print.ppOp2 " =" Print.ppString ppThm ("neg",neg),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    60
            Print.ppString "}"]);
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    61
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    62
  fun ppRefl tm = Print.sequence (Print.addBreak 1) (Term.pp tm);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    63
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    64
  fun ppEquality (lit,path,res) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    65
      Print.sequence (Print.addBreak 1)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    66
        (Print.blockProgram Print.Inconsistent 1
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    67
           [Print.ppString "{",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    68
            Print.ppOp2 " =" Print.ppString Literal.pp ("lit",lit),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    69
            Print.ppString ",",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    70
            Print.addBreak 1,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    71
            Print.ppOp2 " =" Print.ppString Term.ppPath ("path",path),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    72
            Print.ppString ",",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    73
            Print.addBreak 1,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    74
            Print.ppOp2 " =" Print.ppString Term.pp ("res",res),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    75
            Print.ppString "}"]);
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    76
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    77
  fun ppInf ppAxiom ppThm inf =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    78
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    79
        val infString = Thm.inferenceTypeToString (inferenceType inf)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    80
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    81
        Print.block Print.Inconsistent 2
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    82
          (Print.sequence
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    83
             (Print.ppString infString)
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    84
             (case inf of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    85
                Axiom cl => ppAxiom cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    86
              | Assume x => ppAssume x
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    87
              | Subst x => ppSubst ppThm x
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    88
              | Resolve x => ppResolve ppThm x
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    89
              | Refl x => ppRefl x
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    90
              | Equality x => ppEquality x))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    91
      end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    92
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    93
  fun ppAxiom cl =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    94
      Print.sequence
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    95
        (Print.addBreak 1)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    96
        (Print.ppMap
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    97
           LiteralSet.toList
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    98
           (Print.ppBracket "{" "}" (Print.ppOpList "," Literal.pp)) cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    99
in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   100
  val ppInference = ppInf ppAxiom Thm.pp;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   101
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   102
  fun pp prf =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   103
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   104
        fun thmString n = "(" ^ Int.toString n ^ ")"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   105
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   106
        val prf = enumerate prf
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   107
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   108
        fun ppThm th =
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   109
            Print.ppString
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   110
            let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   111
              val cl = Thm.clause th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   112
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   113
              fun pred (_,(th',_)) = LiteralSet.equal (Thm.clause th') cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   114
            in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   115
              case List.find pred prf of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   116
                NONE => "(?)"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   117
              | SOME (n,_) => thmString n
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   118
            end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   119
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   120
        fun ppStep (n,(th,inf)) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   121
            let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   122
              val s = thmString n
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   123
            in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   124
              Print.sequence
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   125
                (Print.blockProgram Print.Consistent (1 + size s)
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   126
                   [Print.ppString (s ^ " "),
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   127
                    Thm.pp th,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   128
                    Print.addBreak 2,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   129
                    Print.ppBracket "[" "]" (ppInf (K Print.skip) ppThm) inf])
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   130
                Print.addNewline
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   131
            end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   132
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   133
        Print.blockProgram Print.Consistent 0
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   134
          [Print.ppString "START OF PROOF",
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   135
           Print.addNewline,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   136
           Print.program (map ppStep prf),
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   137
           Print.ppString "END OF PROOF"]
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   138
      end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   139
(*MetisDebug
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   140
      handle Error err => raise Bug ("Proof.pp: shouldn't fail:\n" ^ err);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   141
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   142
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   143
end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   144
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   145
val inferenceToString = Print.toString ppInference;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   146
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   147
val toString = Print.toString pp;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   148
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   149
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   150
(* Reconstructing single inferences.                                         *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   151
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   152
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   153
fun parents (Axiom _) = []
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   154
  | parents (Assume _) = []
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   155
  | parents (Subst (_,th)) = [th]
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   156
  | parents (Resolve (_,th,th')) = [th,th']
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   157
  | parents (Refl _) = []
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   158
  | parents (Equality _) = [];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   159
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   160
fun inferenceToThm (Axiom cl) = Thm.axiom cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   161
  | inferenceToThm (Assume atm) = Thm.assume (true,atm)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   162
  | inferenceToThm (Subst (sub,th)) = Thm.subst sub th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   163
  | inferenceToThm (Resolve (atm,th,th')) = Thm.resolve (true,atm) th th'
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   164
  | inferenceToThm (Refl tm) = Thm.refl tm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   165
  | inferenceToThm (Equality (lit,path,r)) = Thm.equality lit path r;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   166
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   167
local
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   168
  fun reconstructSubst cl cl' =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   169
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   170
        fun recon [] =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   171
            let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   172
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   173
              val () = Print.trace LiteralSet.pp "reconstructSubst: cl" cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   174
              val () = Print.trace LiteralSet.pp "reconstructSubst: cl'" cl'
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   175
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   176
            in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   177
              raise Bug "can't reconstruct Subst rule"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   178
            end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   179
          | recon (([],sub) :: others) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   180
            if LiteralSet.equal (LiteralSet.subst sub cl) cl' then sub
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   181
            else recon others
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   182
          | recon ((lit :: lits, sub) :: others) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   183
            let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   184
              fun checkLit (lit',acc) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   185
                  case total (Literal.match sub lit) lit' of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   186
                    NONE => acc
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   187
                  | SOME sub => (lits,sub) :: acc
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   188
            in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   189
              recon (LiteralSet.foldl checkLit others cl')
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   190
            end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   191
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   192
        Subst.normalize (recon [(LiteralSet.toList cl, Subst.empty)])
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   193
      end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   194
(*MetisDebug
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   195
      handle Error err =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   196
        raise Bug ("Proof.recontructSubst: shouldn't fail:\n" ^ err);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   197
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   198
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   199
  fun reconstructResolvant cl1 cl2 cl =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   200
      (if not (LiteralSet.subset cl1 cl) then
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   201
         LiteralSet.pick (LiteralSet.difference cl1 cl)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   202
       else if not (LiteralSet.subset cl2 cl) then
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   203
         Literal.negate (LiteralSet.pick (LiteralSet.difference cl2 cl))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   204
       else
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   205
         (* A useless resolution, but we must reconstruct it anyway *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   206
         let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   207
           val cl1' = LiteralSet.negate cl1
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   208
           and cl2' = LiteralSet.negate cl2
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   209
           val lits = LiteralSet.intersectList [cl1,cl1',cl2,cl2']
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   210
         in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   211
           if not (LiteralSet.null lits) then LiteralSet.pick lits
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   212
           else raise Bug "can't reconstruct Resolve rule"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   213
         end)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   214
(*MetisDebug
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   215
      handle Error err =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   216
        raise Bug ("Proof.recontructResolvant: shouldn't fail:\n" ^ err);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   217
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   218
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   219
  fun reconstructEquality cl =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   220
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   221
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   222
        val () = Print.trace LiteralSet.pp "Proof.reconstructEquality: cl" cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   223
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   224
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   225
        fun sync s t path (f,a) (f',a') =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   226
            if not (Name.equal f f' andalso length a = length a') then NONE
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   227
            else
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   228
              let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   229
                val itms = enumerate (zip a a')
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   230
              in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   231
                case List.filter (not o uncurry Term.equal o snd) itms of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   232
                  [(i,(tm,tm'))] =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   233
                  let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   234
                    val path = i :: path
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   235
                  in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   236
                    if Term.equal tm s andalso Term.equal tm' t then
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   237
                      SOME (rev path)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   238
                    else
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   239
                      case (tm,tm') of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   240
                        (Term.Fn f_a, Term.Fn f_a') => sync s t path f_a f_a'
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   241
                      | _ => NONE
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   242
                  end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   243
                | _ => NONE
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   244
              end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   245
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   246
        fun recon (neq,(pol,atm),(pol',atm')) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   247
            if pol = pol' then NONE
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   248
            else
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   249
              let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   250
                val (s,t) = Literal.destNeq neq
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   251
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   252
                val path =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   253
                    if not (Term.equal s t) then sync s t [] atm atm'
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   254
                    else if not (Atom.equal atm atm') then NONE
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   255
                    else Atom.find (Term.equal s) atm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   256
              in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   257
                case path of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   258
                  SOME path => SOME ((pol',atm),path,t)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   259
                | NONE => NONE
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   260
              end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   261
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   262
        val candidates =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   263
            case List.partition Literal.isNeq (LiteralSet.toList cl) of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   264
              ([l1],[l2,l3]) => [(l1,l2,l3),(l1,l3,l2)]
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   265
            | ([l1,l2],[l3]) => [(l1,l2,l3),(l1,l3,l2),(l2,l1,l3),(l2,l3,l1)]
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   266
            | ([l1],[l2]) => [(l1,l1,l2),(l1,l2,l1)]
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   267
            | _ => raise Bug "reconstructEquality: malformed"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   268
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   269
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   270
        val ppCands =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   271
            Print.ppList (Print.ppTriple Literal.pp Literal.pp Literal.pp)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   272
        val () = Print.trace ppCands
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   273
                   "Proof.reconstructEquality: candidates" candidates
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   274
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   275
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   276
        case first recon candidates of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   277
          SOME info => info
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   278
        | NONE => raise Bug "can't reconstruct Equality rule"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   279
      end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   280
(*MetisDebug
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   281
      handle Error err =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   282
        raise Bug ("Proof.recontructEquality: shouldn't fail:\n" ^ err);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   283
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   284
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   285
  fun reconstruct cl (Thm.Axiom,[]) = Axiom cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   286
    | reconstruct cl (Thm.Assume,[]) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   287
      (case LiteralSet.findl Literal.positive cl of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   288
         SOME (_,atm) => Assume atm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   289
       | NONE => raise Bug "malformed Assume inference")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   290
    | reconstruct cl (Thm.Subst,[th]) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   291
      Subst (reconstructSubst (Thm.clause th) cl, th)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   292
    | reconstruct cl (Thm.Resolve,[th1,th2]) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   293
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   294
        val cl1 = Thm.clause th1
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   295
        and cl2 = Thm.clause th2
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   296
        val (pol,atm) = reconstructResolvant cl1 cl2 cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   297
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   298
        if pol then Resolve (atm,th1,th2) else Resolve (atm,th2,th1)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   299
      end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   300
    | reconstruct cl (Thm.Refl,[]) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   301
      (case LiteralSet.findl (K true) cl of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   302
         SOME lit => Refl (Literal.destRefl lit)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   303
       | NONE => raise Bug "malformed Refl inference")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   304
    | reconstruct cl (Thm.Equality,[]) = Equality (reconstructEquality cl)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   305
    | reconstruct _ _ = raise Bug "malformed inference";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   306
in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   307
  fun thmToInference th =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   308
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   309
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   310
        val () = Print.trace Thm.pp "Proof.thmToInference: th" th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   311
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   312
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   313
        val cl = Thm.clause th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   314
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   315
        val thmInf = Thm.inference th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   316
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   317
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   318
        val ppThmInf = Print.ppPair Thm.ppInferenceType (Print.ppList Thm.pp)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   319
        val () = Print.trace ppThmInf "Proof.thmToInference: thmInf" thmInf
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   320
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   321
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   322
        val inf = reconstruct cl thmInf
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   323
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   324
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   325
        val () = Print.trace ppInference "Proof.thmToInference: inf" inf
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   326
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   327
(*MetisDebug
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   328
        val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   329
            let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   330
              val th' = inferenceToThm inf
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   331
            in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   332
              if LiteralSet.equal (Thm.clause th') cl then ()
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   333
              else
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   334
                raise
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   335
                  Bug
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   336
                    ("Proof.thmToInference: bad inference reconstruction:" ^
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   337
                     "\n  th = " ^ Thm.toString th ^
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   338
                     "\n  inf = " ^ inferenceToString inf ^
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   339
                     "\n  inf th = " ^ Thm.toString th')
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   340
            end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   341
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   342
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   343
        inf
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   344
      end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   345
(*MetisDebug
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   346
      handle Error err =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   347
        raise Bug ("Proof.thmToInference: shouldn't fail:\n" ^ err);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   348
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   349
end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   350
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   351
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   352
(* Reconstructing whole proofs.                                              *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   353
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   354
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   355
local
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   356
  val emptyThms : Thm.thm LiteralSetMap.map = LiteralSetMap.new ();
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   357
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   358
  fun addThms (th,ths) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   359
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   360
        val cl = Thm.clause th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   361
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   362
        if LiteralSetMap.inDomain cl ths then ths
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   363
        else
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   364
          let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   365
            val (_,pars) = Thm.inference th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   366
            val ths = List.foldl addThms ths pars
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   367
          in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   368
            if LiteralSetMap.inDomain cl ths then ths
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   369
            else LiteralSetMap.insert ths (cl,th)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   370
          end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   371
      end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   372
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   373
  fun mkThms th = addThms (th,emptyThms);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   374
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   375
  fun addProof (th,(ths,acc)) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   376
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   377
        val cl = Thm.clause th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   378
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   379
        case LiteralSetMap.peek ths cl of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   380
          NONE => (ths,acc)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   381
        | SOME th =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   382
          let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   383
            val (_,pars) = Thm.inference th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   384
            val (ths,acc) = List.foldl addProof (ths,acc) pars
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   385
            val ths = LiteralSetMap.delete ths cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   386
            val acc = (th, thmToInference th) :: acc
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   387
          in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   388
            (ths,acc)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   389
          end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   390
      end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   391
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   392
  fun mkProof ths th =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   393
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   394
        val (ths,acc) = addProof (th,(ths,[]))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   395
(*MetisTrace4
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   396
        val () = Print.trace Print.ppInt "Proof.proof: unnecessary clauses" (LiteralSetMap.size ths)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   397
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   398
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   399
        rev acc
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   400
      end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   401
in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   402
  fun proof th =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   403
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   404
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   405
        val () = Print.trace Thm.pp "Proof.proof: th" th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   406
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   407
        val ths = mkThms th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   408
        val infs = mkProof ths th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   409
(*MetisTrace3
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   410
        val () = Print.trace Print.ppInt "Proof.proof: size" (length infs)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   411
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   412
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   413
        infs
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   414
      end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   415
end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   416
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   417
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   418
(* Free variables.                                                           *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   419
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   420
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   421
fun freeIn v =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   422
    let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   423
      fun free th_inf =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   424
          case th_inf of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   425
            (_, Axiom lits) => LiteralSet.freeIn v lits
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   426
          | (_, Assume atm) => Atom.freeIn v atm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   427
          | (th, Subst _) => Thm.freeIn v th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   428
          | (_, Resolve _) => false
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   429
          | (_, Refl tm) => Term.freeIn v tm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   430
          | (_, Equality (lit,_,tm)) =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   431
            Literal.freeIn v lit orelse Term.freeIn v tm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   432
    in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   433
      List.exists free
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   434
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   435
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   436
val freeVars =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   437
    let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   438
      fun inc (th_inf,set) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   439
          NameSet.union set
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   440
          (case th_inf of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   441
             (_, Axiom lits) => LiteralSet.freeVars lits
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   442
           | (_, Assume atm) => Atom.freeVars atm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   443
           | (th, Subst _) => Thm.freeVars th
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   444
           | (_, Resolve _) => NameSet.empty
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   445
           | (_, Refl tm) => Term.freeVars tm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   446
           | (_, Equality (lit,_,tm)) =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   447
             NameSet.union (Literal.freeVars lit) (Term.freeVars tm))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   448
    in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   449
      List.foldl inc NameSet.empty
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   450
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   451
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   452
end