src/Pure/Proof/proof_syntax.ML
author wenzelm
Wed, 06 Jul 2005 20:00:27 +0200
changeset 16721 e2427ea379a9
parent 16425 2427be27cc60
child 16866 cde0b6864924
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     1
(*  Title:      Pure/Proof/proof_syntax.ML
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
11539
0f17da240450 tuned headers;
wenzelm
parents: 11522
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     4
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     5
Function for parsing and printing proof terms.
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     6
*)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     7
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     8
signature PROOF_SYNTAX =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     9
sig
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    10
  val proofT : typ
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    11
  val add_proof_syntax : theory -> theory
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    12
  val disambiguate_names : theory -> Proofterm.proof ->
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    13
    Proofterm.proof * Proofterm.proof Symtab.table
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    14
  val proof_of_term : theory -> Proofterm.proof Symtab.table ->
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    15
    bool -> term -> Proofterm.proof
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    16
  val term_of_proof : Proofterm.proof -> term
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    17
  val cterm_of_proof : theory -> Proofterm.proof -> cterm * (cterm -> Proofterm.proof)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    18
  val read_term : theory -> typ -> string -> term
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    19
  val read_proof : theory -> bool -> string -> Proofterm.proof
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    20
  val pretty_proof : theory -> Proofterm.proof -> Pretty.T
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    21
  val pretty_proof_of : bool -> thm -> Pretty.T
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    22
  val print_proof_of : bool -> thm -> unit
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    23
end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    24
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    25
structure ProofSyntax : PROOF_SYNTAX =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    26
struct
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    27
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    28
open Proofterm;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    29
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    30
(**** add special syntax for embedding proof terms ****)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    31
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    32
val proofT = Type ("proof", []);
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    33
val paramT = Type ("param", []);
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    34
val paramsT = Type ("params", []);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    35
val idtT = Type ("idt", []);
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 13530
diff changeset
    36
val aT = TFree ("'a", []);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    37
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    38
(** constants for theorems and axioms **)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    39
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    40
fun add_proof_atom_consts names thy =
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    41
  thy
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    42
  |> Theory.absolute_path
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    43
  |> Theory.add_consts_i (map (fn name => (name, proofT, NoSyn)) names);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    44
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    45
(** constants for application and abstraction **)
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    46
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    47
fun add_proof_syntax thy =
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    48
  thy
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    49
  |> Theory.copy
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    50
  |> Theory.root_path
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    51
  |> Theory.add_defsort_i []
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    52
  |> Theory.add_types [("proof", 0, NoSyn)]
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    53
  |> Theory.add_consts_i
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    54
      [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)),
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    55
       ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)),
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    56
       ("Abst", (aT --> proofT) --> proofT, NoSyn),
13199
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
    57
       ("AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn),
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
    58
       ("Hyp", propT --> proofT, NoSyn),
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
    59
       ("Oracle", propT --> proofT, NoSyn),
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
    60
       ("MinProof", proofT, Delimfix "?")]
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    61
  |> Theory.add_nonterminals ["param", "params"]
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    62
  |> Theory.add_syntax_i
11640
be1bc3b88480 Tuned indentation of abstractions.
berghofe
parents: 11614
diff changeset
    63
      [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1Lam _./ _)", [0, 3], 3)),
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    64
       ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    65
       ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    66
       ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)),
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    67
       ("", paramT --> paramT, Delimfix "'(_')"),
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    68
       ("", idtT --> paramsT, Delimfix "_"),
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    69
       ("", paramT --> paramsT, Delimfix "_")]
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    70
  |> Theory.add_modesyntax_i ("xsymbols", true)
11640
be1bc3b88480 Tuned indentation of abstractions.
berghofe
parents: 11614
diff changeset
    71
      [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<Lambda>_./ _)", [0, 3], 3)),
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    72
       ("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    73
       ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))]
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    74
  |> Theory.add_modesyntax_i ("latex", false)
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    75
      [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<^bold>\\<lambda>_./ _)", [0, 3], 3))]
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    76
  |> Theory.add_trrules_i (map Syntax.ParsePrintRule
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    77
      [(Syntax.mk_appl (Constant "_Lam")
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    78
          [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"],
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    79
        Syntax.mk_appl (Constant "_Lam")
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    80
          [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]),
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    81
       (Syntax.mk_appl (Constant "_Lam")
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    82
          [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"],
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    83
        Syntax.mk_appl (Constant "AbsP") [Variable "A",
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    84
          (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]),
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    85
       (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"],
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    86
        Syntax.mk_appl (Constant "Abst")
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    87
          [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    88
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    89
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    90
(**** create unambiguous theorem names ****)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    91
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    92
fun disambiguate_names thy prf =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    93
  let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    94
    val thms = thms_of_proof Symtab.empty prf;
16350
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
    95
    val thms' = map (apsnd (#prop o rep_thm)) (PureThy.all_thms_of thy);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    96
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    97
    val tab = Symtab.foldl (fn (tab, (key, ps)) =>
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    98
      let val prop = getOpt (assoc (thms', key), Bound 0)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
    99
      in fst (foldr (fn ((prop', prf), x as (tab, i)) =>
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   100
        if prop <> prop' then
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   101
          (Symtab.update ((key ^ "_" ^ string_of_int i, prf), tab), i+1)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   102
        else x) (tab, 1) ps)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   103
      end) (Symtab.empty, thms);
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   104
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   105
    fun rename (Abst (s, T, prf)) = Abst (s, T, rename prf)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   106
      | rename (AbsP (s, t, prf)) = AbsP (s, t, rename prf)
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   107
      | rename (prf1 %% prf2) = rename prf1 %% rename prf2
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   108
      | rename (prf % t) = rename prf % t
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   109
      | rename (prf' as PThm ((s, tags), prf, prop, Ts)) =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   110
          let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   111
            val prop' = getOpt (assoc (thms', s), Bound 0);
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   112
            val ps = map fst (valOf (Symtab.lookup (thms, s))) \ prop'
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   113
          in if prop = prop' then prf' else
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   114
            PThm ((s ^ "_" ^ string_of_int (length ps - find_index_eq prop ps), tags),
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   115
              prf, prop, Ts)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   116
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   117
      | rename prf = prf
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   118
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   119
  in (rename prf, tab) end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   120
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   121
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   122
(**** translation between proof terms and pure terms ****)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   123
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   124
fun proof_of_term thy tab ty =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   125
  let
16350
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   126
    val thms = PureThy.all_thms_of thy;
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   127
    val axms = Theory.all_axioms_of thy;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   128
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   129
    fun mk_term t = (if ty then I else map_term_types (K dummyT))
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   130
      (Term.no_dummy_patterns t);
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   131
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   132
    fun prf_of [] (Bound i) = PBound i
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   133
      | prf_of Ts (Const (s, Type ("proof", _))) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   134
          change_type (if ty then SOME Ts else NONE)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   135
            (case NameSpace.unpack s of
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   136
               "axm" :: xs =>
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   137
                 let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   138
                   val name = NameSpace.pack xs;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   139
                   val prop = (case assoc (axms, name) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   140
                       SOME prop => prop
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   141
                     | NONE => error ("Unknown axiom " ^ quote name))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   142
                 in PAxm (name, prop, NONE) end
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   143
             | "thm" :: xs =>
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   144
                 let val name = NameSpace.pack xs;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   145
                 in (case assoc (thms, name) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   146
                     SOME thm => fst (strip_combt (Thm.proof_of thm))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   147
                   | NONE => (case Symtab.lookup (tab, name) of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   148
                         SOME prf => prf
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   149
                       | NONE => error ("Unknown theorem " ^ quote name)))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   150
                 end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   151
             | _ => error ("Illegal proof constant name: " ^ quote s))
13199
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
   152
      | prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   153
      | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   154
      | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   155
          Abst (s, if ty then SOME T else NONE,
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   156
            incr_pboundvars (~1) 0 (prf_of [] prf))
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   157
      | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   158
          AbsP (s, case t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   159
                Const ("dummy_pattern", _) => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   160
              | _ $ Const ("dummy_pattern", _) => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   161
              | _ => SOME (mk_term t),
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   162
            incr_pboundvars 0 (~1) (prf_of [] prf))
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   163
      | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   164
          prf_of [] prf1 %% prf_of [] prf2
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   165
      | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   166
          prf_of (T::Ts) prf
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   167
      | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   168
          (case t of Const ("dummy_pattern", _) => NONE | _ => SOME (mk_term t))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   169
      | prf_of _ t = error ("Not a proof term:\n" ^
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   170
          Sign.string_of_term thy t)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   171
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   172
  in prf_of [] end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   173
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   174
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   175
val AbsPt = Const ("AbsP", [propT, proofT --> proofT] ---> proofT);
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   176
val AppPt = Const ("AppP", [proofT, proofT] ---> proofT);
13199
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
   177
val Hypt = Const ("Hyp", propT --> proofT);
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
   178
val Oraclet = Const ("Oracle", propT --> proofT);
18402c1f76bf Added constants for Hyp, Oracle and MinProof.
berghofe
parents: 12909
diff changeset
   179
val MinProoft = Const ("MinProof", proofT);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   180
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   181
val mk_tyapp = Library.foldl (fn (prf, T) => Const ("Appt",
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   182
  [proofT, itselfT T] ---> proofT) $ prf $ Logic.mk_type T);
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   183
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   184
fun term_of _ (PThm ((name, _), _, _, NONE)) =
16195
wenzelm
parents: 16182
diff changeset
   185
      Const (NameSpace.append "thm" name, proofT)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   186
  | term_of _ (PThm ((name, _), _, _, SOME Ts)) =
16195
wenzelm
parents: 16182
diff changeset
   187
      mk_tyapp (Const (NameSpace.append "thm" name, proofT), Ts)
wenzelm
parents: 16182
diff changeset
   188
  | term_of _ (PAxm (name, _, NONE)) = Const (NameSpace.append "axm" name, proofT)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   189
  | term_of _ (PAxm (name, _, SOME Ts)) =
16195
wenzelm
parents: 16182
diff changeset
   190
      mk_tyapp (Const (NameSpace.append "axm" name, proofT), Ts)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   191
  | term_of _ (PBound i) = Bound i
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   192
  | term_of Ts (Abst (s, opT, prf)) = 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   193
      let val T = getOpt (opT,dummyT)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   194
      in Const ("Abst", (T --> proofT) --> proofT) $
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   195
        Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf))
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   196
      end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   197
  | term_of Ts (AbsP (s, t, prf)) =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   198
      AbsPt $ getOpt (t, Const ("dummy_pattern", propT)) $
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   199
        Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf))
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   200
  | term_of Ts (prf1 %% prf2) =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   201
      AppPt $ term_of Ts prf1 $ term_of Ts prf2
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   202
  | term_of Ts (prf % opt) = 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   203
      let val t = getOpt (opt, Const ("dummy_pattern", dummyT))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   204
      in Const ("Appt",
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   205
        [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   206
          term_of Ts prf $ t
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   207
      end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   208
  | term_of Ts (Hyp t) = Hypt $ t
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   209
  | term_of Ts (Oracle (_, t, _)) = Oraclet $ t
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   210
  | term_of Ts (MinProof _) = MinProoft;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   211
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   212
val term_of_proof = term_of [];
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   213
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   214
fun cterm_of_proof thy prf =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   215
  let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   216
    val (prf', tab) = disambiguate_names thy prf;
16350
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   217
    val thm_names = filter_out (equal "")
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   218
      (map fst (PureThy.all_thms_of thy) @ map fst (Symtab.dest tab));
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   219
    val axm_names = map fst (Theory.all_axioms_of thy);
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   220
    val thy' = thy
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   221
      |> add_proof_syntax
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   222
      |> add_proof_atom_consts
16195
wenzelm
parents: 16182
diff changeset
   223
        (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   224
  in
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   225
    (cterm_of thy' (term_of_proof prf'),
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   226
     proof_of_term thy tab true o Thm.term_of)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   227
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   228
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   229
fun read_term thy =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   230
  let
16350
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   231
    val thm_names = filter_out (equal "") (map fst (PureThy.all_thms_of thy));
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   232
    val axm_names = map fst (Theory.all_axioms_of thy);
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   233
    val thy' = thy
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   234
      |> add_proof_syntax
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   235
      |> add_proof_atom_consts
16195
wenzelm
parents: 16182
diff changeset
   236
        (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   237
  in
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   238
    fn T => fn s => Thm.term_of (read_cterm thy' (s, T))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   239
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   240
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   241
fun read_proof thy =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   242
  let val rd = read_term thy proofT
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   243
  in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   244
    (fn ty => fn s => proof_of_term thy Symtab.empty ty (Logic.varify (rd s)))
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   245
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   246
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   247
fun pretty_proof thy prf =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   248
  let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   249
    val thm_names = map fst (Symtab.dest (thms_of_proof Symtab.empty prf)) \ "";
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   250
    val axm_names = map fst (Symtab.dest (axms_of_proof Symtab.empty prf));
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   251
    val thy' = thy |>
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   252
      add_proof_syntax |>
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   253
      add_proof_atom_consts
16195
wenzelm
parents: 16182
diff changeset
   254
        (map (NameSpace.append "thm") thm_names @ map (NameSpace.append "axm") axm_names)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   255
  in
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   256
    Sign.pretty_term thy' (term_of_proof prf)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   257
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   258
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   259
fun pretty_proof_of full thm =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   260
  let
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   261
    val {thy, der = (_, prf), prop, ...} = rep_thm thm;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   262
    val prf' = (case strip_combt (fst (strip_combP prf)) of
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   263
        (PThm (_, prf', prop', _), _) => if prop=prop' then prf' else prf
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   264
      | _ => prf)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   265
  in
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   266
    pretty_proof thy
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   267
      (if full then Reconstruct.reconstruct_proof thy prop prf' else prf')
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   268
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   269
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   270
val print_proof_of = Pretty.writeln oo pretty_proof_of;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   271
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   272
end;