| author | wenzelm | 
| Fri, 04 Oct 2024 00:00:50 +0200 | |
| changeset 81110 | 08165a4e105d | 
| parent 80868 | 0ed02f473cf9 | 
| permissions | -rw-r--r-- | 
| 28405 | 1 | (* Title: Pure/context_position.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 28409 | 4 | Context position visibility flag. | 
| 28405 | 5 | *) | 
| 6 | ||
| 7 | signature CONTEXT_POSITION = | |
| 8 | sig | |
| 55672 
5e25cc741ab9
support for completion within the formal context;
 wenzelm parents: 
52700diff
changeset | 9 | val is_visible_generic: Context.generic -> bool | 
| 28405 | 10 | val is_visible: Proof.context -> bool | 
| 52700 | 11 | val is_visible_global: theory -> bool | 
| 60858 | 12 | val set_visible_generic: bool -> Context.generic -> Context.generic | 
| 28405 | 13 | val set_visible: bool -> Proof.context -> Proof.context | 
| 52700 | 14 | val set_visible_global: bool -> theory -> theory | 
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 15 | val is_really_visible: Proof.context -> bool | 
| 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 16 | val not_really: Proof.context -> Proof.context | 
| 63030 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 wenzelm parents: 
60858diff
changeset | 17 | val restore_visible_generic: Context.generic -> Context.generic -> Context.generic | 
| 28405 | 18 | val restore_visible: Proof.context -> Proof.context -> Proof.context | 
| 52700 | 19 | val restore_visible_global: theory -> theory -> theory | 
| 76804 | 20 | val reports_enabled0: unit -> bool | 
| 71675 | 21 | val reports_enabled_generic: Context.generic -> bool | 
| 22 | val reports_enabled: Proof.context -> bool | |
| 23 | val reports_enabled_global: theory -> bool | |
| 55923 | 24 | val is_reported_generic: Context.generic -> Position.T -> bool | 
| 25 | val is_reported: Proof.context -> Position.T -> bool | |
| 71674 | 26 | val is_reported_global: theory -> Position.T -> bool | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46874diff
changeset | 27 | val report_generic: Context.generic -> Position.T -> Markup.T -> unit | 
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 28 | val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 29 | val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 30 | val report: Proof.context -> Position.T -> Markup.T -> unit | 
| 48767 | 31 | val reports_text: Proof.context -> Position.report_text list -> unit | 
| 71674 | 32 | val reports_generic: Context.generic -> Position.report list -> unit | 
| 44736 | 33 | val reports: Proof.context -> Position.report list -> unit | 
| 71674 | 34 | val reports_global: theory -> Position.report list -> unit | 
| 28405 | 35 | end; | 
| 36 | ||
| 33383 | 37 | structure Context_Position: CONTEXT_POSITION = | 
| 28405 | 38 | struct | 
| 39 | ||
| 71674 | 40 | (* visible context *) | 
| 41 | ||
| 71680 | 42 | val really = Config.declare_bool ("really", Position.none) (K true);
 | 
| 43 | val visible = Config.declare_bool ("visible", Position.none) (K true);
 | |
| 47813 | 44 | |
| 71680 | 45 | val is_visible_generic = Config.apply_generic visible; | 
| 46 | val is_visible = Config.apply visible; | |
| 47 | val is_visible_global = Config.apply_global visible; | |
| 48 | ||
| 49 | val set_visible_generic = Config.put_generic visible; | |
| 50 | val set_visible = Config.put visible; | |
| 51 | val set_visible_global = Config.put_global visible; | |
| 28405 | 52 | |
| 71680 | 53 | val is_really = Config.apply really; | 
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 54 | fun is_really_visible ctxt = is_really ctxt andalso is_visible ctxt; | 
| 71680 | 55 | val not_really = Config.put really false; | 
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 56 | |
| 71680 | 57 | fun restore_visible_generic context = | 
| 58 | Config.restore_generic really context #> | |
| 59 | Config.restore_generic visible context; | |
| 60 | ||
| 61 | val restore_visible = Context.proof_map o restore_visible_generic o Context.Proof; | |
| 62 | val restore_visible_global = Context.theory_map o restore_visible_generic o Context.Theory; | |
| 38831 
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
 wenzelm parents: 
38481diff
changeset | 63 | |
| 71674 | 64 | |
| 65 | (* PIDE reports *) | |
| 66 | ||
| 80868 
0ed02f473cf9
more robust reports: ensure that markup is actually present;
 wenzelm parents: 
76804diff
changeset | 67 | fun reports_enabled0 () = Options.default_bool "pide_reports" andalso Print_Mode.PIDE_enabled (); | 
| 76804 | 68 | fun reports_enabled_generic context = reports_enabled0 () andalso is_visible_generic context; | 
| 71675 | 69 | val reports_enabled = reports_enabled_generic o Context.Proof; | 
| 70 | val reports_enabled_global = reports_enabled_generic o Context.Theory; | |
| 71 | ||
| 72 | fun is_reported_generic context pos = | |
| 73 | reports_enabled_generic context andalso Position.is_reported pos; | |
| 71674 | 74 | |
| 75 | val is_reported = is_reported_generic o Context.Proof; | |
| 76 | val is_reported_global = is_reported_generic o Context.Theory; | |
| 55923 | 77 | |
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46874diff
changeset | 78 | fun report_generic context pos markup = | 
| 55923 | 79 | if is_reported_generic context pos then | 
| 56333 
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
 wenzelm parents: 
56294diff
changeset | 80 | Output.report [Position.reported_text pos markup ""] | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46874diff
changeset | 81 | else (); | 
| 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46874diff
changeset | 82 | |
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 83 | fun reported_text ctxt pos markup txt = | 
| 55923 | 84 | if is_reported ctxt pos then Position.reported_text pos markup txt else ""; | 
| 38481 
81ec258c4cd3
Output_Position.report_text -- markup with potential "arguments";
 wenzelm parents: 
38237diff
changeset | 85 | |
| 56333 
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
 wenzelm parents: 
56294diff
changeset | 86 | fun report_text ctxt pos markup txt = Output.report [reported_text ctxt pos markup txt]; | 
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 87 | fun report ctxt pos markup = report_text ctxt pos markup ""; | 
| 28405 | 88 | |
| 71675 | 89 | fun reports_text ctxt reps = | 
| 90 | if reports_enabled ctxt then Position.reports_text reps else (); | |
| 91 | ||
| 92 | fun reports_generic context reps = | |
| 93 | if reports_enabled_generic context then Position.reports reps else (); | |
| 71674 | 94 | |
| 95 | val reports = reports_generic o Context.Proof; | |
| 96 | val reports_global = reports_generic o Context.Theory; | |
| 44736 | 97 | |
| 28405 | 98 | end; |