| author | wenzelm | 
| Mon, 10 Aug 2015 17:49:36 +0200 | |
| changeset 60878 | 1f0d2bbcf38b | 
| parent 60858 | 7bf2188a0998 | 
| child 63030 | 4e7eff53bee7 | 
| 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: 
52700 
diff
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: 
56333 
diff
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: 
56333 
diff
changeset
 | 
16  | 
val not_really: Proof.context -> Proof.context  | 
| 28405 | 17  | 
val restore_visible: Proof.context -> Proof.context -> Proof.context  | 
| 52700 | 18  | 
val restore_visible_global: theory -> theory -> theory  | 
| 55923 | 19  | 
val is_reported_generic: Context.generic -> Position.T -> bool  | 
20  | 
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: 
46874 
diff
changeset
 | 
21  | 
val report_generic: Context.generic -> Position.T -> Markup.T -> unit  | 
| 
39507
 
839873937ddd
tuned signature of (Context_)Position.report variants;
 
wenzelm 
parents: 
39440 
diff
changeset
 | 
22  | 
val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string  | 
| 
 
839873937ddd
tuned signature of (Context_)Position.report variants;
 
wenzelm 
parents: 
39440 
diff
changeset
 | 
23  | 
val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit  | 
| 
 
839873937ddd
tuned signature of (Context_)Position.report variants;
 
wenzelm 
parents: 
39440 
diff
changeset
 | 
24  | 
val report: Proof.context -> Position.T -> Markup.T -> unit  | 
| 48767 | 25  | 
val reports_text: Proof.context -> Position.report_text list -> unit  | 
| 44736 | 26  | 
val reports: Proof.context -> Position.report list -> unit  | 
| 28405 | 27  | 
end;  | 
28  | 
||
| 33383 | 29  | 
structure Context_Position: CONTEXT_POSITION =  | 
| 28405 | 30  | 
struct  | 
31  | 
||
| 47813 | 32  | 
structure Data = Generic_Data  | 
33  | 
(  | 
|
| 
57858
 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
34  | 
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: 
56333 
diff
changeset
 | 
35  | 
val empty: T = (NONE, NONE);  | 
| 47813 | 36  | 
val extend = I;  | 
| 
57858
 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
37  | 
fun merge ((a, b), (a', b')) : T = (merge_options (a, a'), merge_options (b, b'));  | 
| 47813 | 38  | 
);  | 
39  | 
||
| 
57858
 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
40  | 
val is_visible_generic = the_default true o snd o Data.get;  | 
| 47813 | 41  | 
val is_visible = is_visible_generic o Context.Proof;  | 
| 52700 | 42  | 
val is_visible_global = is_visible_generic o Context.Theory;  | 
| 28405 | 43  | 
|
| 60858 | 44  | 
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: 
56333 
diff
changeset
 | 
45  | 
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: 
56333 
diff
changeset
 | 
46  | 
val set_visible_global = Context.theory_map o Data.map o apsnd o K o SOME;  | 
| 52700 | 47  | 
|
| 
57858
 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
48  | 
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: 
56333 
diff
changeset
 | 
49  | 
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: 
56333 
diff
changeset
 | 
50  | 
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: 
56333 
diff
changeset
 | 
51  | 
|
| 
 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
52  | 
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: 
56333 
diff
changeset
 | 
53  | 
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: 
38481 
diff
changeset
 | 
54  | 
|
| 55923 | 55  | 
fun is_reported_generic context pos = is_visible_generic context andalso Position.is_reported pos;  | 
56  | 
fun is_reported ctxt pos = is_visible ctxt andalso Position.is_reported pos;  | 
|
57  | 
||
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46874 
diff
changeset
 | 
58  | 
fun report_generic context pos markup =  | 
| 55923 | 59  | 
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: 
56294 
diff
changeset
 | 
60  | 
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: 
46874 
diff
changeset
 | 
61  | 
else ();  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46874 
diff
changeset
 | 
62  | 
|
| 
39507
 
839873937ddd
tuned signature of (Context_)Position.report variants;
 
wenzelm 
parents: 
39440 
diff
changeset
 | 
63  | 
fun reported_text ctxt pos markup txt =  | 
| 55923 | 64  | 
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: 
38237 
diff
changeset
 | 
65  | 
|
| 
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: 
56294 
diff
changeset
 | 
66  | 
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: 
39440 
diff
changeset
 | 
67  | 
fun report ctxt pos markup = report_text ctxt pos markup "";  | 
| 28405 | 68  | 
|
| 48767 | 69  | 
fun reports_text ctxt reps = if is_visible ctxt then Position.reports_text reps else ();  | 
| 44736 | 70  | 
fun reports ctxt reps = if is_visible ctxt then Position.reports reps else ();  | 
71  | 
||
| 28405 | 72  | 
end;  |