| author | wenzelm | 
| Fri, 30 Dec 2011 17:45:13 +0100 | |
| changeset 46058 | 9a790f4a72be | 
| parent 44736 | c2a3f1c84179 | 
| child 46874 | 993c413746f4 | 
| 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 | |
| 9 | val is_visible: Proof.context -> bool | |
| 10 | val set_visible: bool -> Proof.context -> Proof.context | |
| 11 | val restore_visible: Proof.context -> Proof.context -> Proof.context | |
| 38831 
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
 wenzelm parents: 
38481diff
changeset | 12 |   val if_visible: Proof.context -> ('a -> unit) -> 'a -> unit
 | 
| 41470 
890b25753bf7
added Context_Position.is_visible_proof, which treats a global theory as invisible (unlike the default);
 wenzelm parents: 
39508diff
changeset | 13 |   val if_visible_proof: Context.generic -> ('a -> unit) -> 'a -> unit
 | 
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 14 | val reported_text: Proof.context -> Position.T -> Markup.T -> string -> string | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 15 | val report_text: Proof.context -> Position.T -> Markup.T -> string -> unit | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 16 | val report: Proof.context -> Position.T -> Markup.T -> unit | 
| 44736 | 17 | val reports: Proof.context -> Position.report list -> unit | 
| 28405 | 18 | end; | 
| 19 | ||
| 33383 | 20 | structure Context_Position: CONTEXT_POSITION = | 
| 28405 | 21 | struct | 
| 22 | ||
| 39508 
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
 wenzelm parents: 
39507diff
changeset | 23 | val visible = Config.bool (Config.declare "Context_Position.visible" (K (Config.Bool true))); | 
| 
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
 wenzelm parents: 
39507diff
changeset | 24 | fun is_visible ctxt = Config.get ctxt visible; | 
| 
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
 wenzelm parents: 
39507diff
changeset | 25 | val set_visible = Config.put visible; | 
| 28405 | 26 | val restore_visible = set_visible o is_visible; | 
| 27 | ||
| 38831 
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
 wenzelm parents: 
38481diff
changeset | 28 | fun if_visible ctxt f x = if is_visible ctxt then f x else (); | 
| 
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
 wenzelm parents: 
38481diff
changeset | 29 | |
| 41470 
890b25753bf7
added Context_Position.is_visible_proof, which treats a global theory as invisible (unlike the default);
 wenzelm parents: 
39508diff
changeset | 30 | fun if_visible_proof (Context.Proof ctxt) f x = if_visible ctxt f x | 
| 
890b25753bf7
added Context_Position.is_visible_proof, which treats a global theory as invisible (unlike the default);
 wenzelm parents: 
39508diff
changeset | 31 | | if_visible_proof _ _ _ = (); | 
| 
890b25753bf7
added Context_Position.is_visible_proof, which treats a global theory as invisible (unlike the default);
 wenzelm parents: 
39508diff
changeset | 32 | |
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 33 | fun reported_text ctxt pos markup txt = | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 34 | if is_visible ctxt then Position.reported_text pos markup txt else ""; | 
| 38481 
81ec258c4cd3
Output_Position.report_text -- markup with potential "arguments";
 wenzelm parents: 
38237diff
changeset | 35 | |
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 36 | 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: 
39440diff
changeset | 37 | fun report ctxt pos markup = report_text ctxt pos markup ""; | 
| 28405 | 38 | |
| 44736 | 39 | fun reports ctxt reps = if is_visible ctxt then Position.reports reps else (); | 
| 40 | ||
| 28405 | 41 | end; |