| author | wenzelm | 
| Thu, 05 Jan 2017 16:40:11 +0100 | |
| changeset 64804 | b2b05fdff3a7 | 
| parent 63030 | 4e7eff53bee7 | 
| child 71674 | 48ff625687f5 | 
| 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 | 
| 55923 | 20 | val is_reported_generic: Context.generic -> Position.T -> bool | 
| 21 | val is_reported: Proof.context -> Position.T -> bool | |
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46874diff
changeset | 22 | val report_generic: Context.generic -> Position.T -> Markup.T -> unit | 
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 23 | val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 24 | val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 25 | val report: Proof.context -> Position.T -> Markup.T -> unit | 
| 48767 | 26 | val reports_text: Proof.context -> Position.report_text list -> unit | 
| 44736 | 27 | val reports: Proof.context -> Position.report list -> unit | 
| 28405 | 28 | end; | 
| 29 | ||
| 33383 | 30 | structure Context_Position: CONTEXT_POSITION = | 
| 28405 | 31 | struct | 
| 32 | ||
| 47813 | 33 | structure Data = Generic_Data | 
| 34 | ( | |
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 35 | type T = bool option * bool option; (*really, visible*) | 
| 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 36 | val empty: T = (NONE, NONE); | 
| 47813 | 37 | val extend = I; | 
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 38 | fun merge ((a, b), (a', b')) : T = (merge_options (a, a'), merge_options (b, b')); | 
| 47813 | 39 | ); | 
| 40 | ||
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 41 | val is_visible_generic = the_default true o snd o Data.get; | 
| 47813 | 42 | val is_visible = is_visible_generic o Context.Proof; | 
| 52700 | 43 | val is_visible_global = is_visible_generic o Context.Theory; | 
| 28405 | 44 | |
| 60858 | 45 | val set_visible_generic = Data.map o apsnd o K o SOME; | 
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 46 | val set_visible = Context.proof_map o Data.map o apsnd o K o SOME; | 
| 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 47 | val set_visible_global = Context.theory_map o Data.map o apsnd o K o SOME; | 
| 52700 | 48 | |
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 49 | val is_really = the_default true o fst o Data.get o Context.Proof; | 
| 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 50 | fun is_really_visible ctxt = is_really ctxt andalso is_visible ctxt; | 
| 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 51 | val not_really = Context.proof_map (Data.map (apfst (K (SOME false)))); | 
| 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 52 | |
| 63030 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 wenzelm parents: 
60858diff
changeset | 53 | val restore_visible_generic = Data.put o Data.get; | 
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 54 | val restore_visible = Context.proof_map o Data.put o Data.get o Context.Proof; | 
| 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56333diff
changeset | 55 | val restore_visible_global = Context.theory_map o Data.put o Data.get o Context.Theory; | 
| 38831 
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
 wenzelm parents: 
38481diff
changeset | 56 | |
| 55923 | 57 | fun is_reported_generic context pos = is_visible_generic context andalso Position.is_reported pos; | 
| 58 | fun is_reported ctxt pos = is_visible ctxt andalso Position.is_reported pos; | |
| 59 | ||
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46874diff
changeset | 60 | fun report_generic context pos markup = | 
| 55923 | 61 | 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 | 62 | 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 | 63 | else (); | 
| 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46874diff
changeset | 64 | |
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 65 | fun reported_text ctxt pos markup txt = | 
| 55923 | 66 | 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 | 67 | |
| 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 | 68 | 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 | 69 | fun report ctxt pos markup = report_text ctxt pos markup ""; | 
| 28405 | 70 | |
| 48767 | 71 | fun reports_text ctxt reps = if is_visible ctxt then Position.reports_text reps else (); | 
| 44736 | 72 | fun reports ctxt reps = if is_visible ctxt then Position.reports reps else (); | 
| 73 | ||
| 28405 | 74 | end; |