| author | wenzelm | 
| Thu, 09 Jun 2011 17:51:49 +0200 | |
| changeset 43326 | 47cf4bc789aa | 
| parent 42406 | 05f2468d6b36 | 
| child 45156 | a9b6c2ea7bec | 
| 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) | |
| 37227 
bdd8dd217b1f
- Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
 berghofe parents: 
36610diff
changeset | 14 | 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: 
36610diff
changeset | 15 | val read_proof: theory -> bool -> bool -> string -> Proofterm.proof | 
| 17078 | 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 | ||
| 33388 | 22 | structure Proof_Syntax : PROOF_SYNTAX = | 
| 11522 | 23 | struct | 
| 24 | ||
| 25 | (**** add special syntax for embedding proof terms ****) | |
| 26 | ||
| 27 | val proofT = Type ("proof", []);
 | |
| 11614 | 28 | val paramT = Type ("param", []);
 | 
| 29 | val paramsT = Type ("params", []);
 | |
| 11522 | 30 | val idtT = Type ("idt", []);
 | 
| 24848 | 31 | val aT = TFree (Name.aT, []); | 
| 11522 | 32 | |
| 33 | (** constants for theorems and axioms **) | |
| 34 | ||
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 35 | fun add_proof_atom_consts names thy = | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 36 | thy | 
| 30435 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 wenzelm parents: 
30364diff
changeset | 37 | |> Sign.root_path | 
| 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 wenzelm parents: 
30364diff
changeset | 38 | |> Sign.add_consts_i (map (fn name => (Binding.qualified_name name, proofT, NoSyn)) names); | 
| 11522 | 39 | |
| 40 | (** constants for application and abstraction **) | |
| 11614 | 41 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 42 | fun add_proof_syntax thy = | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 43 | thy | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 44 | |> Theory.copy | 
| 22796 | 45 | |> Sign.root_path | 
| 36449 | 46 | |> Sign.set_defsort [] | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 47 | |> Sign.add_types_global [(Binding.name "proof", 0, NoSyn)] | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 48 | |> fold (snd oo Sign.declare_const_global) | 
| 35122 | 49 |       [((Binding.name "Appt", [proofT, aT] ---> proofT), Mixfix ("(1_ %/ _)", [4, 5], 4)),
 | 
| 50 |        ((Binding.name "AppP", [proofT, proofT] ---> proofT), Mixfix ("(1_ %%/ _)", [4, 5], 4)),
 | |
| 51 | ((Binding.name "Abst", (aT --> proofT) --> proofT), NoSyn), | |
| 52 | ((Binding.name "AbsP", [propT, proofT --> proofT] ---> proofT), NoSyn), | |
| 53 | ((Binding.name "Hyp", propT --> proofT), NoSyn), | |
| 54 | ((Binding.name "Oracle", propT --> proofT), NoSyn), | |
| 55 | ((Binding.name "OfClass", (Term.a_itselfT --> propT) --> proofT), NoSyn), | |
| 56 | ((Binding.name "MinProof", proofT), Delimfix "?")] | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 57 | |> Sign.add_nonterminals_global [Binding.name "param", Binding.name "params"] | 
| 22796 | 58 | |> Sign.add_syntax_i | 
| 11640 | 59 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1Lam _./ _)", [0, 3], 3)),
 | 
| 11614 | 60 |        ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
 | 
| 61 |        ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
 | |
| 62 |        ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)),
 | |
| 63 |        ("", paramT --> paramT, Delimfix "'(_')"),
 | |
| 64 |        ("", idtT --> paramsT, Delimfix "_"),
 | |
| 65 |        ("", paramT --> paramsT, Delimfix "_")]
 | |
| 35122 | 66 | |> Sign.add_modesyntax_i (Symbol.xsymbolsN, true) | 
| 11640 | 67 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<Lambda>_./ _)", [0, 3], 3)),
 | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42224diff
changeset | 68 |        (Lexicon.mark_const "Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
 | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42224diff
changeset | 69 |        (Lexicon.mark_const "AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))]
 | 
| 22796 | 70 |   |> Sign.add_modesyntax_i ("latex", false)
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 71 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<^bold>\\<lambda>_./ _)", [0, 3], 3))]
 | 
| 42204 | 72 | |> Sign.add_trrules (map Syntax.Parse_Print_Rule | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 73 | [(Ast.mk_appl (Ast.Constant "_Lam") | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 74 | [Ast.mk_appl (Ast.Constant "_Lam0") | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 75 | [Ast.Variable "l", Ast.Variable "m"], Ast.Variable "A"], | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 76 | Ast.mk_appl (Ast.Constant "_Lam") | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 77 | [Ast.Variable "l", | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 78 | Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "m", Ast.Variable "A"]]), | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 79 | (Ast.mk_appl (Ast.Constant "_Lam") | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 80 | [Ast.mk_appl (Ast.Constant "_Lam1") | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 81 | [Ast.Variable "x", Ast.Variable "A"], Ast.Variable "B"], | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42224diff
changeset | 82 | Ast.mk_appl (Ast.Constant (Lexicon.mark_const "AbsP")) [Ast.Variable "A", | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 83 | (Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "B"])]), | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 84 | (Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "x", Ast.Variable "A"], | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42224diff
changeset | 85 | Ast.mk_appl (Ast.Constant (Lexicon.mark_const "Abst")) | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 86 | [(Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "A"])])]); | 
| 11522 | 87 | |
| 88 | ||
| 89 | (**** translation between proof terms and pure terms ****) | |
| 90 | ||
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 91 | fun proof_of_term thy ty = | 
| 11522 | 92 | let | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 93 | val thms = Global_Theory.all_thms_of thy; | 
| 16350 | 94 | val axms = Theory.all_axioms_of thy; | 
| 11522 | 95 | |
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
19618diff
changeset | 96 | fun mk_term t = (if ty then I else map_types (K dummyT)) | 
| 11614 | 97 | (Term.no_dummy_patterns t); | 
| 98 | ||
| 11522 | 99 | fun prf_of [] (Bound i) = PBound i | 
| 100 |       | prf_of Ts (Const (s, Type ("proof", _))) =
 | |
| 37310 | 101 | Proofterm.change_type (if ty then SOME Ts else NONE) | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 102 | (case Long_Name.explode s of | 
| 11614 | 103 | "axm" :: xs => | 
| 11522 | 104 | let | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 105 | val name = Long_Name.implode xs; | 
| 17223 | 106 | val prop = (case AList.lookup (op =) axms name of | 
| 15531 | 107 | SOME prop => prop | 
| 108 |                      | NONE => error ("Unknown axiom " ^ quote name))
 | |
| 109 | in PAxm (name, prop, NONE) end | |
| 11614 | 110 | | "thm" :: xs => | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 111 | let val name = Long_Name.implode xs; | 
| 17223 | 112 | in (case AList.lookup (op =) thms name of | 
| 37310 | 113 | SOME thm => | 
| 114 | fst (Proofterm.strip_combt (fst (Proofterm.strip_combP (Thm.proof_of thm)))) | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 115 |                    | NONE => error ("Unknown theorem " ^ quote name))
 | 
| 11522 | 116 | end | 
| 117 |              | _ => error ("Illegal proof constant name: " ^ quote s))
 | |
| 31943 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 wenzelm parents: 
31903diff
changeset | 118 |       | prf_of Ts (Const ("OfClass", _) $ Const (c_class, _)) =
 | 
| 31903 | 119 | (case try Logic.class_of_const c_class of | 
| 120 | SOME c => | |
| 37310 | 121 | Proofterm.change_type (if ty then SOME Ts else NONE) | 
| 31943 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 wenzelm parents: 
31903diff
changeset | 122 | (OfClass (TVar ((Name.aT, 0), []), c)) | 
| 31903 | 123 |           | NONE => error ("Bad class constant: " ^ quote c_class))
 | 
| 13199 | 124 |       | prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop
 | 
| 11522 | 125 |       | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
 | 
| 126 |       | 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 | 127 | if T = proofT then | 
| 
1fcfcdcba53c
Added well-formedness check to Abst case in function prf_of.
 berghofe parents: 
24848diff
changeset | 128 |             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 | 129 | else Abst (s, if ty then SOME T else NONE, | 
| 37310 | 130 | Proofterm.incr_pboundvars (~1) 0 (prf_of [] prf)) | 
| 11522 | 131 |       | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
 | 
| 11614 | 132 | AbsP (s, case t of | 
| 15531 | 133 |                 Const ("dummy_pattern", _) => NONE
 | 
| 134 |               | _ $ Const ("dummy_pattern", _) => NONE
 | |
| 135 | | _ => SOME (mk_term t), | |
| 37310 | 136 | Proofterm.incr_pboundvars 0 (~1) (prf_of [] prf)) | 
| 11522 | 137 |       | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
 | 
| 11614 | 138 | prf_of [] prf1 %% prf_of [] prf2 | 
| 11522 | 139 |       | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
 | 
| 140 | prf_of (T::Ts) prf | |
| 11614 | 141 |       | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %
 | 
| 15531 | 142 |           (case t of Const ("dummy_pattern", _) => NONE | _ => SOME (mk_term t))
 | 
| 11522 | 143 |       | 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 | 144 | Syntax.string_of_term_global thy t) | 
| 11522 | 145 | |
| 146 | in prf_of [] end; | |
| 147 | ||
| 148 | ||
| 149 | val AbsPt = Const ("AbsP", [propT, proofT --> proofT] ---> proofT);
 | |
| 150 | val AppPt = Const ("AppP", [proofT, proofT] ---> proofT);
 | |
| 13199 | 151 | val Hypt = Const ("Hyp", propT --> proofT);
 | 
| 152 | val Oraclet = Const ("Oracle", propT --> proofT);
 | |
| 31943 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 wenzelm parents: 
31903diff
changeset | 153 | val OfClasst = Const ("OfClass", (Term.itselfT dummyT --> propT) --> proofT);
 | 
| 13199 | 154 | val MinProoft = Const ("MinProof", proofT);
 | 
| 11522 | 155 | |
| 19473 | 156 | val mk_tyapp = fold (fn T => fn prf => Const ("Appt",
 | 
| 19391 | 157 | [proofT, Term.itselfT T] ---> proofT) $ prf $ Logic.mk_type T); | 
| 11522 | 158 | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 159 | fun term_of _ (PThm (_, ((name, _, NONE), _))) = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 160 | Const (Long_Name.append "thm" name, proofT) | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 161 | | term_of _ (PThm (_, ((name, _, SOME Ts), _))) = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 162 | mk_tyapp Ts (Const (Long_Name.append "thm" name, proofT)) | 
| 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 163 | | term_of _ (PAxm (name, _, NONE)) = Const (Long_Name.append "axm" name, proofT) | 
| 15531 | 164 | | term_of _ (PAxm (name, _, SOME Ts)) = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 165 | mk_tyapp Ts (Const (Long_Name.append "axm" name, proofT)) | 
| 31943 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 wenzelm parents: 
31903diff
changeset | 166 | | term_of _ (OfClass (T, c)) = | 
| 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 wenzelm parents: 
31903diff
changeset | 167 | mk_tyapp [T] (OfClasst $ Const (Logic.const_of_class c, Term.itselfT dummyT --> propT)) | 
| 11522 | 168 | | term_of _ (PBound i) = Bound i | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 169 | | term_of Ts (Abst (s, opT, prf)) = | 
| 18939 | 170 | let val T = the_default dummyT opT | 
| 11522 | 171 |       in Const ("Abst", (T --> proofT) --> proofT) $
 | 
| 37310 | 172 | Abs (s, T, term_of (T::Ts) (Proofterm.incr_pboundvars 1 0 prf)) | 
| 11522 | 173 | end | 
| 174 | | term_of Ts (AbsP (s, t, prf)) = | |
| 18939 | 175 | AbsPt $ the_default (Term.dummy_pattern propT) t $ | 
| 37310 | 176 | Abs (s, proofT, term_of (proofT::Ts) (Proofterm.incr_pboundvars 0 1 prf)) | 
| 11614 | 177 | | term_of Ts (prf1 %% prf2) = | 
| 11522 | 178 | AppPt $ term_of Ts prf1 $ term_of Ts prf2 | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 179 | | term_of Ts (prf % opt) = | 
| 18939 | 180 | let val t = the_default (Term.dummy_pattern dummyT) opt | 
| 11522 | 181 |       in Const ("Appt",
 | 
| 182 | [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $ | |
| 183 | term_of Ts prf $ t | |
| 184 | end | |
| 185 | | term_of Ts (Hyp t) = Hypt $ t | |
| 186 | | term_of Ts (Oracle (_, t, _)) = Oraclet $ t | |
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 187 | | term_of Ts MinProof = MinProoft; | 
| 11522 | 188 | |
| 189 | val term_of_proof = term_of []; | |
| 190 | ||
| 191 | fun cterm_of_proof thy prf = | |
| 192 | let | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 193 | val thm_names = map fst (Global_Theory.all_thms_of thy); | 
| 16350 | 194 | val axm_names = map fst (Theory.all_axioms_of thy); | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 195 | val thy' = thy | 
| 
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 | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 198 | (map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names); | 
| 11522 | 199 | in | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 200 | (cterm_of thy' (term_of_proof prf), proof_of_term thy true o Thm.term_of) | 
| 11522 | 201 | end; | 
| 202 | ||
| 37227 
bdd8dd217b1f
- Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
 berghofe parents: 
36610diff
changeset | 203 | fun read_term thy topsort = | 
| 11522 | 204 | let | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 205 | val thm_names = filter_out (fn s => s = "") (map fst (Global_Theory.all_thms_of thy)); | 
| 16350 | 206 | val axm_names = map fst (Theory.all_axioms_of thy); | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 207 | val ctxt = thy | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 208 | |> add_proof_syntax | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 209 | |> add_proof_atom_consts | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 210 | (map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names) | 
| 42360 | 211 | |> Proof_Context.init_global | 
| 212 | |> Proof_Context.allow_dummies | |
| 213 | |> Proof_Context.set_mode Proof_Context.mode_schematic | |
| 42406 
05f2468d6b36
eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
 wenzelm parents: 
42375diff
changeset | 214 | |> topsort ? (Config.put Type_Infer_Context.const_sorts false #> Proof_Context.set_defsort []); | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 215 | in | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 216 | fn ty => fn s => | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 217 | (if ty = propT then Syntax.parse_prop else Syntax.parse_term) ctxt s | 
| 39288 | 218 | |> Type.constraint ty |> Syntax.check_term ctxt | 
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 219 | end; | 
| 11522 | 220 | |
| 37227 
bdd8dd217b1f
- Added extra flag to read_term and read_proof functions that allows to parse (proof)terms in which
 berghofe parents: 
36610diff
changeset | 221 | 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: 
36610diff
changeset | 222 | 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: 
35262diff
changeset | 223 | in fn ty => fn s => proof_of_term thy ty (Logic.varify_global (rd s)) end; | 
| 11522 | 224 | |
| 17078 | 225 | fun proof_syntax prf = | 
| 11522 | 226 | let | 
| 37310 | 227 | val thm_names = Symtab.keys (Proofterm.fold_proof_atoms true | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 228 | (fn PThm (_, ((name, _, _), _)) => if name <> "" then Symtab.update (name, ()) else I | 
| 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 229 | | _ => I) [prf] Symtab.empty); | 
| 37310 | 230 | val axm_names = Symtab.keys (Proofterm.fold_proof_atoms true | 
| 28807 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 wenzelm parents: 
28375diff
changeset | 231 | (fn PAxm (name, _, _) => Symtab.update (name, ()) | _ => I) [prf] Symtab.empty); | 
| 11522 | 232 | in | 
| 17078 | 233 | add_proof_syntax #> | 
| 234 | add_proof_atom_consts | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 235 | (map (Long_Name.append "thm") thm_names @ map (Long_Name.append "axm") axm_names) | 
| 11522 | 236 | end; | 
| 237 | ||
| 17078 | 238 | fun proof_of full thm = | 
| 239 | let | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
25245diff
changeset | 240 | val thy = Thm.theory_of_thm thm; | 
| 17078 | 241 | val prop = Thm.full_prop_of thm; | 
| 28814 | 242 | val prf = Thm.proof_of thm; | 
| 37310 | 243 | val prf' = | 
| 244 | (case Proofterm.strip_combt (fst (Proofterm.strip_combP prf)) of | |
| 245 | (PThm (_, ((_, prop', _), body)), _) => | |
| 246 | if prop = prop' then Proofterm.join_proof body else prf | |
| 17078 | 247 | | _ => prf) | 
| 248 | in if full then Reconstruct.reconstruct_proof thy prop prf' else prf' end; | |
| 249 | ||
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 250 | fun pretty_proof ctxt prf = | 
| 42360 | 251 | Proof_Context.pretty_term_abbrev | 
| 252 | (Proof_Context.transfer_syntax (proof_syntax prf (Proof_Context.theory_of ctxt)) ctxt) | |
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 253 | (term_of_proof prf); | 
| 17078 | 254 | |
| 27260 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 255 | fun pretty_proof_of ctxt full th = | 
| 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 wenzelm parents: 
26939diff
changeset | 256 | pretty_proof ctxt (proof_of full th); | 
| 11522 | 257 | |
| 258 | end; |