author | wenzelm |
Wed, 05 Mar 2014 19:52:28 +0100 | |
changeset 55923 | 4bdae9403baf |
parent 55672 | 5e25cc741ab9 |
child 56294 | 85911b8a6868 |
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 |
12 |
val if_visible: Proof.context -> ('a -> unit) -> 'a -> unit |
|
13 |
val if_visible_global: theory -> ('a -> unit) -> 'a -> unit |
|
28405 | 14 |
val set_visible: bool -> Proof.context -> Proof.context |
52700 | 15 |
val set_visible_global: bool -> theory -> theory |
28405 | 16 |
val restore_visible: Proof.context -> Proof.context -> Proof.context |
52700 | 17 |
val restore_visible_global: theory -> theory -> theory |
55923 | 18 |
val is_reported_generic: Context.generic -> Position.T -> bool |
19 |
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
|
20 |
val report_generic: Context.generic -> Position.T -> Markup.T -> unit |
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
21 |
val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string |
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
22 |
val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit |
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
23 |
val report: Proof.context -> Position.T -> Markup.T -> unit |
48767 | 24 |
val reports_text: Proof.context -> Position.report_text list -> unit |
44736 | 25 |
val reports: Proof.context -> Position.report list -> unit |
28405 | 26 |
end; |
27 |
||
33383 | 28 |
structure Context_Position: CONTEXT_POSITION = |
28405 | 29 |
struct |
30 |
||
47813 | 31 |
structure Data = Generic_Data |
32 |
( |
|
33 |
type T = bool option; |
|
34 |
val empty: T = NONE; |
|
35 |
val extend = I; |
|
36 |
fun merge (x, y): T = if is_some x then x else y; |
|
37 |
); |
|
38 |
||
39 |
val is_visible_generic = the_default true o Data.get; |
|
40 |
val is_visible = is_visible_generic o Context.Proof; |
|
52700 | 41 |
val is_visible_global = is_visible_generic o Context.Theory; |
28405 | 42 |
|
38831
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents:
38481
diff
changeset
|
43 |
fun if_visible ctxt f x = if is_visible ctxt then f x else (); |
52700 | 44 |
fun if_visible_global thy f x = if is_visible_global thy then f x else (); |
45 |
||
46 |
val set_visible = Context.proof_map o Data.put o SOME; |
|
47 |
val set_visible_global = Context.theory_map o Data.put o SOME; |
|
48 |
||
49 |
val restore_visible = set_visible o is_visible; |
|
50 |
val restore_visible_global = set_visible_global o is_visible_global; |
|
38831
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents:
38481
diff
changeset
|
51 |
|
55923 | 52 |
fun is_reported_generic context pos = is_visible_generic context andalso Position.is_reported pos; |
53 |
fun is_reported ctxt pos = is_visible ctxt andalso Position.is_reported pos; |
|
54 |
||
47005
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46874
diff
changeset
|
55 |
fun report_generic context pos markup = |
55923 | 56 |
if is_reported_generic context pos then |
47005
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46874
diff
changeset
|
57 |
Output.report (Position.reported_text pos markup "") |
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46874
diff
changeset
|
58 |
else (); |
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46874
diff
changeset
|
59 |
|
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
60 |
fun reported_text ctxt pos markup txt = |
55923 | 61 |
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
|
62 |
|
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
63 |
fun report_text ctxt pos markup txt = Output.report (reported_text ctxt pos markup txt); |
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
64 |
fun report ctxt pos markup = report_text ctxt pos markup ""; |
28405 | 65 |
|
48767 | 66 |
fun reports_text ctxt reps = if is_visible ctxt then Position.reports_text reps else (); |
44736 | 67 |
fun reports ctxt reps = if is_visible ctxt then Position.reports reps else (); |
68 |
||
28405 | 69 |
end; |