| author | blanchet | 
| Wed, 15 Sep 2010 19:55:26 +0200 | |
| changeset 39430 | afb4d5c672bd | 
| parent 38757 | 2b3e054ae6fc | 
| child 39557 | fe5722fce758 | 
| permissions | -rw-r--r-- | 
| 20894 | 1  | 
(* Title: Pure/Isar/theory_target.ML  | 
2  | 
Author: Makarius  | 
|
| 30437 | 3  | 
Author: Florian Haftmann, TU Muenchen  | 
| 20894 | 4  | 
|
| 
38350
 
480b2de9927c
renamed Theory_Target to the more appropriate Named_Target
 
haftmann 
parents: 
38349 
diff
changeset
 | 
5  | 
Targets for theory, locale and class.  | 
| 20894 | 6  | 
*)  | 
7  | 
||
| 
38350
 
480b2de9927c
renamed Theory_Target to the more appropriate Named_Target
 
haftmann 
parents: 
38349 
diff
changeset
 | 
8  | 
signature NAMED_TARGET =  | 
| 20894 | 9  | 
sig  | 
| 
38389
 
d7d915bae307
Named_Target.init: empty string represents theory target
 
haftmann 
parents: 
38388 
diff
changeset
 | 
10  | 
val init: string -> theory -> local_theory  | 
| 38388 | 11  | 
val theory_init: theory -> local_theory  | 
| 
38391
 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 
haftmann 
parents: 
38389 
diff
changeset
 | 
12  | 
val reinit: local_theory -> local_theory -> local_theory  | 
| 35205 | 13  | 
val context_cmd: xstring -> theory -> local_theory  | 
| 
38391
 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 
haftmann 
parents: 
38389 
diff
changeset
 | 
14  | 
  val peek: local_theory -> {target: string, is_locale: bool, is_class: bool} option
 | 
| 20894 | 15  | 
end;  | 
16  | 
||
| 
38350
 
480b2de9927c
renamed Theory_Target to the more appropriate Named_Target
 
haftmann 
parents: 
38349 
diff
changeset
 | 
17  | 
structure Named_Target: NAMED_TARGET =  | 
| 20894 | 18  | 
struct  | 
19  | 
||
| 21006 | 20  | 
(* context data *)  | 
21  | 
||
| 
38338
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
22  | 
datatype target = Target of {target: string, is_locale: bool, is_class: bool};
 | 
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
23  | 
|
| 
38338
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
24  | 
fun make_target target is_locale is_class =  | 
| 
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
25  | 
  Target {target = target, is_locale = is_locale, is_class = is_class};
 | 
| 25291 | 26  | 
|
| 38378 | 27  | 
val global_target = Target {target = "", is_locale = false, is_class = false};
 | 
28  | 
||
| 
38389
 
d7d915bae307
Named_Target.init: empty string represents theory target
 
haftmann 
parents: 
38388 
diff
changeset
 | 
29  | 
fun named_target _ "" = global_target  | 
| 
 
d7d915bae307
Named_Target.init: empty string represents theory target
 
haftmann 
parents: 
38388 
diff
changeset
 | 
30  | 
| named_target thy locale =  | 
| 38378 | 31  | 
if Locale.defined thy locale  | 
| 
38379
 
67d71449e85b
more convenient split of class modules: class and class_declaration
 
haftmann 
parents: 
38378 
diff
changeset
 | 
32  | 
      then Target {target = locale, is_locale = true, is_class = Class.is_class thy locale}
 | 
| 38378 | 33  | 
      else error ("No such locale: " ^ quote locale);
 | 
| 
25012
 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 
wenzelm 
parents: 
25002 
diff
changeset
 | 
34  | 
|
| 33519 | 35  | 
structure Data = Proof_Data  | 
| 21006 | 36  | 
(  | 
| 
38391
 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 
haftmann 
parents: 
38389 
diff
changeset
 | 
37  | 
type T = target option;  | 
| 
 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 
haftmann 
parents: 
38389 
diff
changeset
 | 
38  | 
fun init _ = NONE;  | 
| 21006 | 39  | 
);  | 
40  | 
||
| 
38391
 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 
haftmann 
parents: 
38389 
diff
changeset
 | 
41  | 
val peek = Option.map (fn Target args => args) o Data.get;  | 
| 21006 | 42  | 
|
43  | 
||
| 
33456
 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 
wenzelm 
parents: 
33282 
diff
changeset
 | 
44  | 
(* generic declarations *)  | 
| 
 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 
wenzelm 
parents: 
33282 
diff
changeset
 | 
45  | 
|
| 38305 | 46  | 
fun locale_declaration locale { syntax, pervasive } decl lthy =
 | 
| 24838 | 47  | 
let  | 
| 38305 | 48  | 
val add = if syntax then Locale.add_syntax_declaration else Locale.add_declaration;  | 
49  | 
val locale_decl = Morphism.transform (Local_Theory.target_morphism lthy) decl;  | 
|
| 24838 | 50  | 
in  | 
| 38305 | 51  | 
lthy  | 
| 
38341
 
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
 
haftmann 
parents: 
38339 
diff
changeset
 | 
52  | 
|> pervasive ? Generic_Target.theory_declaration decl  | 
| 38305 | 53  | 
|> Local_Theory.target (add locale locale_decl)  | 
| 24838 | 54  | 
end;  | 
55  | 
||
| 38305 | 56  | 
fun target_declaration (Target {target, ...}) { syntax, pervasive } =
 | 
| 
38341
 
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
 
haftmann 
parents: 
38339 
diff
changeset
 | 
57  | 
if target = "" then Generic_Target.theory_declaration  | 
| 38305 | 58  | 
  else locale_declaration target { syntax = syntax, pervasive = pervasive };
 | 
| 24838 | 59  | 
|
| 
21611
 
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
 
wenzelm 
parents: 
21594 
diff
changeset
 | 
60  | 
|
| 38297 | 61  | 
(* consts in locales *)  | 
62  | 
||
63  | 
fun locale_const (Target {target, is_class, ...}) (prmode as (mode, _)) ((b, mx), rhs) phi =
 | 
|
64  | 
let  | 
|
65  | 
val b' = Morphism.binding phi b;  | 
|
66  | 
val rhs' = Morphism.term phi rhs;  | 
|
67  | 
val arg = (b', Term.close_schematic_term rhs');  | 
|
68  | 
val same_shape = Term.aconv_untyped (rhs, rhs');  | 
|
69  | 
(* FIXME workaround based on educated guess *)  | 
|
70  | 
val prefix' = Binding.prefix_of b';  | 
|
71  | 
val is_canonical_class = is_class andalso  | 
|
72  | 
(Binding.eq_name (b, b')  | 
|
73  | 
andalso not (null prefix')  | 
|
| 
38379
 
67d71449e85b
more convenient split of class modules: class and class_declaration
 
haftmann 
parents: 
38378 
diff
changeset
 | 
74  | 
andalso List.last prefix' = (Class.class_prefix target, false)  | 
| 38297 | 75  | 
orelse same_shape);  | 
76  | 
in  | 
|
77  | 
not is_canonical_class ?  | 
|
78  | 
(Context.mapping_result  | 
|
79  | 
(Sign.add_abbrev Print_Mode.internal arg)  | 
|
80  | 
(ProofContext.add_abbrev Print_Mode.internal arg)  | 
|
81  | 
#-> (fn (lhs' as Const (d, _), _) =>  | 
|
82  | 
same_shape ?  | 
|
| 38310 | 83  | 
(Context.mapping  | 
84  | 
(Sign.revert_abbrev mode d) (ProofContext.revert_abbrev mode d) #>  | 
|
| 38297 | 85  | 
Morphism.form (ProofContext.target_notation true prmode [(lhs', mx)]))))  | 
86  | 
end;  | 
|
87  | 
||
| 38305 | 88  | 
fun locale_const_declaration (ta as Target {target, ...}) prmode arg =
 | 
89  | 
  locale_declaration target { syntax = true, pervasive = false } (locale_const ta prmode arg);
 | 
|
90  | 
||
| 38297 | 91  | 
|
| 
38318
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
92  | 
(* define *)  | 
| 
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
93  | 
|
| 
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
94  | 
fun locale_foundation ta (((b, U), mx), (b_def, rhs)) (type_params, term_params) =  | 
| 
38341
 
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
 
haftmann 
parents: 
38339 
diff
changeset
 | 
95  | 
Generic_Target.theory_foundation (((b, U), NoSyn), (b_def, rhs)) (type_params, term_params)  | 
| 
38318
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
96  | 
#-> (fn (lhs, def) => locale_const_declaration ta Syntax.mode_default ((b, mx), lhs)  | 
| 
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
97  | 
#> pair (lhs, def))  | 
| 
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
98  | 
|
| 
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
99  | 
fun class_foundation (ta as Target {target, ...})
 | 
| 
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
100  | 
(((b, U), mx), (b_def, rhs)) (type_params, term_params) =  | 
| 
38341
 
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
 
haftmann 
parents: 
38339 
diff
changeset
 | 
101  | 
Generic_Target.theory_foundation (((b, U), NoSyn), (b_def, rhs)) (type_params, term_params)  | 
| 
38318
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
102  | 
#-> (fn (lhs, def) => locale_const_declaration ta Syntax.mode_default ((b, NoSyn), lhs)  | 
| 
38619
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
103  | 
#> Class.const target ((b, mx), (type_params, lhs))  | 
| 
38318
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
104  | 
#> pair (lhs, def))  | 
| 
 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 
haftmann 
parents: 
38314 
diff
changeset
 | 
105  | 
|
| 38349 | 106  | 
fun target_foundation (ta as Target {target, is_locale, is_class, ...}) =
 | 
107  | 
if is_class then class_foundation ta  | 
|
108  | 
else if is_locale then locale_foundation ta  | 
|
109  | 
else Generic_Target.theory_foundation;  | 
|
| 
38303
 
ad4b59e9d0f9
factored out foundation of `define` into separate function
 
haftmann 
parents: 
38302 
diff
changeset
 | 
110  | 
|
| 24939 | 111  | 
|
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
112  | 
(* notes *)  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
113  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
114  | 
fun locale_notes locale kind global_facts local_facts lthy =  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
115  | 
let  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
116  | 
val global_facts' = Attrib.map_facts (K I) global_facts;  | 
| 38310 | 117  | 
val local_facts' = Element.facts_map  | 
118  | 
(Element.morph_ctxt (Local_Theory.target_morphism lthy)) local_facts;  | 
|
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
119  | 
in  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
120  | 
lthy  | 
| 
38757
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
121  | 
|> Local_Theory.background_theory (PureThy.note_thmss kind global_facts' #> snd)  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
122  | 
|> Local_Theory.target (Locale.add_thmss locale kind local_facts')  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
123  | 
end  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
124  | 
|
| 
38319
 
1cfc2b128e33
separate initialisation for overloading and instantiation target
 
haftmann 
parents: 
38318 
diff
changeset
 | 
125  | 
fun target_notes (ta as Target {target, is_locale, ...}) =
 | 
| 
 
1cfc2b128e33
separate initialisation for overloading and instantiation target
 
haftmann 
parents: 
38318 
diff
changeset
 | 
126  | 
if is_locale then locale_notes target  | 
| 
38341
 
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
 
haftmann 
parents: 
38339 
diff
changeset
 | 
127  | 
else fn kind => fn global_facts => fn _ => Generic_Target.theory_notes kind global_facts;  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
128  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
129  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
130  | 
(* abbrev *)  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
131  | 
|
| 
38757
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
132  | 
fun locale_abbrev ta prmode ((b, mx), t) xs =  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
133  | 
Local_Theory.background_theory_result  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
134  | 
(Sign.add_abbrev Print_Mode.internal (b, t)) #->  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
135  | 
(fn (lhs, _) => locale_const_declaration ta prmode  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
136  | 
((b, mx), Term.list_comb (Logic.unvarify_global lhs, xs)));  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
137  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
138  | 
fun target_abbrev (ta as Target {target, is_locale, is_class, ...})
 | 
| 
38312
 
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
 
haftmann 
parents: 
38311 
diff
changeset
 | 
139  | 
prmode (b, mx) (global_rhs, t') xs lthy =  | 
| 
38757
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
140  | 
if is_locale then  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
141  | 
lthy  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
142  | 
|> locale_abbrev ta prmode ((b, if is_class then NoSyn else mx), global_rhs) xs  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
143  | 
|> is_class ? Class.abbrev target prmode ((b, mx), t')  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
144  | 
else  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
145  | 
lthy  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
146  | 
|> Generic_Target.theory_abbrev prmode ((b, mx), global_rhs);  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
147  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
148  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
149  | 
(* pretty *)  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
150  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
151  | 
fun pretty_thy ctxt target is_class =  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
152  | 
let  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
153  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
154  | 
val target_name =  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
155  | 
[Pretty.command (if is_class then "class" else "locale"),  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
156  | 
        Pretty.str (" " ^ Locale.extern thy target)];
 | 
| 38310 | 157  | 
val fixes = map (fn (x, T) => (Binding.name x, SOME T, NoSyn))  | 
158  | 
(#1 (ProofContext.inferred_fixes ctxt));  | 
|
159  | 
val assumes = map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])]))  | 
|
160  | 
(Assumption.all_assms_of ctxt);  | 
|
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
161  | 
val elems =  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
162  | 
(if null fixes then [] else [Element.Fixes fixes]) @  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
163  | 
(if null assumes then [] else [Element.Assumes assumes]);  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
164  | 
in  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
165  | 
if target = "" then []  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
166  | 
else if null elems then [Pretty.block target_name]  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
167  | 
else [Pretty.block (Pretty.fbreaks (Pretty.block (target_name @ [Pretty.str " ="]) ::  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
168  | 
map (Pretty.chunks o Element.pretty_ctxt ctxt) elems))]  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
169  | 
end;  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
170  | 
|
| 
38338
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
171  | 
fun pretty (Target {target, is_class, ...}) ctxt =
 | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
172  | 
Pretty.block [Pretty.command "theory", Pretty.brk 1,  | 
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
173  | 
Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::  | 
| 
38338
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
174  | 
pretty_thy ctxt target is_class;  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
175  | 
|
| 
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
176  | 
|
| 38378 | 177  | 
(* init *)  | 
| 25291 | 178  | 
|
179  | 
local  | 
|
| 20894 | 180  | 
|
| 38378 | 181  | 
fun init_context (Target {target, is_locale, is_class}) =
 | 
| 
38338
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
182  | 
if not is_locale then ProofContext.init_global  | 
| 29576 | 183  | 
else if not is_class then Locale.init target  | 
| 
38379
 
67d71449e85b
more convenient split of class modules: class and class_declaration
 
haftmann 
parents: 
38378 
diff
changeset
 | 
184  | 
else Class.init target;  | 
| 25269 | 185  | 
|
| 
38338
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
186  | 
fun init_target (ta as Target {target, ...}) thy =
 | 
| 38247 | 187  | 
thy  | 
| 38378 | 188  | 
|> init_context ta  | 
| 
38391
 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 
haftmann 
parents: 
38389 
diff
changeset
 | 
189  | 
|> Data.put (SOME ta)  | 
| 38247 | 190  | 
|> Local_Theory.init NONE (Long_Name.base_name target)  | 
| 38349 | 191  | 
     {define = Generic_Target.define (target_foundation ta),
 | 
| 
38319
 
1cfc2b128e33
separate initialisation for overloading and instantiation target
 
haftmann 
parents: 
38318 
diff
changeset
 | 
192  | 
notes = Generic_Target.notes (target_notes ta),  | 
| 
 
1cfc2b128e33
separate initialisation for overloading and instantiation target
 
haftmann 
parents: 
38318 
diff
changeset
 | 
193  | 
abbrev = Generic_Target.abbrev (target_abbrev ta),  | 
| 38310 | 194  | 
declaration = fn pervasive => target_declaration ta  | 
195  | 
        { syntax = false, pervasive = pervasive },
 | 
|
196  | 
syntax_declaration = fn pervasive => target_declaration ta  | 
|
197  | 
        { syntax = true, pervasive = pervasive },
 | 
|
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
38307 
diff
changeset
 | 
198  | 
pretty = pretty ta,  | 
| 
38338
 
0e0e1fd9cc03
remove overloading and instantiation from data slot
 
haftmann 
parents: 
38319 
diff
changeset
 | 
199  | 
exit = Local_Theory.target_of};  | 
| 38247 | 200  | 
|
| 25291 | 201  | 
in  | 
| 20894 | 202  | 
|
| 38247 | 203  | 
fun init target thy = init_target (named_target thy target) thy;  | 
| 25269 | 204  | 
|
| 
38757
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
205  | 
fun reinit lthy =  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
206  | 
(case peek lthy of  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
207  | 
    SOME {target, ...} => init target o Local_Theory.exit_global
 | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
208  | 
| NONE => error "Not in a named target");  | 
| 
38391
 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 
haftmann 
parents: 
38389 
diff
changeset
 | 
209  | 
|
| 38388 | 210  | 
val theory_init = init_target global_target;  | 
211  | 
||
| 
38389
 
d7d915bae307
Named_Target.init: empty string represents theory target
 
haftmann 
parents: 
38388 
diff
changeset
 | 
212  | 
fun context_cmd "-" thy = init "" thy  | 
| 
 
d7d915bae307
Named_Target.init: empty string represents theory target
 
haftmann 
parents: 
38388 
diff
changeset
 | 
213  | 
| context_cmd target thy = init (Locale.intern thy target) thy;  | 
| 33282 | 214  | 
|
| 20894 | 215  | 
end;  | 
| 25291 | 216  | 
|
217  | 
end;  |