src/Pure/context_position.ML
author haftmann
Sun, 18 Jan 2009 21:12:06 +0100
changeset 29554 7e5e5ebb7bf7
parent 28409 a1feb819b0a2
child 29606 fedb8be05f24
permissions -rw-r--r--
added churn script

(*  Title:      Pure/context_position.ML
    ID:         $Id$
    Author:     Makarius

Context position visibility flag.
*)

signature CONTEXT_POSITION =
sig
  val is_visible: Proof.context -> bool
  val set_visible: bool -> Proof.context -> Proof.context
  val restore_visible: Proof.context -> Proof.context -> Proof.context
  val report_visible: Proof.context -> Markup.T -> Position.T -> unit
end;

structure ContextPosition: CONTEXT_POSITION =
struct

structure Data = ProofDataFun
(
  type T = bool;
  fun init _ = true;
);

val is_visible = Data.get;
val set_visible = Data.put;
val restore_visible = set_visible o is_visible;

fun report_visible ctxt markup pos =
  if is_visible ctxt then Position.report markup pos else ();

end;