| author | wenzelm | 
| Wed, 05 Dec 2012 20:24:49 +0100 | |
| changeset 50377 | fe4bc5b2abb4 | 
| parent 48102 | 9ed089bcad93 | 
| child 52103 | fb577a13abbd | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: Pure/Isar/named_target.ML | 
| 20894 | 2 | Author: Makarius | 
| 30437 | 3 | Author: Florian Haftmann, TU Muenchen | 
| 20894 | 4 | |
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47061diff
changeset | 5 | Targets for theory, locale, class -- at the bottom the nested structure. | 
| 20894 | 6 | *) | 
| 7 | ||
| 38350 
480b2de9927c
renamed Theory_Target to the more appropriate Named_Target
 haftmann parents: 
38349diff
changeset | 8 | signature NAMED_TARGET = | 
| 20894 | 9 | sig | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47061diff
changeset | 10 |   val peek: local_theory -> {target: string, is_locale: bool, is_class: bool} option
 | 
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 11 | val init: (local_theory -> local_theory) -> string -> theory -> local_theory | 
| 38388 | 12 | val theory_init: theory -> local_theory | 
| 38391 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 haftmann parents: 
38389diff
changeset | 13 | val reinit: local_theory -> local_theory -> local_theory | 
| 45488 
6d71d9e52369
pass positions for named targets, for formal links in the document model;
 wenzelm parents: 
45390diff
changeset | 14 | val context_cmd: xstring * Position.T -> theory -> local_theory | 
| 20894 | 15 | end; | 
| 16 | ||
| 38350 
480b2de9927c
renamed Theory_Target to the more appropriate Named_Target
 haftmann parents: 
38349diff
changeset | 17 | structure Named_Target: NAMED_TARGET = | 
| 20894 | 18 | struct | 
| 19 | ||
| 21006 | 20 | (* context data *) | 
| 21 | ||
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 22 | datatype target = | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 23 |   Target of {target: string, is_locale: bool, is_class: bool,
 | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 24 | before_exit: local_theory -> local_theory}; | 
| 25012 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 wenzelm parents: 
25002diff
changeset | 25 | |
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 26 | fun make_target target is_locale is_class before_exit = | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 27 |   Target {target = target, is_locale = is_locale, is_class = is_class,
 | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 28 | before_exit = before_exit}; | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 29 | |
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 30 | fun named_target _ "" before_exit = make_target "" false false before_exit | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 31 | | named_target thy locale before_exit = | 
| 38378 | 32 | if Locale.defined thy locale | 
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 33 | then make_target locale true (Class.is_class thy locale) before_exit | 
| 38378 | 34 |       else error ("No such locale: " ^ quote locale);
 | 
| 25012 
448af76a1ba3
pass explicit target record -- more informative peek operation;
 wenzelm parents: 
25002diff
changeset | 35 | |
| 33519 | 36 | structure Data = Proof_Data | 
| 21006 | 37 | ( | 
| 38391 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 haftmann parents: 
38389diff
changeset | 38 | type T = target option; | 
| 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 haftmann parents: 
38389diff
changeset | 39 | fun init _ = NONE; | 
| 21006 | 40 | ); | 
| 41 | ||
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 42 | val peek = | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 43 |   Data.get #> Option.map (fn Target {target, is_locale, is_class, ...} =>
 | 
| 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 44 |     {target = target, is_locale = is_locale, is_class = is_class});
 | 
| 21006 | 45 | |
| 46 | ||
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 47 | (* consts in locale *) | 
| 24838 | 48 | |
| 47282 | 49 | fun locale_const (Target {target, is_class, ...}) prmode ((b, mx), rhs) =
 | 
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 50 | Generic_Target.locale_declaration target true (fn phi => | 
| 45352 | 51 | let | 
| 52 | val b' = Morphism.binding phi b; | |
| 53 | val rhs' = Morphism.term phi rhs; | |
| 54 | val same_shape = Term.aconv_untyped (rhs, rhs'); | |
| 38297 | 55 | |
| 45352 | 56 | (* FIXME workaround based on educated guess *) | 
| 57 | val prefix' = Binding.prefix_of b'; | |
| 58 | val is_canonical_class = is_class andalso | |
| 59 | (Binding.eq_name (b, b') | |
| 60 | andalso not (null prefix') | |
| 61 | andalso List.last prefix' = (Class.class_prefix target, false) | |
| 62 | orelse same_shape); | |
| 47282 | 63 | in | 
| 47288 
b79bf8288b29
clarified generic_const vs. close_schematic_term;
 wenzelm parents: 
47286diff
changeset | 64 | not is_canonical_class ? Generic_Target.generic_const same_shape prmode ((b', mx), rhs') | 
| 47282 | 65 | end) #> | 
| 66 | (fn lthy => lthy |> Generic_Target.const_declaration | |
| 67 | (fn level => level <> 0 andalso level <> Local_Theory.level lthy) prmode ((b, mx), rhs)); | |
| 38305 | 68 | |
| 38297 | 69 | |
| 38318 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 haftmann parents: 
38314diff
changeset | 70 | (* define *) | 
| 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 haftmann parents: 
38314diff
changeset | 71 | |
| 47289 | 72 | fun locale_foundation ta (((b, U), mx), (b_def, rhs)) params = | 
| 73 | Generic_Target.background_foundation (((b, U), NoSyn), (b_def, rhs)) params | |
| 45352 | 74 | #-> (fn (lhs, def) => locale_const ta Syntax.mode_default ((b, mx), lhs) | 
| 47282 | 75 | #> pair (lhs, def)); | 
| 38318 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 haftmann parents: 
38314diff
changeset | 76 | |
| 47289 | 77 | fun class_foundation (ta as Target {target, ...}) (((b, U), mx), (b_def, rhs)) params =
 | 
| 78 | Generic_Target.background_foundation (((b, U), NoSyn), (b_def, rhs)) params | |
| 45352 | 79 | #-> (fn (lhs, def) => locale_const ta Syntax.mode_default ((b, NoSyn), lhs) | 
| 48102 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 80 | #> Class.const target ((b, mx), (#1 params, #2 params, lhs)) | 
| 47282 | 81 | #> pair (lhs, def)); | 
| 38318 
1fade69519ab
different foundations for different targets; simplified syntax handling of abbreviations
 haftmann parents: 
38314diff
changeset | 82 | |
| 39639 | 83 | fun target_foundation (ta as Target {is_locale, is_class, ...}) =
 | 
| 38349 | 84 | if is_class then class_foundation ta | 
| 85 | else if is_locale then locale_foundation ta | |
| 47284 
0e65b6a016dc
clarified background_foundation vs. theory_foundation (with const_declaration);
 wenzelm parents: 
47283diff
changeset | 86 | else Generic_Target.theory_foundation; | 
| 38303 
ad4b59e9d0f9
factored out foundation of `define` into separate function
 haftmann parents: 
38302diff
changeset | 87 | |
| 24939 | 88 | |
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 89 | (* notes *) | 
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 90 | |
| 39639 | 91 | fun target_notes (Target {target, is_locale, ...}) =
 | 
| 47250 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 wenzelm parents: 
47247diff
changeset | 92 | if is_locale then Generic_Target.locale_notes target | 
| 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 wenzelm parents: 
47247diff
changeset | 93 | else Generic_Target.theory_notes; | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 94 | |
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 95 | |
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 96 | (* abbrev *) | 
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 97 | |
| 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: 
38619diff
changeset | 98 | 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: 
38619diff
changeset | 99 | 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: 
38619diff
changeset | 100 | (Sign.add_abbrev Print_Mode.internal (b, t)) #-> | 
| 47286 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 wenzelm parents: 
47284diff
changeset | 101 | (fn (lhs, _) => | 
| 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 wenzelm parents: 
47284diff
changeset | 102 | locale_const ta prmode ((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: 
38307diff
changeset | 103 | |
| 47286 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 wenzelm parents: 
47284diff
changeset | 104 | fun target_abbrev (ta as Target {target, is_locale, is_class, ...}) prmode (b, mx) (t, 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: 
38619diff
changeset | 105 | 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: 
38619diff
changeset | 106 | lthy | 
| 47286 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 wenzelm parents: 
47284diff
changeset | 107 | |> locale_abbrev ta prmode ((b, if is_class then NoSyn else mx), t) xs | 
| 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: 
38619diff
changeset | 108 | |> is_class ? Class.abbrev target prmode ((b, mx), t') | 
| 47286 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 wenzelm parents: 
47284diff
changeset | 109 | else lthy |> Generic_Target.theory_abbrev prmode (b, mx) (t, t') xs; | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 110 | |
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 111 | |
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 112 | (* declaration *) | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 113 | |
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 114 | fun target_declaration (Target {target, ...}) {syntax, pervasive} decl lthy =
 | 
| 47279 
4bab649dedf0
clarified standard_declaration vs. theory_declaration;
 wenzelm parents: 
47278diff
changeset | 115 | if target = "" then Generic_Target.theory_declaration decl lthy | 
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 116 | else | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 117 | lthy | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 118 | |> pervasive ? Generic_Target.background_declaration decl | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 119 | |> Generic_Target.locale_declaration target syntax decl | 
| 47282 | 120 | |> (fn lthy' => lthy' |> Generic_Target.standard_declaration (fn level => level <> 0) decl); | 
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 121 | |
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47081diff
changeset | 122 | |
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 123 | (* pretty *) | 
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 124 | |
| 39639 | 125 | fun pretty (Target {target, is_locale, is_class, ...}) ctxt =
 | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 126 | let | 
| 42360 | 127 | val thy = Proof_Context.theory_of ctxt; | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 128 | val target_name = | 
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 129 | [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: 
38307diff
changeset | 130 |         Pretty.str (" " ^ Locale.extern thy target)];
 | 
| 45353 | 131 | val fixes = | 
| 132 | map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) | |
| 133 | (#1 (Proof_Context.inferred_fixes ctxt)); | |
| 134 | val assumes = | |
| 135 | map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])])) | |
| 136 | (Assumption.all_assms_of ctxt); | |
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 137 | val elems = | 
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 138 | (if null fixes then [] else [Element.Fixes fixes]) @ | 
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 139 | (if null assumes then [] else [Element.Assumes assumes]); | 
| 40782 | 140 | val body_elems = | 
| 141 | if not is_locale then [] | |
| 39639 | 142 | else if null elems then [Pretty.block target_name] | 
| 143 | else [Pretty.block (Pretty.fbreaks (Pretty.block (target_name @ [Pretty.str " ="]) :: | |
| 40782 | 144 | map (Pretty.chunks o Element.pretty_ctxt ctxt) elems))]; | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 145 | in | 
| 39639 | 146 | Pretty.block [Pretty.command "theory", Pretty.brk 1, | 
| 42360 | 147 | Pretty.str (Context.theory_name (Proof_Context.theory_of ctxt))] :: body_elems | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 148 | end; | 
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 149 | |
| 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
38307diff
changeset | 150 | |
| 38378 | 151 | (* init *) | 
| 25291 | 152 | |
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 153 | fun init_context (Target {target, is_locale, is_class, ...}) =
 | 
| 42360 | 154 | if not is_locale then Proof_Context.init_global | 
| 29576 | 155 | else if not is_class then Locale.init target | 
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38378diff
changeset | 156 | else Class.init target; | 
| 25269 | 157 | |
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 158 | fun init before_exit target thy = | 
| 39639 | 159 | let | 
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 160 | val ta = named_target thy target before_exit; | 
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
47000diff
changeset | 161 | val naming = Sign.naming_of thy | 
| 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
47000diff
changeset | 162 | |> Name_Space.mandatory_path (Long_Name.base_name target); | 
| 39639 | 163 | in | 
| 164 | thy | |
| 165 | |> init_context ta | |
| 166 | |> Data.put (SOME ta) | |
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
47000diff
changeset | 167 | |> Local_Theory.init naming | 
| 39639 | 168 |        {define = Generic_Target.define (target_foundation ta),
 | 
| 169 | notes = Generic_Target.notes (target_notes ta), | |
| 170 | abbrev = Generic_Target.abbrev (target_abbrev ta), | |
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
45290diff
changeset | 171 | declaration = target_declaration ta, | 
| 39639 | 172 | pretty = pretty ta, | 
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 173 | exit = Local_Theory.target_of o before_exit} | 
| 39639 | 174 | end; | 
| 38247 | 175 | |
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
40782diff
changeset | 176 | val theory_init = init I ""; | 
| 25269 | 177 | |
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47061diff
changeset | 178 | val reinit = | 
| 47069 | 179 | Local_Theory.assert_bottom true #> Data.get #> the #> | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47061diff
changeset | 180 |   (fn Target {target, before_exit, ...} => Local_Theory.exit_global #> init before_exit target);
 | 
| 38391 
ba1cc1815ce1
named target is optional; explicit Name_Target.reinit
 haftmann parents: 
38389diff
changeset | 181 | |
| 47274 | 182 | fun context_cmd ("-", _) thy = theory_init thy
 | 
| 45488 
6d71d9e52369
pass positions for named targets, for formal links in the document model;
 wenzelm parents: 
45390diff
changeset | 183 | | context_cmd target thy = init I (Locale.check thy target) thy; | 
| 33282 | 184 | |
| 20894 | 185 | end; |