| author | wenzelm | 
| Thu, 08 Nov 2007 14:51:31 +0100 | |
| changeset 25346 | 7f2e3292e3dd | 
| parent 25345 | dd5b851f8ef0 | 
| child 25353 | 17f04d987f37 | 
| permissions | -rw-r--r-- | 
| 5819 | 1  | 
(* Title: Pure/Isar/proof_context.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
4  | 
||
| 19001 | 5  | 
The key concept of Isar proof contexts: elevates primitive local  | 
6  | 
reasoning Gamma |- phi to a structured concept, with generic context  | 
|
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
7  | 
elements. See also structure Variable and Assumption.  | 
| 5819 | 8  | 
*)  | 
9  | 
||
10  | 
signature PROOF_CONTEXT =  | 
|
11  | 
sig  | 
|
| 20310 | 12  | 
val theory_of: Proof.context -> theory  | 
| 24812 | 13  | 
val tsig_of: Proof.context -> Type.tsig  | 
| 20310 | 14  | 
val init: theory -> Proof.context  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
15  | 
type mode  | 
| 24501 | 16  | 
val mode_default: mode  | 
17  | 
val mode_stmt: mode  | 
|
18  | 
val mode_pattern: mode  | 
|
19  | 
val mode_schematic: mode  | 
|
20  | 
val mode_abbrev: mode  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
21  | 
val set_mode: mode -> Proof.context -> Proof.context  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
22  | 
val get_mode: Proof.context -> mode  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
23  | 
val restore_mode: Proof.context -> Proof.context -> Proof.context  | 
| 
21667
 
ce813b82c88b
add_notation: permissive about undeclared consts;
 
wenzelm 
parents: 
21648 
diff
changeset
 | 
24  | 
val set_stmt: bool -> Proof.context -> Proof.context  | 
| 
 
ce813b82c88b
add_notation: permissive about undeclared consts;
 
wenzelm 
parents: 
21648 
diff
changeset
 | 
25  | 
val naming_of: Proof.context -> NameSpace.naming  | 
| 20310 | 26  | 
val full_name: Proof.context -> bstring -> string  | 
27  | 
val consts_of: Proof.context -> Consts.T  | 
|
| 21183 | 28  | 
val const_syntax_name: Proof.context -> string -> string  | 
| 24752 | 29  | 
val the_const_constraint: Proof.context -> string -> typ  | 
| 20784 | 30  | 
val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context  | 
| 20310 | 31  | 
val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context  | 
| 22358 | 32  | 
val theorems_of: Proof.context -> thm list NameSpace.table  | 
| 20310 | 33  | 
val fact_index_of: Proof.context -> FactIndex.T  | 
34  | 
val transfer: theory -> Proof.context -> Proof.context  | 
|
| 20367 | 35  | 
val theory: (theory -> theory) -> Proof.context -> Proof.context  | 
36  | 
val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context  | 
|
| 21728 | 37  | 
val pretty_term_abbrev: Proof.context -> term -> Pretty.T  | 
| 22874 | 38  | 
val pretty_thm_legacy: thm -> Pretty.T  | 
| 20310 | 39  | 
val pretty_thm: Proof.context -> thm -> Pretty.T  | 
40  | 
val pretty_thms: Proof.context -> thm list -> Pretty.T  | 
|
41  | 
val pretty_fact: Proof.context -> string * thm list -> Pretty.T  | 
|
42  | 
val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T  | 
|
43  | 
val pretty_proof_of: Proof.context -> bool -> thm -> Pretty.T  | 
|
44  | 
val string_of_thm: Proof.context -> thm -> string  | 
|
45  | 
val read_typ: Proof.context -> string -> typ  | 
|
46  | 
val read_typ_syntax: Proof.context -> string -> typ  | 
|
47  | 
val read_typ_abbrev: Proof.context -> string -> typ  | 
|
48  | 
val cert_typ: Proof.context -> typ -> typ  | 
|
49  | 
val cert_typ_syntax: Proof.context -> typ -> typ  | 
|
50  | 
val cert_typ_abbrev: Proof.context -> typ -> typ  | 
|
51  | 
val get_skolem: Proof.context -> string -> string  | 
|
52  | 
val revert_skolem: Proof.context -> string -> string  | 
|
53  | 
val revert_skolems: Proof.context -> term -> term  | 
|
| 25328 | 54  | 
val infer_type: Proof.context -> string -> typ  | 
55  | 
val inferred_param: string -> Proof.context -> (string * typ) * Proof.context  | 
|
56  | 
val inferred_fixes: Proof.context -> (string * typ) list * Proof.context  | 
|
57  | 
val read_tyname: Proof.context -> string -> typ  | 
|
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
58  | 
val read_const_proper: Proof.context -> string -> term  | 
| 25328 | 59  | 
val read_const: Proof.context -> string -> term  | 
| 22763 | 60  | 
val decode_term: Proof.context -> term -> term  | 
| 24684 | 61  | 
val read_term_pattern: Proof.context -> string -> term  | 
62  | 
val read_term_schematic: Proof.context -> string -> term  | 
|
63  | 
val read_term_abbrev: Proof.context -> string -> term  | 
|
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
64  | 
val expand_abbrevs: Proof.context -> term -> term  | 
| 20310 | 65  | 
val read_termTs_schematic: Proof.context -> (string -> bool) -> (indexname -> typ option)  | 
| 
14720
 
ceff6d4fb836
cleanup up read functions, include liberal versions;
 
wenzelm 
parents: 
14707 
diff
changeset
 | 
66  | 
-> (indexname -> sort option) -> string list -> (string * typ) list  | 
| 
 
ceff6d4fb836
cleanup up read functions, include liberal versions;
 
wenzelm 
parents: 
14707 
diff
changeset
 | 
67  | 
-> term list * (indexname * typ) list  | 
| 20310 | 68  | 
val cert_term: Proof.context -> term -> term  | 
69  | 
val cert_prop: Proof.context -> term -> term  | 
|
70  | 
val goal_export: Proof.context -> Proof.context -> thm list -> thm list  | 
|
71  | 
val export: Proof.context -> Proof.context -> thm list -> thm list  | 
|
| 21531 | 72  | 
val export_morphism: Proof.context -> Proof.context -> morphism  | 
| 20310 | 73  | 
val add_binds: (indexname * string option) list -> Proof.context -> Proof.context  | 
74  | 
val add_binds_i: (indexname * term option) list -> Proof.context -> Proof.context  | 
|
75  | 
val auto_bind_goal: term list -> Proof.context -> Proof.context  | 
|
76  | 
val auto_bind_facts: term list -> Proof.context -> Proof.context  | 
|
77  | 
val match_bind: bool -> (string list * string) list -> Proof.context -> term list * Proof.context  | 
|
78  | 
val match_bind_i: bool -> (term list * term) list -> Proof.context -> term list * Proof.context  | 
|
79  | 
val read_propp: Proof.context * (string * string list) list list  | 
|
80  | 
-> Proof.context * (term * term list) list list  | 
|
81  | 
val cert_propp: Proof.context * (term * term list) list list  | 
|
82  | 
-> Proof.context * (term * term list) list list  | 
|
83  | 
val read_propp_schematic: Proof.context * (string * string list) list list  | 
|
84  | 
-> Proof.context * (term * term list) list list  | 
|
85  | 
val cert_propp_schematic: Proof.context * (term * term list) list list  | 
|
86  | 
-> Proof.context * (term * term list) list list  | 
|
87  | 
val bind_propp: Proof.context * (string * string list) list list  | 
|
88  | 
-> Proof.context * (term list list * (Proof.context -> Proof.context))  | 
|
89  | 
val bind_propp_i: Proof.context * (term * term list) list list  | 
|
90  | 
-> Proof.context * (term list list * (Proof.context -> Proof.context))  | 
|
91  | 
val bind_propp_schematic: Proof.context * (string * string list) list list  | 
|
92  | 
-> Proof.context * (term list list * (Proof.context -> Proof.context))  | 
|
93  | 
val bind_propp_schematic_i: Proof.context * (term * term list) list list  | 
|
94  | 
-> Proof.context * (term list list * (Proof.context -> Proof.context))  | 
|
| 18042 | 95  | 
val fact_tac: thm list -> int -> tactic  | 
| 20310 | 96  | 
val some_fact_tac: Proof.context -> int -> tactic  | 
97  | 
val get_thm: Proof.context -> thmref -> thm  | 
|
98  | 
val get_thm_closure: Proof.context -> thmref -> thm  | 
|
99  | 
val get_thms: Proof.context -> thmref -> thm list  | 
|
100  | 
val get_thms_closure: Proof.context -> thmref -> thm list  | 
|
101  | 
val valid_thms: Proof.context -> string * thm list -> bool  | 
|
102  | 
val lthms_containing: Proof.context -> FactIndex.spec -> (string * thm list) list  | 
|
| 22352 | 103  | 
val add_path: string -> Proof.context -> Proof.context  | 
| 20310 | 104  | 
val no_base_names: Proof.context -> Proof.context  | 
105  | 
val qualified_names: Proof.context -> Proof.context  | 
|
106  | 
val sticky_prefix: string -> Proof.context -> Proof.context  | 
|
107  | 
val restore_naming: Proof.context -> Proof.context -> Proof.context  | 
|
| 21728 | 108  | 
val reset_naming: Proof.context -> Proof.context  | 
| 20310 | 109  | 
val hide_thms: bool -> string list -> Proof.context -> Proof.context  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
110  | 
val put_thms: string * thm list option -> Proof.context -> Proof.context  | 
| 
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
111  | 
val note_thmss: string ->  | 
| 18728 | 112  | 
((bstring * attribute list) * (thmref * attribute list) list) list ->  | 
| 20310 | 113  | 
Proof.context -> (bstring * thm list) list * Proof.context  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
114  | 
val note_thmss_i: string ->  | 
| 18728 | 115  | 
((bstring * attribute list) * (thm list * attribute list) list) list ->  | 
| 20310 | 116  | 
Proof.context -> (bstring * thm list) list * Proof.context  | 
117  | 
val read_vars: (string * string option * mixfix) list -> Proof.context ->  | 
|
118  | 
(string * typ option * mixfix) list * Proof.context  | 
|
119  | 
val cert_vars: (string * typ option * mixfix) list -> Proof.context ->  | 
|
120  | 
(string * typ option * mixfix) list * Proof.context  | 
|
121  | 
val read_vars_legacy: (string * string option * mixfix) list -> Proof.context ->  | 
|
122  | 
(string * typ option * mixfix) list * Proof.context  | 
|
123  | 
val cert_vars_legacy: (string * typ option * mixfix) list -> Proof.context ->  | 
|
124  | 
(string * typ option * mixfix) list * Proof.context  | 
|
125  | 
val add_fixes: (string * string option * mixfix) list ->  | 
|
126  | 
Proof.context -> string list * Proof.context  | 
|
127  | 
val add_fixes_i: (string * typ option * mixfix) list ->  | 
|
128  | 
Proof.context -> string list * Proof.context  | 
|
129  | 
val add_fixes_legacy: (string * typ option * mixfix) list ->  | 
|
130  | 
Proof.context -> string list * Proof.context  | 
|
131  | 
val auto_fixes: Proof.context * (term list list * 'a) -> Proof.context * (term list list * 'a)  | 
|
132  | 
val bind_fixes: string list -> Proof.context -> (term -> term) * Proof.context  | 
|
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
133  | 
val add_assms: Assumption.export ->  | 
| 19585 | 134  | 
((string * attribute list) * (string * string list) list) list ->  | 
| 20310 | 135  | 
Proof.context -> (bstring * thm list) list * Proof.context  | 
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
136  | 
val add_assms_i: Assumption.export ->  | 
| 19585 | 137  | 
((string * attribute list) * (term * term list) list) list ->  | 
| 20310 | 138  | 
Proof.context -> (bstring * thm list) list * Proof.context  | 
139  | 
val add_cases: bool -> (string * RuleCases.T option) list -> Proof.context -> Proof.context  | 
|
140  | 
val apply_case: RuleCases.T -> Proof.context -> (string * term list) list * Proof.context  | 
|
141  | 
val get_case: Proof.context -> string -> string option list -> RuleCases.T  | 
|
| 24949 | 142  | 
val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context  | 
143  | 
val target_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism ->  | 
|
| 21744 | 144  | 
Context.generic -> Context.generic  | 
| 24767 | 145  | 
val add_const_constraint: string * typ option -> Proof.context -> Proof.context  | 
| 24778 | 146  | 
val add_abbrev: string -> Markup.property list ->  | 
147  | 
bstring * term -> Proof.context -> (term * term) * Proof.context  | 
|
| 25052 | 148  | 
val revert_abbrev: string -> string -> Proof.context -> Proof.context  | 
| 10810 | 149  | 
val verbose: bool ref  | 
150  | 
  val setmp_verbose: ('a -> 'b) -> 'a -> 'b
 | 
|
| 20310 | 151  | 
val print_syntax: Proof.context -> unit  | 
| 21728 | 152  | 
val print_abbrevs: Proof.context -> unit  | 
| 20310 | 153  | 
val print_binds: Proof.context -> unit  | 
154  | 
val print_lthms: Proof.context -> unit  | 
|
155  | 
val print_cases: Proof.context -> unit  | 
|
156  | 
val debug: bool ref  | 
|
| 10810 | 157  | 
val prems_limit: int ref  | 
| 20310 | 158  | 
val pretty_ctxt: Proof.context -> Pretty.T list  | 
159  | 
val pretty_context: Proof.context -> Pretty.T list  | 
|
| 5819 | 160  | 
end;  | 
161  | 
||
| 16540 | 162  | 
structure ProofContext: PROOF_CONTEXT =  | 
| 5819 | 163  | 
struct  | 
164  | 
||
| 16540 | 165  | 
val theory_of = Context.theory_of_proof;  | 
| 19001 | 166  | 
val tsig_of = Sign.tsig_of o theory_of;  | 
167  | 
||
| 16540 | 168  | 
val init = Context.init_proof;  | 
| 12057 | 169  | 
|
| 7270 | 170  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
171  | 
(** inner syntax mode **)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
172  | 
|
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
173  | 
datatype mode =  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
174  | 
Mode of  | 
| 24486 | 175  | 
   {stmt: bool,                (*inner statement mode*)
 | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
176  | 
pattern: bool, (*pattern binding schematic variables*)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
177  | 
schematic: bool, (*term referencing loose schematic variables*)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
178  | 
abbrev: bool}; (*abbrev mode -- no normalization*)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
179  | 
|
| 24486 | 180  | 
fun make_mode (stmt, pattern, schematic, abbrev) =  | 
181  | 
  Mode {stmt = stmt, pattern = pattern, schematic = schematic, abbrev = abbrev};
 | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
182  | 
|
| 24486 | 183  | 
val mode_default = make_mode (false, false, false, false);  | 
184  | 
val mode_stmt = make_mode (true, false, false, false);  | 
|
185  | 
val mode_pattern = make_mode (false, true, false, false);  | 
|
186  | 
val mode_schematic = make_mode (false, false, true, false);  | 
|
187  | 
val mode_abbrev = make_mode (false, false, false, true);  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
188  | 
|
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
189  | 
|
| 5819 | 190  | 
|
| 16540 | 191  | 
(** Isar proof context information **)  | 
| 5819 | 192  | 
|
| 16540 | 193  | 
datatype ctxt =  | 
194  | 
Ctxt of  | 
|
| 
24675
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
195  | 
   {mode: mode,                                       (*inner syntax mode*)
 | 
| 
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
196  | 
naming: NameSpace.naming, (*local naming conventions*)  | 
| 
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
197  | 
syntax: LocalSyntax.T, (*local syntax*)  | 
| 25039 | 198  | 
consts: Consts.T * Consts.T, (*local/global consts*)  | 
| 
24675
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
199  | 
thms: thm list NameSpace.table * FactIndex.T, (*local thms*)  | 
| 
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
200  | 
cases: (string * (RuleCases.T * bool)) list}; (*local contexts*)  | 
| 5819 | 201  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
202  | 
fun make_ctxt (mode, naming, syntax, consts, thms, cases) =  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
203  | 
  Ctxt {mode = mode, naming = naming, syntax = syntax,
 | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
204  | 
consts = consts, thms = thms, cases = cases};  | 
| 5819 | 205  | 
|
| 
19079
 
9a7678a0736d
added put_thms_internal: local_naming, no fact index;
 
wenzelm 
parents: 
19062 
diff
changeset
 | 
206  | 
val local_naming = NameSpace.default_naming |> NameSpace.add_path "local";  | 
| 
 
9a7678a0736d
added put_thms_internal: local_naming, no fact index;
 
wenzelm 
parents: 
19062 
diff
changeset
 | 
207  | 
|
| 16540 | 208  | 
structure ContextData = ProofDataFun  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
209  | 
(  | 
| 16540 | 210  | 
type T = ctxt;  | 
211  | 
fun init thy =  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
212  | 
make_ctxt (mode_default, local_naming, LocalSyntax.init thy,  | 
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
213  | 
(Sign.consts_of thy, Sign.consts_of thy),  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
214  | 
(NameSpace.empty_table, FactIndex.empty), []);  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
215  | 
);  | 
| 5819 | 216  | 
|
| 16540 | 217  | 
fun rep_context ctxt = ContextData.get ctxt |> (fn Ctxt args => args);  | 
| 5819 | 218  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
219  | 
fun map_context f =  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
220  | 
  ContextData.map (fn Ctxt {mode, naming, syntax, consts, thms, cases} =>
 | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
221  | 
make_ctxt (f (mode, naming, syntax, consts, thms, cases)));  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
222  | 
|
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
223  | 
fun set_mode mode = map_context (fn (_, naming, syntax, consts, thms, cases) =>  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
224  | 
(mode, naming, syntax, consts, thms, cases));  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
225  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
226  | 
fun map_mode f =  | 
| 24486 | 227  | 
  map_context (fn (Mode {stmt, pattern, schematic, abbrev}, naming, syntax, consts, thms, cases) =>
 | 
228  | 
(make_mode (f (stmt, pattern, schematic, abbrev)), naming, syntax, consts, thms, cases));  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
229  | 
|
| 19001 | 230  | 
fun map_naming f =  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
231  | 
map_context (fn (mode, naming, syntax, consts, thms, cases) =>  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
232  | 
(mode, f naming, syntax, consts, thms, cases));  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
233  | 
|
| 19001 | 234  | 
fun map_syntax f =  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
235  | 
map_context (fn (mode, naming, syntax, consts, thms, cases) =>  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
236  | 
(mode, naming, f syntax, consts, thms, cases));  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
237  | 
|
| 19001 | 238  | 
fun map_consts f =  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
239  | 
map_context (fn (mode, naming, syntax, consts, thms, cases) =>  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
240  | 
(mode, naming, syntax, f consts, thms, cases));  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
241  | 
|
| 19001 | 242  | 
fun map_thms f =  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
243  | 
map_context (fn (mode, naming, syntax, consts, thms, cases) =>  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
244  | 
(mode, naming, syntax, consts, f thms, cases));  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
245  | 
|
| 19001 | 246  | 
fun map_cases f =  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
247  | 
map_context (fn (mode, naming, syntax, consts, thms, cases) =>  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
248  | 
(mode, naming, syntax, consts, thms, f cases));  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
249  | 
|
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
250  | 
val get_mode = #mode o rep_context;  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
251  | 
fun restore_mode ctxt = set_mode (get_mode ctxt);  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
252  | 
|
| 24486 | 253  | 
fun set_stmt stmt =  | 
254  | 
map_mode (fn (_, pattern, schematic, abbrev) => (stmt, pattern, schematic, abbrev));  | 
|
| 19001 | 255  | 
|
256  | 
val naming_of = #naming o rep_context;  | 
|
| 19387 | 257  | 
val full_name = NameSpace.full o naming_of;  | 
| 5819 | 258  | 
|
| 16540 | 259  | 
val syntax_of = #syntax o rep_context;  | 
| 19001 | 260  | 
val syn_of = LocalSyntax.syn_of o syntax_of;  | 
| 19543 | 261  | 
val set_syntax_mode = map_syntax o LocalSyntax.set_mode;  | 
262  | 
val restore_syntax_mode = map_syntax o LocalSyntax.restore_mode o syntax_of;  | 
|
| 19001 | 263  | 
|
| 25052 | 264  | 
val consts_of = #1 o #consts o rep_context;  | 
| 21183 | 265  | 
val const_syntax_name = Consts.syntax_name o consts_of;  | 
| 24752 | 266  | 
val the_const_constraint = Consts.the_constraint o consts_of;  | 
| 5819 | 267  | 
|
| 16540 | 268  | 
val thms_of = #thms o rep_context;  | 
| 22358 | 269  | 
val theorems_of = #1 o thms_of;  | 
| 19001 | 270  | 
val fact_index_of = #2 o thms_of;  | 
| 16540 | 271  | 
|
272  | 
val cases_of = #cases o rep_context;  | 
|
| 5819 | 273  | 
|
274  | 
||
| 20367 | 275  | 
(* theory transfer *)  | 
| 12093 | 276  | 
|
| 19001 | 277  | 
fun transfer_syntax thy =  | 
278  | 
map_syntax (LocalSyntax.rebuild thy) #>  | 
|
| 25039 | 279  | 
map_consts (fn consts as (local_consts, global_consts) =>  | 
| 
19033
 
24e251657e56
consts: maintain thy version for efficient transfer;
 
wenzelm 
parents: 
19019 
diff
changeset
 | 
280  | 
let val thy_consts = Sign.consts_of thy in  | 
| 
 
24e251657e56
consts: maintain thy version for efficient transfer;
 
wenzelm 
parents: 
19019 
diff
changeset
 | 
281  | 
if Consts.eq_consts (thy_consts, global_consts) then consts  | 
| 25039 | 282  | 
else (Consts.merge (local_consts, thy_consts), thy_consts)  | 
| 
19033
 
24e251657e56
consts: maintain thy version for efficient transfer;
 
wenzelm 
parents: 
19019 
diff
changeset
 | 
283  | 
end);  | 
| 17072 | 284  | 
|
| 19001 | 285  | 
fun transfer thy = Context.transfer_proof thy #> transfer_syntax thy;  | 
| 17072 | 286  | 
|
| 20367 | 287  | 
fun theory f ctxt = transfer (f (theory_of ctxt)) ctxt;  | 
288  | 
||
289  | 
fun theory_result f ctxt =  | 
|
290  | 
let val (res, thy') = f (theory_of ctxt)  | 
|
291  | 
in (res, ctxt |> transfer thy') end;  | 
|
| 19019 | 292  | 
|
| 12093 | 293  | 
|
294  | 
||
| 14828 | 295  | 
(** pretty printing **)  | 
296  | 
||
| 24922 | 297  | 
fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt);  | 
| 14828 | 298  | 
|
| 22874 | 299  | 
fun pretty_thm_legacy th =  | 
300  | 
let val thy = Thm.theory_of_thm th  | 
|
| 24961 | 301  | 
in Display.pretty_thm_aux (Syntax.pp (init thy)) true false [] th end;  | 
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
302  | 
|
| 17451 | 303  | 
fun pretty_thm ctxt th =  | 
| 22874 | 304  | 
let val asms = map Thm.term_of (Assumption.assms_of ctxt)  | 
| 24961 | 305  | 
in Display.pretty_thm_aux (Syntax.pp ctxt) false true asms th end;  | 
| 14828 | 306  | 
|
307  | 
fun pretty_thms ctxt [th] = pretty_thm ctxt th  | 
|
308  | 
| pretty_thms ctxt ths = Pretty.blk (0, Pretty.fbreaks (map (pretty_thm ctxt) ths));  | 
|
309  | 
||
310  | 
fun pretty_fact ctxt ("", ths) = pretty_thms ctxt ths
 | 
|
311  | 
| pretty_fact ctxt (a, [th]) =  | 
|
312  | 
Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, pretty_thm ctxt th]  | 
|
313  | 
| pretty_fact ctxt (a, ths) =  | 
|
314  | 
Pretty.block (Pretty.fbreaks (Pretty.str (a ^ ":") :: map (pretty_thm ctxt) ths));  | 
|
315  | 
||
| 17072 | 316  | 
fun pretty_proof ctxt prf =  | 
| 21728 | 317  | 
pretty_term_abbrev (ctxt |> transfer_syntax (ProofSyntax.proof_syntax prf (theory_of ctxt)))  | 
| 17072 | 318  | 
(ProofSyntax.term_of_proof prf);  | 
319  | 
||
320  | 
fun pretty_proof_of ctxt full th =  | 
|
321  | 
pretty_proof ctxt (ProofSyntax.proof_of full th);  | 
|
322  | 
||
| 
17860
 
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
 
wenzelm 
parents: 
17756 
diff
changeset
 | 
323  | 
val string_of_thm = Pretty.string_of oo pretty_thm;  | 
| 
 
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
 
wenzelm 
parents: 
17756 
diff
changeset
 | 
324  | 
|
| 14828 | 325  | 
|
326  | 
||
| 5819 | 327  | 
(** prepare types **)  | 
328  | 
||
| 24277 | 329  | 
(* read_typ *)  | 
330  | 
||
331  | 
fun read_typ_mode mode ctxt s =  | 
|
| 24486 | 332  | 
Syntax.read_typ (Type.set_mode mode ctxt) s;  | 
| 24277 | 333  | 
|
334  | 
val read_typ = read_typ_mode Type.mode_default;  | 
|
335  | 
val read_typ_syntax = read_typ_mode Type.mode_syntax;  | 
|
336  | 
val read_typ_abbrev = read_typ_mode Type.mode_abbrev;  | 
|
337  | 
||
338  | 
||
339  | 
(* cert_typ *)  | 
|
340  | 
||
341  | 
fun cert_typ_mode mode ctxt T =  | 
|
342  | 
Sign.certify_typ_mode mode (theory_of ctxt) T  | 
|
343  | 
handle TYPE (msg, _, _) => error msg;  | 
|
344  | 
||
345  | 
val cert_typ = cert_typ_mode Type.mode_default;  | 
|
346  | 
val cert_typ_syntax = cert_typ_mode Type.mode_syntax;  | 
|
347  | 
val cert_typ_abbrev = cert_typ_mode Type.mode_abbrev;  | 
|
348  | 
||
349  | 
||
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
350  | 
|
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
351  | 
(** prepare variables **)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
352  | 
|
| 7679 | 353  | 
(* internalize Skolem constants *)  | 
354  | 
||
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
355  | 
val lookup_skolem = AList.lookup (op =) o Variable.fixes_of;  | 
| 18187 | 356  | 
fun get_skolem ctxt x = the_default x (lookup_skolem ctxt x);  | 
| 7679 | 357  | 
|
| 18678 | 358  | 
fun no_skolem internal x =  | 
| 
20086
 
94ca946fb689
adapted to more efficient Name/Variable implementation;
 
wenzelm 
parents: 
20049 
diff
changeset
 | 
359  | 
if can Name.dest_skolem x then  | 
| 18678 | 360  | 
    error ("Illegal reference to internal Skolem constant: " ^ quote x)
 | 
| 
20086
 
94ca946fb689
adapted to more efficient Name/Variable implementation;
 
wenzelm 
parents: 
20049 
diff
changeset
 | 
361  | 
else if not internal andalso can Name.dest_internal x then  | 
| 18678 | 362  | 
    error ("Illegal reference to internal variable: " ^ quote x)
 | 
| 7679 | 363  | 
else x;  | 
364  | 
||
365  | 
||
| 
20244
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
366  | 
(* revert Skolem constants -- approximation only! *)  | 
| 18255 | 367  | 
|
| 
20244
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
368  | 
fun revert_skolem ctxt =  | 
| 9133 | 369  | 
let  | 
| 
20244
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
370  | 
val rev_fixes = map Library.swap (Variable.fixes_of ctxt);  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
371  | 
val revert = AList.lookup (op =) rev_fixes;  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
372  | 
in  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
373  | 
fn x =>  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
374  | 
(case revert x of  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
375  | 
SOME x' => x'  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
376  | 
| NONE => perhaps (try Name.dest_skolem) x)  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
377  | 
end;  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
378  | 
|
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
379  | 
fun revert_skolems ctxt =  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
380  | 
let  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
381  | 
val revert = revert_skolem ctxt;  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
382  | 
fun reverts (Free (x, T)) = Free (revert x, T)  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
383  | 
| reverts (t $ u) = reverts t $ reverts u  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
384  | 
| reverts (Abs (x, T, t)) = Abs (x, T, reverts t)  | 
| 
 
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
 
wenzelm 
parents: 
20234 
diff
changeset
 | 
385  | 
| reverts a = a;  | 
| 24495 | 386  | 
in reverts end;  | 
| 9133 | 387  | 
|
| 8096 | 388  | 
|
| 18187 | 389  | 
|
| 5819 | 390  | 
(** prepare terms and propositions **)  | 
391  | 
||
| 25328 | 392  | 
(* inferred types of parameters *)  | 
393  | 
||
394  | 
fun infer_type ctxt x =  | 
|
395  | 
Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt))  | 
|
396  | 
(Free (x, dummyT)));  | 
|
397  | 
||
398  | 
fun inferred_param x ctxt =  | 
|
399  | 
let val T = infer_type ctxt x  | 
|
400  | 
in ((x, T), ctxt |> Variable.declare_term (Free (x, T))) end;  | 
|
401  | 
||
402  | 
fun inferred_fixes ctxt =  | 
|
403  | 
fold_map inferred_param (rev (map #2 (Variable.fixes_of ctxt))) ctxt;  | 
|
404  | 
||
405  | 
||
406  | 
(* type and constant names *)  | 
|
407  | 
||
408  | 
(* type and constant names *)  | 
|
409  | 
||
410  | 
fun read_tyname ctxt c =  | 
|
411  | 
if Syntax.is_tid c then  | 
|
412  | 
TFree (c, the_default (Sign.defaultS (theory_of ctxt)) (Variable.def_sort ctxt (c, ~1)))  | 
|
413  | 
else  | 
|
414  | 
let  | 
|
415  | 
val thy = theory_of ctxt;  | 
|
416  | 
val d = Sign.intern_type thy c;  | 
|
417  | 
in Type (d, replicate (Sign.arity_number thy d) dummyT) end;  | 
|
418  | 
||
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
419  | 
fun read_const_proper ctxt c =  | 
| 25328 | 420  | 
(case Variable.lookup_const ctxt c of  | 
421  | 
SOME d => Const (d, Consts.type_scheme (consts_of ctxt) d handle TYPE (msg, _, _) => error msg)  | 
|
422  | 
| NONE => Consts.read_const (consts_of ctxt) c);  | 
|
423  | 
||
424  | 
fun read_const ctxt c =  | 
|
425  | 
(case (lookup_skolem ctxt c, Variable.is_const ctxt c) of  | 
|
426  | 
(SOME x, false) => Free (x, infer_type ctxt x)  | 
|
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
427  | 
| _ => read_const_proper ctxt c);  | 
| 25328 | 428  | 
|
429  | 
||
| 24684 | 430  | 
(* read_term *)  | 
431  | 
||
432  | 
fun read_term_mode mode ctxt = Syntax.read_term (set_mode mode ctxt);  | 
|
433  | 
||
434  | 
val read_term_pattern = read_term_mode mode_pattern;  | 
|
435  | 
val read_term_schematic = read_term_mode mode_schematic;  | 
|
436  | 
val read_term_abbrev = read_term_mode mode_abbrev;  | 
|
437  | 
||
438  | 
||
| 19001 | 439  | 
(* local abbreviations *)  | 
| 5819 | 440  | 
|
| 24501 | 441  | 
local  | 
442  | 
||
| 19371 | 443  | 
fun certify_consts ctxt =  | 
| 
24675
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
444  | 
  let val Mode {abbrev, ...} = get_mode ctxt
 | 
| 24961 | 445  | 
in Consts.certify (Syntax.pp ctxt) (tsig_of ctxt) (not abbrev) (consts_of ctxt) end;  | 
| 19001 | 446  | 
|
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
447  | 
fun reject_schematic (Var (xi, _)) =  | 
| 22678 | 448  | 
      error ("Unbound schematic variable: " ^ Term.string_of_vname xi)
 | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
449  | 
| reject_schematic (Abs (_, _, t)) = reject_schematic t  | 
| 
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
450  | 
| reject_schematic (t $ u) = (reject_schematic t; reject_schematic u)  | 
| 
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
451  | 
| reject_schematic _ = ();  | 
| 5819 | 452  | 
|
| 24495 | 453  | 
fun expand_binds ctxt =  | 
454  | 
  let val Mode {pattern, schematic, ...} = get_mode ctxt in
 | 
|
455  | 
if pattern then I  | 
|
456  | 
else Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic)  | 
|
457  | 
end;  | 
|
| 5819 | 458  | 
|
| 24501 | 459  | 
in  | 
460  | 
||
461  | 
fun expand_abbrevs ctxt = certify_consts ctxt #> expand_binds ctxt;  | 
|
462  | 
||
463  | 
end;  | 
|
464  | 
||
| 5819 | 465  | 
|
| 24922 | 466  | 
fun contract_abbrevs ctxt t =  | 
467  | 
let  | 
|
468  | 
val thy = theory_of ctxt;  | 
|
469  | 
val consts = consts_of ctxt;  | 
|
470  | 
    val Mode {abbrev, ...} = get_mode ctxt;
 | 
|
471  | 
in  | 
|
| 25052 | 472  | 
if abbrev orelse print_mode_active "no_abbrevs" orelse not (can Term.type_of t) then t  | 
473  | 
else t |> Pattern.rewrite_term thy (Consts.abbrevs_of consts (print_mode_value () @ [""])) []  | 
|
| 24922 | 474  | 
end;  | 
475  | 
||
476  | 
||
| 24518 | 477  | 
(* patterns *)  | 
478  | 
||
479  | 
fun prepare_patternT ctxt =  | 
|
480  | 
  let val Mode {pattern, schematic, ...} = get_mode ctxt in
 | 
|
481  | 
if pattern orelse schematic then I  | 
|
482  | 
else Term.map_atyps  | 
|
483  | 
(fn T as TVar (xi, _) =>  | 
|
484  | 
if not (TypeInfer.is_param xi)  | 
|
485  | 
            then error ("Illegal schematic type variable: " ^ Term.string_of_vname xi)
 | 
|
486  | 
else T  | 
|
487  | 
| T => T)  | 
|
488  | 
end;  | 
|
489  | 
||
| 22712 | 490  | 
|
| 24505 | 491  | 
local  | 
| 6550 | 492  | 
|
| 24767 | 493  | 
fun prepare_dummies ts = #1 (fold_map Term.replace_dummy_patterns ts 1);  | 
| 6762 | 494  | 
|
| 18678 | 495  | 
fun reject_dummies t = Term.no_dummy_patterns t  | 
496  | 
handle TERM _ => error "Illegal dummy pattern(s) in term";  | 
|
| 6550 | 497  | 
|
| 24505 | 498  | 
in  | 
499  | 
||
| 24684 | 500  | 
fun prepare_patterns ctxt =  | 
| 24518 | 501  | 
  let val Mode {pattern, ...} = get_mode ctxt in
 | 
| 24767 | 502  | 
TypeInfer.fixate_params (Variable.names_of ctxt) #>  | 
503  | 
pattern ? Variable.polymorphic ctxt #>  | 
|
| 24684 | 504  | 
(map o Term.map_types) (prepare_patternT ctxt) #>  | 
| 24767 | 505  | 
(if pattern then prepare_dummies else map reject_dummies)  | 
| 24505 | 506  | 
end;  | 
507  | 
||
508  | 
end;  | 
|
509  | 
||
| 6550 | 510  | 
|
| 22763 | 511  | 
(* decoding raw terms (syntax trees) *)  | 
512  | 
||
513  | 
fun legacy_intern_skolem ctxt internal def_type x = (* FIXME cleanup this mess *)  | 
|
514  | 
let  | 
|
515  | 
val sko = lookup_skolem ctxt x;  | 
|
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
516  | 
val is_const = can (read_const_proper ctxt) x;  | 
| 
25319
 
074d41176558
read_const/legacy_intern_skolem: cover consts within the local scope;
 
wenzelm 
parents: 
25168 
diff
changeset
 | 
517  | 
val is_scoped_const = Variable.is_const ctxt x;  | 
| 
 
074d41176558
read_const/legacy_intern_skolem: cover consts within the local scope;
 
wenzelm 
parents: 
25168 
diff
changeset
 | 
518  | 
val is_free = (is_some sko orelse not is_const) andalso not is_scoped_const;  | 
| 22763 | 519  | 
val is_internal = internal x;  | 
520  | 
val is_declared = is_some (def_type (x, ~1));  | 
|
521  | 
val res =  | 
|
522  | 
if is_free andalso not is_internal then (no_skolem false x; sko)  | 
|
| 23361 | 523  | 
else ((no_skolem false x; ()) handle ERROR msg =>  | 
524  | 
legacy_feature (msg ^ ContextPosition.str_of ctxt);  | 
|
| 22763 | 525  | 
if is_internal andalso is_declared then SOME x else NONE);  | 
| 
25319
 
074d41176558
read_const/legacy_intern_skolem: cover consts within the local scope;
 
wenzelm 
parents: 
25168 
diff
changeset
 | 
526  | 
in  | 
| 
 
074d41176558
read_const/legacy_intern_skolem: cover consts within the local scope;
 
wenzelm 
parents: 
25168 
diff
changeset
 | 
527  | 
if is_some res then res  | 
| 
 
074d41176558
read_const/legacy_intern_skolem: cover consts within the local scope;
 
wenzelm 
parents: 
25168 
diff
changeset
 | 
528  | 
else if is_declared andalso not is_scoped_const then SOME x else NONE  | 
| 
 
074d41176558
read_const/legacy_intern_skolem: cover consts within the local scope;
 
wenzelm 
parents: 
25168 
diff
changeset
 | 
529  | 
end;  | 
| 22763 | 530  | 
|
| 24371 | 531  | 
fun term_context ctxt =  | 
532  | 
let val thy = theory_of ctxt in  | 
|
533  | 
   {get_sort = Sign.get_sort thy (Variable.def_sort ctxt),
 | 
|
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
534  | 
map_const = try (#1 o Term.dest_Const o read_const_proper ctxt),  | 
| 24371 | 535  | 
map_free = legacy_intern_skolem ctxt (K false) (Variable.def_type ctxt false),  | 
536  | 
map_type = Sign.intern_tycons thy,  | 
|
537  | 
map_sort = Sign.intern_sort thy}  | 
|
538  | 
end;  | 
|
539  | 
||
| 22763 | 540  | 
fun decode_term ctxt =  | 
| 24371 | 541  | 
  let val {get_sort, map_const, map_free, map_type, map_sort} = term_context ctxt
 | 
| 22763 | 542  | 
in Syntax.decode_term get_sort map_const map_free map_type map_sort end;  | 
543  | 
||
544  | 
||
| 24687 | 545  | 
(* read terms -- legacy *)  | 
| 5819 | 546  | 
|
| 10554 | 547  | 
local  | 
548  | 
||
| 24687 | 549  | 
fun read_def_termTs freeze pp syn ctxt (types, sorts, used) fixed sTs =  | 
550  | 
Sign.read_def_terms' pp (Sign.is_logtype (theory_of ctxt)) syn (consts_of ctxt) fixed  | 
|
551  | 
ctxt (types, sorts) used freeze sTs;  | 
|
552  | 
||
| 15531 | 553  | 
fun append_env e1 e2 x = (case e2 x of NONE => e1 x | some => some);  | 
| 
14720
 
ceff6d4fb836
cleanup up read functions, include liberal versions;
 
wenzelm 
parents: 
14707 
diff
changeset
 | 
554  | 
|
| 24687 | 555  | 
fun gen_read' read app schematic ctxt0 internal more_types more_sorts more_used s =  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
13629 
diff
changeset
 | 
556  | 
let  | 
| 24687 | 557  | 
val ctxt = if schematic then set_mode mode_schematic ctxt0 else ctxt0;  | 
558  | 
val types = append_env (Variable.def_type ctxt false) more_types;  | 
|
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
559  | 
val sorts = append_env (Variable.def_sort ctxt) more_sorts;  | 
| 20163 | 560  | 
val used = fold Name.declare more_used (Variable.names_of ctxt);  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
13629 
diff
changeset
 | 
561  | 
in  | 
| 24961 | 562  | 
(read (Syntax.pp ctxt) (syn_of ctxt) ctxt (types, sorts, used)  | 
| 22763 | 563  | 
(legacy_intern_skolem ctxt internal types) s  | 
| 18678 | 564  | 
handle TERM (msg, _) => error msg)  | 
| 24505 | 565  | 
|> app (expand_abbrevs ctxt)  | 
| 24684 | 566  | 
|> app (singleton (prepare_patterns ctxt))  | 
| 
14720
 
ceff6d4fb836
cleanup up read functions, include liberal versions;
 
wenzelm 
parents: 
14707 
diff
changeset
 | 
567  | 
end;  | 
| 
 
ceff6d4fb836
cleanup up read functions, include liberal versions;
 
wenzelm 
parents: 
14707 
diff
changeset
 | 
568  | 
|
| 10554 | 569  | 
in  | 
570  | 
||
| 24687 | 571  | 
val read_termTs_schematic = gen_read' (read_def_termTs false) (apfst o map) true;  | 
| 8096 | 572  | 
|
| 10554 | 573  | 
end;  | 
574  | 
||
| 5819 | 575  | 
|
576  | 
(* certify terms *)  | 
|
577  | 
||
| 10554 | 578  | 
local  | 
579  | 
||
| 24684 | 580  | 
fun gen_cert prop ctxt t =  | 
581  | 
t  | 
|
582  | 
|> expand_abbrevs ctxt  | 
|
| 24961 | 583  | 
|> (fn t' => #1 (Sign.certify' prop (Syntax.pp ctxt) false (consts_of ctxt) (theory_of ctxt) t')  | 
| 24684 | 584  | 
handle TYPE (msg, _, _) => error msg  | 
585  | 
| TERM (msg, _) => error msg);  | 
|
| 16501 | 586  | 
|
| 10554 | 587  | 
in  | 
| 8096 | 588  | 
|
| 24684 | 589  | 
val cert_term = gen_cert false;  | 
590  | 
val cert_prop = gen_cert true;  | 
|
| 10554 | 591  | 
|
592  | 
end;  | 
|
| 5819 | 593  | 
|
594  | 
||
| 24495 | 595  | 
(* type checking/inference *)  | 
| 22701 | 596  | 
|
| 
25168
 
2650a4a6ad3e
parse_term: invoke full Syntax.check_term, not just standard_infer_types;
 
wenzelm 
parents: 
25159 
diff
changeset
 | 
597  | 
local  | 
| 
 
2650a4a6ad3e
parse_term: invoke full Syntax.check_term, not just standard_infer_types;
 
wenzelm 
parents: 
25159 
diff
changeset
 | 
598  | 
|
| 24495 | 599  | 
fun standard_infer_types ctxt ts =  | 
600  | 
  let val Mode {pattern, ...} = get_mode ctxt in
 | 
|
| 24961 | 601  | 
TypeInfer.infer_types (Syntax.pp ctxt) (tsig_of ctxt) (Syntax.check_typs ctxt)  | 
| 24495 | 602  | 
(try (Consts.the_constraint (consts_of ctxt))) (Variable.def_type ctxt pattern)  | 
| 24767 | 603  | 
(Variable.names_of ctxt) (Variable.maxidx_of ctxt) NONE (map (rpair dummyT) ts)  | 
| 24495 | 604  | 
|> #1 |> map #1  | 
605  | 
handle TYPE (msg, _, _) => error msg  | 
|
606  | 
end;  | 
|
| 24486 | 607  | 
|
| 24518 | 608  | 
fun standard_typ_check ctxt =  | 
609  | 
map (cert_typ_mode (Type.get_mode ctxt) ctxt) #>  | 
|
610  | 
map (prepare_patternT ctxt);  | 
|
611  | 
||
| 
24769
 
1372969969e0
standard_term_check: include expand_abbrevs (back again);
 
wenzelm 
parents: 
24767 
diff
changeset
 | 
612  | 
fun standard_term_check ctxt =  | 
| 
 
1372969969e0
standard_term_check: include expand_abbrevs (back again);
 
wenzelm 
parents: 
24767 
diff
changeset
 | 
613  | 
standard_infer_types ctxt #>  | 
| 
 
1372969969e0
standard_term_check: include expand_abbrevs (back again);
 
wenzelm 
parents: 
24767 
diff
changeset
 | 
614  | 
map (expand_abbrevs ctxt);  | 
| 24922 | 615  | 
|
616  | 
fun standard_term_uncheck ctxt =  | 
|
617  | 
map (contract_abbrevs ctxt);  | 
|
618  | 
||
| 25060 | 619  | 
fun add eq what f = Context.add_setup  | 
620  | 
(Context.theory_map (what (fn xs => fn ctxt =>  | 
|
621  | 
let val xs' = f ctxt xs in if eq_list eq (xs, xs') then NONE else SOME (xs', ctxt) end)));  | 
|
| 24767 | 622  | 
|
| 24518 | 623  | 
in  | 
624  | 
||
| 25060 | 625  | 
val _ = add (op =) (Syntax.add_typ_check 0 "standard") standard_typ_check;  | 
626  | 
val _ = add (op aconv) (Syntax.add_term_check 0 "standard") standard_term_check;  | 
|
627  | 
val _ = add (op aconv) (Syntax.add_term_check 100 "fixate") prepare_patterns;  | 
|
| 24922 | 628  | 
|
| 25060 | 629  | 
val _ = add (op aconv) (Syntax.add_term_uncheck 0 "standard") standard_term_uncheck;  | 
| 22701 | 630  | 
|
| 24518 | 631  | 
end;  | 
| 22701 | 632  | 
|
633  | 
||
| 9553 | 634  | 
|
| 24767 | 635  | 
(** inner syntax operations **)  | 
| 24371 | 636  | 
|
637  | 
local  | 
|
638  | 
||
639  | 
fun parse_sort ctxt str =  | 
|
640  | 
Syntax.standard_parse_sort ctxt (syn_of ctxt) (Sign.intern_sort (theory_of ctxt)) str;  | 
|
641  | 
||
642  | 
fun parse_typ ctxt str =  | 
|
643  | 
let  | 
|
644  | 
val thy = ProofContext.theory_of ctxt;  | 
|
645  | 
val get_sort = Sign.get_sort thy (Variable.def_sort ctxt);  | 
|
646  | 
val T = Sign.intern_tycons thy  | 
|
647  | 
(Syntax.standard_parse_typ ctxt (syn_of ctxt) get_sort (Sign.intern_sort thy) str);  | 
|
648  | 
in T end  | 
|
649  | 
  handle ERROR msg => cat_error msg ("The error(s) above occurred in type " ^ quote str);
 | 
|
650  | 
||
651  | 
fun parse_term T ctxt str =  | 
|
652  | 
let  | 
|
653  | 
val thy = theory_of ctxt;  | 
|
| 
25168
 
2650a4a6ad3e
parse_term: invoke full Syntax.check_term, not just standard_infer_types;
 
wenzelm 
parents: 
25159 
diff
changeset
 | 
654  | 
val (T', _) = TypeInfer.paramify_dummies T 0;  | 
| 
 
2650a4a6ad3e
parse_term: invoke full Syntax.check_term, not just standard_infer_types;
 
wenzelm 
parents: 
25159 
diff
changeset
 | 
655  | 
fun check t = Exn.Result (Syntax.check_term ctxt (TypeInfer.constrain T' t))  | 
| 
 
2650a4a6ad3e
parse_term: invoke full Syntax.check_term, not just standard_infer_types;
 
wenzelm 
parents: 
25159 
diff
changeset
 | 
656  | 
handle ERROR msg => Exn.Exn (ERROR msg);  | 
| 24371 | 657  | 
    val {get_sort, map_const, map_free, map_type, map_sort} = term_context ctxt;
 | 
| 24961 | 658  | 
val read = Syntax.standard_parse_term (Syntax.pp ctxt) check get_sort  | 
| 
25168
 
2650a4a6ad3e
parse_term: invoke full Syntax.check_term, not just standard_infer_types;
 
wenzelm 
parents: 
25159 
diff
changeset
 | 
659  | 
map_const map_free map_type map_sort ctxt (Sign.is_logtype thy) (syn_of ctxt) T';  | 
| 24371 | 660  | 
in read str end;  | 
661  | 
||
| 24922 | 662  | 
|
663  | 
fun unparse_sort ctxt S =  | 
|
664  | 
Syntax.standard_unparse_sort ctxt (syn_of ctxt) (Sign.extern_sort (theory_of ctxt) S);  | 
|
665  | 
||
666  | 
fun unparse_typ ctxt T =  | 
|
667  | 
Syntax.standard_unparse_typ ctxt (syn_of ctxt) (Sign.extern_typ (theory_of ctxt) T);  | 
|
668  | 
||
669  | 
fun unparse_term ctxt t =  | 
|
670  | 
let  | 
|
671  | 
val thy = theory_of ctxt;  | 
|
672  | 
val syntax = syntax_of ctxt;  | 
|
673  | 
val consts = consts_of ctxt;  | 
|
674  | 
in  | 
|
675  | 
t  | 
|
676  | 
|> Sign.extern_term (Consts.extern_early consts) thy  | 
|
677  | 
|> LocalSyntax.extern_term syntax  | 
|
678  | 
|> Syntax.standard_unparse_term (Consts.extern consts) ctxt (LocalSyntax.syn_of syntax)  | 
|
679  | 
(Context.exists_name Context.CPureN thy)  | 
|
680  | 
end;  | 
|
681  | 
||
| 24371 | 682  | 
in  | 
683  | 
||
| 24767 | 684  | 
val _ = Syntax.install_operations  | 
685  | 
  {parse_sort = parse_sort,
 | 
|
686  | 
parse_typ = parse_typ,  | 
|
687  | 
parse_term = parse_term dummyT,  | 
|
688  | 
parse_prop = parse_term propT,  | 
|
| 24922 | 689  | 
unparse_sort = unparse_sort,  | 
690  | 
unparse_typ = unparse_typ,  | 
|
691  | 
unparse_term = unparse_term};  | 
|
| 24371 | 692  | 
|
693  | 
end;  | 
|
694  | 
||
695  | 
||
696  | 
||
| 21610 | 697  | 
(** export results **)  | 
| 21531 | 698  | 
|
| 20310 | 699  | 
fun common_export is_goal inner outer =  | 
700  | 
map (Assumption.export is_goal inner outer) #>  | 
|
701  | 
Variable.export inner outer;  | 
|
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
702  | 
|
| 20310 | 703  | 
val goal_export = common_export true;  | 
704  | 
val export = common_export false;  | 
|
| 12704 | 705  | 
|
| 21531 | 706  | 
fun export_morphism inner outer =  | 
707  | 
Assumption.export_morphism inner outer $>  | 
|
708  | 
Variable.export_morphism inner outer;  | 
|
709  | 
||
710  | 
||
| 
15758
 
07e382399a96
binds/thms: do not store options, but delete from table;
 
wenzelm 
parents: 
15750 
diff
changeset
 | 
711  | 
|
| 5819 | 712  | 
(** bindings **)  | 
713  | 
||
| 8096 | 714  | 
(* simult_matches *)  | 
715  | 
||
| 19867 | 716  | 
fun simult_matches ctxt (t, pats) =  | 
717  | 
(case Seq.pull (Unify.matchers (theory_of ctxt) (map (rpair t) pats)) of  | 
|
718  | 
NONE => error "Pattern match failed!"  | 
|
719  | 
| SOME (env, _) => map (apsnd snd) (Envir.alist_of env));  | 
|
| 8096 | 720  | 
|
721  | 
||
722  | 
(* add_binds(_i) *)  | 
|
| 5819 | 723  | 
|
| 7925 | 724  | 
local  | 
725  | 
||
| 
16031
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
726  | 
fun gen_bind prep (xi as (x, _), raw_t) ctxt =  | 
| 
24511
 
69d270cc7e4f
removed obsolete read/cert variations (cf. Syntax.read/check);
 
wenzelm 
parents: 
24505 
diff
changeset
 | 
727  | 
ctxt  | 
| 
24675
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
728  | 
|> Variable.add_binds [(xi, Option.map (prep (set_mode mode_default ctxt)) raw_t)];  | 
| 5819 | 729  | 
|
| 10810 | 730  | 
in  | 
731  | 
||
| 20330 | 732  | 
fun drop_schematic (b as (xi, SOME t)) = if Term.exists_subterm is_Var t then (xi, NONE) else b  | 
| 10554 | 733  | 
| drop_schematic b = b;  | 
734  | 
||
| 
24511
 
69d270cc7e4f
removed obsolete read/cert variations (cf. Syntax.read/check);
 
wenzelm 
parents: 
24505 
diff
changeset
 | 
735  | 
val add_binds = fold (gen_bind Syntax.read_term);  | 
| 
16031
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
736  | 
val add_binds_i = fold (gen_bind cert_term);  | 
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
737  | 
|
| 16458 | 738  | 
fun auto_bind f ts ctxt = ctxt |> add_binds_i (map drop_schematic (f (theory_of ctxt) ts));  | 
| 12147 | 739  | 
val auto_bind_goal = auto_bind AutoBind.goal;  | 
740  | 
val auto_bind_facts = auto_bind AutoBind.facts;  | 
|
| 7925 | 741  | 
|
742  | 
end;  | 
|
| 5819 | 743  | 
|
744  | 
||
| 8096 | 745  | 
(* match_bind(_i) *)  | 
| 5819 | 746  | 
|
| 8096 | 747  | 
local  | 
748  | 
||
| 24684 | 749  | 
fun gen_bind prep_terms gen raw_binds ctxt =  | 
| 5819 | 750  | 
let  | 
| 24684 | 751  | 
fun prep_bind (raw_pats, t) ctxt1 =  | 
752  | 
let  | 
|
753  | 
val T = Term.fastype_of t;  | 
|
754  | 
val ctxt2 = Variable.declare_term t ctxt1;  | 
|
755  | 
val pats = prep_terms (set_mode mode_pattern ctxt2) T raw_pats;  | 
|
756  | 
val binds = simult_matches ctxt2 (t, pats);  | 
|
757  | 
in (binds, ctxt2) end;  | 
|
| 7670 | 758  | 
|
| 24686 | 759  | 
val ts = prep_terms ctxt dummyT (map snd raw_binds);  | 
760  | 
val (binds, ctxt') = apfst flat (fold_map prep_bind (map fst raw_binds ~~ ts) ctxt);  | 
|
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
761  | 
val binds' =  | 
| 
19916
 
3bbb9cc5d4f1
export: simultaneous facts, refer to Variable.export;
 
wenzelm 
parents: 
19897 
diff
changeset
 | 
762  | 
if gen then map #1 binds ~~ Variable.exportT_terms ctxt' ctxt (map #2 binds)  | 
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
763  | 
else binds;  | 
| 15531 | 764  | 
val binds'' = map (apsnd SOME) binds';  | 
| 18310 | 765  | 
val ctxt'' =  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
766  | 
tap (Variable.warn_extra_tfrees ctxt)  | 
| 18310 | 767  | 
(if gen then  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
768  | 
ctxt (*sic!*) |> fold Variable.declare_term (map #2 binds') |> add_binds_i binds''  | 
| 18310 | 769  | 
else ctxt' |> add_binds_i binds'');  | 
770  | 
in (ts, ctxt'') end;  | 
|
| 8096 | 771  | 
|
772  | 
in  | 
|
| 5935 | 773  | 
|
| 24684 | 774  | 
fun read_terms ctxt T =  | 
775  | 
map (Syntax.parse_term ctxt #> TypeInfer.constrain T) #> Syntax.check_terms ctxt;  | 
|
776  | 
||
777  | 
val match_bind = gen_bind read_terms;  | 
|
778  | 
val match_bind_i = gen_bind (fn ctxt => fn _ => map (cert_term ctxt));  | 
|
| 8096 | 779  | 
|
780  | 
end;  | 
|
| 5935 | 781  | 
|
782  | 
||
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
783  | 
(* propositions with patterns *)  | 
| 5935 | 784  | 
|
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
785  | 
local  | 
| 8096 | 786  | 
|
| 24684 | 787  | 
fun prep_propp mode prep_props (context, args) =  | 
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
788  | 
let  | 
| 19585 | 789  | 
fun prep (_, raw_pats) (ctxt, prop :: props) =  | 
| 24684 | 790  | 
let val ctxt' = Variable.declare_term prop ctxt  | 
791  | 
in ((prop, prep_props (set_mode mode_pattern ctxt') raw_pats), (ctxt', props)) end;  | 
|
792  | 
||
| 
17860
 
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
 
wenzelm 
parents: 
17756 
diff
changeset
 | 
793  | 
val (propp, (context', _)) = (fold_map o fold_map) prep args  | 
| 24684 | 794  | 
(context, prep_props (set_mode mode context) (maps (map fst) args));  | 
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
795  | 
in (context', propp) end;  | 
| 5935 | 796  | 
|
| 24684 | 797  | 
fun gen_bind_propp mode parse_prop (ctxt, raw_args) =  | 
| 8096 | 798  | 
let  | 
| 24684 | 799  | 
val (ctxt', args) = prep_propp mode parse_prop (ctxt, raw_args);  | 
| 19585 | 800  | 
val binds = flat (flat (map (map (simult_matches ctxt')) args));  | 
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
801  | 
val propss = map (map #1) args;  | 
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
802  | 
|
| 10554 | 803  | 
(*generalize result: context evaluated now, binds added later*)  | 
| 
19916
 
3bbb9cc5d4f1
export: simultaneous facts, refer to Variable.export;
 
wenzelm 
parents: 
19897 
diff
changeset
 | 
804  | 
val gen = Variable.exportT_terms ctxt' ctxt;  | 
| 15531 | 805  | 
fun gen_binds c = c |> add_binds_i (map #1 binds ~~ map SOME (gen (map #2 binds)));  | 
806  | 
in (ctxt' |> add_binds_i (map (apsnd SOME) binds), (propss, gen_binds)) end;  | 
|
| 8096 | 807  | 
|
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
808  | 
in  | 
| 
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
809  | 
|
| 24684 | 810  | 
val read_propp = prep_propp mode_default Syntax.read_props;  | 
811  | 
val cert_propp = prep_propp mode_default (map o cert_prop);  | 
|
812  | 
val read_propp_schematic = prep_propp mode_schematic Syntax.read_props;  | 
|
813  | 
val cert_propp_schematic = prep_propp mode_schematic (map o cert_prop);  | 
|
| 10554 | 814  | 
|
| 24684 | 815  | 
val bind_propp = gen_bind_propp mode_default Syntax.read_props;  | 
816  | 
val bind_propp_i = gen_bind_propp mode_default (map o cert_prop);  | 
|
817  | 
val bind_propp_schematic = gen_bind_propp mode_schematic Syntax.read_props;  | 
|
818  | 
val bind_propp_schematic_i = gen_bind_propp mode_schematic (map o cert_prop);  | 
|
| 6789 | 819  | 
|
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
820  | 
end;  | 
| 
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
821  | 
|
| 6789 | 822  | 
|
| 5819 | 823  | 
|
824  | 
(** theorems **)  | 
|
825  | 
||
| 18042 | 826  | 
(* fact_tac *)  | 
827  | 
||
| 18122 | 828  | 
fun comp_incr_tac [] _ st = no_tac st  | 
829  | 
| comp_incr_tac (th :: ths) i st =  | 
|
830  | 
(Goal.compose_hhf_tac (Drule.incr_indexes st th) i APPEND comp_incr_tac ths i) st;  | 
|
| 18042 | 831  | 
|
| 21687 | 832  | 
fun fact_tac facts = Goal.norm_hhf_tac THEN' comp_incr_tac facts;  | 
| 18122 | 833  | 
|
834  | 
fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) =>  | 
|
| 18042 | 835  | 
let  | 
| 19001 | 836  | 
val index = fact_index_of ctxt;  | 
| 18042 | 837  | 
val facts = FactIndex.could_unify index (Term.strip_all_body goal);  | 
838  | 
in fact_tac facts i end);  | 
|
839  | 
||
840  | 
||
| 6091 | 841  | 
(* get_thm(s) *)  | 
| 5819 | 842  | 
|
| 18042 | 843  | 
fun retrieve_thms _ pick ctxt (Fact s) =  | 
| 16501 | 844  | 
let  | 
| 
24511
 
69d270cc7e4f
removed obsolete read/cert variations (cf. Syntax.read/check);
 
wenzelm 
parents: 
24505 
diff
changeset
 | 
845  | 
val prop = Syntax.read_prop (set_mode mode_default ctxt) s  | 
| 
 
69d270cc7e4f
removed obsolete read/cert variations (cf. Syntax.read/check);
 
wenzelm 
parents: 
24505 
diff
changeset
 | 
846  | 
|> singleton (Variable.polymorphic ctxt);  | 
| 
 
69d270cc7e4f
removed obsolete read/cert variations (cf. Syntax.read/check);
 
wenzelm 
parents: 
24505 
diff
changeset
 | 
847  | 
val th = Goal.prove ctxt [] [] prop (K (ALLGOALS (some_fact_tac ctxt)))  | 
| 18678 | 848  | 
handle ERROR msg => cat_error msg "Failed to retrieve literal fact.";  | 
| 18042 | 849  | 
in pick "" [th] end  | 
850  | 
| retrieve_thms from_thy pick ctxt xthmref =  | 
|
851  | 
let  | 
|
852  | 
val thy = theory_of ctxt;  | 
|
| 19001 | 853  | 
val (space, tab) = #1 (thms_of ctxt);  | 
| 16501 | 854  | 
val thmref = PureThy.map_name_of_thmref (NameSpace.intern space) xthmref;  | 
855  | 
val name = PureThy.name_of_thmref thmref;  | 
|
| 24012 | 856  | 
val thms =  | 
857  | 
if name = "" then [Thm.transfer thy Drule.dummy_thm]  | 
|
858  | 
else  | 
|
859  | 
(case Symtab.lookup tab name of  | 
|
860  | 
SOME ths => map (Thm.transfer thy) (PureThy.select_thm thmref ths)  | 
|
861  | 
| NONE => from_thy thy xthmref);  | 
|
862  | 
in pick name thms end;  | 
|
| 5819 | 863  | 
|
| 9566 | 864  | 
val get_thm = retrieve_thms PureThy.get_thms PureThy.single_thm;  | 
865  | 
val get_thm_closure = retrieve_thms PureThy.get_thms_closure PureThy.single_thm;  | 
|
866  | 
val get_thms = retrieve_thms PureThy.get_thms (K I);  | 
|
867  | 
val get_thms_closure = retrieve_thms PureThy.get_thms_closure (K I);  | 
|
| 5819 | 868  | 
|
869  | 
||
| 
16031
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
870  | 
(* valid_thms *)  | 
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
871  | 
|
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
872  | 
fun valid_thms ctxt (name, ths) =  | 
| 18678 | 873  | 
(case try (fn () => get_thms ctxt (Name name)) () of  | 
| 
16031
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
874  | 
NONE => false  | 
| 16147 | 875  | 
| SOME ths' => Thm.eq_thms (ths, ths'));  | 
| 
16031
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
876  | 
|
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
877  | 
|
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
878  | 
(* lthms_containing *)  | 
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
879  | 
|
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
880  | 
fun lthms_containing ctxt spec =  | 
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
881  | 
FactIndex.find (fact_index_of ctxt) spec  | 
| 21569 | 882  | 
|> map (fn (name, ths) =>  | 
883  | 
if valid_thms ctxt (name, ths) then (name, ths)  | 
|
884  | 
else (NameSpace.hidden (if name = "" then "unnamed" else name), ths));  | 
|
| 
16031
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
885  | 
|
| 
 
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
 
wenzelm 
parents: 
15979 
diff
changeset
 | 
886  | 
|
| 
13425
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
887  | 
(* name space operations *)  | 
| 
12309
 
03e9287be350
name space for local thms (export cond_extern, qualified);
 
wenzelm 
parents: 
12291 
diff
changeset
 | 
888  | 
|
| 22352 | 889  | 
val add_path = map_naming o NameSpace.add_path;  | 
| 
19062
 
0fd52e819c24
replaced qualified_force_prefix to sticky_prefix;
 
wenzelm 
parents: 
19033 
diff
changeset
 | 
890  | 
val no_base_names = map_naming NameSpace.no_base_names;  | 
| 16147 | 891  | 
val qualified_names = map_naming NameSpace.qualified_names;  | 
| 
19062
 
0fd52e819c24
replaced qualified_force_prefix to sticky_prefix;
 
wenzelm 
parents: 
19033 
diff
changeset
 | 
892  | 
val sticky_prefix = map_naming o NameSpace.sticky_prefix;  | 
| 
 
0fd52e819c24
replaced qualified_force_prefix to sticky_prefix;
 
wenzelm 
parents: 
19033 
diff
changeset
 | 
893  | 
val restore_naming = map_naming o K o naming_of;  | 
| 21728 | 894  | 
val reset_naming = map_naming (K local_naming);  | 
| 
12309
 
03e9287be350
name space for local thms (export cond_extern, qualified);
 
wenzelm 
parents: 
12291 
diff
changeset
 | 
895  | 
|
| 19001 | 896  | 
fun hide_thms fully names = map_thms (fn ((space, tab), index) =>  | 
897  | 
((fold (NameSpace.hide fully) names space, tab), index));  | 
|
| 
13425
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
898  | 
|
| 
12309
 
03e9287be350
name space for local thms (export cond_extern, qualified);
 
wenzelm 
parents: 
12291 
diff
changeset
 | 
899  | 
|
| 
17360
 
fa1f262dbc4e
added add_view, export_view (supercedes adhoc view arguments);
 
wenzelm 
parents: 
17221 
diff
changeset
 | 
900  | 
(* put_thms *)  | 
| 5819 | 901  | 
|
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
902  | 
fun update_thms _ ("", NONE) ctxt = ctxt
 | 
| 21648 | 903  | 
  | update_thms opts ("", SOME ths) ctxt = ctxt |> map_thms (fn (facts, index) =>
 | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
904  | 
let  | 
| 21648 | 905  | 
        val index' = uncurry FactIndex.add_local opts ("", ths) index;
 | 
| 19001 | 906  | 
in (facts, index') end)  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
907  | 
| update_thms _ (bname, NONE) ctxt = ctxt |> map_thms (fn ((space, tab), index) =>  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
908  | 
let  | 
| 19387 | 909  | 
val name = full_name ctxt bname;  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
910  | 
val tab' = Symtab.delete_safe name tab;  | 
| 19001 | 911  | 
in ((space, tab'), index) end)  | 
| 21648 | 912  | 
| update_thms opts (bname, SOME ths) ctxt = ctxt |> map_thms (fn ((space, tab), index) =>  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
913  | 
let  | 
| 19387 | 914  | 
val name = full_name ctxt bname;  | 
| 19001 | 915  | 
val space' = NameSpace.declare (naming_of ctxt) name space;  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
916  | 
val tab' = Symtab.update (name, ths) tab;  | 
| 21648 | 917  | 
val index' = uncurry FactIndex.add_local opts (name, ths) index;  | 
| 19001 | 918  | 
in ((space', tab'), index') end);  | 
| 5819 | 919  | 
|
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
920  | 
fun put_thms thms ctxt =  | 
| 21648 | 921  | 
ctxt |> map_naming (K local_naming) |> update_thms (true, false) thms |> restore_naming ctxt;  | 
| 
19079
 
9a7678a0736d
added put_thms_internal: local_naming, no fact index;
 
wenzelm 
parents: 
19062 
diff
changeset
 | 
922  | 
|
| 7606 | 923  | 
|
| 14564 | 924  | 
(* note_thmss *)  | 
| 5819 | 925  | 
|
| 12711 | 926  | 
local  | 
| 16147 | 927  | 
|
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
928  | 
fun gen_note_thmss get k = fold_map (fn ((bname, more_attrs), raw_facts) => fn ctxt =>  | 
| 5819 | 929  | 
let  | 
| 
21643
 
bdf3e74727df
note_thmss: added kind tag and non-official name;
 
wenzelm 
parents: 
21622 
diff
changeset
 | 
930  | 
val name = full_name ctxt bname;  | 
| 
 
bdf3e74727df
note_thmss: added kind tag and non-official name;
 
wenzelm 
parents: 
21622 
diff
changeset
 | 
931  | 
val facts = PureThy.name_thmss false name (map (apfst (get ctxt)) raw_facts);  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
932  | 
fun app (th, attrs) x =  | 
| 
21643
 
bdf3e74727df
note_thmss: added kind tag and non-official name;
 
wenzelm 
parents: 
21622 
diff
changeset
 | 
933  | 
swap (foldl_map (Thm.proof_attributes (attrs @ more_attrs @ [PureThy.kind k])) (x, th));  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
934  | 
val (res, ctxt') = fold_map app facts ctxt;  | 
| 
21643
 
bdf3e74727df
note_thmss: added kind tag and non-official name;
 
wenzelm 
parents: 
21622 
diff
changeset
 | 
935  | 
val thms = PureThy.name_thms false false name (flat res);  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
936  | 
    val Mode {stmt, ...} = get_mode ctxt;
 | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
937  | 
in ((bname, thms), ctxt' |> update_thms (stmt, true) (bname, SOME thms)) end);  | 
| 12711 | 938  | 
|
939  | 
in  | 
|
940  | 
||
| 21622 | 941  | 
fun note_thmss k = gen_note_thmss get_thms k;  | 
942  | 
fun note_thmss_i k = gen_note_thmss (K I) k;  | 
|
| 15696 | 943  | 
|
| 12711 | 944  | 
end;  | 
| 9196 | 945  | 
|
| 5819 | 946  | 
|
947  | 
||
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
948  | 
(** parameters **)  | 
| 
17360
 
fa1f262dbc4e
added add_view, export_view (supercedes adhoc view arguments);
 
wenzelm 
parents: 
17221 
diff
changeset
 | 
949  | 
|
| 8096 | 950  | 
(* variables *)  | 
951  | 
||
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
952  | 
fun declare_var (x, opt_T, mx) ctxt =  | 
| 22701 | 953  | 
let val T = (case opt_T of SOME T => T | NONE => Syntax.mixfixT mx)  | 
| 20163 | 954  | 
in ((x, T, mx), ctxt |> Variable.declare_constraints (Free (x, T))) end;  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
955  | 
|
| 10381 | 956  | 
local  | 
957  | 
||
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
958  | 
fun prep_vars prep_typ internal legacy =  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
959  | 
fold_map (fn (raw_x, raw_T, raw_mx) => fn ctxt =>  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
960  | 
let  | 
| 19371 | 961  | 
val (x, mx) = Syntax.const_mixfix raw_x raw_mx;  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
962  | 
val _ =  | 
| 18678 | 963  | 
if not legacy andalso not (Syntax.is_identifier (no_skolem internal x)) then  | 
964  | 
          error ("Illegal variable name: " ^ quote x)
 | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
965  | 
else ();  | 
| 12504 | 966  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
967  | 
fun cond_tvars T =  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
968  | 
if internal then T  | 
| 18678 | 969  | 
else Type.no_tvars T handle TYPE (msg, _, _) => error msg;  | 
| 24277 | 970  | 
val opt_T = Option.map (cond_tvars o cert_typ ctxt o prep_typ ctxt) raw_T;  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
971  | 
val var = (x, opt_T, mx);  | 
| 19001 | 972  | 
in (var, ctxt |> declare_var var |> #2) end);  | 
| 8096 | 973  | 
|
| 10381 | 974  | 
in  | 
975  | 
||
| 24277 | 976  | 
val read_vars = prep_vars Syntax.parse_typ false false;  | 
977  | 
val cert_vars = prep_vars (K I) true false;  | 
|
978  | 
val read_vars_legacy = prep_vars Syntax.parse_typ true true;  | 
|
979  | 
val cert_vars_legacy = prep_vars (K I) true true;  | 
|
| 8096 | 980  | 
|
| 10381 | 981  | 
end;  | 
982  | 
||
| 8096 | 983  | 
|
| 19681 | 984  | 
(* authentic constants *)  | 
| 19663 | 985  | 
|
| 
25159
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
986  | 
local  | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
987  | 
|
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
988  | 
fun const_ast_tr intern ctxt [Syntax.Variable c] =  | 
| 19681 | 989  | 
let  | 
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
990  | 
val Const (c', _) = read_const_proper ctxt c;  | 
| 
25159
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
991  | 
val d = if intern then const_syntax_name ctxt c' else c;  | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
992  | 
in Syntax.Constant d end  | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
993  | 
  | const_ast_tr _ _ asts = raise Syntax.AST ("const_ast_tr", asts);
 | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
994  | 
|
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
995  | 
in  | 
| 19681 | 996  | 
|
997  | 
val _ = Context.add_setup  | 
|
998  | 
(Sign.add_syntax  | 
|
999  | 
   [("_context_const", "id => 'a", Delimfix "CONST _"),
 | 
|
| 
25159
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
1000  | 
    ("_context_const", "longid => 'a", Delimfix "CONST _"),
 | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
1001  | 
    ("_context_xconst", "id => 'a", Delimfix "XCONST _"),
 | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
1002  | 
    ("_context_xconst", "longid => 'a", Delimfix "XCONST _")] #>
 | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
1003  | 
Sign.add_advanced_trfuns  | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
1004  | 
    ([("_context_const", const_ast_tr true), ("_context_xconst", const_ast_tr false)], [], [], []));
 | 
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
1005  | 
|
| 
 
1822da5446bc
added XCONST syntax (keeps original spelling of const);
 
wenzelm 
parents: 
25133 
diff
changeset
 | 
1006  | 
end;  | 
| 19681 | 1007  | 
|
| 19663 | 1008  | 
|
| 21744 | 1009  | 
(* notation *)  | 
1010  | 
||
| 24949 | 1011  | 
local  | 
1012  | 
||
1013  | 
fun const_syntax ctxt (Free (x, T), mx) = SOME (true, (x, T, mx))  | 
|
1014  | 
| const_syntax ctxt (Const (c, _), mx) =  | 
|
1015  | 
Option.map (pair false) (try (Consts.syntax (consts_of ctxt)) (c, mx))  | 
|
1016  | 
| const_syntax _ _ = NONE;  | 
|
1017  | 
||
1018  | 
in  | 
|
| 21744 | 1019  | 
|
| 24949 | 1020  | 
fun notation add mode args ctxt =  | 
1021  | 
ctxt |> map_syntax  | 
|
1022  | 
(LocalSyntax.update_modesyntax (theory_of ctxt) add mode (map_filter (const_syntax ctxt) args));  | 
|
1023  | 
||
1024  | 
fun target_notation add mode args phi =  | 
|
| 24982 | 1025  | 
let val args' = filter (fn (t, _) => Type.similar_types (t, Morphism.term phi t)) args;  | 
| 24949 | 1026  | 
in Context.mapping (Sign.notation add mode args') (notation add mode args') end;  | 
1027  | 
||
1028  | 
end;  | 
|
| 21744 | 1029  | 
|
1030  | 
||
| 24767 | 1031  | 
(* local constants *)  | 
1032  | 
||
1033  | 
fun add_const_constraint (c, opt_T) ctxt =  | 
|
1034  | 
let  | 
|
1035  | 
fun prepT raw_T =  | 
|
1036  | 
let val T = cert_typ ctxt raw_T  | 
|
1037  | 
in cert_term ctxt (Const (c, T)); T end;  | 
|
| 25039 | 1038  | 
in ctxt |> (map_consts o apfst) (Consts.constrain (c, Option.map prepT opt_T)) end;  | 
| 19001 | 1039  | 
|
| 24778 | 1040  | 
fun add_abbrev mode tags (c, raw_t) ctxt =  | 
| 19001 | 1041  | 
let  | 
| 
24675
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
1042  | 
val t0 = cert_term (ctxt |> set_mode mode_abbrev) raw_t  | 
| 21681 | 1043  | 
      handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote c);
 | 
| 
20008
 
8d9d770e1f06
add_abbrevs/polymorphic: Variable.exportT_terms avoids over-generalization;
 
wenzelm 
parents: 
19916 
diff
changeset
 | 
1044  | 
val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0];  | 
| 
21807
 
a59f083632a7
add_abbrev: removed Assumption.add_assms (danger of inconsistent naming);
 
wenzelm 
parents: 
21803 
diff
changeset
 | 
1045  | 
val ((lhs, rhs), consts') = consts_of ctxt  | 
| 24961 | 1046  | 
|> Consts.abbreviate (Syntax.pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode tags (c, t);  | 
| 19001 | 1047  | 
in  | 
1048  | 
ctxt  | 
|
| 25039 | 1049  | 
|> (map_consts o apfst) (K consts')  | 
| 21803 | 1050  | 
|> Variable.declare_term rhs  | 
1051  | 
|> pair (lhs, rhs)  | 
|
| 21704 | 1052  | 
end;  | 
| 19001 | 1053  | 
|
| 25052 | 1054  | 
fun revert_abbrev mode c = (map_consts o apfst) (Consts.revert_abbrev mode c);  | 
1055  | 
||
| 19001 | 1056  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1057  | 
(* fixes *)  | 
| 5819 | 1058  | 
|
| 8096 | 1059  | 
local  | 
1060  | 
||
| 19001 | 1061  | 
fun prep_mixfix (x, T, mx) =  | 
| 19019 | 1062  | 
if mx <> NoSyn andalso mx <> Structure andalso  | 
| 
20086
 
94ca946fb689
adapted to more efficient Name/Variable implementation;
 
wenzelm 
parents: 
20049 
diff
changeset
 | 
1063  | 
(can Name.dest_internal x orelse can Name.dest_skolem x) then  | 
| 19001 | 1064  | 
    error ("Illegal mixfix syntax for internal/skolem constant " ^ quote x)
 | 
1065  | 
else (true, (x, T, mx));  | 
|
1066  | 
||
| 18844 | 1067  | 
fun gen_fixes prep raw_vars ctxt =  | 
| 8096 | 1068  | 
let  | 
| 22712 | 1069  | 
val (vars, _) = prep raw_vars ctxt;  | 
1070  | 
val (xs', ctxt') = Variable.add_fixes (map #1 vars) ctxt;  | 
|
| 8096 | 1071  | 
in  | 
| 22712 | 1072  | 
ctxt'  | 
| 19001 | 1073  | 
|> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars)  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
1074  | 
|-> (map_syntax o LocalSyntax.add_syntax (theory_of ctxt) o map prep_mixfix)  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1075  | 
|> pair xs'  | 
| 8096 | 1076  | 
end;  | 
| 5819 | 1077  | 
|
| 8096 | 1078  | 
in  | 
| 7679 | 1079  | 
|
| 18844 | 1080  | 
val add_fixes = gen_fixes read_vars;  | 
1081  | 
val add_fixes_i = gen_fixes cert_vars;  | 
|
1082  | 
val add_fixes_legacy = gen_fixes cert_vars_legacy;  | 
|
| 8096 | 1083  | 
|
1084  | 
end;  | 
|
| 5819 | 1085  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1086  | 
|
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1087  | 
(* fixes vs. frees *)  | 
| 12016 | 1088  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1089  | 
fun auto_fixes (arg as (ctxt, (propss, x))) =  | 
| 
21370
 
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
1090  | 
((fold o fold) Variable.auto_fixes propss ctxt, (propss, x));  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1091  | 
|
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1092  | 
fun bind_fixes xs ctxt =  | 
| 
9291
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1093  | 
let  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1094  | 
val (_, ctxt') = ctxt |> add_fixes_i (map (fn x => (x, NONE, NoSyn)) xs);  | 
| 
9291
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1095  | 
fun bind (t as Free (x, T)) =  | 
| 18340 | 1096  | 
if member (op =) xs x then  | 
| 15531 | 1097  | 
(case lookup_skolem ctxt' x of SOME x' => Free (x', T) | NONE => t)  | 
| 
9291
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1098  | 
else t  | 
| 
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1099  | 
| bind (t $ u) = bind t $ bind u  | 
| 
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1100  | 
| bind (Abs (x, T, t)) = Abs (x, T, bind t)  | 
| 
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1101  | 
| bind a = a;  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1102  | 
in (bind, ctxt') end;  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1103  | 
|
| 
9291
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1104  | 
|
| 
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1105  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1106  | 
(** assumptions **)  | 
| 18187 | 1107  | 
|
| 20209 | 1108  | 
local  | 
1109  | 
||
1110  | 
fun gen_assms prepp exp args ctxt =  | 
|
1111  | 
let  | 
|
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1112  | 
val cert = Thm.cterm_of (theory_of ctxt);  | 
| 20209 | 1113  | 
val (propss, ctxt1) = swap (prepp (ctxt, map snd args));  | 
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1114  | 
val _ = Variable.warn_extra_tfrees ctxt ctxt1;  | 
| 
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1115  | 
val (premss, ctxt2) = fold_burrow (Assumption.add_assms exp o map cert) propss ctxt1;  | 
| 
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1116  | 
in  | 
| 
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1117  | 
ctxt2  | 
| 
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1118  | 
|> auto_bind_facts (flat propss)  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1119  | 
|> put_thms (AutoBind.premsN, SOME (Assumption.prems_of ctxt2))  | 
| 
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1120  | 
|> note_thmss_i Thm.assumptionK (map fst args ~~ map (map (fn th => ([th], []))) premss)  | 
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1121  | 
end;  | 
| 20209 | 1122  | 
|
1123  | 
in  | 
|
1124  | 
||
1125  | 
val add_assms = gen_assms (apsnd #1 o bind_propp);  | 
|
1126  | 
val add_assms_i = gen_assms (apsnd #1 o bind_propp_i);  | 
|
1127  | 
||
1128  | 
end;  | 
|
1129  | 
||
1130  | 
||
| 5819 | 1131  | 
|
| 
8373
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1132  | 
(** cases **)  | 
| 
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1133  | 
|
| 16147 | 1134  | 
local  | 
1135  | 
||
| 16668 | 1136  | 
fun rem_case name = remove (fn (x: string, (y, _)) => x = y) name;  | 
| 16147 | 1137  | 
|
| 18476 | 1138  | 
fun add_case _ ("", _) cases = cases
 | 
1139  | 
| add_case _ (name, NONE) cases = rem_case name cases  | 
|
1140  | 
| add_case is_proper (name, SOME c) cases = (name, (c, is_proper)) :: rem_case name cases;  | 
|
| 16147 | 1141  | 
|
| 18678 | 1142  | 
fun prep_case name fxs c =  | 
| 18609 | 1143  | 
let  | 
1144  | 
fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys  | 
|
1145  | 
| replace [] ys = ys  | 
|
| 18678 | 1146  | 
      | replace (_ :: _) [] = error ("Too many parameters for case " ^ quote name);
 | 
| 18609 | 1147  | 
    val RuleCases.Case {fixes, assumes, binds, cases} = c;
 | 
1148  | 
val fixes' = replace fxs fixes;  | 
|
1149  | 
val binds' = map drop_schematic binds;  | 
|
1150  | 
in  | 
|
1151  | 
if null (fold (Term.add_tvarsT o snd) fixes []) andalso  | 
|
1152  | 
null (fold (fold Term.add_vars o snd) assumes []) then  | 
|
1153  | 
        RuleCases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases}
 | 
|
| 18678 | 1154  | 
    else error ("Illegal schematic variable(s) in case " ^ quote name)
 | 
| 18609 | 1155  | 
end;  | 
1156  | 
||
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1157  | 
fun fix (x, T) ctxt =  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1158  | 
let  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1159  | 
val (bind, ctxt') = bind_fixes [x] ctxt;  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1160  | 
val t = bind (Free (x, T));  | 
| 20163 | 1161  | 
in (t, ctxt' |> Variable.declare_constraints t) end;  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1162  | 
|
| 16147 | 1163  | 
in  | 
1164  | 
||
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1165  | 
fun add_cases is_proper = map_cases o fold (add_case is_proper);  | 
| 18609 | 1166  | 
|
1167  | 
fun case_result c ctxt =  | 
|
1168  | 
let  | 
|
1169  | 
    val RuleCases.Case {fixes, ...} = c;
 | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1170  | 
val (ts, ctxt') = ctxt |> fold_map fix fixes;  | 
| 
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1171  | 
    val RuleCases.Case {assumes, binds, cases, ...} = RuleCases.apply ts c;
 | 
| 18609 | 1172  | 
in  | 
1173  | 
ctxt'  | 
|
| 
18699
 
f3bfe81b6e58
case_result: drop_schematic, i.e. be permissive about illegal binds;
 
wenzelm 
parents: 
18678 
diff
changeset
 | 
1174  | 
|> add_binds_i (map drop_schematic binds)  | 
| 18609 | 1175  | 
|> add_cases true (map (apsnd SOME) cases)  | 
1176  | 
|> pair (assumes, (binds, cases))  | 
|
1177  | 
end;  | 
|
1178  | 
||
1179  | 
val apply_case = apfst fst oo case_result;  | 
|
1180  | 
||
| 16540 | 1181  | 
fun get_case ctxt name xs =  | 
| 17184 | 1182  | 
(case AList.lookup (op =) (cases_of ctxt) name of  | 
| 18678 | 1183  | 
    NONE => error ("Unknown case: " ^ quote name)
 | 
1184  | 
| SOME (c, _) => prep_case name xs c);  | 
|
| 
8373
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1185  | 
|
| 16147 | 1186  | 
end;  | 
| 
8373
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1187  | 
|
| 
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1188  | 
|
| 
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1189  | 
|
| 10810 | 1190  | 
(** print context information **)  | 
1191  | 
||
| 20310 | 1192  | 
val debug = ref false;  | 
1193  | 
||
| 10810 | 1194  | 
val verbose = ref false;  | 
1195  | 
fun verb f x = if ! verbose then f (x ()) else [];  | 
|
1196  | 
||
1197  | 
fun setmp_verbose f x = Library.setmp verbose true f x;  | 
|
1198  | 
||
1199  | 
||
| 
12072
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1200  | 
(* local syntax *)  | 
| 
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1201  | 
|
| 12093 | 1202  | 
val print_syntax = Syntax.print_syntax o syn_of;  | 
| 
12072
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1203  | 
|
| 
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1204  | 
|
| 21728 | 1205  | 
(* abbreviations *)  | 
| 18971 | 1206  | 
|
| 21728 | 1207  | 
fun pretty_abbrevs show_globals ctxt =  | 
| 18971 | 1208  | 
let  | 
| 25052 | 1209  | 
val ((space, consts), (_, globals)) =  | 
| 
19033
 
24e251657e56
consts: maintain thy version for efficient transfer;
 
wenzelm 
parents: 
19019 
diff
changeset
 | 
1210  | 
pairself (#constants o Consts.dest) (#consts (rep_context ctxt));  | 
| 21803 | 1211  | 
fun add_abbr (_, (_, NONE)) = I  | 
1212  | 
| add_abbr (c, (T, SOME (t, _))) =  | 
|
| 21728 | 1213  | 
if not show_globals andalso Symtab.defined globals c then I  | 
1214  | 
else cons (c, Logic.mk_equals (Const (c, T), t));  | 
|
| 21803 | 1215  | 
val abbrevs = NameSpace.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));  | 
| 18971 | 1216  | 
in  | 
1217  | 
if null abbrevs andalso not (! verbose) then []  | 
|
| 21728 | 1218  | 
else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]  | 
| 18971 | 1219  | 
end;  | 
1220  | 
||
| 21728 | 1221  | 
val print_abbrevs = Pretty.writeln o Pretty.chunks o pretty_abbrevs true;  | 
1222  | 
||
| 18971 | 1223  | 
|
| 10810 | 1224  | 
(* term bindings *)  | 
1225  | 
||
| 16540 | 1226  | 
fun pretty_binds ctxt =  | 
| 10810 | 1227  | 
let  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
1228  | 
val binds = Variable.binds_of ctxt;  | 
| 21728 | 1229  | 
fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t));  | 
| 10810 | 1230  | 
in  | 
| 
15758
 
07e382399a96
binds/thms: do not store options, but delete from table;
 
wenzelm 
parents: 
15750 
diff
changeset
 | 
1231  | 
if Vartab.is_empty binds andalso not (! verbose) then []  | 
| 
 
07e382399a96
binds/thms: do not store options, but delete from table;
 
wenzelm 
parents: 
15750 
diff
changeset
 | 
1232  | 
else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))]  | 
| 10810 | 1233  | 
end;  | 
1234  | 
||
1235  | 
val print_binds = Pretty.writeln o Pretty.chunks o pretty_binds;  | 
|
1236  | 
||
1237  | 
||
1238  | 
(* local theorems *)  | 
|
1239  | 
||
| 16540 | 1240  | 
fun pretty_lthms ctxt =  | 
| 20012 | 1241  | 
let  | 
1242  | 
val props = FactIndex.props (fact_index_of ctxt);  | 
|
1243  | 
val facts =  | 
|
1244  | 
      (if null props then I else cons ("unnamed", props))
 | 
|
1245  | 
(NameSpace.extern_table (#1 (thms_of ctxt)));  | 
|
1246  | 
in  | 
|
1247  | 
if null facts andalso not (! verbose) then []  | 
|
1248  | 
else [Pretty.big_list "facts:" (map (pretty_fact ctxt) facts)]  | 
|
1249  | 
end;  | 
|
| 10810 | 1250  | 
|
| 12057 | 1251  | 
val print_lthms = Pretty.writeln o Pretty.chunks o pretty_lthms;  | 
| 10810 | 1252  | 
|
1253  | 
||
1254  | 
(* local contexts *)  | 
|
1255  | 
||
| 16540 | 1256  | 
fun pretty_cases ctxt =  | 
| 10810 | 1257  | 
let  | 
| 24922 | 1258  | 
val prt_term = Syntax.pretty_term ctxt;  | 
| 12057 | 1259  | 
|
| 10810 | 1260  | 
fun prt_let (xi, t) = Pretty.block  | 
| 10818 | 1261  | 
[Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1,  | 
| 10810 | 1262  | 
Pretty.quote (prt_term t)];  | 
1263  | 
||
| 
13425
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
1264  | 
fun prt_asm (a, ts) = Pretty.block (Pretty.breaks  | 
| 
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
1265  | 
((if a = "" then [] else [Pretty.str (a ^ ":")]) @ map (Pretty.quote o prt_term) ts));  | 
| 
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
1266  | 
|
| 10810 | 1267  | 
fun prt_sect _ _ _ [] = []  | 
1268  | 
| prt_sect s sep prt xs = [Pretty.block (Pretty.breaks (Pretty.str s ::  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19422 
diff
changeset
 | 
1269  | 
flat (Library.separate sep (map (Library.single o prt) xs))))];  | 
| 10810 | 1270  | 
|
| 18609 | 1271  | 
fun prt_case (name, (fixes, (asms, (lets, cs)))) = Pretty.block (Pretty.fbreaks  | 
| 10810 | 1272  | 
(Pretty.str (name ^ ":") ::  | 
| 11915 | 1273  | 
prt_sect "fix" [] (Pretty.str o fst) fixes @  | 
| 10810 | 1274  | 
prt_sect "let" [Pretty.str "and"] prt_let  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19422 
diff
changeset
 | 
1275  | 
(map_filter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @  | 
| 
13425
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
1276  | 
(if forall (null o #2) asms then []  | 
| 18609 | 1277  | 
else prt_sect "assume" [Pretty.str "and"] prt_asm asms) @  | 
1278  | 
prt_sect "subcases:" [] (Pretty.str o fst) cs));  | 
|
| 16540 | 1279  | 
|
| 18476 | 1280  | 
fun add_case (_, (_, false)) = I  | 
| 18609 | 1281  | 
      | add_case (name, (c as RuleCases.Case {fixes, ...}, true)) =
 | 
1282  | 
cons (name, (fixes, #1 (case_result c ctxt)));  | 
|
| 18476 | 1283  | 
val cases = fold add_case (cases_of ctxt) [];  | 
| 10810 | 1284  | 
in  | 
1285  | 
if null cases andalso not (! verbose) then []  | 
|
| 18476 | 1286  | 
else [Pretty.big_list "cases:" (map prt_case cases)]  | 
| 10810 | 1287  | 
end;  | 
1288  | 
||
1289  | 
val print_cases = Pretty.writeln o Pretty.chunks o pretty_cases;  | 
|
1290  | 
||
1291  | 
||
| 12057 | 1292  | 
(* core context *)  | 
| 10810 | 1293  | 
|
| 20367 | 1294  | 
val prems_limit = ref ~1;  | 
| 10810 | 1295  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1296  | 
fun pretty_ctxt ctxt =  | 
| 20310 | 1297  | 
if ! prems_limit < 0 andalso not (! debug) then []  | 
1298  | 
else  | 
|
1299  | 
let  | 
|
| 24922 | 1300  | 
val prt_term = Syntax.pretty_term ctxt;  | 
| 12057 | 1301  | 
|
| 20310 | 1302  | 
(*structures*)  | 
1303  | 
val structs = LocalSyntax.structs_of (syntax_of ctxt);  | 
|
1304  | 
val prt_structs = if null structs then []  | 
|
1305  | 
else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 ::  | 
|
1306  | 
Pretty.commas (map Pretty.str structs))];  | 
|
| 12093 | 1307  | 
|
| 20310 | 1308  | 
(*fixes*)  | 
1309  | 
fun prt_fix (x, x') =  | 
|
1310  | 
if x = x' then Pretty.str x  | 
|
1311  | 
else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')];  | 
|
1312  | 
val fixes =  | 
|
1313  | 
rev (filter_out ((can Name.dest_internal orf member (op =) structs) o #1)  | 
|
1314  | 
(Variable.fixes_of ctxt));  | 
|
1315  | 
val prt_fixes = if null fixes then []  | 
|
1316  | 
else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 ::  | 
|
1317  | 
Pretty.commas (map prt_fix fixes))];  | 
|
| 12057 | 1318  | 
|
| 20310 | 1319  | 
(*prems*)  | 
1320  | 
val prems = Assumption.prems_of ctxt;  | 
|
1321  | 
val len = length prems;  | 
|
| 20367 | 1322  | 
val suppressed = len - ! prems_limit;  | 
| 20310 | 1323  | 
val prt_prems = if null prems then []  | 
| 20367 | 1324  | 
else [Pretty.big_list "prems:" ((if suppressed <= 0 then [] else [Pretty.str "..."]) @  | 
1325  | 
map (pretty_thm ctxt) (Library.drop (suppressed, prems)))];  | 
|
| 20310 | 1326  | 
in prt_structs @ prt_fixes @ prt_prems end;  | 
| 10810 | 1327  | 
|
1328  | 
||
1329  | 
(* main context *)  | 
|
1330  | 
||
| 16540 | 1331  | 
fun pretty_context ctxt =  | 
| 10810 | 1332  | 
let  | 
| 24922 | 1333  | 
val prt_term = Syntax.pretty_term ctxt;  | 
1334  | 
val prt_typ = Syntax.pretty_typ ctxt;  | 
|
1335  | 
val prt_sort = Syntax.pretty_sort ctxt;  | 
|
| 10810 | 1336  | 
|
1337  | 
(*theory*)  | 
|
| 12057 | 1338  | 
val pretty_thy = Pretty.block  | 
| 17384 | 1339  | 
[Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)];  | 
| 10810 | 1340  | 
|
1341  | 
(*defaults*)  | 
|
1342  | 
fun prt_atom prt prtT (x, X) = Pretty.block  | 
|
1343  | 
[prt x, Pretty.str " ::", Pretty.brk 1, prtT X];  | 
|
1344  | 
||
1345  | 
fun prt_var (x, ~1) = prt_term (Syntax.free x)  | 
|
1346  | 
| prt_var xi = prt_term (Syntax.var xi);  | 
|
1347  | 
||
1348  | 
fun prt_varT (x, ~1) = prt_typ (TFree (x, []))  | 
|
1349  | 
| prt_varT xi = prt_typ (TVar (xi, []));  | 
|
1350  | 
||
1351  | 
val prt_defT = prt_atom prt_var prt_typ;  | 
|
1352  | 
val prt_defS = prt_atom prt_varT prt_sort;  | 
|
| 16540 | 1353  | 
|
| 20163 | 1354  | 
val (types, sorts) = Variable.constraints_of ctxt;  | 
| 10810 | 1355  | 
in  | 
| 18609 | 1356  | 
verb single (K pretty_thy) @  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1357  | 
pretty_ctxt ctxt @  | 
| 21728 | 1358  | 
verb (pretty_abbrevs false) (K ctxt) @  | 
| 10810 | 1359  | 
verb pretty_binds (K ctxt) @  | 
| 12057 | 1360  | 
verb pretty_lthms (K ctxt) @  | 
| 10810 | 1361  | 
verb pretty_cases (K ctxt) @  | 
| 18609 | 1362  | 
verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @  | 
| 20163 | 1363  | 
verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts)))  | 
| 10810 | 1364  | 
end;  | 
1365  | 
||
| 5819 | 1366  | 
end;  |