| author | wenzelm | 
| Tue, 27 Jul 2021 15:31:54 +0200 | |
| changeset 74072 | dc98bb7a439b | 
| parent 71777 | 3875815f5967 | 
| child 74330 | d882abae3379 | 
| 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 add_proof_syntax: theory -> theory  | 
| 70980 | 10  | 
val term_of_proof: proof -> term  | 
| 
28807
 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 
wenzelm 
parents: 
28375 
diff
changeset
 | 
11  | 
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
 | 
12  | 
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
 | 
13  | 
val read_proof: theory -> bool -> bool -> string -> Proofterm.proof  | 
| 17078 | 14  | 
val proof_syntax: Proofterm.proof -> theory -> theory  | 
| 70449 | 15  | 
val proof_of: bool -> thm -> Proofterm.proof  | 
| 
27260
 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
16  | 
val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T  | 
| 71010 | 17  | 
val pretty_proof_boxes_of: Proof.context ->  | 
18  | 
    {full: bool, preproc: theory -> proof -> proof} -> thm -> Pretty.T
 | 
|
| 71088 | 19  | 
  val standard_proof_of: {full: bool, expand_name: Proofterm.thm_header -> string option} ->
 | 
20  | 
thm -> Proofterm.proof  | 
|
21  | 
val pretty_standard_proof_of: Proof.context -> bool -> thm -> Pretty.T  | 
|
| 11522 | 22  | 
end;  | 
23  | 
||
| 33388 | 24  | 
structure Proof_Syntax : PROOF_SYNTAX =  | 
| 11522 | 25  | 
struct  | 
26  | 
||
27  | 
(**** add special syntax for embedding proof terms ****)  | 
|
28  | 
||
| 70980 | 29  | 
val proofT = Type ("Pure.proof", []);
 | 
| 
70388
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
30  | 
|
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
31  | 
local  | 
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
32  | 
|
| 11614 | 33  | 
val paramT = Type ("param", []);
 | 
34  | 
val paramsT = Type ("params", []);
 | 
|
| 11522 | 35  | 
val idtT = Type ("idt", []);
 | 
| 70387 | 36  | 
val aT = Term.aT [];  | 
| 11522 | 37  | 
|
| 
70388
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
38  | 
fun mixfix (sy, ps, p) = Mixfix (Input.string sy, ps, p, Position.no_range);  | 
| 11522 | 39  | 
|
| 
70388
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
40  | 
in  | 
| 62752 | 41  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16350 
diff
changeset
 | 
42  | 
fun add_proof_syntax thy =  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16350 
diff
changeset
 | 
43  | 
thy  | 
| 22796 | 44  | 
|> Sign.root_path  | 
| 36449 | 45  | 
|> Sign.set_defsort []  | 
| 56436 | 46  | 
|> Sign.add_nonterminals_global  | 
| 64556 | 47  | 
    [Binding.make ("param", \<^here>),
 | 
48  | 
     Binding.make ("params", \<^here>)]
 | 
|
| 56240 | 49  | 
|> Sign.add_syntax Syntax.mode_default  | 
| 62752 | 50  | 
    [("_Lam", [paramsT, proofT] ---> proofT, mixfix ("(1\<^bold>\<lambda>_./ _)", [0, 3], 3)),
 | 
51  | 
     ("_Lam0", [paramT, paramsT] ---> paramsT, mixfix ("_/ _", [1, 0], 0)),
 | 
|
52  | 
     ("_Lam0", [idtT, paramsT] ---> paramsT, mixfix ("_/ _", [1, 0], 0)),
 | 
|
53  | 
     ("_Lam1", [idtT, propT] ---> paramT, mixfix ("_: _", [0, 0], 0)),
 | 
|
| 62761 | 54  | 
     ("", paramT --> paramT, Mixfix.mixfix "'(_')"),
 | 
55  | 
     ("", idtT --> paramsT, Mixfix.mixfix "_"),
 | 
|
| 
70388
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
56  | 
     ("", paramT --> paramsT, Mixfix.mixfix "_"),
 | 
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
57  | 
     (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
 | 
58  | 
     (Lexicon.mark_const "Pure.AppP", [proofT, proofT] ---> proofT, mixfix ("(1_ \<bullet>/ _)", [4, 5], 4)),
 | 
| 70407 | 59  | 
(Lexicon.mark_const "Pure.MinProof", proofT, Mixfix.mixfix "\<^bold>?")]  | 
| 42204 | 60  | 
|> Sign.add_trrules (map Syntax.Parse_Print_Rule  | 
| 56436 | 61  | 
[(Ast.mk_appl (Ast.Constant "_Lam")  | 
62  | 
[Ast.mk_appl (Ast.Constant "_Lam0")  | 
|
63  | 
[Ast.Variable "l", Ast.Variable "m"], Ast.Variable "A"],  | 
|
64  | 
Ast.mk_appl (Ast.Constant "_Lam")  | 
|
65  | 
[Ast.Variable "l",  | 
|
66  | 
Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "m", Ast.Variable "A"]]),  | 
|
67  | 
(Ast.mk_appl (Ast.Constant "_Lam")  | 
|
68  | 
[Ast.mk_appl (Ast.Constant "_Lam1")  | 
|
69  | 
[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
 | 
70  | 
Ast.mk_appl (Ast.Constant (Lexicon.mark_const "Pure.AbsP")) [Ast.Variable "A",  | 
| 56436 | 71  | 
(Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "B"])]),  | 
72  | 
(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
 | 
73  | 
Ast.mk_appl (Ast.Constant (Lexicon.mark_const "Pure.Abst"))  | 
| 56436 | 74  | 
[(Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "A"])])]);  | 
| 11522 | 75  | 
|
| 
70388
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
76  | 
end;  | 
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
77  | 
|
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
78  | 
|
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
79  | 
(** constants for theorems and axioms **)  | 
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
80  | 
|
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
81  | 
fun add_proof_atom_consts names thy =  | 
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
82  | 
thy  | 
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
83  | 
|> Sign.root_path  | 
| 
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
84  | 
|> 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
 | 
85  | 
|
| 11522 | 86  | 
|
| 70980 | 87  | 
|
88  | 
(** proof terms as pure terms **)  | 
|
89  | 
||
90  | 
(* term_of_proof *)  | 
|
91  | 
||
92  | 
local  | 
|
93  | 
||
94  | 
val AbsPt = Const ("Pure.AbsP", propT --> (proofT --> proofT) --> proofT);
 | 
|
95  | 
val AppPt = Const ("Pure.AppP", proofT --> proofT --> proofT);
 | 
|
96  | 
val Hypt = Const ("Pure.Hyp", propT --> proofT);
 | 
|
97  | 
val Oraclet = Const ("Pure.Oracle", propT --> proofT);
 | 
|
98  | 
val MinProoft = Const ("Pure.MinProof", proofT);
 | 
|
99  | 
||
100  | 
fun AppT T prf =  | 
|
101  | 
  Const ("Pure.Appt", proofT --> Term.itselfT T --> proofT) $ prf $ Logic.mk_type T;
 | 
|
102  | 
||
| 
71777
 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 
wenzelm 
parents: 
71090 
diff
changeset
 | 
103  | 
fun PClasst (T, c) =  | 
| 70980 | 104  | 
let val U = Term.itselfT T --> propT  | 
| 
71777
 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 
wenzelm 
parents: 
71090 
diff
changeset
 | 
105  | 
  in Const ("Pure.PClass", U --> proofT) $ Const (Logic.const_of_class c, U) end;
 | 
| 70980 | 106  | 
|
107  | 
fun term_of _ (PThm ({serial = i, name, types = Ts, ...}, _)) =
 | 
|
108  | 
fold AppT (these Ts)  | 
|
109  | 
(Const (Long_Name.append "thm" (if name = "" then string_of_int i else name), proofT))  | 
|
110  | 
| term_of _ (PAxm (name, _, Ts)) =  | 
|
111  | 
fold AppT (these Ts) (Const (Long_Name.append "axm" name, proofT))  | 
|
| 
71777
 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 
wenzelm 
parents: 
71090 
diff
changeset
 | 
112  | 
| term_of _ (PClass (T, c)) = AppT T (PClasst (T, c))  | 
| 70980 | 113  | 
| term_of _ (PBound i) = Bound i  | 
114  | 
| term_of Ts (Abst (s, opT, prf)) =  | 
|
115  | 
let val T = the_default dummyT opT in  | 
|
116  | 
        Const ("Pure.Abst", (T --> proofT) --> proofT) $
 | 
|
117  | 
Abs (s, T, term_of (T::Ts) (Proofterm.incr_pboundvars 1 0 prf))  | 
|
118  | 
end  | 
|
119  | 
| term_of Ts (AbsP (s, t, prf)) =  | 
|
120  | 
AbsPt $ the_default Term.dummy_prop t $  | 
|
121  | 
Abs (s, proofT, term_of (proofT::Ts) (Proofterm.incr_pboundvars 0 1 prf))  | 
|
122  | 
| term_of Ts (prf1 %% prf2) =  | 
|
123  | 
AppPt $ term_of Ts prf1 $ term_of Ts prf2  | 
|
124  | 
| term_of Ts (prf % opt) =  | 
|
125  | 
let  | 
|
126  | 
val t = the_default Term.dummy opt;  | 
|
127  | 
val T = fastype_of1 (Ts, t) handle TERM _ => dummyT;  | 
|
128  | 
      in Const ("Pure.Appt", proofT --> T --> proofT) $ term_of Ts prf $ t end
 | 
|
129  | 
| term_of _ (Hyp t) = Hypt $ t  | 
|
130  | 
| term_of _ (Oracle (_, t, _)) = Oraclet $ t  | 
|
131  | 
| term_of _ MinProof = MinProoft;  | 
|
132  | 
||
133  | 
in  | 
|
134  | 
||
135  | 
val term_of_proof = term_of [];  | 
|
136  | 
||
137  | 
end;  | 
|
138  | 
||
139  | 
||
140  | 
(* proof_of_term *)  | 
|
| 11522 | 141  | 
|
| 
28807
 
9f3ecb4aaac2
proof_of_term: removed obsolete disambiguisation table;
 
wenzelm 
parents: 
28375 
diff
changeset
 | 
142  | 
fun proof_of_term thy ty =  | 
| 11522 | 143  | 
let  | 
| 56161 | 144  | 
val thms = Global_Theory.all_thms_of thy true;  | 
| 16350 | 145  | 
val axms = Theory.all_axioms_of thy;  | 
| 11522 | 146  | 
|
| 
20548
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
19618 
diff
changeset
 | 
147  | 
fun mk_term t = (if ty then I else map_types (K dummyT))  | 
| 11614 | 148  | 
(Term.no_dummy_patterns t);  | 
149  | 
||
| 11522 | 150  | 
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
 | 
151  | 
      | prf_of Ts (Const (s, Type ("Pure.proof", _))) =
 | 
| 70417 | 152  | 
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
 | 
153  | 
(case Long_Name.explode s of  | 
| 11614 | 154  | 
"axm" :: xs =>  | 
| 11522 | 155  | 
let  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30344 
diff
changeset
 | 
156  | 
val name = Long_Name.implode xs;  | 
| 17223 | 157  | 
val prop = (case AList.lookup (op =) axms name of  | 
| 15531 | 158  | 
SOME prop => prop  | 
159  | 
                     | NONE => error ("Unknown axiom " ^ quote name))
 | 
|
160  | 
in PAxm (name, prop, NONE) end  | 
|
| 11614 | 161  | 
| "thm" :: xs =>  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30344 
diff
changeset
 | 
162  | 
let val name = Long_Name.implode xs;  | 
| 17223 | 163  | 
in (case AList.lookup (op =) thms name of  | 
| 37310 | 164  | 
SOME thm =>  | 
165  | 
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
 | 
166  | 
                   | NONE => error ("Unknown theorem " ^ quote name))
 | 
| 11522 | 167  | 
end  | 
168  | 
             | _ => error ("Illegal proof constant name: " ^ quote s))
 | 
|
| 
71777
 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 
wenzelm 
parents: 
71090 
diff
changeset
 | 
169  | 
      | prf_of Ts (Const ("Pure.PClass", _) $ Const (c_class, _)) =
 | 
| 31903 | 170  | 
(case try Logic.class_of_const c_class of  | 
171  | 
SOME c =>  | 
|
| 70417 | 172  | 
Proofterm.change_types (if ty then SOME Ts else NONE)  | 
| 
71777
 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 
wenzelm 
parents: 
71090 
diff
changeset
 | 
173  | 
(PClass (TVar ((Name.aT, 0), []), c))  | 
| 31903 | 174  | 
          | 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
 | 
175  | 
      | 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
 | 
176  | 
      | 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
 | 
177  | 
      | 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
 | 
178  | 
if T = proofT then  | 
| 
 
1fcfcdcba53c
Added well-formedness check to Abst case in function prf_of.
 
berghofe 
parents: 
24848 
diff
changeset
 | 
179  | 
            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
 | 
180  | 
else Abst (s, if ty then SOME T else NONE,  | 
| 37310 | 181  | 
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
 | 
182  | 
      | prf_of [] (Const ("Pure.AbsP", _) $ t $ Abs (s, _, prf)) =
 | 
| 11614 | 183  | 
AbsP (s, case t of  | 
| 56241 | 184  | 
                Const ("Pure.dummy_pattern", _) => NONE
 | 
185  | 
              | _ $ Const ("Pure.dummy_pattern", _) => NONE
 | 
|
| 15531 | 186  | 
| _ => SOME (mk_term t),  | 
| 37310 | 187  | 
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
 | 
188  | 
      | prf_of [] (Const ("Pure.AppP", _) $ prf1 $ prf2) =
 | 
| 11614 | 189  | 
prf_of [] prf1 %% prf_of [] prf2  | 
| 70418 | 190  | 
      | prf_of Ts (Const ("Pure.Appt", _) $ prf $ Const ("Pure.type", Type ("itself", [T]))) =
 | 
| 11522 | 191  | 
prf_of (T::Ts) prf  | 
| 
70388
 
e31271559de8
global declaration of abstract syntax for proof terms, with qualified names;
 
wenzelm 
parents: 
70387 
diff
changeset
 | 
192  | 
      | prf_of [] (Const ("Pure.Appt", _) $ prf $ t) = prf_of [] prf %
 | 
| 56241 | 193  | 
          (case t of Const ("Pure.dummy_pattern", _) => NONE | _ => SOME (mk_term t))
 | 
| 11522 | 194  | 
      | 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
 | 
195  | 
Syntax.string_of_term_global thy t)  | 
| 11522 | 196  | 
|
197  | 
in prf_of [] end;  | 
|
198  | 
||
199  | 
||
| 
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
 | 
200  | 
fun read_term thy topsort =  | 
| 11522 | 201  | 
let  | 
| 56161 | 202  | 
val thm_names = filter_out (fn s => s = "") (map fst (Global_Theory.all_thms_of thy true));  | 
| 16350 | 203  | 
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
 | 
204  | 
val ctxt = thy  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16350 
diff
changeset
 | 
205  | 
|> add_proof_syntax  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16350 
diff
changeset
 | 
206  | 
|> add_proof_atom_consts  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30344 
diff
changeset
 | 
207  | 
(map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names)  | 
| 42360 | 208  | 
|> Proof_Context.init_global  | 
209  | 
|> Proof_Context.allow_dummies  | 
|
210  | 
|> 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
 | 
211  | 
|> topsort ?  | 
| 
 
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
 
wenzelm 
parents: 
62922 
diff
changeset
 | 
212  | 
(Proof_Context.set_defsort [] #>  | 
| 
 
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
 
wenzelm 
parents: 
62922 
diff
changeset
 | 
213  | 
Config.put Type_Infer.object_logic false #>  | 
| 
 
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
 
wenzelm 
parents: 
62922 
diff
changeset
 | 
214  | 
Config.put Type_Infer_Context.const_sorts false);  | 
| 
27260
 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
215  | 
in  | 
| 
 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
216  | 
fn ty => fn s =>  | 
| 
 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 
wenzelm 
parents: 
26939 
diff
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: 
26939 
diff
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: 
36610 
diff
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: 
36610 
diff
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: 
35262 
diff
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  | 
| 70493 | 228  | 
      (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
 | 
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: 
28375 
diff
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: 
30344 
diff
changeset
 | 
235  | 
(map (Long_Name.append "thm") thm_names @ map (Long_Name.append "axm") axm_names)  | 
| 11522 | 236  | 
end;  | 
237  | 
||
| 70449 | 238  | 
fun proof_of full thm =  | 
| 17078 | 239  | 
let  | 
| 70449 | 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;  | 
| 70813 | 243  | 
in  | 
244  | 
(case fst (Proofterm.strip_combt (fst (Proofterm.strip_combP prf))) of  | 
|
245  | 
      PThm ({prop = prop', ...}, thm_body) =>
 | 
|
246  | 
if prop = prop' then Proofterm.thm_body_proof_raw thm_body else prf  | 
|
247  | 
| _ => prf)  | 
|
248  | 
|> full ? Proofterm.reconstruct_proof thy prop  | 
|
249  | 
end;  | 
|
| 17078 | 250  | 
|
| 
27260
 
17d617c6b026
moved ProofContext.pretty_proof to ProofSyntax.pretty_proof;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
251  | 
fun pretty_proof ctxt prf =  | 
| 42360 | 252  | 
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
 | 
253  | 
(Proof_Context.transfer (proof_syntax prf (Proof_Context.theory_of ctxt)) ctxt)  | 
| 70980 | 254  | 
(term_of_proof prf);  | 
| 17078 | 255  | 
|
| 71010 | 256  | 
fun pretty_proof_boxes_of ctxt {full, preproc} thm =
 | 
| 70979 | 257  | 
let  | 
258  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
259  | 
val selection =  | 
|
260  | 
      {included = Proofterm.this_id (Thm.derivation_id thm),
 | 
|
261  | 
excluded = is_some o Global_Theory.lookup_thm_id thy}  | 
|
262  | 
in  | 
|
263  | 
Proofterm.proof_boxes selection [Thm.proof_of thm]  | 
|
264  | 
    |> map (fn ({serial = i, pos, prop, ...}, proof) =>
 | 
|
265  | 
let  | 
|
266  | 
val proof' = proof  | 
|
| 71010 | 267  | 
|> Proofterm.reconstruct_proof thy prop  | 
268  | 
|> preproc thy  | 
|
269  | 
|> not full ? Proofterm.shrink_proof  | 
|
| 70979 | 270  | 
|> Proofterm.forall_intr_variables prop;  | 
271  | 
val prop' = prop  | 
|
272  | 
|> Proofterm.forall_intr_variables_term;  | 
|
273  | 
val name = Long_Name.append "thm" (string_of_int i);  | 
|
274  | 
in  | 
|
275  | 
Pretty.item  | 
|
276  | 
[Pretty.str (name ^ Position.here_list pos ^ ":"), Pretty.brk 1,  | 
|
277  | 
Syntax.pretty_term ctxt prop', Pretty.fbrk, pretty_proof ctxt proof']  | 
|
278  | 
end)  | 
|
279  | 
|> Pretty.chunks  | 
|
280  | 
end;  | 
|
281  | 
||
| 71088 | 282  | 
|
283  | 
(* standardized proofs *)  | 
|
284  | 
||
285  | 
fun standard_proof_of {full, expand_name} thm =
 | 
|
286  | 
let val thy = Thm.theory_of_thm thm in  | 
|
287  | 
Thm.reconstruct_proof_of thm  | 
|
288  | 
|> Proofterm.expand_proof thy expand_name  | 
|
| 
71090
 
06c6495fb1d0
retain type information from reconstruct_proof, notably for Export_Theory.export_thm;
 
wenzelm 
parents: 
71088 
diff
changeset
 | 
289  | 
|> Proofterm.rewrite_proof thy ([], Proof_Rewrite_Rules.rprocs true)  | 
| 71088 | 290  | 
|> Proofterm.no_thm_proofs  | 
291  | 
|> not full ? Proofterm.shrink_proof  | 
|
292  | 
end;  | 
|
293  | 
||
294  | 
fun pretty_standard_proof_of ctxt full thm =  | 
|
295  | 
  pretty_proof ctxt (standard_proof_of {full = full, expand_name = Thm.expand_name thm} thm);
 | 
|
296  | 
||
| 11522 | 297  | 
end;  |