src/Pure/Proof/proof_syntax.ML
author wenzelm
Wed, 09 Oct 2019 22:22:17 +0200
changeset 70813 83b7d1927fda
parent 70493 a9053fa30909
child 70839 2136e4670ad2
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
11539
0f17da240450 tuned headers;
wenzelm
parents: 11522
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     3
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     4
Function for parsing and printing proof terms.
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     5
*)
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
signature PROOF_SYNTAX =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     8
sig
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
     9
  val add_proof_syntax: theory -> theory
28807
9f3ecb4aaac2 proof_of_term: removed obsolete disambiguisation table;
wenzelm
parents: 28375
diff changeset
    10
  val proof_of_term: theory -> bool -> term -> Proofterm.proof
37227
bdd8dd217b1f - Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
berghofe
parents: 36610
diff changeset
    11
  val read_term: theory -> bool -> typ -> string -> term
bdd8dd217b1f - Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
berghofe
parents: 36610
diff changeset
    12
  val read_proof: theory -> bool -> bool -> string -> Proofterm.proof
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
    13
  val proof_syntax: Proofterm.proof -> theory -> theory
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
    14
  val proof_of: bool -> thm -> Proofterm.proof
27260
17d617c6b026 moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
wenzelm
parents: 26939
diff changeset
    15
  val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T
64986
b81a048960a3 more uniform use of Reconstruct.clean_proof_of;
wenzelm
parents: 64556
diff changeset
    16
  val pretty_clean_proof_of: Proof.context -> bool -> thm -> Pretty.T
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    17
end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    18
33388
d64545e6cba5 modernized structure Proof_Syntax;
wenzelm
parents: 31943
diff changeset
    19
structure Proof_Syntax : PROOF_SYNTAX =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    20
struct
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    21
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    22
(**** add special syntax for embedding proof terms ****)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    23
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    24
val proofT = Proofterm.proofT;
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    25
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    26
local
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    27
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    28
val paramT = Type ("param", []);
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    29
val paramsT = Type ("params", []);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    30
val idtT = Type ("idt", []);
70387
wenzelm
parents: 67649
diff changeset
    31
val aT = Term.aT [];
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    32
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    33
fun mixfix (sy, ps, p) = Mixfix (Input.string sy, ps, p, Position.no_range);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    34
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    35
in
62752
d09d71223e7a more position information for type mixfix;
wenzelm
parents: 61957
diff changeset
    36
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    37
fun add_proof_syntax thy =
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
    38
  thy
22796
34c316d7b630 renamed some old names Theory.xxx to Sign.xxx;
wenzelm
parents: 22675
diff changeset
    39
  |> Sign.root_path
36449
78721f3adb13 modernized/simplified Sign.set_defsort;
wenzelm
parents: 35845
diff changeset
    40
  |> Sign.set_defsort []
56436
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    41
  |> Sign.add_nonterminals_global
64556
851ae0e7b09c more symbols;
wenzelm
parents: 62958
diff changeset
    42
    [Binding.make ("param", \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 62958
diff changeset
    43
     Binding.make ("params", \<^here>)]
56240
938c6c7e10eb tuned signature;
wenzelm
parents: 56239
diff changeset
    44
  |> Sign.add_syntax Syntax.mode_default
62752
d09d71223e7a more position information for type mixfix;
wenzelm
parents: 61957
diff changeset
    45
    [("_Lam", [paramsT, proofT] ---> proofT, mixfix ("(1\<^bold>\<lambda>_./ _)", [0, 3], 3)),
d09d71223e7a more position information for type mixfix;
wenzelm
parents: 61957
diff changeset
    46
     ("_Lam0", [paramT, paramsT] ---> paramsT, mixfix ("_/ _", [1, 0], 0)),
d09d71223e7a more position information for type mixfix;
wenzelm
parents: 61957
diff changeset
    47
     ("_Lam0", [idtT, paramsT] ---> paramsT, mixfix ("_/ _", [1, 0], 0)),
d09d71223e7a more position information for type mixfix;
wenzelm
parents: 61957
diff changeset
    48
     ("_Lam1", [idtT, propT] ---> paramT, mixfix ("_: _", [0, 0], 0)),
62761
5c672b22dcc2 clarified simple mixfix;
wenzelm
parents: 62752
diff changeset
    49
     ("", paramT --> paramT, Mixfix.mixfix "'(_')"),
5c672b22dcc2 clarified simple mixfix;
wenzelm
parents: 62752
diff changeset
    50
     ("", idtT --> paramsT, Mixfix.mixfix "_"),
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    51
     ("", paramT --> paramsT, Mixfix.mixfix "_"),
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    52
     (Lexicon.mark_const "Pure.Appt", [proofT, aT] ---> proofT, mixfix ("(1_ \<cdot>/ _)", [4, 5], 4)),
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    53
     (Lexicon.mark_const "Pure.AppP", [proofT, proofT] ---> proofT, mixfix ("(1_ \<bullet>/ _)", [4, 5], 4)),
70407
e8558735961a clarified syntax;
wenzelm
parents: 70390
diff changeset
    54
     (Lexicon.mark_const "Pure.MinProof", proofT, Mixfix.mixfix "\<^bold>?")]
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 39557
diff changeset
    55
  |> Sign.add_trrules (map Syntax.Parse_Print_Rule
56436
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    56
    [(Ast.mk_appl (Ast.Constant "_Lam")
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    57
        [Ast.mk_appl (Ast.Constant "_Lam0")
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    58
          [Ast.Variable "l", Ast.Variable "m"], Ast.Variable "A"],
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    59
      Ast.mk_appl (Ast.Constant "_Lam")
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    60
        [Ast.Variable "l",
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    61
          Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "m", Ast.Variable "A"]]),
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    62
     (Ast.mk_appl (Ast.Constant "_Lam")
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    63
        [Ast.mk_appl (Ast.Constant "_Lam1")
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    64
          [Ast.Variable "x", Ast.Variable "A"], Ast.Variable "B"],
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    65
      Ast.mk_appl (Ast.Constant (Lexicon.mark_const "Pure.AbsP")) [Ast.Variable "A",
56436
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    66
        (Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "B"])]),
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    67
     (Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "x", Ast.Variable "A"],
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    68
      Ast.mk_appl (Ast.Constant (Lexicon.mark_const "Pure.Abst"))
56436
30ccec1e82fb more source positions;
wenzelm
parents: 56243
diff changeset
    69
        [(Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "A"])])]);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    70
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    71
end;
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    72
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    73
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    74
(** constants for theorems and axioms **)
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    75
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    76
fun add_proof_atom_consts names thy =
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    77
  thy
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    78
  |> Sign.root_path
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    79
  |> Sign.add_consts (map (fn name => (Binding.qualified_name name, proofT, NoSyn)) names);
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    80
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    81
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    82
(**** translation between proof terms and pure terms ****)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    83
28807
9f3ecb4aaac2 proof_of_term: removed obsolete disambiguisation table;
wenzelm
parents: 28375
diff changeset
    84
fun proof_of_term thy ty =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    85
  let
56161
300f613060b0 tuned signature;
wenzelm
parents: 55725
diff changeset
    86
    val thms = Global_Theory.all_thms_of thy true;
16350
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
    87
    val axms = Theory.all_axioms_of thy;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    88
20548
8ef25fe585a8 renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents: 19618
diff changeset
    89
    fun mk_term t = (if ty then I else map_types (K dummyT))
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    90
      (Term.no_dummy_patterns t);
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    91
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    92
    fun prf_of [] (Bound i) = PBound i
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
    93
      | prf_of Ts (Const (s, Type ("Pure.proof", _))) =
70417
eb6ff14767cd tuned signature;
wenzelm
parents: 70412
diff changeset
    94
          Proofterm.change_types (if ty then SOME Ts else NONE)
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30344
diff changeset
    95
            (case Long_Name.explode s of
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
    96
               "axm" :: xs =>
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    97
                 let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30344
diff changeset
    98
                   val name = Long_Name.implode xs;
17223
430edc6b7826 curried_lookup/update;
wenzelm
parents: 17078
diff changeset
    99
                   val prop = (case AList.lookup (op =) axms name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   100
                       SOME prop => prop
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   101
                     | NONE => error ("Unknown axiom " ^ quote name))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   102
                 in PAxm (name, prop, NONE) end
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   103
             | "thm" :: xs =>
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30344
diff changeset
   104
                 let val name = Long_Name.implode xs;
17223
430edc6b7826 curried_lookup/update;
wenzelm
parents: 17078
diff changeset
   105
                 in (case AList.lookup (op =) thms name of
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37236
diff changeset
   106
                     SOME thm =>
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37236
diff changeset
   107
                      fst (Proofterm.strip_combt (fst (Proofterm.strip_combP (Thm.proof_of thm))))
28807
9f3ecb4aaac2 proof_of_term: removed obsolete disambiguisation table;
wenzelm
parents: 28375
diff changeset
   108
                   | NONE => error ("Unknown theorem " ^ quote name))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   109
                 end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   110
             | _ => error ("Illegal proof constant name: " ^ quote s))
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   111
      | prf_of Ts (Const ("Pure.OfClass", _) $ Const (c_class, _)) =
31903
c5221dbc40f6 added pro-forma proof constructor Inclass;
wenzelm
parents: 30435
diff changeset
   112
          (case try Logic.class_of_const c_class of
c5221dbc40f6 added pro-forma proof constructor Inclass;
wenzelm
parents: 30435
diff changeset
   113
            SOME c =>
70417
eb6ff14767cd tuned signature;
wenzelm
parents: 70412
diff changeset
   114
              Proofterm.change_types (if ty then SOME Ts else NONE)
31943
5e960a0780a2 renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
wenzelm
parents: 31903
diff changeset
   115
                (OfClass (TVar ((Name.aT, 0), []), c))
31903
c5221dbc40f6 added pro-forma proof constructor Inclass;
wenzelm
parents: 30435
diff changeset
   116
          | NONE => error ("Bad class constant: " ^ quote c_class))
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   117
      | prf_of Ts (Const ("Pure.Hyp", _) $ prop) = Hyp prop
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   118
      | prf_of Ts (v as Var ((_, Type ("Pure.proof", _)))) = Hyp v
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   119
      | prf_of [] (Const ("Pure.Abst", _) $ Abs (s, T, prf)) =
25245
1fcfcdcba53c Added well-formedness check to Abst case in function prf_of.
berghofe
parents: 24848
diff changeset
   120
          if T = proofT then
1fcfcdcba53c Added well-formedness check to Abst case in function prf_of.
berghofe
parents: 24848
diff changeset
   121
            error ("Term variable abstraction may not bind proof variable " ^ quote s)
1fcfcdcba53c Added well-formedness check to Abst case in function prf_of.
berghofe
parents: 24848
diff changeset
   122
          else Abst (s, if ty then SOME T else NONE,
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37236
diff changeset
   123
            Proofterm.incr_pboundvars (~1) 0 (prf_of [] prf))
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   124
      | prf_of [] (Const ("Pure.AbsP", _) $ t $ Abs (s, _, prf)) =
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   125
          AbsP (s, case t of
56241
029246729dc0 more qualified names;
wenzelm
parents: 56240
diff changeset
   126
                Const ("Pure.dummy_pattern", _) => NONE
029246729dc0 more qualified names;
wenzelm
parents: 56240
diff changeset
   127
              | _ $ Const ("Pure.dummy_pattern", _) => NONE
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   128
              | _ => SOME (mk_term t),
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37236
diff changeset
   129
            Proofterm.incr_pboundvars 0 (~1) (prf_of [] prf))
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   130
      | prf_of [] (Const ("Pure.AppP", _) $ prf1 $ prf2) =
11614
3131fa12d425 - Tuned syntax
berghofe
parents: 11539
diff changeset
   131
          prf_of [] prf1 %% prf_of [] prf2
70418
d23cfb85438a proper argument type (amending 42fbb6abed5a);
wenzelm
parents: 70417
diff changeset
   132
      | prf_of Ts (Const ("Pure.Appt", _) $ prf $ Const ("Pure.type", Type ("itself", [T]))) =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   133
          prf_of (T::Ts) prf
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   134
      | prf_of [] (Const ("Pure.Appt", _) $ prf $ t) = prf_of [] prf %
56241
029246729dc0 more qualified names;
wenzelm
parents: 56240
diff changeset
   135
          (case t of Const ("Pure.dummy_pattern", _) => NONE | _ => SOME (mk_term t))
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   136
      | prf_of _ t = error ("Not a proof term:\n" ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26626
diff changeset
   137
          Syntax.string_of_term_global thy t)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   138
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   139
  in prf_of [] end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   140
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   141
37227
bdd8dd217b1f - Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
berghofe
parents: 36610
diff changeset
   142
fun read_term thy topsort =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   143
  let
56161
300f613060b0 tuned signature;
wenzelm
parents: 55725
diff changeset
   144
    val thm_names = filter_out (fn s => s = "") (map fst (Global_Theory.all_thms_of thy true));
16350
caa9b780ad91 Theory.all_axioms_of, PureThy.all_thms_of;
wenzelm
parents: 16195
diff changeset
   145
    val axm_names = map fst (Theory.all_axioms_of thy);
27260
17d617c6b026 moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
wenzelm
parents: 26939
diff changeset
   146
    val ctxt = thy
16425
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   147
      |> add_proof_syntax
2427be27cc60 accomodate identification of type Sign.sg and theory;
wenzelm
parents: 16350
diff changeset
   148
      |> add_proof_atom_consts
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30344
diff changeset
   149
        (map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names)
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42290
diff changeset
   150
      |> Proof_Context.init_global
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42290
diff changeset
   151
      |> Proof_Context.allow_dummies
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42290
diff changeset
   152
      |> Proof_Context.set_mode Proof_Context.mode_schematic
62958
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62922
diff changeset
   153
      |> topsort ?
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62922
diff changeset
   154
        (Proof_Context.set_defsort [] #>
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62922
diff changeset
   155
         Config.put Type_Infer.object_logic false #>
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62922
diff changeset
   156
         Config.put Type_Infer_Context.const_sorts false);
27260
17d617c6b026 moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
wenzelm
parents: 26939
diff changeset
   157
  in
17d617c6b026 moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
wenzelm
parents: 26939
diff changeset
   158
    fn ty => fn s =>
17d617c6b026 moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
wenzelm
parents: 26939
diff changeset
   159
      (if ty = propT then Syntax.parse_prop else Syntax.parse_term) ctxt s
39288
f1ae2493d93f eliminated aliases of Type.constraint;
wenzelm
parents: 37310
diff changeset
   160
      |> Type.constraint ty |> Syntax.check_term ctxt
27260
17d617c6b026 moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
wenzelm
parents: 26939
diff changeset
   161
  end;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   162
37227
bdd8dd217b1f - Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
berghofe
parents: 36610
diff changeset
   163
fun read_proof thy topsort =
bdd8dd217b1f - Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
berghofe
parents: 36610
diff changeset
   164
  let val rd = read_term thy topsort proofT
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35262
diff changeset
   165
  in fn ty => fn s => proof_of_term thy ty (Logic.varify_global (rd s)) end;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   166
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
   167
fun proof_syntax prf =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   168
  let
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37236
diff changeset
   169
    val thm_names = Symtab.keys (Proofterm.fold_proof_atoms true
70493
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   170
      (fn PThm ({name, ...}, _) => if name <> "" then Symtab.update (name, ()) else I
28807
9f3ecb4aaac2 proof_of_term: removed obsolete disambiguisation table;
wenzelm
parents: 28375
diff changeset
   171
        | _ => I) [prf] Symtab.empty);
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37236
diff changeset
   172
    val axm_names = Symtab.keys (Proofterm.fold_proof_atoms true
28807
9f3ecb4aaac2 proof_of_term: removed obsolete disambiguisation table;
wenzelm
parents: 28375
diff changeset
   173
      (fn PAxm (name, _, _) => Symtab.update (name, ()) | _ => I) [prf] Symtab.empty);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   174
  in
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
   175
    add_proof_syntax #>
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
   176
    add_proof_atom_consts
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30344
diff changeset
   177
      (map (Long_Name.append "thm") thm_names @ map (Long_Name.append "axm") axm_names)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   178
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   179
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   180
fun proof_of full thm =
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
   181
  let
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   182
    val thy = Thm.theory_of_thm thm;
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
   183
    val prop = Thm.full_prop_of thm;
28814
463c9e9111ae clarified Thm.proof_body_of vs. Thm.proof_of;
wenzelm
parents: 28807
diff changeset
   184
    val prf = Thm.proof_of thm;
70813
wenzelm
parents: 70493
diff changeset
   185
  in
wenzelm
parents: 70493
diff changeset
   186
    (case fst (Proofterm.strip_combt (fst (Proofterm.strip_combP prf))) of
wenzelm
parents: 70493
diff changeset
   187
      PThm ({prop = prop', ...}, thm_body) =>
wenzelm
parents: 70493
diff changeset
   188
        if prop = prop' then Proofterm.thm_body_proof_raw thm_body else prf
wenzelm
parents: 70493
diff changeset
   189
    | _ => prf)
wenzelm
parents: 70493
diff changeset
   190
    |> full ? Proofterm.reconstruct_proof thy prop
wenzelm
parents: 70493
diff changeset
   191
  end;
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
   192
27260
17d617c6b026 moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
wenzelm
parents: 26939
diff changeset
   193
fun pretty_proof ctxt prf =
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42290
diff changeset
   194
  Proof_Context.pretty_term_abbrev
55725
9d605a21d7ec prefer standard Proof_Context.transfer, with theory stamp transfer (should now work thanks to purely functional theory, without Theory.copy etc.);
wenzelm
parents: 52788
diff changeset
   195
    (Proof_Context.transfer (proof_syntax prf (Proof_Context.theory_of ctxt)) ctxt)
70388
e31271559de8 global declaration of abstract syntax for proof terms, with qualified names;
wenzelm
parents: 70387
diff changeset
   196
    (Proofterm.term_of_proof prf);
17078
db9d24c8b439 export proof_syntax, proof_of;
wenzelm
parents: 17019
diff changeset
   197
64986
b81a048960a3 more uniform use of Reconstruct.clean_proof_of;
wenzelm
parents: 64556
diff changeset
   198
fun pretty_clean_proof_of ctxt full thm =
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   199
  pretty_proof ctxt (Thm.clean_proof_of full thm);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   200
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   201
end;