| author | wenzelm | 
| Thu, 20 Jul 2023 12:29:57 +0200 | |
| changeset 78416 | f26eba6281b1 | 
| parent 76680 | e95b9c9e17ff | 
| child 80462 | 7a1f9e571046 | 
| permissions | -rw-r--r-- | 
| 60749 | 1  | 
/* Title: Pure/Tools/debugger.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Interactive debugger for Isabelle/ML.  | 
|
5  | 
*/  | 
|
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
9  | 
||
| 61016 | 10  | 
import scala.collection.immutable.SortedMap  | 
11  | 
||
12  | 
||
| 75393 | 13  | 
object Debugger {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
14  | 
/* thread context */  | 
| 60834 | 15  | 
|
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
16  | 
case object Update  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
17  | 
|
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
18  | 
sealed case class Debug_State(pos: Position.T, function: String)  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
19  | 
type Threads = SortedMap[String, List[Debug_State]]  | 
| 60842 | 20  | 
|
| 75393 | 21  | 
  sealed case class Context(thread_name: String, debug_states: List[Debug_State], index: Int = 0) {
 | 
| 61010 | 22  | 
def size: Int = debug_states.length + 1  | 
23  | 
def reset: Context = copy(index = 0)  | 
|
| 71601 | 24  | 
def select(i: Int): Context = copy(index = i + 1)  | 
| 61010 | 25  | 
|
26  | 
def thread_state: Option[Debug_State] = debug_states.headOption  | 
|
27  | 
||
28  | 
def stack_state: Option[Debug_State] =  | 
|
| 76604 | 29  | 
if (1 <= index && index <= debug_states.length) Some(debug_states(index - 1))  | 
| 61010 | 30  | 
else None  | 
31  | 
||
| 61014 | 32  | 
def debug_index: Option[Int] =  | 
| 61010 | 33  | 
if (stack_state.isDefined) Some(index - 1)  | 
34  | 
else if (debug_states.nonEmpty) Some(0)  | 
|
35  | 
else None  | 
|
36  | 
||
37  | 
def debug_state: Option[Debug_State] = stack_state orElse thread_state  | 
|
| 61014 | 38  | 
def debug_position: Option[Position.T] = debug_state.map(_.pos)  | 
| 61010 | 39  | 
|
40  | 
override def toString: String =  | 
|
41  | 
      stack_state match {
 | 
|
42  | 
case None => thread_name  | 
|
43  | 
case Some(d) => d.function  | 
|
44  | 
}  | 
|
45  | 
}  | 
|
46  | 
||
47  | 
||
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
48  | 
/* debugger state */  | 
| 61016 | 49  | 
|
| 60835 | 50  | 
sealed case class State(  | 
| 76605 | 51  | 
active: Set[AnyRef] = Set.empty, // active views  | 
| 
60932
 
13ee73f57c85
allow to break running threads at next possible breakpoint (simplified version of former option, see f3039309702e);
 
wenzelm 
parents: 
60912 
diff
changeset
 | 
52  | 
break: Boolean = false, // break at next possible breakpoint  | 
| 
 
13ee73f57c85
allow to break running threads at next possible breakpoint (simplified version of former option, see f3039309702e);
 
wenzelm 
parents: 
60912 
diff
changeset
 | 
53  | 
active_breakpoints: Set[Long] = Set.empty, // explicit breakpoint state  | 
| 61016 | 54  | 
threads: Threads = SortedMap.empty, // thread name ~> stack of debug states  | 
| 61014 | 55  | 
focus: Map[String, Context] = Map.empty, // thread name ~> focus  | 
| 75393 | 56  | 
output: Map[String, Command.Results] = Map.empty // thread name ~> output messages  | 
57  | 
  ) {
 | 
|
| 
60932
 
13ee73f57c85
allow to break running threads at next possible breakpoint (simplified version of former option, see f3039309702e);
 
wenzelm 
parents: 
60912 
diff
changeset
 | 
58  | 
def set_break(b: Boolean): State = copy(break = b)  | 
| 
 
13ee73f57c85
allow to break running threads at next possible breakpoint (simplified version of former option, see f3039309702e);
 
wenzelm 
parents: 
60912 
diff
changeset
 | 
59  | 
|
| 76605 | 60  | 
def is_active: Boolean = active.nonEmpty  | 
61  | 
def register_active(id: AnyRef): State = copy(active = active + id)  | 
|
62  | 
def unregister_active(id: AnyRef): State = copy(active = active - id)  | 
|
| 60876 | 63  | 
|
| 75393 | 64  | 
    def toggle_breakpoint(breakpoint: Long): (Boolean, State) = {
 | 
| 60876 | 65  | 
val active_breakpoints1 =  | 
| 
60880
 
fa958e24ff24
set breakpoint state on ML side, relying on stable situation within the PIDE editing queue;
 
wenzelm 
parents: 
60878 
diff
changeset
 | 
66  | 
if (active_breakpoints(breakpoint)) active_breakpoints - breakpoint  | 
| 75446 | 67  | 
else active_breakpoints + breakpoint  | 
| 
60880
 
fa958e24ff24
set breakpoint state on ML side, relying on stable situation within the PIDE editing queue;
 
wenzelm 
parents: 
60878 
diff
changeset
 | 
68  | 
(active_breakpoints1(breakpoint), copy(active_breakpoints = active_breakpoints1))  | 
| 60876 | 69  | 
}  | 
70  | 
||
| 60842 | 71  | 
def get_thread(thread_name: String): List[Debug_State] =  | 
72  | 
threads.getOrElse(thread_name, Nil)  | 
|
73  | 
||
| 75393 | 74  | 
    def update_thread(thread_name: String, debug_states: List[Debug_State]): State = {
 | 
| 
61019
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
75  | 
val threads1 =  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
76  | 
if (debug_states.nonEmpty) threads + (thread_name -> debug_states)  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
77  | 
else threads - thread_name  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
78  | 
val focus1 =  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
79  | 
        focus.get(thread_name) match {
 | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
80  | 
case Some(c) if debug_states.nonEmpty =>  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
81  | 
focus + (thread_name -> Context(thread_name, debug_states))  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
82  | 
case _ => focus - thread_name  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
83  | 
}  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
84  | 
copy(threads = threads1, focus = focus1)  | 
| 
 
7ce030f14aa9
reset focus after thread update (with new debug_states);
 
wenzelm 
parents: 
61018 
diff
changeset
 | 
85  | 
}  | 
| 60842 | 86  | 
|
| 61014 | 87  | 
def set_focus(c: Context): State = copy(focus = focus + (c.thread_name -> c))  | 
88  | 
||
| 60842 | 89  | 
def get_output(thread_name: String): Command.Results =  | 
90  | 
output.getOrElse(thread_name, Command.Results.empty)  | 
|
91  | 
||
92  | 
def add_output(thread_name: String, entry: Command.Results.Entry): State =  | 
|
93  | 
copy(output = output + (thread_name -> (get_output(thread_name) + entry)))  | 
|
94  | 
||
95  | 
def clear_output(thread_name: String): State =  | 
|
96  | 
copy(output = output - thread_name)  | 
|
| 60834 | 97  | 
}  | 
98  | 
||
| 60902 | 99  | 
|
100  | 
/* protocol handler */  | 
|
101  | 
||
| 75393 | 102  | 
  class Handler(session: Session) extends Session.Protocol_Handler {
 | 
| 
65222
 
fb8253564483
more robust debugger initialization, e.g. required for GUI components before actual session startup;
 
wenzelm 
parents: 
65219 
diff
changeset
 | 
103  | 
val debugger: Debugger = new Debugger(session)  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
104  | 
|
| 75393 | 105  | 
    private def debugger_state(msg: Prover.Protocol_Output): Boolean = {
 | 
| 60842 | 106  | 
      msg.properties match {
 | 
107  | 
case Markup.Debugger_State(thread_name) =>  | 
|
| 73563 | 108  | 
val msg_body = Symbol.decode_yxml_failsafe(msg.text)  | 
| 75393 | 109  | 
          val debug_states = {
 | 
| 60842 | 110  | 
import XML.Decode._  | 
| 
65344
 
b99283eed13c
clarified YXML vs. symbol encoding: operate on whole message;
 
wenzelm 
parents: 
65223 
diff
changeset
 | 
111  | 
            list(pair(properties, string))(msg_body).map({
 | 
| 60842 | 112  | 
case (pos, function) => Debug_State(pos, function)  | 
113  | 
})  | 
|
114  | 
}  | 
|
| 
65222
 
fb8253564483
more robust debugger initialization, e.g. required for GUI components before actual session startup;
 
wenzelm 
parents: 
65219 
diff
changeset
 | 
115  | 
debugger.update_thread(thread_name, debug_states)  | 
| 60842 | 116  | 
true  | 
117  | 
case _ => false  | 
|
118  | 
}  | 
|
119  | 
}  | 
|
120  | 
||
| 75393 | 121  | 
    private def debugger_output(msg: Prover.Protocol_Output): Boolean = {
 | 
| 60830 | 122  | 
      msg.properties match {
 | 
| 60835 | 123  | 
case Markup.Debugger_Output(thread_name) =>  | 
| 73563 | 124  | 
          Symbol.decode_yxml_failsafe(msg.text) match {
 | 
| 60834 | 125  | 
case List(XML.Elem(Markup(name, props @ Markup.Serial(i)), body)) =>  | 
| 76680 | 126  | 
val message = Protocol.make_message(body, name, props = props)  | 
| 
73031
 
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
 
wenzelm 
parents: 
72212 
diff
changeset
 | 
127  | 
debugger.add_output(thread_name, i -> session.cache.elem(message))  | 
| 60834 | 128  | 
true  | 
129  | 
case _ => false  | 
|
130  | 
}  | 
|
| 60830 | 131  | 
case _ => false  | 
132  | 
}  | 
|
133  | 
}  | 
|
134  | 
||
| 75440 | 135  | 
override val functions: Session.Protocol_Functions =  | 
| 65219 | 136  | 
List(  | 
| 71601 | 137  | 
Markup.DEBUGGER_STATE -> debugger_state,  | 
138  | 
Markup.DEBUGGER_OUTPUT -> debugger_output)  | 
|
| 60749 | 139  | 
}  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
140  | 
}  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
141  | 
|
| 75393 | 142  | 
class Debugger private(session: Session) {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
143  | 
/* debugger state */  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
144  | 
|
| 
65223
 
844c067bc3d4
more robust startup, despite remaining race condition of debugger.is_active vs. session.is_ready;
 
wenzelm 
parents: 
65222 
diff
changeset
 | 
145  | 
private val state = Synchronized(Debugger.State())  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
146  | 
|
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
147  | 
private val delay_update =  | 
| 71704 | 148  | 
    Delay.first(session.output_delay) {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
149  | 
session.debugger_updates.post(Debugger.Update)  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
150  | 
}  | 
| 60765 | 151  | 
|
152  | 
||
153  | 
/* protocol commands */  | 
|
154  | 
||
| 75393 | 155  | 
  def update_thread(thread_name: String, debug_states: List[Debugger.Debug_State]): Unit = {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
156  | 
state.change(_.update_thread(thread_name, debug_states))  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
157  | 
delay_update.invoke()  | 
| 60910 | 158  | 
}  | 
159  | 
||
| 75393 | 160  | 
  def add_output(thread_name: String, entry: Command.Results.Entry): Unit = {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
161  | 
state.change(_.add_output(thread_name, entry))  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
162  | 
delay_update.invoke()  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
163  | 
}  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
164  | 
|
| 
65223
 
844c067bc3d4
more robust startup, despite remaining race condition of debugger.is_active vs. session.is_ready;
 
wenzelm 
parents: 
65222 
diff
changeset
 | 
165  | 
def is_active(): Boolean = session.is_ready && state.value.is_active  | 
| 
 
844c067bc3d4
more robust startup, despite remaining race condition of debugger.is_active vs. session.is_ready;
 
wenzelm 
parents: 
65222 
diff
changeset
 | 
166  | 
|
| 76604 | 167  | 
def ready(): Unit =  | 
168  | 
    if (is_active()) session.protocol_command("Debugger.init")
 | 
|
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
169  | 
|
| 76605 | 170  | 
def init(id: AnyRef): Unit =  | 
| 75394 | 171  | 
    state.change { st =>
 | 
| 76605 | 172  | 
val st1 = st.register_active(id)  | 
| 76604 | 173  | 
      if (session.is_ready && !st.is_active && st1.is_active) {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
174  | 
        session.protocol_command("Debugger.init")
 | 
| 76604 | 175  | 
}  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
176  | 
st1  | 
| 75394 | 177  | 
}  | 
| 60889 | 178  | 
|
| 76605 | 179  | 
def exit(id: AnyRef): Unit =  | 
| 75394 | 180  | 
    state.change { st =>
 | 
| 76605 | 181  | 
val st1 = st.unregister_active(id)  | 
| 76604 | 182  | 
      if (session.is_ready && st.is_active && !st1.is_active) {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
183  | 
        session.protocol_command("Debugger.exit")
 | 
| 76604 | 184  | 
}  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
185  | 
st1  | 
| 75394 | 186  | 
}  | 
| 60876 | 187  | 
|
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
188  | 
def is_break(): Boolean = state.value.break  | 
| 75393 | 189  | 
  def set_break(b: Boolean): Unit = {
 | 
| 75394 | 190  | 
    state.change { st =>
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
191  | 
val st1 = st.set_break(b)  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
192  | 
      session.protocol_command("Debugger.break", b.toString)
 | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
193  | 
st1  | 
| 75394 | 194  | 
}  | 
| 
60932
 
13ee73f57c85
allow to break running threads at next possible breakpoint (simplified version of former option, see f3039309702e);
 
wenzelm 
parents: 
60912 
diff
changeset
 | 
195  | 
delay_update.invoke()  | 
| 
 
13ee73f57c85
allow to break running threads at next possible breakpoint (simplified version of former option, see f3039309702e);
 
wenzelm 
parents: 
60912 
diff
changeset
 | 
196  | 
}  | 
| 
 
13ee73f57c85
allow to break running threads at next possible breakpoint (simplified version of former option, see f3039309702e);
 
wenzelm 
parents: 
60912 
diff
changeset
 | 
197  | 
|
| 75393 | 198  | 
  def active_breakpoint_state(breakpoint: Long): Option[Boolean] = {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
199  | 
val st = state.value  | 
| 
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
200  | 
if (st.is_active) Some(st.active_breakpoints(breakpoint)) else None  | 
| 60876 | 201  | 
}  | 
202  | 
||
| 60882 | 203  | 
def breakpoint_state(breakpoint: Long): Boolean =  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
204  | 
state.value.active_breakpoints(breakpoint)  | 
| 60882 | 205  | 
|
| 75393 | 206  | 
  def toggle_breakpoint(command: Command, breakpoint: Long): Unit = {
 | 
| 75394 | 207  | 
    state.change { st =>
 | 
| 75393 | 208  | 
val (breakpoint_state, st1) = st.toggle_breakpoint(breakpoint)  | 
209  | 
session.protocol_command(  | 
|
210  | 
"Debugger.breakpoint",  | 
|
211  | 
Symbol.encode(command.node_name.node),  | 
|
212  | 
Document_ID(command.id),  | 
|
213  | 
Value.Long(breakpoint),  | 
|
214  | 
Value.Boolean(breakpoint_state))  | 
|
215  | 
st1  | 
|
| 75394 | 216  | 
}  | 
| 
60878
 
1f0d2bbcf38b
added action to toggle breakpoints (on editor side);
 
wenzelm 
parents: 
60876 
diff
changeset
 | 
217  | 
}  | 
| 60876 | 218  | 
|
| 75393 | 219  | 
  def status(focus: Option[Debugger.Context]): (Debugger.Threads, List[XML.Tree]) = {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
220  | 
val st = state.value  | 
| 61018 | 221  | 
val output =  | 
222  | 
      focus match {
 | 
|
223  | 
case None => Nil  | 
|
224  | 
case Some(c) =>  | 
|
225  | 
          (for {
 | 
|
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
226  | 
(thread_name, results) <- st.output  | 
| 61018 | 227  | 
if thread_name == c.thread_name  | 
228  | 
(_, tree) <- results.iterator  | 
|
229  | 
} yield tree).toList  | 
|
230  | 
}  | 
|
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
231  | 
(st.threads, output)  | 
| 61018 | 232  | 
}  | 
| 61011 | 233  | 
|
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
234  | 
def focus(): List[Debugger.Context] = state.value.focus.toList.map(_._2)  | 
| 75393 | 235  | 
  def set_focus(c: Debugger.Context): Unit = {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
236  | 
state.change(_.set_focus(c))  | 
| 61008 | 237  | 
delay_update.invoke()  | 
238  | 
}  | 
|
| 60875 | 239  | 
|
| 60854 | 240  | 
def input(thread_name: String, msg: String*): Unit =  | 
| 73565 | 241  | 
    session.protocol_command_args("Debugger.input", thread_name :: msg.toList)
 | 
| 60854 | 242  | 
|
| 60899 | 243  | 
def continue(thread_name: String): Unit = input(thread_name, "continue")  | 
| 60869 | 244  | 
def step(thread_name: String): Unit = input(thread_name, "step")  | 
245  | 
def step_over(thread_name: String): Unit = input(thread_name, "step_over")  | 
|
246  | 
def step_out(thread_name: String): Unit = input(thread_name, "step_out")  | 
|
| 60856 | 247  | 
|
| 75393 | 248  | 
  def clear_output(thread_name: String): Unit = {
 | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
249  | 
state.change(_.clear_output(thread_name))  | 
| 60902 | 250  | 
delay_update.invoke()  | 
| 60901 | 251  | 
}  | 
252  | 
||
| 75393 | 253  | 
  def eval(c: Debugger.Context, SML: Boolean, context: String, expression: String): Unit = {
 | 
| 75394 | 254  | 
    state.change { st =>
 | 
| 61014 | 255  | 
input(c.thread_name, "eval", c.debug_index.getOrElse(0).toString,  | 
| 61010 | 256  | 
SML.toString, Symbol.encode(context), Symbol.encode(expression))  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
257  | 
st.clear_output(c.thread_name)  | 
| 75394 | 258  | 
}  | 
| 60902 | 259  | 
delay_update.invoke()  | 
| 60856 | 260  | 
}  | 
| 60897 | 261  | 
|
| 75393 | 262  | 
  def print_vals(c: Debugger.Context, SML: Boolean, context: String): Unit = {
 | 
| 61014 | 263  | 
require(c.debug_index.isDefined)  | 
| 61010 | 264  | 
|
| 75394 | 265  | 
    state.change { st =>
 | 
| 61014 | 266  | 
input(c.thread_name, "print_vals", c.debug_index.getOrElse(0).toString,  | 
| 61010 | 267  | 
SML.toString, Symbol.encode(context))  | 
| 
65213
 
51c0f094dc02
proper local debugger state, depending on session;
 
wenzelm 
parents: 
63805 
diff
changeset
 | 
268  | 
st.clear_output(c.thread_name)  | 
| 75394 | 269  | 
}  | 
| 60902 | 270  | 
delay_update.invoke()  | 
| 60897 | 271  | 
}  | 
| 60749 | 272  | 
}  |