equal
deleted
inserted
replaced
13 val if_visible_global: theory -> ('a -> unit) -> 'a -> unit |
13 val if_visible_global: theory -> ('a -> unit) -> 'a -> unit |
14 val set_visible: bool -> Proof.context -> Proof.context |
14 val set_visible: bool -> Proof.context -> Proof.context |
15 val set_visible_global: bool -> theory -> theory |
15 val set_visible_global: bool -> theory -> theory |
16 val restore_visible: Proof.context -> Proof.context -> Proof.context |
16 val restore_visible: Proof.context -> Proof.context -> Proof.context |
17 val restore_visible_global: theory -> theory -> theory |
17 val restore_visible_global: theory -> theory -> theory |
|
18 val is_reported_generic: Context.generic -> Position.T -> bool |
|
19 val is_reported: Proof.context -> Position.T -> bool |
18 val report_generic: Context.generic -> Position.T -> Markup.T -> unit |
20 val report_generic: Context.generic -> Position.T -> Markup.T -> unit |
19 val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string |
21 val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string |
20 val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit |
22 val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit |
21 val report: Proof.context -> Position.T -> Markup.T -> unit |
23 val report: Proof.context -> Position.T -> Markup.T -> unit |
22 val reports_text: Proof.context -> Position.report_text list -> unit |
24 val reports_text: Proof.context -> Position.report_text list -> unit |
45 val set_visible_global = Context.theory_map o Data.put o SOME; |
47 val set_visible_global = Context.theory_map o Data.put o SOME; |
46 |
48 |
47 val restore_visible = set_visible o is_visible; |
49 val restore_visible = set_visible o is_visible; |
48 val restore_visible_global = set_visible_global o is_visible_global; |
50 val restore_visible_global = set_visible_global o is_visible_global; |
49 |
51 |
|
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 |
50 fun report_generic context pos markup = |
55 fun report_generic context pos markup = |
51 if is_visible_generic context then |
56 if is_reported_generic context pos then |
52 Output.report (Position.reported_text pos markup "") |
57 Output.report (Position.reported_text pos markup "") |
53 else (); |
58 else (); |
54 |
59 |
55 fun reported_text ctxt pos markup txt = |
60 fun reported_text ctxt pos markup txt = |
56 if is_visible ctxt then Position.reported_text pos markup txt else ""; |
61 if is_reported ctxt pos then Position.reported_text pos markup txt else ""; |
57 |
62 |
58 fun report_text ctxt pos markup txt = Output.report (reported_text ctxt pos markup txt); |
63 fun report_text ctxt pos markup txt = Output.report (reported_text ctxt pos markup txt); |
59 fun report ctxt pos markup = report_text ctxt pos markup ""; |
64 fun report ctxt pos markup = report_text ctxt pos markup ""; |
60 |
65 |
61 fun reports_text ctxt reps = if is_visible ctxt then Position.reports_text reps else (); |
66 fun reports_text ctxt reps = if is_visible ctxt then Position.reports_text reps else (); |