src/Pure/Tools/simplifier_trace.ML
author wenzelm
Mon, 25 Oct 2021 17:26:27 +0200
changeset 74576 0b43d42cfde7
parent 74561 8e6c973003c8
child 78725 3c02ad5a1586
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Tools/simplifier_trace.ML
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
     2
    Author:     Lars Hupel
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
     3
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
     4
Interactive Simplifier trace.
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
     5
*)
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
     6
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
     7
signature SIMPLIFIER_TRACE =
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
     8
sig
62983
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
     9
  val disable: Proof.context -> Proof.context
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    10
  val add_term_breakpoint: term -> Context.generic -> Context.generic
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    11
  val add_thm_breakpoint: thm -> Context.generic -> Context.generic
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
    12
end
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
    13
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
    14
structure Simplifier_Trace: SIMPLIFIER_TRACE =
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
    15
struct
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
    16
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
    17
(** context data **)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    18
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    19
datatype mode = Disabled | Normal | Full
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    20
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    21
fun merge_modes Disabled m = m
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    22
  | merge_modes Normal Full = Full
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    23
  | merge_modes Normal _ = Normal
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    24
  | merge_modes Full _ = Full
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    25
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    26
val empty_breakpoints =
55560
7ac8f013321c proper term equality;
wenzelm
parents: 55553
diff changeset
    27
  (Item_Net.init (op aconv) single,
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    28
   Item_Net.init eq_rrule (single o Thm.full_prop_of o #thm))
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    29
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    30
fun merge_breakpoints ((term_bs1, thm_bs1), (term_bs2, thm_bs2)) =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    31
  (Item_Net.merge (term_bs1, term_bs2),
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    32
   Item_Net.merge (thm_bs1, thm_bs2))
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    33
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    34
structure Data = Generic_Data
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    35
(
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    36
  type T =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    37
    {max_depth: int,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    38
     mode: mode,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    39
     interactive: bool,
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    40
     memory: bool,
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    41
     parent: int,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    42
     breakpoints: term Item_Net.T * rrule Item_Net.T}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    43
  val empty =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    44
    {max_depth = 10,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    45
     mode = Disabled,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    46
     interactive = false,
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    47
     memory = true,
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    48
     parent = 0,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    49
     breakpoints = empty_breakpoints}
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
    50
  fun merge
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
    51
    ({max_depth = max_depth1, mode = mode1, interactive = interactive1,
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
    52
      memory = memory1, breakpoints = breakpoints1, ...}: T,
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
    53
     {max_depth = max_depth2, mode = mode2, interactive = interactive2,
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
    54
      memory = memory2, breakpoints = breakpoints2, ...}: T) =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    55
    {max_depth = Int.max (max_depth1, max_depth2),
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    56
     mode = merge_modes mode1 mode2,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    57
     interactive = interactive1 orelse interactive2,
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    58
     memory = memory1 andalso memory2,
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    59
     parent = 0,
55335
8192d3acadbe made SML/NJ happy
Lars Hupel <lars.hupel@mytum.de>
parents: 55316
diff changeset
    60
     breakpoints = merge_breakpoints (breakpoints1, breakpoints2)}: T
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    61
)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    62
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    63
val get_data = Data.get o Context.Proof
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    64
val put_data = Context.proof_map o Data.put
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    65
62983
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
    66
val disable =
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
    67
  Config.put simp_trace false #>
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
    68
  (Context.proof_map o Data.map)
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
    69
    (fn {max_depth, mode = _, interactive, parent, memory, breakpoints} =>
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
    70
      {max_depth = max_depth, mode = Disabled, interactive = interactive, parent = parent,
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
    71
        memory = memory, breakpoints = breakpoints});
ba9072b303a2 avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
wenzelm
parents: 62505
diff changeset
    72
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    73
val get_breakpoints = #breakpoints o get_data
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    74
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    75
fun map_breakpoints f =
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    76
  Data.map
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    77
    (fn {max_depth, mode, interactive, parent, memory, breakpoints} =>
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    78
      {max_depth = max_depth,
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    79
       mode = mode,
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    80
       interactive = interactive,
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    81
       memory = memory,
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
    82
       parent = parent,
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    83
       breakpoints = f breakpoints})
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    84
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    85
fun add_term_breakpoint term =
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    86
  map_breakpoints (apfst (Item_Net.update term))
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    87
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    88
fun add_thm_breakpoint thm context =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    89
  let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    90
    val rrules = mk_rrules (Context.proof_of context) [thm]
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    91
  in
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    92
    map_breakpoints (apsnd (fold Item_Net.update rrules)) context
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    93
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    94
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    95
fun check_breakpoint (term, rrule) ctxt =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    96
  let
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    97
    val thy = Proof_Context.theory_of ctxt
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
    98
    val (term_bs, thm_bs) = get_breakpoints ctxt
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
    99
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   100
    val term_matches =
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   101
      filter (fn pat => Pattern.matches thy (pat, term))
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   102
        (Item_Net.retrieve_matching term_bs term)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   103
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   104
    val thm_matches =
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   105
      exists (eq_rrule o pair rrule)
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   106
        (Item_Net.retrieve_matching thm_bs (Thm.full_prop_of (#thm rrule)))
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   107
  in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   108
    (term_matches, thm_matches)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   109
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   110
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   111
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   112
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   113
(** config and attributes **)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   114
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   115
fun config raw_mode interactive max_depth memory =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   116
  let
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   117
    val mode =
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   118
      (case raw_mode of
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   119
        "normal" => Normal
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   120
      | "full" => Full
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   121
      | _ => error ("Simplifier_Trace.config: unknown mode " ^ raw_mode))
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   122
55946
Lars Hupel <lars.hupel@mytum.de>
parents: 55611
diff changeset
   123
    val update = Data.map (fn {parent, breakpoints, ...} =>
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   124
      {max_depth = max_depth,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   125
       mode = mode,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   126
       interactive = interactive,
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   127
       memory = memory,
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   128
       parent = parent,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   129
       breakpoints = breakpoints})
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   130
  in Thm.declaration_attribute (K update) end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   131
57041
aceaca232177 consolidate "break_thm" and "break_term" attributes into "simp_break";
Lars Hupel <lars.hupel@mytum.de>
parents: 56333
diff changeset
   132
fun breakpoint terms =
aceaca232177 consolidate "break_thm" and "break_term" attributes into "simp_break";
Lars Hupel <lars.hupel@mytum.de>
parents: 56333
diff changeset
   133
  Thm.declaration_attribute (fn thm => add_thm_breakpoint thm o fold add_term_breakpoint terms)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   134
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   135
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   136
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   137
(** tracing state **)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   138
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   139
val futures =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   140
  Synchronized.var "Simplifier_Trace.futures" (Inttab.empty: string future Inttab.table)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   141
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   142
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   143
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   144
(** markup **)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   145
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   146
fun output_result (id, data) =
56333
38f1422ef473 support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents: 55946
diff changeset
   147
  Output.result (Markup.serial_properties id) [data]
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   148
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   149
val parentN = "parent"
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   150
val textN = "text"
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   151
val memoryN = "memory"
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   152
val successN = "success"
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   153
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   154
type payload =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   155
  {props: Properties.T,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   156
   pretty: Pretty.T}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   157
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   158
fun empty_payload () : payload =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   159
  {props = [], pretty = Pretty.str ""}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   160
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   161
fun mk_generic_result markup text triggered (payload : unit -> payload) ctxt =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   162
  let
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   163
    val {mode, interactive, memory, parent, ...} = get_data ctxt
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   164
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   165
    val eligible =
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   166
      (case mode of
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   167
        Disabled => false
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   168
      | Normal => triggered
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   169
      | Full => true)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   170
55553
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   171
    val markup' =
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   172
      if markup = Markup.simp_trace_stepN andalso not interactive
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   173
      then Markup.simp_trace_logN
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   174
      else markup
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   175
  in
55946
Lars Hupel <lars.hupel@mytum.de>
parents: 55611
diff changeset
   176
    if not eligible then NONE
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   177
    else
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   178
      let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   179
        val {props = more_props, pretty} = payload ()
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   180
        val props =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   181
          [(textN, text),
63806
c54a53ef1873 clarified modules;
wenzelm
parents: 62983
diff changeset
   182
           (memoryN, Value.print_bool memory),
c54a53ef1873 clarified modules;
wenzelm
parents: 62983
diff changeset
   183
           (parentN, Value.print_int parent)]
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   184
        val data =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   185
          Pretty.string_of (Pretty.markup (markup', props @ more_props) [pretty])
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   186
      in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   187
        SOME (serial (), data)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   188
      end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   189
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   190
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   191
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   192
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   193
(** tracing output **)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   194
57594
037f3b251df5 regular message to refer to Simplifier Trace panel (unused);
wenzelm
parents: 57592
diff changeset
   195
fun see_panel () =
037f3b251df5 regular message to refer to Simplifier Trace panel (unused);
wenzelm
parents: 57592
diff changeset
   196
  let
037f3b251df5 regular message to refer to Simplifier Trace panel (unused);
wenzelm
parents: 57592
diff changeset
   197
    val ((bg1, bg2), en) =
037f3b251df5 regular message to refer to Simplifier Trace panel (unused);
wenzelm
parents: 57592
diff changeset
   198
      YXML.output_markup_elem
037f3b251df5 regular message to refer to Simplifier Trace panel (unused);
wenzelm
parents: 57592
diff changeset
   199
        (Active.make_markup Markup.simp_trace_panelN {implicit = false, properties = []})
57596
3a1b1bda702f refer to Simplifier Trace panel on first invocation;
wenzelm
parents: 57594
diff changeset
   200
  in "See " ^ bg1 ^ bg2 ^ "simplifier trace" ^ en end
57594
037f3b251df5 regular message to refer to Simplifier Trace panel (unused);
wenzelm
parents: 57592
diff changeset
   201
037f3b251df5 regular message to refer to Simplifier Trace panel (unused);
wenzelm
parents: 57592
diff changeset
   202
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   203
fun send_request (result_id, content) =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   204
  let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   205
    fun break () =
56333
38f1422ef473 support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents: 55946
diff changeset
   206
      (Output.protocol_message (Markup.simp_trace_cancel result_id) [];
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   207
       Synchronized.change futures (Inttab.delete_safe result_id))
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   208
    val promise = Future.promise break : string future
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   209
  in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   210
    Synchronized.change futures (Inttab.update_new (result_id, promise));
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   211
    output_result (result_id, content);
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   212
    promise
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   213
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   214
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   215
55335
8192d3acadbe made SML/NJ happy
Lars Hupel <lars.hupel@mytum.de>
parents: 55316
diff changeset
   216
type data = {term: term, thm: thm, unconditional: bool, ctxt: Proof.context, rrule: rrule}
8192d3acadbe made SML/NJ happy
Lars Hupel <lars.hupel@mytum.de>
parents: 55316
diff changeset
   217
8192d3acadbe made SML/NJ happy
Lars Hupel <lars.hupel@mytum.de>
parents: 55316
diff changeset
   218
fun step ({term, thm, unconditional, ctxt, rrule}: data) =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   219
  let
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   220
    val (matching_terms, thm_triggered) = check_breakpoint (term, rrule) ctxt
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   221
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   222
    val {name, ...} = rrule
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   223
    val term_triggered = not (null matching_terms)
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
   224
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   225
    val text =
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   226
      if unconditional then "Apply rewrite rule?"
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   227
      else "Apply conditional rewrite rule?"
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   228
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   229
    fun payload () =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   230
      let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   231
        (* FIXME pretty printing via Proof_Context.pretty_fact *)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   232
        val pretty_thm = Pretty.block
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   233
          [Pretty.str ("Instance of " ^ name ^ ":"),
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   234
           Pretty.brk 1,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   235
           Syntax.pretty_term ctxt (Thm.prop_of thm)]
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   236
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   237
        val pretty_term = Pretty.block
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   238
          [Pretty.str "Trying to rewrite:",
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   239
           Pretty.brk 1,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   240
           Syntax.pretty_term ctxt term]
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   241
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   242
        val pretty_matchings =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   243
          let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   244
            val items = map (Pretty.item o single o Syntax.pretty_term ctxt) matching_terms
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   245
          in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   246
            if not (null matching_terms) then
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   247
              [Pretty.block (Pretty.fbreaks (Pretty.str "Matching terms:" :: items))]
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   248
            else []
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   249
          end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   250
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   251
        val pretty =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   252
          Pretty.chunks ([pretty_thm, pretty_term] @ pretty_matchings)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   253
      in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   254
        {props = [], pretty = pretty}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   255
      end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   256
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   257
    val {max_depth, mode, interactive, memory, breakpoints, ...} = get_data ctxt
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   258
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   259
    fun mk_promise result =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   260
      let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   261
        val result_id = #1 result
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   262
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   263
        fun put mode' interactive' = put_data
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   264
          {max_depth = max_depth,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   265
           mode = mode',
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   266
           interactive = interactive',
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   267
           memory = memory,
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   268
           parent = result_id,
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   269
           breakpoints = breakpoints} ctxt
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   270
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   271
        fun to_response "skip" = NONE
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   272
          | to_response "continue" = SOME (put mode true)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   273
          | to_response "continue_trace" = SOME (put Full true)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   274
          | to_response "continue_passive" = SOME (put mode false)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   275
          | to_response "continue_disable" = SOME (put Disabled false)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   276
          | to_response _ = raise Fail "Simplifier_Trace.step: invalid message"
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   277
      in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   278
        if not interactive then
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   279
          (output_result result; Future.value (SOME (put mode false)))
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   280
        else Future.map to_response (send_request result)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   281
      end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   282
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   283
  in
55553
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   284
    (case mk_generic_result Markup.simp_trace_stepN text
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   285
        (thm_triggered orelse term_triggered) payload ctxt of
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   286
      NONE => Future.value (SOME ctxt)
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   287
    | SOME res => mk_promise res)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   288
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   289
55946
Lars Hupel <lars.hupel@mytum.de>
parents: 55611
diff changeset
   290
fun recurse text depth term ctxt =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   291
  let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   292
    fun payload () =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   293
      {props = [],
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   294
       pretty = Syntax.pretty_term ctxt term}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   295
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   296
    val {max_depth, mode, interactive, memory, breakpoints, ...} = get_data ctxt
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   297
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   298
    fun put result_id = put_data
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   299
      {max_depth = max_depth,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   300
       mode = if depth >= max_depth then Disabled else mode,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   301
       interactive = interactive,
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   302
       memory = memory,
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   303
       parent = result_id,
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   304
       breakpoints = breakpoints} ctxt
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   305
  in
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   306
    (case mk_generic_result Markup.simp_trace_recurseN text true payload ctxt of
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   307
      NONE => put 0
57596
3a1b1bda702f refer to Simplifier Trace panel on first invocation;
wenzelm
parents: 57594
diff changeset
   308
    | SOME res =>
3a1b1bda702f refer to Simplifier Trace panel on first invocation;
wenzelm
parents: 57594
diff changeset
   309
       (if depth = 1 then writeln (see_panel ()) else ();
3a1b1bda702f refer to Simplifier Trace panel on first invocation;
wenzelm
parents: 57594
diff changeset
   310
        output_result res;
3a1b1bda702f refer to Simplifier Trace panel on first invocation;
wenzelm
parents: 57594
diff changeset
   311
        put (#1 res)))
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   312
  end
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   313
55335
8192d3acadbe made SML/NJ happy
Lars Hupel <lars.hupel@mytum.de>
parents: 55316
diff changeset
   314
fun indicate_failure ({term, ctxt, thm, rrule, ...}: data) ctxt' =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   315
  let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   316
    fun payload () =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   317
      let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   318
        val {name, ...} = rrule
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   319
        val pretty_thm =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   320
          (* FIXME pretty printing via Proof_Context.pretty_fact *)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   321
          Pretty.block
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   322
            [Pretty.str ("In an instance of " ^ name ^ ":"),
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   323
             Pretty.brk 1,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   324
             Syntax.pretty_term ctxt (Thm.prop_of thm)]
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   325
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   326
        val pretty_term =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   327
          Pretty.block
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   328
            [Pretty.str "Was trying to rewrite:",
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   329
             Pretty.brk 1,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   330
             Syntax.pretty_term ctxt term]
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   331
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   332
        val pretty =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   333
          Pretty.chunks [pretty_thm, pretty_term]
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   334
      in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   335
        {props = [(successN, "false")], pretty = pretty}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   336
      end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   337
57592
b73d74d0e428 misc tuning and simplification;
wenzelm
parents: 57591
diff changeset
   338
    val {interactive, ...} = get_data ctxt
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   339
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   340
    fun mk_promise result =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   341
      let
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   342
        fun to_response "exit" = false
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   343
          | to_response "redo" =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   344
              (Option.app output_result
55553
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   345
                (mk_generic_result Markup.simp_trace_ignoreN "Ignore" true empty_payload ctxt');
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   346
               true)
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   347
          | to_response _ = raise Fail "Simplifier_Trace.indicate_failure: invalid message"
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   348
      in
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   349
        if not interactive then
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   350
          (output_result result; Future.value false)
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   351
        else Future.map to_response (send_request result)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   352
      end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   353
  in
55553
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   354
    (case mk_generic_result Markup.simp_trace_hintN "Step failed" true payload ctxt' of
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   355
      NONE => Future.value false
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   356
    | SOME res => mk_promise res)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   357
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   358
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   359
fun indicate_success thm ctxt =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   360
  let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   361
    fun payload () =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   362
      {props = [(successN, "true")],
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   363
       pretty = Syntax.pretty_term ctxt (Thm.prop_of thm)}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   364
  in
55553
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   365
    Option.app output_result
99409ccbe04a more standard names for protocol and markup elements;
wenzelm
parents: 55552
diff changeset
   366
      (mk_generic_result Markup.simp_trace_hintN "Successfully rewrote" true payload ctxt)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   367
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   368
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   369
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   370
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   371
(** setup **)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   372
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   373
fun simp_apply args ctxt cont =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   374
  let
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   375
    val {unconditional: bool, term: term, thm: thm, rrule: rrule} = args
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   376
    val data =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   377
      {term = term,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   378
       unconditional = unconditional,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   379
       ctxt = ctxt,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   380
       thm = thm,
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   381
       rrule = rrule}
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   382
  in
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   383
    (case Future.join (step data) of
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   384
      NONE => NONE
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   385
    | SOME ctxt' =>
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   386
        let val res = cont ctxt' in
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   387
          (case res of
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   388
            NONE =>
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   389
              if Future.join (indicate_failure data ctxt') then
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   390
                simp_apply args ctxt cont
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   391
              else NONE
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   392
          | SOME (thm, _) => (indicate_success thm ctxt'; res))
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   393
        end)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   394
  end
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   395
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
   396
val _ = Theory.setup
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54730
diff changeset
   397
  (Simplifier.set_trace_ops
55946
Lars Hupel <lars.hupel@mytum.de>
parents: 55611
diff changeset
   398
    {trace_invoke = fn {depth, term} => recurse "Simplifier invoked" depth term,
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   399
     trace_apply = simp_apply})
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   400
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   401
val _ =
73225
3ab0cedaccad clarified modules: allow early definition of protocol commands;
wenzelm
parents: 72156
diff changeset
   402
  Protocol_Command.define "Simplifier_Trace.reply"
60747
wenzelm
parents: 57596
diff changeset
   403
    (fn [serial_string, reply] =>
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   404
      let
63806
c54a53ef1873 clarified modules;
wenzelm
parents: 62983
diff changeset
   405
        val serial = Value.parse_int serial_string
60747
wenzelm
parents: 57596
diff changeset
   406
        val result =
wenzelm
parents: 57596
diff changeset
   407
          Synchronized.change_result futures
wenzelm
parents: 57596
diff changeset
   408
            (fn tab => (Inttab.lookup tab serial, Inttab.delete_safe serial tab))
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   409
      in
60747
wenzelm
parents: 57596
diff changeset
   410
        (case result of
wenzelm
parents: 57596
diff changeset
   411
          SOME promise => Future.fulfill promise reply
wenzelm
parents: 57596
diff changeset
   412
        | NONE => ()) (* FIXME handle protocol failure, just like in active.ML (!?) *)
62505
9e2a65912111 clarified modules;
wenzelm
parents: 60747
diff changeset
   413
      end handle exn => if Exn.is_interrupt exn then () (*sic!*) else Exn.reraise exn)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   414
55552
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   415
e4907b74a347 tuned whitespace;
wenzelm
parents: 55390
diff changeset
   416
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   417
(** attributes **)
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   418
55560
7ac8f013321c proper term equality;
wenzelm
parents: 55553
diff changeset
   419
val mode_parser =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   420
  Scan.optional
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   421
    (Args.$$$ "mode" |-- Args.$$$ "=" |-- (Args.$$$ "normal" || Args.$$$ "full"))
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   422
    "normal"
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
   423
55560
7ac8f013321c proper term equality;
wenzelm
parents: 55553
diff changeset
   424
val interactive_parser =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   425
  Scan.optional (Args.$$$ "interactive" >> K true) false
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   426
55560
7ac8f013321c proper term equality;
wenzelm
parents: 55553
diff changeset
   427
val memory_parser =
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   428
  Scan.optional (Args.$$$ "no_memory" >> K false) true
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   429
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   430
val depth_parser =
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   431
  Scan.optional (Args.$$$ "depth" |-- Args.$$$ "=" |-- Parse.nat) 10
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   432
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   433
val config_parser =
55390
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   434
  (interactive_parser -- mode_parser -- depth_parser -- memory_parser) >>
36550a4eac5e "no_memory" option for the simplifier trace to bypass memoization
Lars Hupel <lars.hupel@mytum.de>
parents: 55335
diff changeset
   435
    (fn (((interactive, mode), depth), memory) => config mode interactive depth memory)
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
   436
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   437
val _ = Theory.setup
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 63806
diff changeset
   438
  (Attrib.setup \<^binding>\<open>simp_break\<close>
57041
aceaca232177 consolidate "break_thm" and "break_term" attributes into "simp_break";
Lars Hupel <lars.hupel@mytum.de>
parents: 56333
diff changeset
   439
    (Scan.repeat Args.term_pattern >> breakpoint)
aceaca232177 consolidate "break_thm" and "break_term" attributes into "simp_break";
Lars Hupel <lars.hupel@mytum.de>
parents: 56333
diff changeset
   440
    "declaration of a simplifier breakpoint" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 63806
diff changeset
   441
   Attrib.setup \<^binding>\<open>simp_trace_new\<close> (Scan.lift config_parser)
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 54731
diff changeset
   442
    "simplifier trace configuration")
54730
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
   443
de2d99b459b3 skeleton for Simplifier trace by Lars Hupel;
wenzelm
parents:
diff changeset
   444
end