| author | huffman | 
| Wed, 25 Feb 2009 11:29:59 -0800 | |
| changeset 30097 | 57df8626c23b | 
| parent 29635 | 31d14e9fa0da | 
| child 30344 | 10a67c5ddddb | 
| permissions | -rw-r--r-- | 
| 11522 | 1 | (* Title: Pure/Proof/proof_syntax.ML | 
| 11539 | 2 | Author: Stefan Berghofer, TU Muenchen | 
| 11522 | 3 | |
| 4 | Function for parsing and printing proof terms. | |
| 5 | *) | |
| 6 | ||
| 7 | signature PROOF_SYNTAX = | |
| 8 | sig | |
| 17078 | 9 | val proofT: typ | 
| 10 | val add_proof_syntax: theory -> theory | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 11 | val proof_of_term: theory -> bool -> term -> Proofterm.proof | 
| 17078 | 12 | val term_of_proof: Proofterm.proof -> term | 
| 13 | val cterm_of_proof: theory -> Proofterm.proof -> cterm * (cterm -> Proofterm.proof) | |
| 14 | val read_term: theory -> typ -> string -> term | |
| 15 | val read_proof: theory -> bool -> string -> Proofterm.proof | |
| 16 | val proof_syntax: Proofterm.proof -> theory -> theory | |
| 17 | val proof_of: bool -> thm -> Proofterm.proof | |
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 18 | val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 19 | val pretty_proof_of: Proof.context -> bool -> thm -> Pretty.T | 
| 11522 | 20 | end; | 
| 21 | ||
| 22 | structure ProofSyntax : PROOF_SYNTAX = | |
| 23 | struct | |
| 24 | ||
| 25 | open Proofterm; | |
| 26 | ||
| 27 | (**** add special syntax for embedding proof terms ****) | |
| 28 | ||
| 29 | val proofT = Type ("proof", []);
 | |
| 11614 | 30 | val paramT = Type ("param", []);
 | 
| 31 | val paramsT = Type ("params", []);
 | |
| 11522 | 32 | val idtT = Type ("idt", []);
 | 
| 24848 | 33 | val aT = TFree (Name.aT, []); | 
| 11522 | 34 | |
| 35 | (** constants for theorems and axioms **) | |
| 36 | ||
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 37 | fun add_proof_atom_consts names thy = | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 38 | thy | 
| 22796 | 39 | |> Sign.absolute_path | 
| 40 | |> Sign.add_consts_i (map (fn name => (name, proofT, NoSyn)) names); | |
| 11522 | 41 | |
| 42 | (** constants for application and abstraction **) | |
| 11614 | 43 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 44 | fun add_proof_syntax thy = | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 45 | thy | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 46 | |> Theory.copy | 
| 22796 | 47 | |> Sign.root_path | 
| 48 | |> Sign.add_defsort_i [] | |
| 49 |   |> Sign.add_types [("proof", 0, NoSyn)]
 | |
| 50 | |> Sign.add_consts_i | |
| 11614 | 51 |       [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)),
 | 
| 52 |        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)),
 | |
| 11522 | 53 |        ("Abst", (aT --> proofT) --> proofT, NoSyn),
 | 
| 13199 | 54 |        ("AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn),
 | 
| 55 |        ("Hyp", propT --> proofT, NoSyn),
 | |
| 56 |        ("Oracle", propT --> proofT, NoSyn),
 | |
| 57 |        ("MinProof", proofT, Delimfix "?")]
 | |
| 22796 | 58 | |> Sign.add_nonterminals ["param", "params"] | 
| 59 | |> Sign.add_syntax_i | |
| 11640 | 60 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1Lam _./ _)", [0, 3], 3)),
 | 
| 11614 | 61 |        ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
 | 
| 62 |        ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
 | |
| 63 |        ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)),
 | |
| 64 |        ("", paramT --> paramT, Delimfix "'(_')"),
 | |
| 65 |        ("", idtT --> paramsT, Delimfix "_"),
 | |
| 66 |        ("", paramT --> paramsT, Delimfix "_")]
 | |
| 22796 | 67 |   |> Sign.add_modesyntax_i ("xsymbols", true)
 | 
| 11640 | 68 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<Lambda>_./ _)", [0, 3], 3)),
 | 
| 11522 | 69 |        ("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 70 |        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))]
 | 
| 22796 | 71 |   |> Sign.add_modesyntax_i ("latex", false)
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 72 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<^bold>\\<lambda>_./ _)", [0, 3], 3))]
 | 
| 22796 | 73 | |> Sign.add_trrules_i (map Syntax.ParsePrintRule | 
| 11522 | 74 | [(Syntax.mk_appl (Constant "_Lam") | 
| 11614 | 75 | [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"], | 
| 76 | Syntax.mk_appl (Constant "_Lam") | |
| 77 | [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]), | |
| 78 | (Syntax.mk_appl (Constant "_Lam") | |
| 11522 | 79 | [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"], | 
| 80 | Syntax.mk_appl (Constant "AbsP") [Variable "A", | |
| 81 | (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]), | |
| 11614 | 82 | (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"], | 
| 11522 | 83 | Syntax.mk_appl (Constant "Abst") | 
| 11614 | 84 | [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]); | 
| 11522 | 85 | |
| 86 | ||
| 87 | (**** translation between proof terms and pure terms ****) | |
| 88 | ||
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 89 | fun proof_of_term thy ty = | 
| 11522 | 90 | let | 
| 16350 | 91 | val thms = PureThy.all_thms_of thy; | 
| 92 | val axms = Theory.all_axioms_of thy; | |
| 11522 | 93 | |
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
19618diff
changeset | 94 | fun mk_term t = (if ty then I else map_types (K dummyT)) | 
| 11614 | 95 | (Term.no_dummy_patterns t); | 
| 96 | ||
| 11522 | 97 | fun prf_of [] (Bound i) = PBound i | 
| 98 |       | prf_of Ts (Const (s, Type ("proof", _))) =
 | |
| 15531 | 99 | change_type (if ty then SOME Ts else NONE) | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 100 | (case NameSpace.explode s of | 
| 11614 | 101 | "axm" :: xs => | 
| 11522 | 102 | let | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 103 | val name = NameSpace.implode xs; | 
| 17223 | 104 | val prop = (case AList.lookup (op =) axms name of | 
| 15531 | 105 | SOME prop => prop | 
| 106 |                      | NONE => error ("Unknown axiom " ^ quote name))
 | |
| 107 | in PAxm (name, prop, NONE) end | |
| 11614 | 108 | | "thm" :: xs => | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 109 | let val name = NameSpace.implode xs; | 
| 17223 | 110 | in (case AList.lookup (op =) thms name of | 
| 28814 | 111 | SOME thm => fst (strip_combt (Thm.proof_of thm)) | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 112 |                    | NONE => error ("Unknown theorem " ^ quote name))
 | 
| 11522 | 113 | end | 
| 114 |              | _ => error ("Illegal proof constant name: " ^ quote s))
 | |
| 13199 | 115 |       | prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop
 | 
| 11522 | 116 |       | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
 | 
| 117 |       | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) =
 | |
| 25245 
1fcfcdcba53c
Added well-formedness check to Abst case in function prf_of.
 berghofe parents: 
24848diff
changeset | 118 | if T = proofT then | 
| 
1fcfcdcba53c
Added well-formedness check to Abst case in function prf_of.
 berghofe parents: 
24848diff
changeset | 119 |             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: 
24848diff
changeset | 120 | else Abst (s, if ty then SOME T else NONE, | 
| 11522 | 121 | incr_pboundvars (~1) 0 (prf_of [] prf)) | 
| 122 |       | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
 | |
| 11614 | 123 | AbsP (s, case t of | 
| 15531 | 124 |                 Const ("dummy_pattern", _) => NONE
 | 
| 125 |               | _ $ Const ("dummy_pattern", _) => NONE
 | |
| 126 | | _ => SOME (mk_term t), | |
| 11522 | 127 | incr_pboundvars 0 (~1) (prf_of [] prf)) | 
| 128 |       | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
 | |
| 11614 | 129 | prf_of [] prf1 %% prf_of [] prf2 | 
| 11522 | 130 |       | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
 | 
| 131 | prf_of (T::Ts) prf | |
| 11614 | 132 |       | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %
 | 
| 15531 | 133 |           (case t of Const ("dummy_pattern", _) => NONE | _ => SOME (mk_term t))
 | 
| 11522 | 134 |       | prf_of _ t = error ("Not a proof term:\n" ^
 | 
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26626diff
changeset | 135 | Syntax.string_of_term_global thy t) | 
| 11522 | 136 | |
| 137 | in prf_of [] end; | |
| 138 | ||
| 139 | ||
| 140 | val AbsPt = Const ("AbsP", [propT, proofT --> proofT] ---> proofT);
 | |
| 141 | val AppPt = Const ("AppP", [proofT, proofT] ---> proofT);
 | |
| 13199 | 142 | val Hypt = Const ("Hyp", propT --> proofT);
 | 
| 143 | val Oraclet = Const ("Oracle", propT --> proofT);
 | |
| 144 | val MinProoft = Const ("MinProof", proofT);
 | |
| 11522 | 145 | |
| 19473 | 146 | val mk_tyapp = fold (fn T => fn prf => Const ("Appt",
 | 
| 19391 | 147 | [proofT, Term.itselfT T] ---> proofT) $ prf $ Logic.mk_type T); | 
| 11522 | 148 | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 149 | fun term_of _ (PThm (_, ((name, _, NONE), _))) = | 
| 16195 | 150 | Const (NameSpace.append "thm" name, proofT) | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 151 | | term_of _ (PThm (_, ((name, _, SOME Ts), _))) = | 
| 19473 | 152 | mk_tyapp Ts (Const (NameSpace.append "thm" name, proofT)) | 
| 16195 | 153 | | term_of _ (PAxm (name, _, NONE)) = Const (NameSpace.append "axm" name, proofT) | 
| 15531 | 154 | | term_of _ (PAxm (name, _, SOME Ts)) = | 
| 19473 | 155 | mk_tyapp Ts (Const (NameSpace.append "axm" name, proofT)) | 
| 11522 | 156 | | term_of _ (PBound i) = Bound i | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 157 | | term_of Ts (Abst (s, opT, prf)) = | 
| 18939 | 158 | let val T = the_default dummyT opT | 
| 11522 | 159 |       in Const ("Abst", (T --> proofT) --> proofT) $
 | 
| 160 | Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf)) | |
| 161 | end | |
| 162 | | term_of Ts (AbsP (s, t, prf)) = | |
| 18939 | 163 | AbsPt $ the_default (Term.dummy_pattern propT) t $ | 
| 11522 | 164 | Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf)) | 
| 11614 | 165 | | term_of Ts (prf1 %% prf2) = | 
| 11522 | 166 | AppPt $ term_of Ts prf1 $ term_of Ts prf2 | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 167 | | term_of Ts (prf % opt) = | 
| 18939 | 168 | let val t = the_default (Term.dummy_pattern dummyT) opt | 
| 11522 | 169 |       in Const ("Appt",
 | 
| 170 | [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $ | |
| 171 | term_of Ts prf $ t | |
| 172 | end | |
| 173 | | term_of Ts (Hyp t) = Hypt $ t | |
| 174 | | term_of Ts (Oracle (_, t, _)) = Oraclet $ t | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 175 | | term_of Ts MinProof = MinProoft; | 
| 11522 | 176 | |
| 177 | val term_of_proof = term_of []; | |
| 178 | ||
| 179 | fun cterm_of_proof thy prf = | |
| 180 | let | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 181 | val thm_names = map fst (PureThy.all_thms_of thy); | 
| 16350 | 182 | val axm_names = map fst (Theory.all_axioms_of thy); | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 183 | val thy' = thy | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 184 | |> add_proof_syntax | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 185 | |> add_proof_atom_consts | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 186 | (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names); | 
| 11522 | 187 | in | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 188 | (cterm_of thy' (term_of_proof prf), proof_of_term thy true o Thm.term_of) | 
| 11522 | 189 | end; | 
| 190 | ||
| 191 | fun read_term thy = | |
| 192 | let | |
| 28375 | 193 | val thm_names = filter_out (fn s => s = "") (map fst (PureThy.all_thms_of thy)); | 
| 16350 | 194 | val axm_names = map fst (Theory.all_axioms_of thy); | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 195 | val ctxt = thy | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 196 | |> add_proof_syntax | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 197 | |> add_proof_atom_consts | 
| 16195 | 198 | (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names) | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 199 | |> ProofContext.init | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 200 | |> ProofContext.allow_dummies | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 201 | |> ProofContext.set_mode ProofContext.mode_schematic; | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 202 | in | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 203 | fn ty => fn s => | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 204 | (if ty = propT then Syntax.parse_prop else Syntax.parse_term) ctxt s | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 205 | |> TypeInfer.constrain ty |> Syntax.check_term ctxt | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 206 | end; | 
| 11522 | 207 | |
| 208 | fun read_proof thy = | |
| 209 | let val rd = read_term thy proofT | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 210 | in fn ty => fn s => proof_of_term thy ty (Logic.varify (rd s)) end; | 
| 11522 | 211 | |
| 17078 | 212 | fun proof_syntax prf = | 
| 11522 | 213 | let | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 214 | val thm_names = Symtab.keys (fold_proof_atoms true | 
| 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 215 | (fn PThm (_, ((name, _, _), _)) => if name <> "" then Symtab.update (name, ()) else I | 
| 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 216 | | _ => I) [prf] Symtab.empty); | 
| 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 217 | val axm_names = Symtab.keys (fold_proof_atoms true | 
| 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 218 | (fn PAxm (name, _, _) => Symtab.update (name, ()) | _ => I) [prf] Symtab.empty); | 
| 11522 | 219 | in | 
| 17078 | 220 | add_proof_syntax #> | 
| 221 | add_proof_atom_consts | |
| 222 | (map (NameSpace.append "thm") thm_names @ map (NameSpace.append "axm") axm_names) | |
| 11522 | 223 | end; | 
| 224 | ||
| 17078 | 225 | fun proof_of full thm = | 
| 226 | let | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
25245diff
changeset | 227 | val thy = Thm.theory_of_thm thm; | 
| 17078 | 228 | val prop = Thm.full_prop_of thm; | 
| 28814 | 229 | val prf = Thm.proof_of thm; | 
| 17078 | 230 | val prf' = (case strip_combt (fst (strip_combP prf)) of | 
| 29635 
31d14e9fa0da
proof_body: turned lazy into future -- ensures that body is fulfilled eventually, without explicit force;
 wenzelm parents: 
29606diff
changeset | 231 | (PThm (_, ((_, prop', _), body)), _) => if prop = prop' then join_proof body else prf | 
| 17078 | 232 | | _ => prf) | 
| 233 | in if full then Reconstruct.reconstruct_proof thy prop prf' else prf' end; | |
| 234 | ||
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 235 | fun pretty_proof ctxt prf = | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 236 | ProofContext.pretty_term_abbrev | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 237 | (ProofContext.transfer_syntax (proof_syntax prf (ProofContext.theory_of ctxt)) ctxt) | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 238 | (term_of_proof prf); | 
| 17078 | 239 | |
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 240 | fun pretty_proof_of ctxt full th = | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 241 | pretty_proof ctxt (proof_of full th); | 
| 11522 | 242 | |
| 243 | end; |