| author | wenzelm | 
| Thu, 18 Sep 2008 19:39:44 +0200 | |
| changeset 28290 | 4cc2b6046258 | 
| parent 28115 | cd0d170d4dc6 | 
| child 28311 | b86feb50ca58 | 
| permissions | -rw-r--r-- | 
| 20894 | 1  | 
(* Title: Pure/Isar/theory_target.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Makarius  | 
|
4  | 
||
| 25542 | 5  | 
Common theory/locale/class/instantiation/overloading targets.  | 
| 20894 | 6  | 
*)  | 
7  | 
||
8  | 
signature THEORY_TARGET =  | 
|
9  | 
sig  | 
|
| 25462 | 10  | 
  val peek: local_theory -> {target: string, is_locale: bool,
 | 
| 25864 | 11  | 
is_class: bool, instantiation: string list * (string * sort) list * sort,  | 
12  | 
overloading: (string * (string * typ) * bool) list}  | 
|
| 25269 | 13  | 
val init: string option -> theory -> local_theory  | 
| 
24935
 
a033971c63a0
removed LocalTheory.defs/target_morphism operations;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
14  | 
val begin: string -> Proof.context -> local_theory  | 
| 25291 | 15  | 
val context: xstring -> theory -> local_theory  | 
| 25864 | 16  | 
val instantiation: string list * (string * sort) list * sort -> theory -> local_theory  | 
17  | 
val overloading: (string * (string * typ) * bool) list -> theory -> local_theory  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
18  | 
val overloading_cmd: (string * string * bool) list -> theory -> local_theory  | 
| 20894 | 19  | 
end;  | 
20  | 
||
21  | 
structure TheoryTarget: THEORY_TARGET =  | 
|
22  | 
struct  | 
|
23  | 
||
| 21006 | 24  | 
(* context data *)  | 
25  | 
||
| 25462 | 26  | 
datatype target = Target of {target: string, is_locale: bool,
 | 
| 25864 | 27  | 
is_class: bool, instantiation: string list * (string * sort) list * sort,  | 
28  | 
overloading: (string * (string * typ) * bool) list};  | 
|
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
29  | 
|
| 25519 | 30  | 
fun make_target target is_locale is_class instantiation overloading =  | 
| 25462 | 31  | 
  Target {target = target, is_locale = is_locale,
 | 
| 25519 | 32  | 
is_class = is_class, instantiation = instantiation, overloading = overloading};  | 
| 25291 | 33  | 
|
| 25542 | 34  | 
val global_target = make_target "" false false ([], [], []) [];  | 
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
35  | 
|
| 21006 | 36  | 
structure Data = ProofDataFun  | 
37  | 
(  | 
|
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
38  | 
type T = target;  | 
| 25291 | 39  | 
fun init _ = global_target;  | 
| 21006 | 40  | 
);  | 
41  | 
||
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
42  | 
val peek = (fn Target args => args) o Data.get;  | 
| 21006 | 43  | 
|
44  | 
||
| 20894 | 45  | 
(* pretty *)  | 
46  | 
||
| 25607 | 47  | 
fun pretty_thy ctxt target is_locale is_class =  | 
| 20894 | 48  | 
let  | 
49  | 
val thy = ProofContext.theory_of ctxt;  | 
|
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
50  | 
val target_name = (if is_class then "class " else "locale ") ^ Locale.extern thy target;  | 
| 28094 | 51  | 
val fixes = map (fn (x, T) => (Name.binding x, SOME T, NoSyn))  | 
52  | 
(#1 (ProofContext.inferred_fixes ctxt));  | 
|
53  | 
val assumes = map (fn A => (Attrib.no_binding, [(Thm.term_of A, [])]))  | 
|
54  | 
(Assumption.assms_of ctxt);  | 
|
| 20894 | 55  | 
val elems =  | 
56  | 
(if null fixes then [] else [Element.Fixes fixes]) @  | 
|
57  | 
(if null assumes then [] else [Element.Assumes assumes]);  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
58  | 
in  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
59  | 
if target = "" then []  | 
| 25607 | 60  | 
else if null elems then [Pretty.str target_name]  | 
61  | 
else [Pretty.big_list (target_name ^ " =")  | 
|
62  | 
(map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)]  | 
|
| 20894 | 63  | 
end;  | 
64  | 
||
| 25607 | 65  | 
fun pretty (Target {target, is_locale, is_class, instantiation, overloading}) ctxt =
 | 
66  | 
Pretty.block [Pretty.str "theory", Pretty.brk 1,  | 
|
67  | 
Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::  | 
|
68  | 
(if not (null overloading) then [Overloading.pretty ctxt]  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
69  | 
else if not (null (#1 instantiation)) then [Class.pretty_instantiation ctxt]  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
70  | 
else pretty_thy ctxt target is_locale is_class);  | 
| 25607 | 71  | 
|
| 20894 | 72  | 
|
| 24838 | 73  | 
(* target declarations *)  | 
74  | 
||
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27315 
diff
changeset
 | 
75  | 
fun target_decl add (Target {target, is_class, ...}) d lthy =
 | 
| 24838 | 76  | 
let  | 
| 
24935
 
a033971c63a0
removed LocalTheory.defs/target_morphism operations;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
77  | 
val d' = Morphism.transform (LocalTheory.target_morphism lthy) d;  | 
| 24838 | 78  | 
val d0 = Morphism.form d';  | 
79  | 
in  | 
|
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
80  | 
if target = "" then  | 
| 24838 | 81  | 
lthy  | 
82  | 
|> LocalTheory.theory (Context.theory_map d0)  | 
|
83  | 
|> LocalTheory.target (Context.proof_map d0)  | 
|
84  | 
else  | 
|
85  | 
lthy  | 
|
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
86  | 
|> LocalTheory.target (add target d')  | 
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27315 
diff
changeset
 | 
87  | 
(*|> is_class ? LocalTheory.raw_theory (Class.declaration target d')*)  | 
| 24838 | 88  | 
end;  | 
89  | 
||
90  | 
val type_syntax = target_decl Locale.add_type_syntax;  | 
|
91  | 
val term_syntax = target_decl Locale.add_term_syntax;  | 
|
92  | 
val declaration = target_decl Locale.add_declaration;  | 
|
93  | 
||
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
94  | 
fun class_target (Target {target, ...}) f =
 | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
95  | 
LocalTheory.raw_theory f #>  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
96  | 
LocalTheory.target (Class.refresh_syntax target);  | 
| 24838 | 97  | 
|
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
98  | 
|
| 20894 | 99  | 
(* notes *)  | 
100  | 
||
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
101  | 
fun import_export_proof ctxt (name, raw_th) =  | 
| 
21594
 
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
 
wenzelm 
parents: 
21585 
diff
changeset
 | 
102  | 
let  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
103  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
104  | 
val thy_ctxt = ProofContext.init thy;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
105  | 
val certT = Thm.ctyp_of thy;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
106  | 
val cert = Thm.cterm_of thy;  | 
| 
21594
 
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
 
wenzelm 
parents: 
21585 
diff
changeset
 | 
107  | 
|
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
108  | 
(*export assumes/defines*)  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
109  | 
val th = Goal.norm_result raw_th;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
110  | 
val (defs, th') = LocalDefs.export ctxt thy_ctxt th;  | 
| 21708 | 111  | 
val concl_conv = MetaSimplifier.rewrite true defs (Thm.cprop_of th);  | 
112  | 
val assms = map (MetaSimplifier.rewrite_rule defs o Thm.assume) (Assumption.assms_of ctxt);  | 
|
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
113  | 
val nprems = Thm.nprems_of th' - Thm.nprems_of th;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
114  | 
|
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
115  | 
(*export fixes*)  | 
| 22692 | 116  | 
val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []);  | 
117  | 
val frees = map Free (Thm.fold_terms Term.add_frees th' []);  | 
|
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
118  | 
val (th'' :: vs) = (th' :: map (Drule.mk_term o cert) (map Logic.mk_type tfrees @ frees))  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
119  | 
|> Variable.export ctxt thy_ctxt  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
120  | 
|> Drule.zero_var_indexes_list;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
121  | 
|
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
122  | 
(*thm definition*)  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
123  | 
val result = PureThy.name_thm true true Position.none name th'';  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
124  | 
|
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
125  | 
(*import fixes*)  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
126  | 
val (tvars, vars) =  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
127  | 
chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs)  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
128  | 
|>> map Logic.dest_type;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
129  | 
|
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
130  | 
val instT = map_filter (fn (TVar v, T) => SOME (v, T) | _ => NONE) (tvars ~~ tfrees);  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
131  | 
val inst = filter (is_Var o fst) (vars ~~ frees);  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
132  | 
val cinstT = map (pairself certT o apfst TVar) instT;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
133  | 
val cinst = map (pairself (cert o Term.map_types (TermSubst.instantiateT instT))) inst;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
134  | 
val result' = Thm.instantiate (cinstT, cinst) result;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
135  | 
|
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
136  | 
(*import assumes/defines*)  | 
| 
21594
 
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
 
wenzelm 
parents: 
21585 
diff
changeset
 | 
137  | 
val assm_tac = FIRST' (map (fn assm => Tactic.compose_tac (false, assm, 0)) assms);  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
138  | 
val result'' =  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
139  | 
(case SINGLE (Seq.INTERVAL assm_tac 1 nprems) result' of  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
140  | 
        NONE => raise THM ("Failed to re-import result", 0, [result'])
 | 
| 21845 | 141  | 
| SOME res => LocalDefs.trans_props ctxt [res, Thm.symmetric concl_conv])  | 
| 21644 | 142  | 
|> Goal.norm_result  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
143  | 
|> PureThy.name_thm false false Position.none name;  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
144  | 
|
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
145  | 
in (result'', result) end;  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
146  | 
|
| 25684 | 147  | 
fun note_local kind facts ctxt =  | 
148  | 
ctxt  | 
|
149  | 
|> ProofContext.qualified_names  | 
|
150  | 
|> ProofContext.note_thmss_i kind facts  | 
|
151  | 
||> ProofContext.restore_naming ctxt;  | 
|
152  | 
||
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27315 
diff
changeset
 | 
153  | 
fun notes (Target {target, is_locale, is_class, ...}) kind facts lthy =
 | 
| 21585 | 154  | 
let  | 
| 
21594
 
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
 
wenzelm 
parents: 
21585 
diff
changeset
 | 
155  | 
val thy = ProofContext.theory_of lthy;  | 
| 
21615
 
1bd558879c44
notes: proper naming of thm proof, activated import_export_proof;
 
wenzelm 
parents: 
21613 
diff
changeset
 | 
156  | 
val full = LocalTheory.full_name lthy;  | 
| 21585 | 157  | 
|
| 
21594
 
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
 
wenzelm 
parents: 
21585 
diff
changeset
 | 
158  | 
val facts' = facts  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
159  | 
|> map (fn (a, bs) =>  | 
| 
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
160  | 
(a, PureThy.burrow_fact (PureThy.name_multi (full (Name.name_of (fst a)))) bs))  | 
| 
21615
 
1bd558879c44
notes: proper naming of thm proof, activated import_export_proof;
 
wenzelm 
parents: 
21613 
diff
changeset
 | 
161  | 
|> PureThy.map_facts (import_export_proof lthy);  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
162  | 
val local_facts = PureThy.map_facts #1 facts'  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
163  | 
|> Attrib.map_facts (Attrib.attribute_i thy);  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
164  | 
val target_facts = PureThy.map_facts #1 facts'  | 
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
165  | 
|> is_locale ? Element.facts_map (Element.morph_ctxt (LocalTheory.target_morphism lthy));  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
166  | 
val global_facts = PureThy.map_facts #2 facts'  | 
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
167  | 
|> Attrib.map_facts (if is_locale then K I else Attrib.attribute_i thy);  | 
| 
21594
 
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
 
wenzelm 
parents: 
21585 
diff
changeset
 | 
168  | 
in  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
169  | 
lthy |> LocalTheory.theory  | 
| 
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
170  | 
(Sign.qualified_names  | 
| 26132 | 171  | 
#> PureThy.note_thmss_grouped kind (LocalTheory.group_of lthy) global_facts #> snd  | 
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
172  | 
#> Sign.restore_naming thy)  | 
| 25684 | 173  | 
|> not is_locale ? LocalTheory.target (note_local kind global_facts #> snd)  | 
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
174  | 
|> is_locale ? LocalTheory.target (Locale.add_thmss target kind target_facts)  | 
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27315 
diff
changeset
 | 
175  | 
(*|> is_class ? LocalTheory.raw_theory (Class.note target kind target_facts #> snd)*)  | 
| 25684 | 176  | 
|> note_local kind local_facts  | 
| 20894 | 177  | 
end;  | 
178  | 
||
179  | 
||
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
180  | 
(* declare_const *)  | 
| 24939 | 181  | 
|
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
182  | 
fun fork_mixfix (Target {is_locale, is_class, ...}) mx =
 | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
183  | 
if not is_locale then (NoSyn, NoSyn, mx)  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
184  | 
else if not is_class then (NoSyn, mx, NoSyn)  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
185  | 
else (mx, NoSyn, NoSyn);  | 
| 25068 | 186  | 
|
| 26132 | 187  | 
fun locale_const (Target {target, is_class, ...}) (prmode as (mode, _)) tags ((c, mx), rhs) phi =
 | 
| 24939 | 188  | 
let  | 
189  | 
val c' = Morphism.name phi c;  | 
|
190  | 
val rhs' = Morphism.term phi rhs;  | 
|
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
191  | 
val name = Name.name_of c;  | 
| 
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
192  | 
val name' = Name.name_of c'  | 
| 
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
193  | 
val legacy_arg = (name', Term.close_schematic_term (Logic.legacy_varify rhs'));  | 
| 
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
194  | 
val arg = (name', Term.close_schematic_term rhs');  | 
| 
25372
 
a718f1ccaf1a
locale_const: suppress in class body as well (prevents qualified printing);
 
wenzelm 
parents: 
25320 
diff
changeset
 | 
195  | 
val similar_body = Type.similar_types (rhs, rhs');  | 
| 
25212
 
9dd61cb753ae
locale_const: in_class workaround prevents additional locale version of class consts;
 
wenzelm 
parents: 
25150 
diff
changeset
 | 
196  | 
(* FIXME workaround based on educated guess *)  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
197  | 
val class_global = name' = NameSpace.qualified (Class.class_prefix target) name;  | 
| 24939 | 198  | 
in  | 
| 
25372
 
a718f1ccaf1a
locale_const: suppress in class body as well (prevents qualified printing);
 
wenzelm 
parents: 
25320 
diff
changeset
 | 
199  | 
not (is_class andalso (similar_body orelse class_global)) ?  | 
| 
25212
 
9dd61cb753ae
locale_const: in_class workaround prevents additional locale version of class consts;
 
wenzelm 
parents: 
25150 
diff
changeset
 | 
200  | 
(Context.mapping_result  | 
| 26132 | 201  | 
(Sign.add_abbrev PrintMode.internal tags legacy_arg)  | 
202  | 
(ProofContext.add_abbrev PrintMode.internal tags arg)  | 
|
| 
25212
 
9dd61cb753ae
locale_const: in_class workaround prevents additional locale version of class consts;
 
wenzelm 
parents: 
25150 
diff
changeset
 | 
203  | 
#-> (fn (lhs' as Const (d, _), _) =>  | 
| 
25372
 
a718f1ccaf1a
locale_const: suppress in class body as well (prevents qualified printing);
 
wenzelm 
parents: 
25320 
diff
changeset
 | 
204  | 
similar_body ?  | 
| 
25212
 
9dd61cb753ae
locale_const: in_class workaround prevents additional locale version of class consts;
 
wenzelm 
parents: 
25150 
diff
changeset
 | 
205  | 
(Context.mapping (Sign.revert_abbrev mode d) (ProofContext.revert_abbrev mode d) #>  | 
| 
 
9dd61cb753ae
locale_const: in_class workaround prevents additional locale version of class consts;
 
wenzelm 
parents: 
25150 
diff
changeset
 | 
206  | 
Morphism.form (ProofContext.target_notation true prmode [(lhs', mx)]))))  | 
| 24939 | 207  | 
end;  | 
208  | 
||
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
209  | 
fun declare_const (ta as Target {target, is_locale, is_class, ...}) depends ((b, T), mx) lthy =
 | 
| 24939 | 210  | 
let  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
211  | 
val c = Name.name_of b;  | 
| 26132 | 212  | 
val tags = LocalTheory.group_position_of lthy;  | 
| 24939 | 213  | 
val xs = filter depends (#1 (ProofContext.inferred_fixes (LocalTheory.target_of lthy)));  | 
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
214  | 
val U = map #2 xs ---> T;  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
215  | 
val (mx1, mx2, mx3) = fork_mixfix ta mx;  | 
| 25519 | 216  | 
    fun syntax_error c = error ("Illegal mixfix syntax for overloaded constant " ^ quote c);
 | 
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
217  | 
val declare_const =  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
218  | 
(case Class.instantiation_param lthy c of  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
219  | 
SOME c' =>  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
220  | 
if mx3 <> NoSyn then syntax_error c'  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25542 
diff
changeset
 | 
221  | 
else LocalTheory.theory_result (AxClass.declare_overloaded (c', U))  | 
| 25485 | 222  | 
##> Class.confirm_declaration c  | 
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
223  | 
| NONE =>  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
224  | 
(case Overloading.operation lthy c of  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
225  | 
SOME (c', _) =>  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
226  | 
if mx3 <> NoSyn then syntax_error c'  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
227  | 
else LocalTheory.theory_result (Overloading.declare (c', U))  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
228  | 
##> Overloading.confirm c  | 
| 
28115
 
cd0d170d4dc6
discontinued local axioms -- too difficult to implement, too easy to produce nonsense;
 
wenzelm 
parents: 
28094 
diff
changeset
 | 
229  | 
| NONE => LocalTheory.theory_result (Sign.declare_const tags ((b, U), mx3))));  | 
| 25485 | 230  | 
val (const, lthy') = lthy |> declare_const;  | 
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
231  | 
val t = Term.list_comb (const, map Free xs);  | 
| 24939 | 232  | 
in  | 
233  | 
lthy'  | 
|
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
234  | 
|> is_locale ? term_syntax ta (locale_const ta Syntax.mode_default tags ((b, mx2), t))  | 
| 26132 | 235  | 
|> is_class ? class_target ta (Class.declare target tags ((c, mx1), t))  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
236  | 
|> LocalDefs.add_def ((b, NoSyn), t)  | 
| 24939 | 237  | 
end;  | 
238  | 
||
239  | 
||
240  | 
(* abbrev *)  | 
|
241  | 
||
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
242  | 
fun abbrev (ta as Target {target, is_locale, is_class, ...}) prmode ((b, mx), t) lthy =
 | 
| 24939 | 243  | 
let  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
244  | 
val c = Name.name_of b;  | 
| 26132 | 245  | 
val tags = LocalTheory.group_position_of lthy;  | 
| 25053 | 246  | 
val thy_ctxt = ProofContext.init (ProofContext.theory_of lthy);  | 
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
247  | 
val target_ctxt = LocalTheory.target_of lthy;  | 
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
248  | 
|
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
249  | 
val (mx1, mx2, mx3) = fork_mixfix ta mx;  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
250  | 
val t' = Assumption.export_term lthy target_ctxt t;  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
251  | 
val xs = map Free (rev (Variable.add_fixed target_ctxt t' []));  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
252  | 
val u = fold_rev lambda xs t';  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
253  | 
val global_rhs =  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
254  | 
singleton (Variable.export_terms (Variable.declare_term u target_ctxt) thy_ctxt) u;  | 
| 25121 | 255  | 
in  | 
256  | 
lthy |>  | 
|
257  | 
(if is_locale then  | 
|
| 26132 | 258  | 
LocalTheory.theory_result (Sign.add_abbrev PrintMode.internal tags (c, global_rhs))  | 
| 25121 | 259  | 
#-> (fn (lhs, _) =>  | 
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
260  | 
let val lhs' = Term.list_comb (Logic.unvarify lhs, xs) in  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
261  | 
term_syntax ta (locale_const ta prmode tags ((b, mx2), lhs')) #>  | 
| 26132 | 262  | 
is_class ? class_target ta (Class.abbrev target prmode tags ((c, mx1), t'))  | 
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
263  | 
end)  | 
| 
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
264  | 
else  | 
| 25121 | 265  | 
LocalTheory.theory  | 
| 26132 | 266  | 
(Sign.add_abbrev (#1 prmode) tags (c, global_rhs) #-> (fn (lhs, _) =>  | 
| 25121 | 267  | 
Sign.notation true prmode [(lhs, mx3)])))  | 
| 26132 | 268  | 
|> ProofContext.add_abbrev PrintMode.internal tags (c, t) |> snd  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
269  | 
|> LocalDefs.fixed_abbrev ((b, NoSyn), t)  | 
| 24939 | 270  | 
end;  | 
271  | 
||
272  | 
||
| 25022 | 273  | 
(* define *)  | 
| 24939 | 274  | 
|
| 25485 | 275  | 
fun define (ta as Target {target, is_locale, is_class, ...})
 | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
276  | 
kind ((b, mx), ((name, atts), rhs)) lthy =  | 
| 24939 | 277  | 
let  | 
| 24987 | 278  | 
val thy = ProofContext.theory_of lthy;  | 
| 24939 | 279  | 
val thy_ctxt = ProofContext.init thy;  | 
280  | 
||
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
281  | 
val c = Name.name_of b;  | 
| 
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
282  | 
val name' = Name.map_name (Thm.def_name_optional c) name;  | 
| 24987 | 283  | 
val (rhs', rhs_conv) =  | 
284  | 
LocalDefs.export_cterm lthy thy_ctxt (Thm.cterm_of thy rhs) |>> Thm.term_of;  | 
|
285  | 
val xs = Variable.add_fixed (LocalTheory.target_of lthy) rhs' [];  | 
|
| 24939 | 286  | 
val T = Term.fastype_of rhs;  | 
287  | 
||
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
288  | 
(*const*)  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
289  | 
val ((lhs, local_def), lthy2) = lthy |> declare_const ta (member (op =) xs) ((b, T), mx);  | 
| 25022 | 290  | 
val (_, lhs') = Logic.dest_equals (Thm.prop_of local_def);  | 
| 24939 | 291  | 
|
292  | 
(*def*)  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
293  | 
val define_const =  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
294  | 
(case Overloading.operation lthy c of  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
295  | 
SOME (_, checked) =>  | 
| 25519 | 296  | 
(fn name => fn (Const (c, _), rhs) => Overloading.define checked name (c, rhs))  | 
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
297  | 
| NONE =>  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
298  | 
if is_none (Class.instantiation_param lthy c)  | 
| 25519 | 299  | 
then (fn name => fn eq => Thm.add_def false false (name, Logic.mk_equals eq))  | 
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
300  | 
else (fn name => fn (Const (c, _), rhs) => AxClass.define_overloaded name (c, rhs)));  | 
| 25022 | 301  | 
val (global_def, lthy3) = lthy2  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27690 
diff
changeset
 | 
302  | 
|> LocalTheory.theory_result (define_const (Name.name_of name') (lhs', rhs'));  | 
| 25485 | 303  | 
val def = LocalDefs.trans_terms lthy3  | 
| 25022 | 304  | 
[(*c == global.c xs*) local_def,  | 
305  | 
(*global.c xs == rhs'*) global_def,  | 
|
| 25485 | 306  | 
(*rhs' == rhs*) Thm.symmetric rhs_conv];  | 
| 24939 | 307  | 
|
| 
25105
 
c9f67836c4d8
internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
 
wenzelm 
parents: 
25096 
diff
changeset
 | 
308  | 
(*note*)  | 
| 24939 | 309  | 
val ([(res_name, [res])], lthy4) = lthy3  | 
| 26132 | 310  | 
|> notes ta kind [((name', atts), [([def], [])])];  | 
| 24939 | 311  | 
in ((lhs, (res_name, res)), lthy4) end;  | 
312  | 
||
313  | 
||
| 25291 | 314  | 
(* init *)  | 
315  | 
||
316  | 
local  | 
|
| 20894 | 317  | 
|
| 25291 | 318  | 
fun init_target _ NONE = global_target  | 
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
319  | 
| init_target thy (SOME target) =  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
320  | 
make_target target true (Class.is_class thy target) ([], [], []) [];  | 
| 25291 | 321  | 
|
| 25519 | 322  | 
fun init_ctxt (Target {target, is_locale, is_class, instantiation, overloading}) =
 | 
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
323  | 
if not (null (#1 instantiation)) then Class.init_instantiation instantiation  | 
| 25519 | 324  | 
else if not (null overloading) then Overloading.init overloading  | 
| 25485 | 325  | 
else if not is_locale then ProofContext.init  | 
326  | 
else if not is_class then Locale.init target  | 
|
327  | 
else Class.init target;  | 
|
| 25269 | 328  | 
|
| 25542 | 329  | 
fun init_lthy (ta as Target {target, instantiation, overloading, ...}) =
 | 
| 25291 | 330  | 
Data.put ta #>  | 
331  | 
LocalTheory.init (NameSpace.base target)  | 
|
332  | 
   {pretty = pretty ta,
 | 
|
333  | 
abbrev = abbrev ta,  | 
|
334  | 
define = define ta,  | 
|
335  | 
notes = notes ta,  | 
|
336  | 
type_syntax = type_syntax ta,  | 
|
337  | 
term_syntax = term_syntax ta,  | 
|
338  | 
declaration = declaration ta,  | 
|
339  | 
reinit = fn lthy => init_lthy_ctxt ta (ProofContext.theory_of lthy),  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
340  | 
exit = LocalTheory.target_of o  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
341  | 
(if not (null (#1 instantiation)) then Class.conclude_instantiation  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
342  | 
else if not (null overloading) then Overloading.conclude  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
343  | 
else I)}  | 
| 25291 | 344  | 
and init_lthy_ctxt ta = init_lthy ta o init_ctxt ta;  | 
| 20894 | 345  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
346  | 
fun gen_overloading prep_const raw_ops thy =  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
347  | 
let  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
348  | 
val ctxt = ProofContext.init thy;  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
349  | 
val ops = raw_ops |> map (fn (name, const, checked) =>  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
350  | 
(name, Term.dest_Const (prep_const ctxt const), checked));  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
351  | 
in thy |> init_lthy_ctxt (make_target "" false false ([], [], []) ops) end;  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
352  | 
|
| 25291 | 353  | 
in  | 
| 20894 | 354  | 
|
| 25291 | 355  | 
fun init target thy = init_lthy_ctxt (init_target thy target) thy;  | 
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
356  | 
fun begin target ctxt = init_lthy (init_target (ProofContext.theory_of ctxt) (SOME target)) ctxt;  | 
| 25269 | 357  | 
|
| 25291 | 358  | 
fun context "-" thy = init NONE thy  | 
359  | 
| context target thy = init (SOME (Locale.intern thy target)) thy;  | 
|
| 20894 | 360  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
361  | 
fun instantiation arities = init_lthy_ctxt (make_target "" false false arities []);  | 
| 25519 | 362  | 
|
| 
26049
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
363  | 
val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const);  | 
| 
 
8186c03194ed
overloading: reduced code redundancy, no xstrings here;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
364  | 
val overloading_cmd = gen_overloading Syntax.read_term;  | 
| 25462 | 365  | 
|
| 20894 | 366  | 
end;  | 
| 25291 | 367  | 
|
368  | 
end;  | 
|
369  |